| author | wenzelm | 
| Tue, 09 Feb 2021 15:40:23 +0100 | |
| changeset 73245 | f69cbb59813e | 
| parent 73030 | 72a8fdfa185d | 
| child 73340 | 0ffcad1f6130 | 
| permissions | -rw-r--r-- | 
| 44698 | 1 | /* Title: Pure/PIDE/yxml.scala | 
| 27930 | 2 | Author: Makarius | 
| 3 | ||
| 44698 | 4 | Efficient text representation of XML trees. Suitable for direct | 
| 5 | inlining into plain text. | |
| 27930 | 6 | */ | 
| 7 | ||
| 8 | package isabelle | |
| 9 | ||
| 10 | ||
| 40453 | 11 | import scala.collection.mutable | 
| 36684 | 12 | |
| 13 | ||
| 32450 | 14 | object YXML | 
| 15 | {
 | |
| 27943 
f34ff5e7728f
replaced Pattern.split by chunks iterator (more efficient, resembles ML version more closely);
 wenzelm parents: 
27930diff
changeset | 16 | /* chunk markers */ | 
| 27930 | 17 | |
| 56661 
ef623f6f036b
avoid octal escape literals -- deprecated in scala-2.11.0;
 wenzelm parents: 
56600diff
changeset | 18 | val X = '\u0005' | 
| 
ef623f6f036b
avoid octal escape literals -- deprecated in scala-2.11.0;
 wenzelm parents: 
56600diff
changeset | 19 | val Y = '\u0006' | 
| 56600 | 20 | |
| 73029 | 21 | val is_X: Char => Boolean = _ == X | 
| 22 | val is_Y: Char => Boolean = _ == Y | |
| 56600 | 23 | |
| 71601 | 24 | val X_string: String = X.toString | 
| 25 | val Y_string: String = Y.toString | |
| 26 | val XY_string: String = X_string + Y_string | |
| 27 | val XYX_string: String = XY_string + X_string | |
| 27930 | 28 | |
| 43782 | 29 | def detect(s: String): Boolean = s.exists(c => c == X || c == Y) | 
| 67820 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
64370diff
changeset | 30 | def detect_elem(s: String): Boolean = s.startsWith(XY_string) | 
| 43782 | 31 | |
| 27930 | 32 | |
| 73029 | 33 | /* string representation */ | 
| 38268 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 34 | |
| 73029 | 35 | def traversal(string: String => Unit, body: XML.Body) | 
| 38268 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 36 |   {
 | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 37 | def tree(t: XML.Tree): Unit = | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 38 |       t match {
 | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 39 | case XML.Elem(Markup(name, atts), ts) => | 
| 71534 | 40 | string(XY_string) | 
| 41 | string(name) | |
| 42 |           for ((a, x) <- atts) { string(Y_string); string(a); string("="); string(x) }
 | |
| 43 | string(X_string) | |
| 38268 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 44 | ts.foreach(tree) | 
| 71534 | 45 | string(XYX_string) | 
| 46 | case XML.Text(text) => string(text) | |
| 38268 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 47 | } | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 48 | body.foreach(tree) | 
| 71534 | 49 | } | 
| 50 | ||
| 51 | def string_of_body(body: XML.Body): String = | |
| 52 |   {
 | |
| 53 | val s = new StringBuilder | |
| 54 | traversal(str => s ++= str, body) | |
| 38268 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 55 | s.toString | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 56 | } | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 57 | |
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 58 | def string_of_tree(tree: XML.Tree): String = string_of_body(List(tree)) | 
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 59 | |
| 
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
 wenzelm parents: 
38267diff
changeset | 60 | |
| 27930 | 61 | /* parsing */ | 
| 62 | ||
| 27993 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 wenzelm parents: 
27971diff
changeset | 63 |   private def err(msg: String) = error("Malformed YXML: " + msg)
 | 
| 27930 | 64 |   private def err_attribute() = err("bad attribute")
 | 
| 65 |   private def err_element() = err("bad element")
 | |
| 66 | private def err_unbalanced(name: String) = | |
| 67 |     if (name == "") err("unbalanced element")
 | |
| 44181 | 68 |     else err("unbalanced element " + quote(name))
 | 
| 27930 | 69 | |
| 73029 | 70 | private def parse_attrib(source: CharSequence): (String, String) = | 
| 71 |   {
 | |
| 27944 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 72 | val s = source.toString | 
| 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 73 |     val i = s.indexOf('=')
 | 
| 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 74 | if (i <= 0) err_attribute() | 
| 38234 
4b610fbb2d83
YXML.parse: refrain from interning, let XML.Cache do it (partially);
 wenzelm parents: 
38231diff
changeset | 75 | (s.substring(0, i), s.substring(i + 1)) | 
| 27944 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 76 | } | 
| 27930 | 77 | |
| 78 | ||
| 73030 | 79 | def parse_body(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body = | 
| 32450 | 80 |   {
 | 
| 27930 | 81 | /* stack operations */ | 
| 82 | ||
| 40453 | 83 | def buffer(): mutable.ListBuffer[XML.Tree] = new mutable.ListBuffer[XML.Tree] | 
| 84 | var stack: List[(Markup, mutable.ListBuffer[XML.Tree])] = List((Markup.Empty, buffer())) | |
| 27930 | 85 | |
| 36684 | 86 | def add(x: XML.Tree) | 
| 87 |     {
 | |
| 73029 | 88 |       (stack: @unchecked) match { case (_, body) :: _ => body += x }
 | 
| 27930 | 89 | } | 
| 90 | ||
| 36684 | 91 | def push(name: String, atts: XML.Attributes) | 
| 92 |     {
 | |
| 27930 | 93 | if (name == "") err_element() | 
| 73030 | 94 | else stack = (cache.markup(Markup(name, atts)), buffer()) :: stack | 
| 36684 | 95 | } | 
| 27930 | 96 | |
| 36684 | 97 | def pop() | 
| 98 |     {
 | |
| 99 |       (stack: @unchecked) match {
 | |
| 73029 | 100 |         case (Markup.Empty, _) :: _ => err_unbalanced("")
 | 
| 101 | case (markup, body) :: pending => | |
| 38230 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 wenzelm parents: 
36685diff
changeset | 102 | stack = pending | 
| 73030 | 103 | add(cache.tree0(XML.Elem(markup, body.toList))) | 
| 36684 | 104 | } | 
| 27930 | 105 | } | 
| 106 | ||
| 107 | ||
| 108 | /* parse chunks */ | |
| 109 | ||
| 56600 | 110 |     for (chunk <- Library.separated_chunks(is_X, source) if chunk.length != 0) {
 | 
| 34099 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 111 | if (chunk.length == 1 && chunk.charAt(0) == Y) pop() | 
| 27945 
d2dc5a1903e8
tuned parse performance: avoid splitting terminal Y chunk;
 wenzelm parents: 
27944diff
changeset | 112 |       else {
 | 
| 56600 | 113 |         Library.separated_chunks(is_Y, chunk).toList match {
 | 
| 34099 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 114 | case ch :: name :: atts if ch.length == 0 => | 
| 38234 
4b610fbb2d83
YXML.parse: refrain from interning, let XML.Cache do it (partially);
 wenzelm parents: 
38231diff
changeset | 115 | push(name.toString, atts.map(parse_attrib)) | 
| 73030 | 116 | case txts => for (txt <- txts) add(cache.tree0(XML.Text(cache.string(txt.toString)))) | 
| 27945 
d2dc5a1903e8
tuned parse performance: avoid splitting terminal Y chunk;
 wenzelm parents: 
27944diff
changeset | 117 | } | 
| 27930 | 118 | } | 
| 119 | } | |
| 42719 | 120 |     (stack: @unchecked) match {
 | 
| 38475 | 121 | case List((Markup.Empty, body)) => body.toList | 
| 38230 
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
 wenzelm parents: 
36685diff
changeset | 122 | case (Markup(name, _), _) :: _ => err_unbalanced(name) | 
| 27930 | 123 | } | 
| 124 | } | |
| 125 | ||
| 73030 | 126 | def parse(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Tree = | 
| 127 |     parse_body(source, cache = cache) match {
 | |
| 27930 | 128 | case List(result) => result | 
| 73028 | 129 | case Nil => XML.no_text | 
| 67820 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
64370diff
changeset | 130 |       case _ => err("multiple XML trees")
 | 
| 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
64370diff
changeset | 131 | } | 
| 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
64370diff
changeset | 132 | |
| 73030 | 133 | def parse_elem(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Tree = | 
| 134 |     parse_body(source, cache = cache) match {
 | |
| 67820 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
64370diff
changeset | 135 | case List(elem: XML.Elem) => elem | 
| 
e30d6368c7c8
clarified argument formats: explicit Unit, allow XML.Elem as well;
 wenzelm parents: 
64370diff
changeset | 136 |       case _ => err("single XML element expected")
 | 
| 27930 | 137 | } | 
| 27960 | 138 | |
| 29521 | 139 | |
| 140 | /* failsafe parsing */ | |
| 141 | ||
| 45666 | 142 | private def markup_broken(source: CharSequence) = | 
| 55551 | 143 | XML.Elem(Markup.Broken, List(XML.Text(source.toString))) | 
| 29521 | 144 | |
| 73030 | 145 | def parse_body_failsafe(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Body = | 
| 32450 | 146 |   {
 | 
| 73030 | 147 |     try { parse_body(source, cache = cache) }
 | 
| 45666 | 148 |     catch { case ERROR(_) => List(markup_broken(source)) }
 | 
| 29521 | 149 | } | 
| 150 | ||
| 73030 | 151 | def parse_failsafe(source: CharSequence, cache: XML.Cache = XML.Cache.none): XML.Tree = | 
| 32450 | 152 |   {
 | 
| 73030 | 153 |     try { parse(source, cache = cache) }
 | 
| 45666 | 154 |     catch { case ERROR(_) => markup_broken(source) }
 | 
| 27960 | 155 | } | 
| 27930 | 156 | } |