author | Fabian Huch <huch@in.tum.de> |
Fri, 17 Nov 2023 09:38:15 +0100 | |
changeset 78982 | be5c078f5292 |
parent 78981 | 47f8533c6887 |
child 78983 | 295aa95cbff9 |
permissions | -rw-r--r-- |
78316 | 1 |
/* Title: Pure/General/toml.scala |
78317 | 2 |
Author: Fabian Huch, TU Muenchen |
78316 | 3 |
|
4 |
Support for TOML: https://toml.io/en/v1.0.0 |
|
78317 | 5 |
*/ |
6 |
||
78316 | 7 |
package isabelle |
8 |
||
9 |
||
10 |
import TOML.Parse_Context.Seen |
|
11 |
||
12 |
import java.lang.Long.parseLong |
|
13 |
import java.lang.{String => Str} |
|
14 |
import java.time.{LocalDate, LocalDateTime, LocalTime, OffsetDateTime} |
|
15 |
||
16 |
import scala.{Boolean => Bool} |
|
17 |
import scala.collection.immutable.ListMap |
|
18 |
import scala.reflect.{ClassTag, classTag} |
|
19 |
import scala.util.Try |
|
20 |
import scala.util.matching.Regex |
|
21 |
import scala.util.parsing.combinator |
|
22 |
import scala.util.parsing.combinator.lexical.Scanners |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
23 |
import scala.util.parsing.input |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
24 |
import scala.util.parsing.input.Positional |
78316 | 25 |
import scala.util.parsing.input.CharArrayReader.EofCh |
26 |
||
27 |
||
28 |
object TOML { |
|
29 |
/* typed representation and access */ |
|
30 |
||
31 |
type Key = Str |
|
32 |
||
33 |
sealed trait T |
|
34 |
case class String(rep: Str) extends T |
|
35 |
case class Integer(rep: Long) extends T |
|
36 |
case class Float(rep: Double) extends T |
|
37 |
case class Boolean(rep: Bool) extends T |
|
38 |
case class Offset_Date_Time(rep: OffsetDateTime) extends T |
|
39 |
case class Local_Date_Time(rep: LocalDateTime) extends T |
|
40 |
case class Local_Date(rep: LocalDate) extends T |
|
41 |
case class Local_Time(rep: LocalTime) extends T |
|
42 |
||
78941 | 43 |
object Scalar { |
44 |
def unapply(t: T): Option[Str] = |
|
45 |
t match { |
|
46 |
case s: String => Some(s.rep) |
|
47 |
case i: Integer => Some(i.rep.toString) |
|
48 |
case f: Float => Some(f.rep.toString) |
|
49 |
case b: Boolean => Some(b.rep.toString) |
|
50 |
case o: Offset_Date_Time => Some(o.rep.toString) |
|
51 |
case l: Local_Date_Time => Some(l.rep.toString) |
|
52 |
case l: Local_Date => Some(l.rep.toString) |
|
53 |
case l: Local_Time => Some(l.rep.toString) |
|
54 |
case _ => None |
|
55 |
} |
|
56 |
} |
|
57 |
||
78316 | 58 |
class Array private(private val rep: List[T]) extends T { |
59 |
override def hashCode(): Int = rep.hashCode() |
|
60 |
override def equals(that: Any): Bool = that match { |
|
61 |
case other: Array => rep == other.rep |
|
62 |
case _ => false |
|
63 |
} |
|
78978
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
64 |
override def toString: Str = "Array(" + length.toString + ")" |
78316 | 65 |
|
66 |
class Values[A](pf: PartialFunction[T, A]) { def values: List[A] = rep.collect(pf).reverse } |
|
67 |
lazy val string = new Values({ case s: String => s }) |
|
68 |
lazy val integer = new Values({ case i: Integer => i }) |
|
69 |
lazy val float = new Values({ case f: Float => f }) |
|
70 |
lazy val boolean = new Values({ case b: Boolean => b }) |
|
71 |
lazy val offset_date_time = new Values({ case o: Offset_Date_Time => o }) |
|
72 |
lazy val local_date_time = new Values({ case l: Local_Date_Time => l }) |
|
73 |
lazy val local_date = new Values({ case l: Local_Date => l }) |
|
74 |
lazy val local_time = new Values({ case l: Local_Time => l }) |
|
75 |
lazy val array = new Values({ case a: Array => a }) |
|
76 |
lazy val table = new Values({ case t: Table => t }) |
|
77 |
lazy val any = new Values({ case t => t }) |
|
78 |
||
79 |
def +(elem: T): Array = new Array(elem :: rep) |
|
80 |
def ++(other: Array): Array = new Array(other.rep ::: rep) |
|
81 |
def length: Int = rep.length |
|
82 |
def is_empty: Bool = rep.isEmpty |
|
83 |
} |
|
84 |
||
85 |
object Array { |
|
86 |
def apply(elems: Iterable[T]): Array = new Array(elems.toList.reverse) |
|
87 |
def apply(elems: T*): Array = Array(elems) |
|
78919 | 88 |
val empty: Array = apply() |
78316 | 89 |
} |
90 |
||
91 |
class Table private(private val rep: Map[Key, T]) extends T { |
|
92 |
override def hashCode(): Int = rep.hashCode() |
|
93 |
override def equals(that: Any): Bool = |
|
94 |
that match { |
|
95 |
case other: Table => rep == other.rep |
|
96 |
case _ => false |
|
97 |
} |
|
78978
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
98 |
override def toString: Str = |
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
99 |
rep.map { |
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
100 |
case (k, t: Table) => k + "(" + t.domain.size + ")" |
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
101 |
case (k, a: Array) => k + "(" + a.length + ")" |
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
102 |
case (k, _) => k |
80fda74a045d
clarified toString for toml objects;
Fabian Huch <huch@in.tum.de>
parents:
78948
diff
changeset
|
103 |
}.mkString("Table(", ", ", ")") |
78316 | 104 |
|
105 |
class Value[A: ClassTag](pf: PartialFunction[T, A]) { |
|
106 |
def values: List[(Key, A)] = |
|
107 |
rep.toList.collect { case (k, v) if pf.isDefinedAt(v) => k -> pf(v) } |
|
108 |
def get(k: Key): Option[A] = rep.get(k).flatMap(v => PartialFunction.condOpt(v)(pf)) |
|
109 |
def apply(k: Key): A = |
|
110 |
rep.get(k) match { |
|
111 |
case Some(v) => PartialFunction.condOpt(v)(pf) match { |
|
112 |
case Some(value) => value |
|
113 |
case None => |
|
114 |
error("Expected" + classTag[A].runtimeClass.getName + |
|
78327 | 115 |
", got " + v.getClass.getSimpleName + " for key " + Format.key(k)) |
78316 | 116 |
} |
78327 | 117 |
case None => error("Key " + Format.key(k) + " does not exist") |
78316 | 118 |
} |
119 |
} |
|
120 |
||
121 |
lazy val string = new Value({ case s: String => s }) |
|
122 |
lazy val integer = new Value({ case i: Integer => i }) |
|
123 |
lazy val float = new Value({ case f: Float => f }) |
|
124 |
lazy val boolean = new Value({ case b: Boolean => b }) |
|
125 |
lazy val offset_date_time = new Value({ case o: Offset_Date_Time => o }) |
|
126 |
lazy val local_date_time = new Value({ case l: Local_Date_Time => l }) |
|
127 |
lazy val local_date = new Value({ case l: Local_Date => l }) |
|
128 |
lazy val local_time = new Value({ case l: Local_Time => l }) |
|
129 |
lazy val array = new Value({ case a: Array => a }) |
|
130 |
lazy val table = new Value({ case t: Table => t }) |
|
131 |
lazy val any = new Value({ case t => t }) |
|
132 |
||
133 |
def +(elem: (Key, T)): Table = { |
|
134 |
val (k, v) = elem |
|
135 |
val v1 = rep.get(k) match { |
|
136 |
case None => v |
|
137 |
case Some(v0) => |
|
138 |
(v0, v) match { |
|
139 |
case (t0: Table, t: Table) => t0 ++ t |
|
140 |
case (a0: Array, a: Array) => a0 ++ a |
|
78979
d5ce982ae60f
allow re-defining keys in toml object (already checked during parse time);
Fabian Huch <huch@in.tum.de>
parents:
78978
diff
changeset
|
141 |
case _ => v |
78316 | 142 |
} |
143 |
} |
|
144 |
new Table(rep + (k -> v1)) |
|
145 |
} |
|
146 |
def -(k: Key): Table = new Table(rep - k) |
|
147 |
def ++(other: Table): Table = other.rep.foldLeft(this)(_ + _) |
|
148 |
def domain: Set[Key] = rep.keySet |
|
149 |
def is_empty: Bool = rep.isEmpty |
|
150 |
} |
|
151 |
||
152 |
object Table { |
|
153 |
def apply(elems: Iterable[(Key, T)]): Table = elems.foldLeft(new Table(ListMap.empty))(_ + _) |
|
154 |
def apply(elems: (Key, T)*): Table = Table(elems) |
|
78919 | 155 |
val empty: Table = apply() |
78316 | 156 |
} |
157 |
||
158 |
||
159 |
/* lexer */ |
|
160 |
||
78609 | 161 |
enum Kind { case KEYWORD, VALUE, STRING, MULTILINE_STRING, LINE_SEP, ERROR } |
78316 | 162 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
163 |
sealed case class Token(kind: Kind, text: Str) extends Positional { |
78316 | 164 |
def is_keyword(name: Str): Bool = kind == Kind.KEYWORD && text == name |
165 |
def is_value: Bool = kind == Kind.VALUE |
|
166 |
def is_string: Bool = kind == Kind.STRING |
|
167 |
def is_multiline_string: Bool = kind == Kind.MULTILINE_STRING |
|
168 |
def is_line_sep: Bool = kind == Kind.LINE_SEP |
|
169 |
} |
|
170 |
||
171 |
object Lexer extends Scanners with Scan.Parsers { |
|
172 |
override type Elem = Char |
|
173 |
type Token = TOML.Token |
|
174 |
||
175 |
def errorToken(msg: Str): Token = Token(Kind.ERROR, msg) |
|
176 |
||
177 |
val white_space: Str = " \t" |
|
178 |
override val whiteSpace: Regex = ("[" + white_space + "]+").r |
|
179 |
override def whitespace: Parser[Any] = rep(comment | many1(character(white_space.contains(_)))) |
|
180 |
||
181 |
def line_sep: Parser[Str] = rep1("\n" | s"\r\n" | EofCh) ^^ (cs => cs.mkString) |
|
182 |
def line_sep_token: Parser[Token] = line_sep ^^ (s => Token(Kind.LINE_SEP, s)) |
|
183 |
||
184 |
def is_control(e: Elem): Bool = |
|
185 |
e <= '\u0008' || ('\u000A' <= e && e <= '\u001F') || e == '\u007F' |
|
186 |
||
187 |
override def comment: Parser[Str] = '#' ~>! many(character(c => !is_control(c))) |
|
188 |
||
189 |
def keyword: Parser[Token] = one(character("{}[],=.".contains)) ^^ (s => Token(Kind.KEYWORD, s)) |
|
190 |
||
191 |
def is_value(c: Elem): Bool = |
|
192 |
Symbol.is_ascii_letter(c) || Symbol.is_ascii_digit(c) || "_-:+".contains(c) |
|
193 |
def value: Parser[Token] = |
|
194 |
many1(character(is_value)) ~ |
|
195 |
opt(' ' ~ many1(character(is_value)) ^^ { case ws ~ s => ws.toString + s }) ~ |
|
196 |
opt('.' ~ many1(character(is_value)) ^^ { case dot ~ s => dot.toString + s}) ^^ |
|
197 |
{ case s ~ ss ~ sd => Token(Kind.VALUE, s + ss.getOrElse("") + sd.getOrElse("")) } |
|
198 |
||
199 |
def string: Parser[Token] = |
|
200 |
multiline_basic_string | basic_string | multiline_literal_string | literal_string |
|
201 |
||
202 |
private def trim(s: Str): Str = |
|
203 |
if (s.startsWith("\n")) s.stripPrefix("\n") else s.stripPrefix("\r\n") |
|
204 |
||
205 |
def basic_string: Parser[Token] = |
|
206 |
'"' ~> rep(basic_string_elem) <~ '"' ^^ (cs => Token(Kind.STRING, cs.mkString)) |
|
207 |
||
208 |
def multiline_basic_string: Parser[Token] = |
|
209 |
"\"\"\"" ~> |
|
210 |
rep(multiline_basic_string_elem | |
|
211 |
("\"\"" | "\"") ~ multiline_basic_string_elem ^^ { case s ~ t => s + t }) ~ |
|
212 |
repeated(character(_ == '"'), 3, 5) ^^ { case cs ~ q => |
|
213 |
Token(Kind.MULTILINE_STRING, trim(cs.mkString + q.drop(3))) } |
|
214 |
||
215 |
private def multiline_basic_string_elem: Parser[Str] = |
|
216 |
('\\' ~ line_sep ~ rep(many1(character(white_space.contains)) | line_sep)) ^^ (_ => "") | |
|
217 |
basic_string_elem ^^ (_.toString) | line_sep |
|
218 |
||
219 |
def literal_string: Parser[Token] = |
|
220 |
'\'' ~> rep(literal_string_elem) <~ '\'' ^^ (cs => Token(Kind.STRING, cs.mkString)) |
|
221 |
||
222 |
def multiline_literal_string: Parser[Token] = |
|
223 |
"'''" ~> |
|
224 |
rep(multiline_literal_string_elem | |
|
225 |
("''" | "'") ~ multiline_literal_string_elem ^^ { case s ~ t => s + t }) ~ |
|
226 |
repeated(character(_ == '\''), 3, 5) ^^ { case cs ~ q => |
|
227 |
Token(Kind.MULTILINE_STRING, trim(cs.mkString + q.drop(3))) } |
|
228 |
||
229 |
private def multiline_literal_string_elem: Parser[Str] = |
|
230 |
line_sep | literal_string_elem ^^ (_.toString) |
|
231 |
||
232 |
private def basic_string_elem: Parser[Elem] = |
|
233 |
elem("", c => !is_control(c) && !"\"\\".contains(c)) | '\\' ~> string_escape |
|
234 |
||
235 |
private def string_escape: Parser[Elem] = |
|
236 |
elem("", "\"\\".contains(_)) | |
|
237 |
elem("", "btnfr".contains(_)) ^^ |
|
238 |
{ case 'b' => '\b' case 't' => '\t' case 'n' => '\n' case 'f' => '\f' case 'r' => '\r' } | |
|
78920 | 239 |
('u' ~> repeated(character(Symbol.is_ascii_hex), 4, 4) | |
240 |
'U' ~> repeated(character(Symbol.is_ascii_hex), 8, 8)) ^^ |
|
78316 | 241 |
(s => java.lang.Integer.parseInt(s, 16).toChar) |
242 |
||
243 |
private def literal_string_elem: Parser[Elem] = elem("", c => !is_control(c) && c != '\'') |
|
244 |
||
245 |
def string_failure: Parser[Token] = ("\"" | "'") ~> failure("Unterminated string") |
|
246 |
||
247 |
def token: Parser[Token] = |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
248 |
positioned( |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
249 |
line_sep_token | keyword | value | string | string_failure | failure("Unrecognized token")) |
78316 | 250 |
} |
251 |
||
252 |
||
253 |
/* parser */ |
|
254 |
||
255 |
trait Parsers extends combinator.Parsers { |
|
256 |
type Elem = Token |
|
78317 | 257 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
258 |
case class Keys(rep: List[Key]) extends Positional { |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
259 |
def last: Keys = Keys(rep.takeRight(1)) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
260 |
def the_last: Key = rep.last |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
261 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
262 |
def head: Keys = Keys(rep.take(1)) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
263 |
def the_head: Key = rep.head |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
264 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
265 |
def length: Int = rep.length |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
266 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
267 |
def +(other: Keys): Keys = Keys(rep ::: other.rep) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
268 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
269 |
def prefixes: Set[Keys] = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
270 |
if (rep.isEmpty) Set.empty |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
271 |
else Range.inclusive(1, length).toSet.map(i => Keys(rep.take(i))) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
272 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
273 |
def splits: List[(Keys, Keys)] = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
274 |
Range.inclusive(0, length).toList.map(n => |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
275 |
(Keys(rep.dropRight(n)), Keys(rep.takeRight(n)))) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
276 |
def split(i: Int): (Keys, Keys) = { |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
277 |
val (rep0, rep1) = rep.splitAt(i) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
278 |
(Keys(rep0), Keys(rep1)) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
279 |
} |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
280 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
281 |
def parent: Keys = Keys(rep.dropRight(1)) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
282 |
def is_child_of(parent: Keys): Bool = rep.startsWith(parent.rep) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
283 |
} |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
284 |
|
78317 | 285 |
|
78316 | 286 |
/* parse structure */ |
78317 | 287 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
288 |
sealed trait V extends Positional |
78316 | 289 |
case class Primitive(t: T) extends V |
290 |
case class Array(rep: List[V]) extends V |
|
291 |
case class Inline_Table(elems: List[(Keys, V)]) extends V |
|
292 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
293 |
sealed trait Def extends Positional |
78316 | 294 |
case class Table(key: Keys, elems: List[(Keys, V)]) extends Def |
295 |
case class Array_Of_Tables(key: Keys, elems: List[(Keys, V)]) extends Def |
|
78317 | 296 |
|
78316 | 297 |
case class File(elems: List[(Keys, V)], defs: List[Def]) |
78317 | 298 |
|
78316 | 299 |
|
300 |
/* top-level syntax structure */ |
|
78317 | 301 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
302 |
def key: Parser[Keys] = positioned(rep1sep(keys, $$$(".")) ^^ (ks => Keys(ks.flatten))) |
78317 | 303 |
|
78316 | 304 |
def string: Parser[String] = |
305 |
elem("string", e => e.is_string || e.is_multiline_string) ^^ (s => String(s.text)) |
|
306 |
def integer: Parser[Integer] = |
|
307 |
(decimal_int | binary_int | octal_int | hexadecimal_int) ^^ Integer.apply |
|
308 |
def float: Parser[Float] = (symbol_float | number_float) ^^ Float.apply |
|
309 |
def boolean: Parser[Boolean] = token("boolean", _.is_value, s => Boolean(Value.Boolean.parse(s))) |
|
310 |
||
311 |
def offset_date_time: Parser[Offset_Date_Time] = |
|
312 |
token("offset date-time", _.is_value, |
|
313 |
s => Offset_Date_Time(OffsetDateTime.parse(s.replace(" ", "T")))) |
|
314 |
def local_date_time: Parser[Local_Date_Time] = |
|
315 |
token("local date-time", _.is_value, |
|
316 |
s => Local_Date_Time(LocalDateTime.parse(s.replace(" ", "T")))) |
|
317 |
def local_date: Parser[Local_Date] = |
|
318 |
token("local date", _.is_value, s => Local_Date(LocalDate.parse(s))) |
|
319 |
def local_time: Parser[Local_Time] = |
|
320 |
token("local time", _.is_value, s => Local_Time(LocalTime.parse(s))) |
|
321 |
||
322 |
def array: Parser[Array] = |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
323 |
positioned( |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
324 |
$$$("[") ~>! repsep(opt(line_sep) ~> toml_value, opt(line_sep) ~ $$$(",")) <~! |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
325 |
opt(line_sep) ~! opt($$$(",")) ~! opt(line_sep) <~! $$$("]") ^^ Array.apply) |
78316 | 326 |
|
327 |
def inline_table: Parser[Inline_Table] = |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
328 |
positioned( |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
329 |
$$$("{") ~>! repsep(pair, $$$(",")) <~! $$$("}") ^^ Inline_Table.apply) |
78317 | 330 |
|
78316 | 331 |
def pair: Parser[(Keys, V)] = (key <~! $$$("=")) ~! toml_value ^^ { case ks ~ v => (ks, v) } |
332 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
333 |
def table: Parser[Table] = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
334 |
positioned( |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
335 |
$$$("[") ~> (key <~! $$$("]") ~! line_sep) ~! content ^^ |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
336 |
{ case key ~ content => Table(key, content) }) |
78316 | 337 |
|
338 |
def array_of_tables: Parser[Array_Of_Tables] = |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
339 |
positioned( |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
340 |
$$$("[") ~ $$$("[") ~>! (key <~! $$$("]") ~! $$$("]") ~! line_sep) ~! content ^^ |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
341 |
{ case key ~ content => Array_Of_Tables(key, content) }) |
78316 | 342 |
|
343 |
def toml: Parser[File] = |
|
344 |
(opt(line_sep) ~>! content ~! rep(table | array_of_tables)) ^^ |
|
345 |
{ case content ~ defs => File(content, defs) } |
|
346 |
||
347 |
||
348 |
/* auxiliary */ |
|
349 |
||
350 |
private def $$$(name: Str): Parser[Token] = elem(name, _.is_keyword(name)) |
|
351 |
private def maybe[A, B](p: Parser[A], f: A => B): Parser[B] = |
|
352 |
p ^^ (a => Try(f(a))) ^? { case util.Success(v) => v } |
|
353 |
private def token[A](name: Str, p: Token => Bool, parser: Str => A): Parser[A] = |
|
354 |
maybe(elem(name, p), s => parser(s.text)) |
|
355 |
private def prefixed[A]( |
|
356 |
prefix: Str, name: Str, p: Str => Bool, parser: Str => A |
|
357 |
): Parser[A] = |
|
358 |
token(name, e => e.is_value && e.text.startsWith(prefix) && p(e.text.stripPrefix(prefix)), |
|
359 |
s => parser(s.stripPrefix(prefix))) |
|
360 |
||
361 |
private def is_key(e: Elem): Bool = e.is_value && !e.text.exists("+: ".contains(_)) |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
362 |
private def keys: Parser[List[Key]] = |
78316 | 363 |
token("string key", _.is_string, List(_)) | token("key", is_key, _.split('.').toList) |
364 |
||
365 |
private def sep_surrounded(s: Str): Bool = |
|
366 |
!s.startsWith("_") && !s.endsWith("_") && s.split('_').forall(_.nonEmpty) |
|
367 |
private def no_leading_zero(s: Str): Bool = { |
|
368 |
val t = s.replaceAll("_", "").takeWhile(_.isDigit) |
|
369 |
t == "0" || !t.startsWith("0") |
|
370 |
} |
|
371 |
||
372 |
private def is_int(s: Str): Bool = |
|
373 |
no_leading_zero(s.replaceAll("[-+]", "")) && sep_surrounded(s.replaceAll("[-+]", "")) |
|
374 |
private def decimal_int: Parser[Long] = |
|
375 |
token("integer", e => e.is_value && is_int(e.text), _.replace("_", "").toLong) |
|
376 |
private def binary_int: Parser[Long] = |
|
377 |
prefixed("0b", "integer", sep_surrounded, s => parseLong(s.replace("_", ""), 2)) |
|
378 |
private def octal_int: Parser[Long] = |
|
379 |
prefixed("0o", "integer", sep_surrounded, s => parseLong(s.replace("_", ""), 8)) |
|
380 |
private def hexadecimal_int: Parser[Long] = |
|
381 |
prefixed("0x", "integer", sep_surrounded, s => parseLong(s.replace("_", ""), 16)) |
|
382 |
||
383 |
private def is_float(s: Str): Bool = |
|
384 |
s.exists(".eE".contains) && s.count("eE".contains) <= 1 && |
|
385 |
no_leading_zero(s.replaceAll("[-+]", "")) && |
|
386 |
sep_surrounded(s.replaceAll("[-+]", "").replaceAll("[.eE][+\\-]?", "_")) |
|
387 |
private def number_float: Parser[Double] = |
|
388 |
token("float", e => e.is_value && is_float(e.text), _.replace("_", "").toDouble) |
|
389 |
private def symbol_float: Parser[Double] = |
|
390 |
token("float", _.is_value, { |
|
391 |
case "inf" | "+inf" => Double.PositiveInfinity |
|
392 |
case "-inf" => Double.NegativeInfinity |
|
393 |
case "nan" | "+nan" | "-nan" => Double.NaN |
|
394 |
}) |
|
395 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
396 |
private def toml_value: Parser[V] = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
397 |
positioned( |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
398 |
(string | float | integer | boolean | offset_date_time | |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
399 |
local_date_time | local_date | local_time) ^^ Primitive.apply | array | inline_table) |
78316 | 400 |
|
401 |
private def line_sep: Parser[Any] = rep1(elem("line sep", _.is_line_sep)) |
|
402 |
||
403 |
private def content: Parser[List[(Keys, V)]] = |
|
404 |
rep((key <~! $$$("=")) ~! toml_value <~! line_sep ^^ { case ks ~ v => ks -> v }) |
|
405 |
||
406 |
||
407 |
/* parse */ |
|
408 |
||
409 |
def parse(input: Str): File = { |
|
410 |
val scanner = new Lexer.Scanner(Scan.char_reader(input + EofCh)) |
|
411 |
val result = phrase(toml)(scanner) |
|
412 |
result match { |
|
413 |
case Success(toml, _) => toml |
|
414 |
case Failure(msg, next) => error("Malformed TOML input at " + next.pos + ": " + msg) |
|
415 |
case Error(msg, next) => error("Error parsing toml at " + next.pos + ": " + msg) |
|
416 |
} |
|
417 |
} |
|
418 |
} |
|
419 |
||
420 |
object Parsers extends Parsers |
|
421 |
||
422 |
||
423 |
/* checking and translating parse structure into toml */ |
|
424 |
||
78982
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
425 |
class Parse_Context private(var seen_tables: Map[Parsers.Keys, Seen], file: Option[Path] = None) { |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
426 |
private def recent_array(ks: Parsers.Keys): (Parsers.Keys, Seen, Parsers.Keys) = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
427 |
ks.splits.collectFirst { |
78316 | 428 |
case (ks1, ks2) if seen_tables.contains(ks1) => (ks1, seen_tables(ks1), ks2) |
429 |
}.get |
|
430 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
431 |
private def check_add(start: Int, ks: Parsers.Keys, elem: Parsers.Def | Parsers.V): Unit = { |
78316 | 432 |
val (to, seen, rest, split) = |
433 |
elem match { |
|
434 |
case _: Parsers.Array_Of_Tables => |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
435 |
val (_, seen, rest) = recent_array(ks.parent) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
436 |
(ks, seen, rest + ks.last, 0) |
78316 | 437 |
case _ => |
438 |
val (to, seen, rest) = recent_array(ks) |
|
439 |
(to, seen, rest, start - to.length) |
|
440 |
} |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
441 |
val (rest0, rest1) = rest.split(split) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
442 |
val implicitly_seen = rest1.parent.prefixes.map(rest0 + _) |
78316 | 443 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
444 |
def error[A](s: Str): A = this.error(s, ks, elem) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
445 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
446 |
seen.inlines.find(rest.is_child_of).foreach(ks => |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
447 |
error("Attempting to update an inline value")) |
78316 | 448 |
|
449 |
val (inlines, tables) = |
|
450 |
elem match { |
|
451 |
case _: Parsers.Primitive => |
|
452 |
(seen.inlines, seen.tables ++ implicitly_seen) |
|
453 |
case _: Parsers.Array => |
|
78327 | 454 |
if (seen_tables.contains(ks)) |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
455 |
error("Attempting to append with an inline array") |
78316 | 456 |
(seen.inlines + rest, seen.tables ++ implicitly_seen) |
457 |
case _: Parsers.Inline_Table => |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
458 |
seen.tables.find(_.is_child_of(rest)).foreach(ks => |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
459 |
error("Attempting to add with an inline table")) |
78316 | 460 |
(seen.inlines + rest, seen.tables ++ implicitly_seen) |
461 |
case _: Parsers.Table => |
|
78327 | 462 |
if (seen.tables.contains(rest)) |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
463 |
error("Attempting to define a table twice") |
78316 | 464 |
(seen.inlines, seen.tables + rest) |
465 |
case _: Parsers.Array_Of_Tables => (Set.empty, Set.empty) |
|
466 |
} |
|
467 |
||
468 |
seen_tables += to -> Seen(inlines, tables) |
|
469 |
} |
|
78982
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
470 |
|
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
471 |
def for_file(file: Path): Parse_Context = new Parse_Context(seen_tables, Some(file)) |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
472 |
|
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
473 |
def error[A](s: Str, ks: Parsers.Keys, elem: Positional): A = { |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
474 |
val file_msg = if_proper(file, " in " + file.get) |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
475 |
isabelle.error( |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
476 |
s + " in table " + Format.keys(ks.rep) + " at line " + elem.pos.line + file_msg) |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
477 |
} |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
478 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
479 |
def check_add_def(ks: Parsers.Keys, defn: Parsers.Def): Unit = check_add(0, ks, defn) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
480 |
def check_add_value(ks0: Parsers.Keys, ks1: Parsers.Keys, v: Parsers.V): Unit = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
481 |
check_add(ks0.length - 1, ks0 + ks1, v) |
78316 | 482 |
} |
483 |
||
484 |
object Parse_Context { |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
485 |
case class Seen(inlines: Set[Parsers.Keys], tables: Set[Parsers.Keys]) |
78316 | 486 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
487 |
def apply(): Parse_Context = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
488 |
new Parse_Context(Map(Parsers.Keys(Nil) -> Seen(Set.empty, Set.empty))) |
78316 | 489 |
} |
490 |
||
78927 | 491 |
def parse(s: Str, context: Parse_Context = Parse_Context()): Table = { |
78316 | 492 |
val file = Parsers.parse(s) |
493 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
494 |
def convert(ks0: Parsers.Keys, ks1: Parsers.Keys, v: Parsers.V): T = { |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
495 |
def to_table(ks: Parsers.Keys, t: T): Table = |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
496 |
ks.parent.rep.foldRight(Table(ks.the_last -> t))((k, v) => Table(k -> v)) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
497 |
def to_toml(v: Parsers.V): (T, Set[Parsers.Keys]) = v match { |
78316 | 498 |
case Parsers.Primitive(t) => (t, Set.empty) |
499 |
case Parsers.Array(rep) => (Array(rep.map(to_toml(_)._1)), Set.empty) |
|
500 |
case Parsers.Inline_Table(elems) => |
|
78327 | 501 |
elems.find(e => elems.count(_._1 == e._1) > 1).foreach((ks, _) => |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
502 |
context.error("Duplicate key " + Format.keys(ks.rep) + " in inline table", ks0 + ks1, v)) |
78316 | 503 |
|
504 |
val (tables, pfxs, key_spaces) = |
|
505 |
elems.map { (ks, v) => |
|
506 |
val (t, keys) = to_toml(v) |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
507 |
(to_table(ks, t), ks.prefixes, keys.map(ks + _) + ks) |
78316 | 508 |
}.unzip3 |
509 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
510 |
for { |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
511 |
pfx <- pfxs |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
512 |
if key_spaces.count(pfx.intersect(_).nonEmpty) > 1 |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
513 |
} context.error("Inline table not self-contained", ks0 + ks1, v) |
78316 | 514 |
|
515 |
(tables.foldLeft(Table())(_ ++ _), pfxs.toSet.flatten ++ key_spaces.toSet.flatten) |
|
516 |
} |
|
517 |
context.check_add_value(ks0, ks1, v) |
|
518 |
to_toml(v)._1 |
|
519 |
} |
|
520 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
521 |
def update(map: Table, ks0: Parsers.Keys, value: T): Table = { |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
522 |
def error[A](s: Str): A = context.error(s, ks0, ks0) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
523 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
524 |
def update_rec(hd: Parsers.Keys, map: Table, ks: Parsers.Keys): Table = { |
78327 | 525 |
val updated = |
526 |
if (ks.length == 1) { |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
527 |
map.any.get(ks.the_head) match { |
78327 | 528 |
case Some(a: Array) => |
529 |
value match { |
|
530 |
case a2: Array => a ++ a2 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
531 |
case _ => error("Table conflicts with previous array of tables") |
78327 | 532 |
} |
533 |
case Some(t: Table) => value match { |
|
534 |
case t2: Table => |
|
535 |
if (t.domain.intersect(t2.domain).nonEmpty) |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
536 |
error("Attempting to redefine existing value in super-table") |
78327 | 537 |
else t ++ t2 |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
538 |
case _ => error("Attempting to redefine a table") |
78316 | 539 |
} |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
540 |
case Some(_) => error("Attempting to redefine a value") |
78327 | 541 |
case None => value |
78316 | 542 |
} |
543 |
} |
|
78327 | 544 |
else { |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
545 |
val (head, tail) = ks.split(1) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
546 |
val hd1 = hd + head |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
547 |
map.any.get(ks.the_head) match { |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
548 |
case Some(t: Table) => update_rec(hd1, t, tail) |
78327 | 549 |
case Some(a: Array) => |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
550 |
Array(a.table.values.dropRight(1) :+ update_rec(hd1, a.table.values.last, tail)) |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
551 |
case Some(_) => error("Attempting to redefine a value") |
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
552 |
case None => update_rec(hd1, Table(), tail) |
78327 | 553 |
} |
78316 | 554 |
} |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
555 |
(map - ks.the_head) + (ks.the_head -> updated) |
78327 | 556 |
} |
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
557 |
|
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
558 |
update_rec(Parsers.Keys(Nil), map, ks0) |
78316 | 559 |
} |
560 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
561 |
def fold(elems: List[(Parsers.Keys, T)]): Table = |
78316 | 562 |
elems.foldLeft(Table()) { case (t0, (ks1, t1)) => update(t0, ks1, t1) } |
563 |
||
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
564 |
val t = fold(file.elems.map((ks, v) => (ks, convert(Parsers.Keys(Nil), ks, v)))) |
78316 | 565 |
file.defs.foldLeft(t) { |
566 |
case (t0, defn@Parsers.Table(ks0, elems)) => |
|
567 |
context.check_add_def(ks0, defn) |
|
568 |
update(t0, ks0, fold(elems.map((ks, v) => (ks, convert(ks0, ks, v))))) |
|
569 |
case (t0, defn@Parsers.Array_Of_Tables(ks0, elems)) => |
|
570 |
context.check_add_def(ks0, defn) |
|
571 |
update(t0, ks0, Array(fold(elems.map((ks, v) => (ks, convert(ks0, ks, v)))))) |
|
572 |
} |
|
573 |
} |
|
574 |
||
78980
a80ee3c97aae
properly concatenate toml files: regular toml rules still apply (e.g., inline values may not be changed), but values defined in one file may be updated in another;
Fabian Huch <huch@in.tum.de>
parents:
78979
diff
changeset
|
575 |
def parse_files(files: Iterable[Path], context: Parse_Context = Parse_Context()): Table = |
78982
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
576 |
files.foldLeft((Table(), context)) { |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
577 |
case ((t, context0), file) => |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
578 |
val context = context0.for_file(file) |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
579 |
(t ++ parse(File.read(file), context), context) |
be5c078f5292
add file information to toml parse context and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78981
diff
changeset
|
580 |
}._1 |
78939 | 581 |
|
78316 | 582 |
|
583 |
/* Format TOML */ |
|
584 |
||
585 |
object Format { |
|
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
586 |
private def basic_string(s: Str): Str = |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
587 |
"\"" + s.iterator.map { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
588 |
case '\b' => "\\b" case '\t' => "\\t" case '\n' => "\\n" case '\f' => "\\f" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
589 |
case '\r' => "\\r" case '"' => "\\\"" case '\\' => "\\\\" case c => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
590 |
if (c <= '\u001f' || c == '\u007f') "\\u%04x".format(c.toInt) else c |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
591 |
}.mkString + "\"" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
592 |
|
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
593 |
private def multiline_basic_string(s: Str): Str = |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
594 |
"\"\"\"\n" + s.iterator.map { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
595 |
case '\b' => "\\b" case '\t' => "\t" case '\n' => "\n" case '\f' => "\\f" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
596 |
case '\r' => "\r" case '"' => "\\\"" case '\\' => "\\\\" case c => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
597 |
if (c <= '\u001f' || c == '\u007f') "\\u%04x".format(c.toInt) else c |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
598 |
}.mkString.replace("\"\"\"", "\"\"\\\"") + "\"\"\"" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
599 |
|
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
600 |
def key(k: Key): Str = { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
601 |
val Bare_Key = """[A-Za-z0-9_-]+""".r |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
602 |
k match { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
603 |
case Bare_Key() => k |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
604 |
case _ => basic_string(k) |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
605 |
} |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
606 |
} |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
607 |
|
78981
47f8533c6887
add position information to toml parser and error messages;
Fabian Huch <huch@in.tum.de>
parents:
78980
diff
changeset
|
608 |
def keys(ks: List[Key]): Str = ks.map(key).mkString(".") |
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
609 |
|
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
610 |
def inline(t: T, indent: Int = 0): Str = { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
611 |
val result = new StringBuilder |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
612 |
|
78917 | 613 |
result ++= Symbol.spaces(indent * 2) |
78941 | 614 |
(t: @unchecked) match { |
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
615 |
case s: String => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
616 |
if (s.rep.contains("\n") && s.rep.length > 20) result ++= multiline_basic_string(s.rep) |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
617 |
else result ++= basic_string(s.rep) |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
618 |
case a: Array => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
619 |
if (a.is_empty) result ++= "[]" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
620 |
else { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
621 |
result ++= "[\n" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
622 |
a.any.values.foreach { elem => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
623 |
result ++= inline(elem, indent + 1) |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
624 |
result ++= ",\n" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
625 |
} |
78917 | 626 |
result ++= Symbol.spaces(indent * 2) |
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
627 |
result += ']' |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
628 |
} |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
629 |
case table: Table => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
630 |
if (table.is_empty) result ++= "{}" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
631 |
else { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
632 |
result ++= "{ " |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
633 |
Library.separate(None, table.any.values.map(Some(_))).foreach { |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
634 |
case None => result ++= ", " |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
635 |
case Some((k, v)) => |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
636 |
result ++= key(k) |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
637 |
result ++= " = " |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
638 |
result ++= inline(v) |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
639 |
} |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
640 |
result ++= " }" |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
641 |
} |
78941 | 642 |
case Scalar(s) => result ++= s |
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
643 |
} |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
644 |
result.toString() |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
645 |
} |
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
646 |
|
78316 | 647 |
def apply(toml: Table): Str = { |
648 |
val result = new StringBuilder |
|
649 |
||
650 |
def inline_values(path: List[Key], t: T): Unit = |
|
651 |
t match { |
|
652 |
case t: Table => t.any.values.foreach { case (k, v1) => inline_values(k :: path, v1) } |
|
653 |
case _ => |
|
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
654 |
result ++= keys(path.reverse) |
78316 | 655 |
result ++= " = " |
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
656 |
result ++= inline(t) |
78316 | 657 |
result += '\n' |
658 |
} |
|
659 |
||
660 |
def is_inline(elem: T): Bool = |
|
661 |
elem match { |
|
662 |
case _: String | _: Integer | _: Float | _: Boolean | _: Offset_Date_Time | |
|
663 |
_: Local_Date_Time | _: Local_Date | _: Local_Time => true |
|
664 |
case a: Array => a.any.values.forall(is_inline) |
|
665 |
case _ => false |
|
666 |
} |
|
667 |
def is_table(elem: T): Bool = elem match { case _: Table => true case _ => false } |
|
668 |
||
669 |
def array(path: List[Key], a: Array): Unit = |
|
670 |
if (a.any.values.forall(is_inline) || !a.any.values.forall(is_table)) |
|
671 |
inline_values(path.take(1), a) |
|
672 |
else a.table.values.foreach { t => |
|
673 |
result ++= "\n[[" |
|
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
674 |
result ++= keys(path.reverse) |
78316 | 675 |
result ++= "]]\n" |
676 |
table(path, t, is_table_entry = true) |
|
677 |
} |
|
678 |
||
679 |
def table(path: List[Key], t: Table, is_table_entry: Bool = false): Unit = { |
|
680 |
val (values, nodes) = t.any.values.partition(kv => is_inline(kv._2)) |
|
681 |
||
682 |
if (!is_table_entry && path.nonEmpty) { |
|
683 |
result ++= "\n[" |
|
78326
1cdc65476c2e
more TOML formatting functions;
Fabian Huch <huch@in.tum.de>
parents:
78317
diff
changeset
|
684 |
result ++= keys(path.reverse) |
78316 | 685 |
result ++= "]\n" |
686 |
} |
|
687 |
||
688 |
values.foreach { case (k, v) => inline_values(List(k), v) } |
|
689 |
nodes.foreach { |
|
690 |
case (k, t: Table) => table(k :: path, t) |
|
691 |
case (k, arr: Array) => array(k :: path, arr) |
|
692 |
case _ => |
|
693 |
} |
|
694 |
} |
|
695 |
||
696 |
table(Nil, toml) |
|
697 |
result.toString |
|
698 |
} |
|
699 |
} |
|
700 |
} |