| author | blanchet | 
| Thu, 29 Jul 2010 22:13:15 +0200 | |
| changeset 38099 | e3bb96b83807 | 
| parent 36685 | 2b3076cfd6dd | 
| child 38230 | ed147003de4b | 
| permissions | -rw-r--r-- | 
| 27930 | 1 | /* Title: Pure/General/yxml.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Efficient text representation of XML trees. | |
| 5 | */ | |
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 36684 | 10 | import scala.collection.mutable.ListBuffer | 
| 11 | ||
| 12 | ||
| 32450 | 13 | object YXML | 
| 14 | {
 | |
| 27943 
f34ff5e7728f
replaced Pattern.split by chunks iterator (more efficient, resembles ML version more closely);
 wenzelm parents: 
27930diff
changeset | 15 | /* chunk markers */ | 
| 27930 | 16 | |
| 17 | private val X = '\5' | |
| 18 | private val Y = '\6' | |
| 27960 | 19 | private val X_string = X.toString | 
| 27945 
d2dc5a1903e8
tuned parse performance: avoid splitting terminal Y chunk;
 wenzelm parents: 
27944diff
changeset | 20 | private val Y_string = Y.toString | 
| 27930 | 21 | |
| 22 | ||
| 34099 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 23 | /* decoding pseudo UTF-8 */ | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 24 | |
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 25 | private class Decode_Chars(decode: String => String, | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 26 | buffer: Array[Byte], start: Int, end: Int) extends CharSequence | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 27 |   {
 | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 28 | def length: Int = end - start | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 29 | def charAt(i: Int): Char = (buffer(start + i).asInstanceOf[Int] & 0xFF).asInstanceOf[Char] | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 30 | def subSequence(i: Int, j: Int): CharSequence = | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 31 | new Decode_Chars(decode, buffer, start + i, start + j) | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 32 | |
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 33 | // toString with adhoc decoding: abuse of CharSequence interface | 
| 34201 
c95dcd12f48a
separate Standard_System (Cygwin/Posix compatibility) vs. Isabelle_System (settings environment etc.);
 wenzelm parents: 
34198diff
changeset | 34 | override def toString: String = decode(Standard_System.decode_permissive_utf8(this)) | 
| 34099 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 35 | } | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 36 | |
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 37 | def decode_chars(decode: String => String, | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 38 | buffer: Array[Byte], start: Int, end: Int): CharSequence = | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 39 |   {
 | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 40 | require(0 <= start && start <= end && end <= buffer.length) | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 41 | new Decode_Chars(decode, buffer, start, end) | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 42 | } | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 43 | |
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 44 | |
| 27930 | 45 | /* parsing */ | 
| 46 | ||
| 27993 
6dd90ef9f927
simplified exceptions: use plain error function / RuntimeException;
 wenzelm parents: 
27971diff
changeset | 47 |   private def err(msg: String) = error("Malformed YXML: " + msg)
 | 
| 27930 | 48 |   private def err_attribute() = err("bad attribute")
 | 
| 49 |   private def err_element() = err("bad element")
 | |
| 50 | private def err_unbalanced(name: String) = | |
| 51 |     if (name == "") err("unbalanced element")
 | |
| 52 |     else err("unbalanced element \"" + name + "\"")
 | |
| 53 | ||
| 27944 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 54 |   private def parse_attrib(source: CharSequence) = {
 | 
| 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 55 | val s = source.toString | 
| 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 56 |     val i = s.indexOf('=')
 | 
| 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 57 | if (i <= 0) err_attribute() | 
| 29563 | 58 | (s.substring(0, i).intern, s.substring(i + 1)) | 
| 27944 
2bf3f30558ed
parse_attrib: more efficient due to indexOf('=');
 wenzelm parents: 
27943diff
changeset | 59 | } | 
| 27930 | 60 | |
| 61 | ||
| 32450 | 62 | def parse_body(source: CharSequence): List[XML.Tree] = | 
| 63 |   {
 | |
| 27930 | 64 | /* stack operations */ | 
| 65 | ||
| 36684 | 66 | def buffer(): ListBuffer[XML.Tree] = new ListBuffer[XML.Tree] | 
| 67 |     var stack: List[((String, XML.Attributes), ListBuffer[XML.Tree])] = List((("", Nil), buffer()))
 | |
| 27930 | 68 | |
| 36684 | 69 | def add(x: XML.Tree) | 
| 70 |     {
 | |
| 71 |       (stack: @unchecked) match { case ((_, body) :: _) => body += x }
 | |
| 27930 | 72 | } | 
| 73 | ||
| 36684 | 74 | def push(name: String, atts: XML.Attributes) | 
| 75 |     {
 | |
| 27930 | 76 | if (name == "") err_element() | 
| 36684 | 77 | else stack = ((name, atts), buffer()) :: stack | 
| 78 | } | |
| 27930 | 79 | |
| 36684 | 80 | def pop() | 
| 81 |     {
 | |
| 82 |       (stack: @unchecked) match {
 | |
| 83 |         case ((("", _), _) :: _) => err_unbalanced("")
 | |
| 84 | case (((name, atts), body) :: pending) => | |
| 85 | stack = pending; add(XML.Elem(name, atts, body.toList)) | |
| 86 | } | |
| 27930 | 87 | } | 
| 88 | ||
| 89 | ||
| 90 | /* parse chunks */ | |
| 91 | ||
| 36685 | 92 |     for (chunk <- Library.chunks(source, X) 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 | 93 | if (chunk.length == 1 && chunk.charAt(0) == Y) pop() | 
| 27945 
d2dc5a1903e8
tuned parse performance: avoid splitting terminal Y chunk;
 wenzelm parents: 
27944diff
changeset | 94 |       else {
 | 
| 36685 | 95 |         Library.chunks(chunk, Y).toList match {
 | 
| 34099 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 96 | case ch :: name :: atts if ch.length == 0 => | 
| 
2541de190d92
added decode_chars, with raw character view on byte buffer and adhoc decoding via toString;
 wenzelm parents: 
32450diff
changeset | 97 | push(name.toString.intern, atts.map(parse_attrib)) | 
| 27945 
d2dc5a1903e8
tuned parse performance: avoid splitting terminal Y chunk;
 wenzelm parents: 
27944diff
changeset | 98 | case txts => for (txt <- txts) add(XML.Text(txt.toString)) | 
| 
d2dc5a1903e8
tuned parse performance: avoid splitting terminal Y chunk;
 wenzelm parents: 
27944diff
changeset | 99 | } | 
| 27930 | 100 | } | 
| 101 | } | |
| 102 |     stack match {
 | |
| 36684 | 103 |       case List((("", _), body)) => body.toList
 | 
| 27930 | 104 | case ((name, _), _) :: _ => err_unbalanced(name) | 
| 105 | } | |
| 106 | } | |
| 107 | ||
| 32450 | 108 | def parse(source: CharSequence): XML.Tree = | 
| 27930 | 109 |     parse_body(source) match {
 | 
| 110 | case List(result) => result | |
| 111 |       case Nil => XML.Text("")
 | |
| 112 |       case _ => err("multiple results")
 | |
| 113 | } | |
| 27960 | 114 | |
| 29521 | 115 | |
| 116 | /* failsafe parsing */ | |
| 117 | ||
| 118 | private def markup_failsafe(source: CharSequence) = | |
| 34119 | 119 | XML.Elem (Markup.MALFORMED, Nil, | 
| 29521 | 120 | List(XML.Text(source.toString.replace(X_string, "\\<^X>").replace(Y_string, "\\<^Y>")))) | 
| 121 | ||
| 32450 | 122 | def parse_body_failsafe(source: CharSequence): List[XML.Tree] = | 
| 123 |   {
 | |
| 29521 | 124 |     try { parse_body(source) }
 | 
| 125 |     catch { case _: RuntimeException => List(markup_failsafe(source)) }
 | |
| 126 | } | |
| 127 | ||
| 32450 | 128 | def parse_failsafe(source: CharSequence): XML.Tree = | 
| 129 |   {
 | |
| 27960 | 130 |     try { parse(source) }
 | 
| 29521 | 131 |     catch { case _: RuntimeException => markup_failsafe(source) }
 | 
| 27960 | 132 | } | 
| 27930 | 133 | } |