| author | wenzelm | 
| Wed, 02 Apr 2025 13:39:12 +0200 | |
| changeset 82413 | a6046b6d23b4 | 
| parent 80441 | c420429fdf4c | 
| 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  | 
||
| 80058 | 91  | 
  class Table private(private val rep: ListMap[Key, T]) extends T {
 | 
| 78316 | 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 +
 | 
|
| 
79003
 
9d1c4824a055
clarified toml keys formatting vs. toString;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79002 
diff
changeset
 | 
115  | 
", got " + v.getClass.getSimpleName + " for key " + quote(k))  | 
| 78316 | 116  | 
}  | 
| 
79003
 
9d1c4824a055
clarified toml keys formatting vs. toString;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79002 
diff
changeset
 | 
117  | 
          case None => error("Key " + quote(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  | 
|
258  | 
||
| 78316 | 259  | 
/* parse structure */  | 
| 78317 | 260  | 
|
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
261  | 
sealed trait V extends Positional  | 
| 78316 | 262  | 
case class Primitive(t: T) extends V  | 
263  | 
case class Array(rep: List[V]) extends V  | 
|
264  | 
case class Inline_Table(elems: List[(Keys, V)]) extends V  | 
|
265  | 
||
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
266  | 
sealed trait Def extends Positional  | 
| 78316 | 267  | 
case class Table(key: Keys, elems: List[(Keys, V)]) extends Def  | 
268  | 
case class Array_Of_Tables(key: Keys, elems: List[(Keys, V)]) extends Def  | 
|
| 78317 | 269  | 
|
| 78316 | 270  | 
case class File(elems: List[(Keys, V)], defs: List[Def])  | 
| 78317 | 271  | 
|
| 78316 | 272  | 
|
273  | 
/* top-level syntax structure */  | 
|
| 78317 | 274  | 
|
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
275  | 
    def key: Parser[Keys] = positioned(rep1sep(keys, $$$(".")) ^^ (_.reduce(_ ++ _)))
 | 
| 78317 | 276  | 
|
| 78316 | 277  | 
def string: Parser[String] =  | 
278  | 
      elem("string", e => e.is_string || e.is_multiline_string) ^^ (s => String(s.text))
 | 
|
279  | 
def integer: Parser[Integer] =  | 
|
280  | 
(decimal_int | binary_int | octal_int | hexadecimal_int) ^^ Integer.apply  | 
|
281  | 
def float: Parser[Float] = (symbol_float | number_float) ^^ Float.apply  | 
|
282  | 
    def boolean: Parser[Boolean] = token("boolean", _.is_value, s => Boolean(Value.Boolean.parse(s)))
 | 
|
283  | 
||
284  | 
def offset_date_time: Parser[Offset_Date_Time] =  | 
|
285  | 
      token("offset date-time", _.is_value,
 | 
|
286  | 
        s => Offset_Date_Time(OffsetDateTime.parse(s.replace(" ", "T"))))
 | 
|
287  | 
def local_date_time: Parser[Local_Date_Time] =  | 
|
288  | 
      token("local date-time", _.is_value,
 | 
|
289  | 
        s => Local_Date_Time(LocalDateTime.parse(s.replace(" ", "T"))))
 | 
|
290  | 
def local_date: Parser[Local_Date] =  | 
|
291  | 
      token("local date", _.is_value, s => Local_Date(LocalDate.parse(s)))
 | 
|
292  | 
def local_time: Parser[Local_Time] =  | 
|
293  | 
      token("local time", _.is_value, s => Local_Time(LocalTime.parse(s)))
 | 
|
294  | 
||
295  | 
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
 | 
296  | 
positioned(  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
297  | 
        $$$("[") ~>! 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
 | 
298  | 
          opt(line_sep) ~! opt($$$(",")) ~! opt(line_sep) <~! $$$("]") ^^ Array.apply)
 | 
| 78316 | 299  | 
|
300  | 
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
 | 
301  | 
positioned(  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
302  | 
        $$$("{") ~>! repsep(pair, $$$(",")) <~! $$$("}") ^^ Inline_Table.apply)
 | 
| 78317 | 303  | 
|
| 78316 | 304  | 
    def pair: Parser[(Keys, V)] = (key <~! $$$("=")) ~! toml_value ^^ { case ks ~ v => (ks, v) }
 | 
305  | 
||
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
306  | 
def table: Parser[Table] =  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
307  | 
positioned(  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
308  | 
        $$$("[") ~> (key <~! $$$("]") ~! line_sep) ~! content ^^
 | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
309  | 
          { case key ~ content => Table(key, content) })
 | 
| 78316 | 310  | 
|
311  | 
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
 | 
312  | 
positioned(  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
313  | 
        $$$("[") ~ $$$("[") ~>! (key <~! $$$("]") ~! $$$("]") ~! line_sep) ~! content ^^
 | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
314  | 
          { case key ~ content => Array_Of_Tables(key, content) })
 | 
| 78316 | 315  | 
|
316  | 
def toml: Parser[File] =  | 
|
317  | 
(opt(line_sep) ~>! content ~! rep(table | array_of_tables)) ^^  | 
|
318  | 
        { case content ~ defs => File(content, defs) }
 | 
|
319  | 
||
320  | 
||
321  | 
/* auxiliary */  | 
|
322  | 
||
323  | 
private def $$$(name: Str): Parser[Token] = elem(name, _.is_keyword(name))  | 
|
324  | 
private def maybe[A, B](p: Parser[A], f: A => B): Parser[B] =  | 
|
325  | 
      p ^^ (a => Try(f(a))) ^? { case util.Success(v) => v }
 | 
|
326  | 
private def token[A](name: Str, p: Token => Bool, parser: Str => A): Parser[A] =  | 
|
327  | 
maybe(elem(name, p), s => parser(s.text))  | 
|
328  | 
private def prefixed[A](  | 
|
329  | 
prefix: Str, name: Str, p: Str => Bool, parser: Str => A  | 
|
330  | 
): Parser[A] =  | 
|
331  | 
token(name, e => e.is_value && e.text.startsWith(prefix) && p(e.text.stripPrefix(prefix)),  | 
|
332  | 
s => parser(s.stripPrefix(prefix)))  | 
|
333  | 
||
334  | 
    private def is_key(e: Elem): Bool = e.is_value && !e.text.exists("+: ".contains(_))
 | 
|
| 79002 | 335  | 
private def keys: Parser[Keys] =  | 
336  | 
      token("string key", _.is_string, Keys.quoted) | token("key", is_key, Keys.dotted)
 | 
|
| 78316 | 337  | 
|
338  | 
private def sep_surrounded(s: Str): Bool =  | 
|
339  | 
      !s.startsWith("_") && !s.endsWith("_") && s.split('_').forall(_.nonEmpty)
 | 
|
340  | 
    private def no_leading_zero(s: Str): Bool = {
 | 
|
341  | 
      val t = s.replaceAll("_", "").takeWhile(_.isDigit)
 | 
|
342  | 
      t == "0" || !t.startsWith("0")
 | 
|
343  | 
}  | 
|
344  | 
||
345  | 
private def is_int(s: Str): Bool =  | 
|
346  | 
      no_leading_zero(s.replaceAll("[-+]", "")) && sep_surrounded(s.replaceAll("[-+]", ""))
 | 
|
347  | 
private def decimal_int: Parser[Long] =  | 
|
348  | 
      token("integer", e => e.is_value && is_int(e.text), _.replace("_", "").toLong)
 | 
|
349  | 
private def binary_int: Parser[Long] =  | 
|
350  | 
      prefixed("0b", "integer", sep_surrounded, s => parseLong(s.replace("_", ""), 2))
 | 
|
351  | 
private def octal_int: Parser[Long] =  | 
|
352  | 
      prefixed("0o", "integer", sep_surrounded, s => parseLong(s.replace("_", ""), 8))
 | 
|
353  | 
private def hexadecimal_int: Parser[Long] =  | 
|
354  | 
      prefixed("0x", "integer", sep_surrounded, s => parseLong(s.replace("_", ""), 16))
 | 
|
355  | 
||
356  | 
private def is_float(s: Str): Bool =  | 
|
357  | 
      s.exists(".eE".contains) && s.count("eE".contains) <= 1 &&
 | 
|
358  | 
        no_leading_zero(s.replaceAll("[-+]", "")) &&
 | 
|
359  | 
        sep_surrounded(s.replaceAll("[-+]", "").replaceAll("[.eE][+\\-]?", "_"))
 | 
|
360  | 
private def number_float: Parser[Double] =  | 
|
361  | 
      token("float", e => e.is_value && is_float(e.text), _.replace("_", "").toDouble)
 | 
|
362  | 
private def symbol_float: Parser[Double] =  | 
|
363  | 
      token("float", _.is_value, {
 | 
|
364  | 
case "inf" | "+inf" => Double.PositiveInfinity  | 
|
365  | 
case "-inf" => Double.NegativeInfinity  | 
|
366  | 
case "nan" | "+nan" | "-nan" => Double.NaN  | 
|
367  | 
})  | 
|
368  | 
||
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
369  | 
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
 | 
370  | 
positioned(  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
371  | 
(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
 | 
372  | 
local_date_time | local_date | local_time) ^^ Primitive.apply | array | inline_table)  | 
| 78316 | 373  | 
|
374  | 
    private def line_sep: Parser[Any] = rep1(elem("line sep", _.is_line_sep))
 | 
|
375  | 
||
376  | 
private def content: Parser[List[(Keys, V)]] =  | 
|
377  | 
      rep((key <~! $$$("=")) ~! toml_value <~! line_sep ^^ { case ks ~ v => ks -> v })
 | 
|
378  | 
}  | 
|
379  | 
||
380  | 
object Parsers extends Parsers  | 
|
381  | 
||
382  | 
||
| 79001 | 383  | 
/* Keys for nested tables */  | 
384  | 
||
| 79002 | 385  | 
  object Keys {
 | 
386  | 
def empty: Keys = new Keys(Nil)  | 
|
387  | 
def quoted(s: Str): Keys = new Keys(List(s))  | 
|
388  | 
    def dotted(s: Str): Keys = new Keys(s.split('.').toList)
 | 
|
389  | 
}  | 
|
390  | 
||
| 
79004
 
39373f2151c4
use toml key operations properly;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79003 
diff
changeset
 | 
391  | 
  class Keys private(private val rep: List[Key]) extends Positional {
 | 
| 79002 | 392  | 
override def hashCode(): Int = rep.hashCode()  | 
393  | 
override def equals(obj: Any): Bool =  | 
|
394  | 
      obj match {
 | 
|
395  | 
case other: Keys => rep == other.rep  | 
|
396  | 
case _ => false  | 
|
397  | 
}  | 
|
398  | 
    override def toString: Str = rep.mkString("Keys(", ".", ")")
 | 
|
399  | 
||
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
400  | 
def first: Keys = new Keys(rep.take(1))  | 
| 79002 | 401  | 
def last: Keys = new Keys(rep.takeRight(1))  | 
| 79006 | 402  | 
def next: Keys = new Keys(rep.drop(1))  | 
403  | 
def parent: Keys = new Keys(rep.dropRight(1))  | 
|
| 79001 | 404  | 
|
| 
79005
 
6201057b98dd
clarified toml keys: more operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79004 
diff
changeset
 | 
405  | 
def the_key: Key = Library.the_single(rep)  | 
| 79001 | 406  | 
|
407  | 
def length: Int = rep.length  | 
|
408  | 
||
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
409  | 
def ++(other: Keys): Keys = new Keys(rep ::: other.rep)  | 
| 79001 | 410  | 
|
| 79006 | 411  | 
def prefixes: Set[Keys] = splits.map(_._1).toSet  | 
412  | 
def splits: List[(Keys, Keys)] = Range.inclusive(0, length).toList.map(split).reverse  | 
|
| 79001 | 413  | 
    def split(i: Int): (Keys, Keys) = {
 | 
414  | 
val (rep0, rep1) = rep.splitAt(i)  | 
|
| 79002 | 415  | 
(new Keys(rep0), new Keys(rep1))  | 
| 79001 | 416  | 
}  | 
417  | 
||
418  | 
def is_child_of(parent: Keys): Bool = rep.startsWith(parent.rep)  | 
|
419  | 
}  | 
|
420  | 
||
421  | 
||
| 78316 | 422  | 
/* checking and translating parse structure into toml */  | 
423  | 
||
| 79001 | 424  | 
  class Parse_Context private(var seen_tables: Map[Keys, Seen], file: Option[Path] = None) {
 | 
425  | 
private def recent_array(ks: Keys): (Keys, Seen, Keys) =  | 
|
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
426  | 
      ks.splits.collectFirst {
 | 
| 78316 | 427  | 
case (ks1, ks2) if seen_tables.contains(ks1) => (ks1, seen_tables(ks1), ks2)  | 
428  | 
}.get  | 
|
429  | 
||
| 79001 | 430  | 
    private def check_add(start: Int, ks: Keys, elem: Parsers.Def | Parsers.V): Unit = {
 | 
| 78316 | 431  | 
val (to, seen, rest, split) =  | 
432  | 
        elem match {
 | 
|
433  | 
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
 | 
434  | 
val (_, seen, rest) = recent_array(ks.parent)  | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
435  | 
(ks, seen, rest ++ ks.last, 0)  | 
| 78316 | 436  | 
case _ =>  | 
437  | 
val (to, seen, rest) = recent_array(ks)  | 
|
438  | 
(to, seen, rest, start - to.length)  | 
|
439  | 
}  | 
|
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
440  | 
val (rest0, rest1) = rest.split(split)  | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
441  | 
val implicitly_seen = rest1.parent.prefixes.map(rest0 ++ _)  | 
| 78316 | 442  | 
|
| 78983 | 443  | 
def error[A](s: Str): A = this.error(s, elem.pos, Some(ks))  | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
444  | 
|
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
445  | 
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
 | 
446  | 
        error("Attempting to update an inline value"))
 | 
| 78316 | 447  | 
|
448  | 
val (inlines, tables) =  | 
|
449  | 
        elem match {
 | 
|
450  | 
case _: Parsers.Primitive =>  | 
|
451  | 
(seen.inlines, seen.tables ++ implicitly_seen)  | 
|
452  | 
case _: Parsers.Array =>  | 
|
| 78327 | 453  | 
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
 | 
454  | 
              error("Attempting to append with an inline array")
 | 
| 78316 | 455  | 
(seen.inlines + rest, seen.tables ++ implicitly_seen)  | 
456  | 
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
 | 
457  | 
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
 | 
458  | 
              error("Attempting to add with an inline table"))
 | 
| 78316 | 459  | 
(seen.inlines + rest, seen.tables ++ implicitly_seen)  | 
460  | 
case _: Parsers.Table =>  | 
|
| 78327 | 461  | 
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
 | 
462  | 
              error("Attempting to define a table twice")
 | 
| 78316 | 463  | 
(seen.inlines, seen.tables + rest)  | 
464  | 
case _: Parsers.Array_Of_Tables => (Set.empty, Set.empty)  | 
|
465  | 
}  | 
|
466  | 
||
467  | 
seen_tables += to -> Seen(inlines, tables)  | 
|
468  | 
}  | 
|
| 
78982
 
be5c078f5292
add file information to toml parse context and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78981 
diff
changeset
 | 
469  | 
|
| 
 
be5c078f5292
add file information to toml parse context and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78981 
diff
changeset
 | 
470  | 
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
 | 
471  | 
|
| 79001 | 472  | 
    def error[A](s: Str, pos: input.Position, key: Option[Keys] = None): A = {
 | 
| 
79003
 
9d1c4824a055
clarified toml keys formatting vs. toString;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79002 
diff
changeset
 | 
473  | 
val key_msg = if_proper(key, " in table " + key.get)  | 
| 
78982
 
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)  | 
| 78983 | 475  | 
isabelle.error(s + key_msg + " at line " + pos.line + file_msg)  | 
| 
78982
 
be5c078f5292
add file information to toml parse context and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78981 
diff
changeset
 | 
476  | 
}  | 
| 
 
be5c078f5292
add file information to toml parse context and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78981 
diff
changeset
 | 
477  | 
|
| 79001 | 478  | 
def check_add_def(ks: Keys, defn: Parsers.Def): Unit = check_add(0, ks, defn)  | 
479  | 
def check_add_value(ks0: Keys, ks1: Keys, v: Parsers.V): Unit =  | 
|
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
480  | 
check_add(ks0.length - 1, ks0 ++ ks1, v)  | 
| 78316 | 481  | 
}  | 
482  | 
||
483  | 
  object Parse_Context {
 | 
|
| 79001 | 484  | 
case class Seen(inlines: Set[Keys], tables: Set[Keys])  | 
| 78316 | 485  | 
|
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
486  | 
def apply(): Parse_Context =  | 
| 79002 | 487  | 
new Parse_Context(Map(Keys.empty -> Seen(Set.empty, Set.empty)))  | 
| 78316 | 488  | 
}  | 
489  | 
||
| 78927 | 490  | 
  def parse(s: Str, context: Parse_Context = Parse_Context()): Table = {
 | 
| 78983 | 491  | 
val file =  | 
| 
79000
 
8f5be65a176b
clarified toml parser interface;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78983 
diff
changeset
 | 
492  | 
      Parsers.phrase(Parsers.toml)(new Lexer.Scanner(Scan.char_reader(s + EofCh))) match {
 | 
| 78983 | 493  | 
case Parsers.Success (toml, _) => toml  | 
494  | 
        case Parsers.Error(msg, next) => context.error("Error parsing toml: " + msg, next.pos)
 | 
|
495  | 
case Parsers.Failure (msg, next) =>  | 
|
496  | 
          context.error("Malformed TOML input: " + msg, next.pos)
 | 
|
497  | 
}  | 
|
| 78316 | 498  | 
|
| 79001 | 499  | 
    def convert(ks0: Keys, ks1: Keys, v: Parsers.V): T = {
 | 
500  | 
def to_table(ks: Keys, t: T): Table =  | 
|
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
501  | 
if (ks.length == 1) Table(ks.first.the_key -> t)  | 
| 
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
502  | 
else Table(ks.first.the_key -> to_table(ks.next, t))  | 
| 
79004
 
39373f2151c4
use toml key operations properly;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79003 
diff
changeset
 | 
503  | 
|
| 79001 | 504  | 
      def to_toml(v: Parsers.V): (T, Set[Keys]) = v match {
 | 
| 78316 | 505  | 
case Parsers.Primitive(t) => (t, Set.empty)  | 
506  | 
case Parsers.Array(rep) => (Array(rep.map(to_toml(_)._1)), Set.empty)  | 
|
507  | 
case Parsers.Inline_Table(elems) =>  | 
|
| 78327 | 508  | 
elems.find(e => elems.count(_._1 == e._1) > 1).foreach((ks, _) =>  | 
| 78983 | 509  | 
context.error(  | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
510  | 
"Duplicate " + ks + " in inline table", v.pos, Some(ks0 ++ ks1)))  | 
| 78316 | 511  | 
|
512  | 
val (tables, pfxs, key_spaces) =  | 
|
513  | 
            elems.map { (ks, v) =>
 | 
|
514  | 
val (t, keys) = to_toml(v)  | 
|
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
515  | 
(to_table(ks, t), ks.prefixes, keys.map(ks ++ _) + ks)  | 
| 78316 | 516  | 
}.unzip3  | 
517  | 
||
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
518  | 
          for {
 | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
519  | 
pfx <- pfxs  | 
| 
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
520  | 
if key_spaces.count(pfx.intersect(_).nonEmpty) > 1  | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
521  | 
          } context.error("Inline table not self-contained", v.pos, Some(ks0 ++ ks1))
 | 
| 78316 | 522  | 
|
523  | 
(tables.foldLeft(Table())(_ ++ _), pfxs.toSet.flatten ++ key_spaces.toSet.flatten)  | 
|
524  | 
}  | 
|
525  | 
context.check_add_value(ks0, ks1, v)  | 
|
526  | 
to_toml(v)._1  | 
|
527  | 
}  | 
|
528  | 
||
| 79001 | 529  | 
    def update(map: Table, ks0: Keys, value: T): Table = {
 | 
| 78983 | 530  | 
def error[A](s: Str): A = context.error(s, ks0.pos, Some(ks0))  | 
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
531  | 
|
| 79001 | 532  | 
      def update_rec(hd: Keys, map: Table, ks: Keys): Table = {
 | 
| 78327 | 533  | 
val updated =  | 
534  | 
          if (ks.length == 1) {
 | 
|
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
535  | 
            map.any.get(ks.first.the_key) match {
 | 
| 78327 | 536  | 
case Some(a: Array) =>  | 
537  | 
                value match {
 | 
|
538  | 
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
 | 
539  | 
                  case _ => error("Table conflicts with previous array of tables")
 | 
| 78327 | 540  | 
}  | 
541  | 
              case Some(t: Table) => value match {
 | 
|
542  | 
case t2: Table =>  | 
|
543  | 
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
 | 
544  | 
                    error("Attempting to redefine existing value in super-table")
 | 
| 78327 | 545  | 
else t ++ t2  | 
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
546  | 
                case _ => error("Attempting to redefine a table")
 | 
| 78316 | 547  | 
}  | 
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
548  | 
              case Some(_) => error("Attempting to redefine a value")
 | 
| 78327 | 549  | 
case None => value  | 
| 78316 | 550  | 
}  | 
551  | 
}  | 
|
| 78327 | 552  | 
          else {
 | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
553  | 
val hd1 = hd ++ ks.first  | 
| 
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
554  | 
            map.any.get(ks.first.the_key) match {
 | 
| 79006 | 555  | 
case Some(t: Table) => update_rec(hd1, t, ks.next)  | 
| 78327 | 556  | 
case Some(a: Array) =>  | 
| 79006 | 557  | 
Array(a.table.values.dropRight(1) :+ update_rec(hd1, a.table.values.last, ks.next))  | 
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
558  | 
              case Some(_) => error("Attempting to redefine a value")
 | 
| 79006 | 559  | 
case None => update_rec(hd1, Table(), ks.next)  | 
| 78327 | 560  | 
}  | 
| 78316 | 561  | 
}  | 
| 
79007
 
eed4ca224c9c
clarified toml keys operations;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79006 
diff
changeset
 | 
562  | 
(map - ks.first.the_key) + (ks.first.the_key -> updated)  | 
| 78327 | 563  | 
}  | 
| 
78981
 
47f8533c6887
add position information to toml parser and error messages;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78980 
diff
changeset
 | 
564  | 
|
| 79002 | 565  | 
update_rec(Keys.empty, map, ks0)  | 
| 78316 | 566  | 
}  | 
567  | 
||
| 79001 | 568  | 
def fold(elems: List[(Keys, T)]): Table =  | 
| 78316 | 569  | 
      elems.foldLeft(Table()) { case (t0, (ks1, t1)) => update(t0, ks1, t1) }
 | 
570  | 
||
| 79002 | 571  | 
val t = fold(file.elems.map((ks, v) => (ks, convert(Keys.empty, ks, v))))  | 
| 78316 | 572  | 
    file.defs.foldLeft(t) {
 | 
573  | 
case (t0, defn@Parsers.Table(ks0, elems)) =>  | 
|
574  | 
context.check_add_def(ks0, defn)  | 
|
575  | 
update(t0, ks0, fold(elems.map((ks, v) => (ks, convert(ks0, ks, v)))))  | 
|
576  | 
case (t0, defn@Parsers.Array_Of_Tables(ks0, elems)) =>  | 
|
577  | 
context.check_add_def(ks0, defn)  | 
|
578  | 
update(t0, ks0, Array(fold(elems.map((ks, v) => (ks, convert(ks0, ks, v))))))  | 
|
579  | 
}  | 
|
580  | 
}  | 
|
581  | 
||
| 
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
 | 
582  | 
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
 | 
583  | 
    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
 | 
584  | 
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
 | 
585  | 
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
 | 
586  | 
(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
 | 
587  | 
}._1  | 
| 78939 | 588  | 
|
| 78316 | 589  | 
|
590  | 
/* Format TOML */  | 
|
591  | 
||
592  | 
  object Format {
 | 
|
| 
78326
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
593  | 
private def basic_string(s: Str): Str =  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
594  | 
      "\"" + 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 + "\""  | 
| 
 
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  | 
private def multiline_basic_string(s: Str): Str =  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
601  | 
      "\"\"\"\n" + s.iterator.map {
 | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
602  | 
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
 | 
603  | 
case '\r' => "\r" case '"' => "\\\"" case '\\' => "\\\\" case c =>  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
604  | 
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
 | 
605  | 
      }.mkString.replace("\"\"\"", "\"\"\\\"") + "\"\"\""
 | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
606  | 
|
| 
79003
 
9d1c4824a055
clarified toml keys formatting vs. toString;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79002 
diff
changeset
 | 
607  | 
    private def key(k: Key): Str = {
 | 
| 
78326
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
608  | 
val Bare_Key = """[A-Za-z0-9_-]+""".r  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
609  | 
      k match {
 | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
610  | 
case Bare_Key() => k  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
611  | 
case _ => basic_string(k)  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
612  | 
}  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
613  | 
}  | 
| 
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
614  | 
|
| 
79003
 
9d1c4824a055
clarified toml keys formatting vs. toString;
 
Fabian Huch <huch@in.tum.de> 
parents: 
79002 
diff
changeset
 | 
615  | 
    private 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
 | 
616  | 
|
| 80441 | 617  | 
private def inline(t: T, indent: Int = 0): Str =  | 
618  | 
      Library.string_builder() { result =>
 | 
|
619  | 
result ++= Symbol.spaces(indent * 2)  | 
|
620  | 
        (t: @unchecked) match {
 | 
|
621  | 
case s: String =>  | 
|
622  | 
            if (s.rep.contains("\n") && s.rep.length > 20) result ++= multiline_basic_string(s.rep)
 | 
|
623  | 
else result ++= basic_string(s.rep)  | 
|
624  | 
case a: Array =>  | 
|
625  | 
if (a.is_empty) result ++= "[]"  | 
|
626  | 
            else {
 | 
|
627  | 
result ++= "[\n"  | 
|
628  | 
              a.any.values.foreach { elem =>
 | 
|
629  | 
result ++= inline(elem, indent + 1)  | 
|
630  | 
result ++= ",\n"  | 
|
631  | 
}  | 
|
632  | 
result ++= Symbol.spaces(indent * 2)  | 
|
633  | 
result += ']'  | 
|
| 
78326
 
1cdc65476c2e
more TOML formatting functions;
 
Fabian Huch <huch@in.tum.de> 
parents: 
78317 
diff
changeset
 | 
634  | 
}  | 
| 80441 | 635  | 
case table: Table =>  | 
636  | 
            if (table.is_empty) result ++= "{}"
 | 
|
637  | 
            else {
 | 
|
638  | 
              result ++= "{ "
 | 
|
639  | 
              Library.separate(None, table.any.values.map(Some(_))).foreach {
 | 
|
640  | 
case None => result ++= ", "  | 
|
641  | 
case Some((k, v)) =>  | 
|
642  | 
result ++= key(k)  | 
|
643  | 
result ++= " = "  | 
|
644  | 
result ++= inline(v)  | 
|
645  | 
}  | 
|
646  | 
result ++= " }"  | 
|
647  | 
}  | 
|
648  | 
case Scalar(s) => result ++= s  | 
|
| 78316 | 649  | 
}  | 
650  | 
}  | 
|
651  | 
||
| 80441 | 652  | 
def apply(toml: Table): Str =  | 
653  | 
      Library.string_builder() { result =>
 | 
|
654  | 
def inline_values(path: List[Key], t: T): Unit =  | 
|
655  | 
          t match {
 | 
|
656  | 
            case t: Table => t.any.values.foreach { case (k, v1) => inline_values(k :: path, v1) }
 | 
|
657  | 
case _ =>  | 
|
658  | 
result ++= keys(path.reverse)  | 
|
659  | 
result ++= " = "  | 
|
660  | 
result ++= inline(t)  | 
|
661  | 
result += '\n'  | 
|
662  | 
}  | 
|
663  | 
||
664  | 
def is_inline(elem: T): Bool =  | 
|
665  | 
          elem match {
 | 
|
666  | 
case _: String | _: Integer | _: Float | _: Boolean | _: Offset_Date_Time |  | 
|
667  | 
_: Local_Date_Time | _: Local_Date | _: Local_Time => true  | 
|
668  | 
case a: Array => a.any.values.forall(is_inline)  | 
|
669  | 
case _ => false  | 
|
670  | 
}  | 
|
671  | 
        def is_table(elem: T): Bool = elem match { case _: Table => true case _ => false }
 | 
|
672  | 
||
673  | 
def array(path: List[Key], a: Array): Unit =  | 
|
674  | 
if (a.any.values.forall(is_inline) || !a.any.values.forall(is_table))  | 
|
675  | 
inline_values(path.take(1), a)  | 
|
676  | 
          else a.table.values.foreach { t =>
 | 
|
677  | 
result ++= "\n[["  | 
|
678  | 
result ++= keys(path.reverse)  | 
|
679  | 
result ++= "]]\n"  | 
|
680  | 
table(path, t, is_table_entry = true)  | 
|
681  | 
}  | 
|
682  | 
||
683  | 
        def table(path: List[Key], t: Table, is_table_entry: Bool = false): Unit = {
 | 
|
684  | 
val (values, nodes) = t.any.values.partition(kv => is_inline(kv._2))  | 
|
685  | 
||
686  | 
          if (!is_table_entry && path.nonEmpty) {
 | 
|
687  | 
result ++= "\n["  | 
|
688  | 
result ++= keys(path.reverse)  | 
|
689  | 
result ++= "]\n"  | 
|
690  | 
}  | 
|
691  | 
||
692  | 
          values.foreach { case (k, v) => inline_values(List(k), v) }
 | 
|
693  | 
          nodes.foreach {
 | 
|
694  | 
case (k, t: Table) => table(k :: path, t)  | 
|
695  | 
case (k, arr: Array) => array(k :: path, arr)  | 
|
696  | 
case _ =>  | 
|
697  | 
}  | 
|
698  | 
}  | 
|
699  | 
||
700  | 
table(Nil, toml)  | 
|
701  | 
}  | 
|
| 78316 | 702  | 
}  | 
703  | 
}  |