author | wenzelm |
Sat, 02 Jan 2021 22:22:34 +0100 | |
changeset 73031 | f93f0597f4fb |
parent 73030 | 72a8fdfa185d |
child 73032 | 72b13af7f266 |
permissions | -rw-r--r-- |
44698 | 1 |
/* Title: Pure/PIDE/xml.scala |
27931 | 2 |
Author: Makarius |
3 |
||
44698 | 4 |
Untyped XML trees and basic data representation. |
27931 | 5 |
*/ |
6 |
||
7 |
package isabelle |
|
8 |
||
55618 | 9 |
|
29203 | 10 |
object XML |
11 |
{ |
|
43767 | 12 |
/** XML trees **/ |
13 |
||
27947 | 14 |
/* datatype representation */ |
15 |
||
65753 | 16 |
type Attribute = Properties.Entry |
43780 | 17 |
type Attributes = Properties.T |
27931 | 18 |
|
57912 | 19 |
sealed abstract class Tree { override def toString: String = string_of_tree(this) } |
64354 | 20 |
type Body = List[Tree] |
21 |
case class Elem(markup: Markup, body: Body) extends Tree |
|
52890 | 22 |
{ |
23 |
def name: String = markup.name |
|
65753 | 24 |
|
64358 | 25 |
def update_attributes(more_attributes: Attributes): Elem = |
26 |
if (more_attributes.isEmpty) this |
|
27 |
else Elem(markup.update_properties(more_attributes), body) |
|
65753 | 28 |
|
65772 | 29 |
def + (att: Attribute): Elem = Elem(markup + att, body) |
52890 | 30 |
} |
29204 | 31 |
case class Text(content: String) extends Tree |
29203 | 32 |
|
66196 | 33 |
def elem(markup: Markup): XML.Elem = XML.Elem(markup, Nil) |
64354 | 34 |
def elem(name: String, body: Body): XML.Elem = XML.Elem(Markup(name, Nil), body) |
35 |
def elem(name: String): XML.Elem = XML.Elem(Markup(name, Nil), Nil) |
|
38267
e50c283dd125
type XML.Body as basic data representation language (Scala version);
wenzelm
parents:
38263
diff
changeset
|
36 |
|
73028 | 37 |
val no_text: Text = Text("") |
69867 | 38 |
val newline: Text = Text("\n") |
39 |
||
29203 | 40 |
|
69805 | 41 |
/* name space */ |
42 |
||
43 |
object Namespace |
|
44 |
{ |
|
45 |
def apply(prefix: String, target: String): Namespace = |
|
46 |
new Namespace(prefix, target) |
|
47 |
} |
|
48 |
||
49 |
final class Namespace private(prefix: String, target: String) |
|
50 |
{ |
|
51 |
def apply(name: String): String = prefix + ":" + name |
|
52 |
val attribute: XML.Attribute = ("xmlns:" + prefix, target) |
|
53 |
||
54 |
override def toString: String = attribute.toString |
|
55 |
} |
|
56 |
||
57 |
||
49650
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
58 |
/* wrapped elements */ |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
59 |
|
60215 | 60 |
val XML_ELEM = "xml_elem" |
61 |
val XML_NAME = "xml_name" |
|
62 |
val XML_BODY = "xml_body" |
|
49650
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
63 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
64 |
object Wrapped_Elem |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
65 |
{ |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
66 |
def apply(markup: Markup, body1: Body, body2: Body): XML.Elem = |
61026 | 67 |
XML.Elem(Markup(XML_ELEM, (XML_NAME, markup.name) :: markup.properties), |
68 |
XML.Elem(Markup(XML_BODY, Nil), body1) :: body2) |
|
49650
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
69 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
70 |
def unapply(tree: Tree): Option[(Markup, Body, Body)] = |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
71 |
tree match { |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
72 |
case |
61026 | 73 |
XML.Elem(Markup(XML_ELEM, (XML_NAME, name) :: props), |
74 |
XML.Elem(Markup(XML_BODY, Nil), body1) :: body2) => |
|
49650
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
75 |
Some(Markup(name, props), body1, body2) |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
76 |
case _ => None |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
77 |
} |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
78 |
} |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
79 |
|
67818
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
80 |
object Root_Elem |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
81 |
{ |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
82 |
def apply(body: Body): XML.Elem = XML.Elem(Markup(XML_ELEM, Nil), body) |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
83 |
def unapply(tree: Tree): Option[Body] = |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
84 |
tree match { |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
85 |
case XML.Elem(Markup(XML_ELEM, Nil), body) => Some(body) |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
86 |
case _ => None |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
87 |
} |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
88 |
} |
2457bea123e4
convenience to represent XML.Body as single XML.Elem;
wenzelm
parents:
67113
diff
changeset
|
89 |
|
49650
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
90 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
91 |
/* traverse text */ |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
92 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
93 |
def traverse_text[A](body: Body)(a: A)(op: (A, String) => A): A = |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
94 |
{ |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
95 |
def traverse(x: A, t: Tree): A = |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
96 |
t match { |
61026 | 97 |
case XML.Wrapped_Elem(_, _, ts) => (x /: ts)(traverse) |
98 |
case XML.Elem(_, ts) => (x /: ts)(traverse) |
|
99 |
case XML.Text(s) => op(x, s) |
|
49650
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
100 |
} |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
101 |
(a /: body)(traverse) |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
102 |
} |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
103 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
104 |
def text_length(body: Body): Int = traverse_text(body)(0) { case (n, s) => n + s.length } |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
105 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
106 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
107 |
/* text content */ |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
108 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
109 |
def content(body: Body): String = |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
110 |
{ |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
111 |
val text = new StringBuilder(text_length(body)) |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
112 |
traverse_text(body)(()) { case (_, s) => text.append(s) } |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
113 |
text.toString |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
114 |
} |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
115 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
116 |
def content(tree: Tree): String = content(List(tree)) |
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
117 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
118 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
119 |
|
9fad6480300d
support for wrapped XML elements, which allows to preserve full markup tree information in to_XML/from_XML conversion;
wenzelm
parents:
49613
diff
changeset
|
120 |
/** string representation **/ |
29203 | 121 |
|
69804 | 122 |
val header: String = "<?xml version=\"1.0\" encoding=\"utf-8\"?>\n" |
123 |
||
65990 | 124 |
def output_char(c: Char, s: StringBuilder) |
125 |
{ |
|
126 |
c match { |
|
127 |
case '<' => s ++= "<" |
|
128 |
case '>' => s ++= ">" |
|
129 |
case '&' => s ++= "&" |
|
130 |
case '"' => s ++= """ |
|
131 |
case '\'' => s ++= "'" |
|
132 |
case _ => s += c |
|
133 |
} |
|
134 |
} |
|
135 |
||
136 |
def output_string(str: String, s: StringBuilder) |
|
137 |
{ |
|
138 |
if (str == null) s ++= str |
|
139 |
else str.iterator.foreach(c => output_char(c, s)) |
|
140 |
} |
|
141 |
||
38268
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
142 |
def string_of_body(body: Body): String = |
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
143 |
{ |
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
144 |
val s = new StringBuilder |
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
145 |
|
65990 | 146 |
def text(txt: String) { output_string(txt, s) } |
65991 | 147 |
def elem(markup: Markup) |
148 |
{ |
|
149 |
s ++= markup.name |
|
150 |
for ((a, b) <- markup.properties) { |
|
151 |
s += ' '; s ++= a; s += '='; s += '"'; text(b); s += '"' |
|
152 |
} |
|
153 |
} |
|
38268
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
154 |
def tree(t: Tree): Unit = |
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
155 |
t match { |
61026 | 156 |
case XML.Elem(markup, Nil) => |
65991 | 157 |
s += '<'; elem(markup); s ++= "/>" |
61026 | 158 |
case XML.Elem(markup, ts) => |
65991 | 159 |
s += '<'; elem(markup); s += '>' |
38268
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
160 |
ts.foreach(tree) |
65991 | 161 |
s ++= "</"; s ++= markup.name; s += '>' |
61026 | 162 |
case XML.Text(txt) => text(txt) |
38268
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
163 |
} |
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
164 |
body.foreach(tree) |
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
165 |
s.toString |
29203 | 166 |
} |
167 |
||
38268
beb86b805590
more uniform XML/YXML string_of_body/string_of_tree;
wenzelm
parents:
38267
diff
changeset
|
168 |
def string_of_tree(tree: XML.Tree): String = string_of_body(List(tree)) |
27941 | 169 |
|
170 |
||
44808 | 171 |
|
68265 | 172 |
/** cache **/ |
34108 | 173 |
|
73024 | 174 |
object Cache |
175 |
{ |
|
176 |
def make( |
|
73031
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
wenzelm
parents:
73030
diff
changeset
|
177 |
xz: XZ.Cache = XZ.Cache.make(), |
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
wenzelm
parents:
73030
diff
changeset
|
178 |
max_string: Int = isabelle.Cache.default_max_string, |
73024 | 179 |
initial_size: Int = isabelle.Cache.default_initial_size): Cache = |
73031
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
wenzelm
parents:
73030
diff
changeset
|
180 |
new Cache(xz, max_string, initial_size) |
68169 | 181 |
|
73031
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
wenzelm
parents:
73030
diff
changeset
|
182 |
val none: Cache = make(XZ.Cache.none, max_string = 0) |
73024 | 183 |
} |
184 |
||
73031
f93f0597f4fb
clarified signature: absorb XZ.Cache into XML.Cache;
wenzelm
parents:
73030
diff
changeset
|
185 |
class Cache private[XML](val xz: XZ.Cache, max_string: Int, initial_size: Int) |
73024 | 186 |
extends isabelle.Cache(max_string, initial_size) |
34108 | 187 |
{ |
68265 | 188 |
protected def cache_props(x: Properties.T): Properties.T = |
44704
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
189 |
{ |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
190 |
if (x.isEmpty) x |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
191 |
else |
34133 | 192 |
lookup(x) match { |
193 |
case Some(y) => y |
|
65903 | 194 |
case None => store(x.map(p => (Library.isolate_substring(p._1).intern, cache_string(p._2)))) |
34133 | 195 |
} |
68265 | 196 |
} |
197 |
||
198 |
protected def cache_markup(x: Markup): Markup = |
|
199 |
{ |
|
44704
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
200 |
lookup(x) match { |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
201 |
case Some(y) => y |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
202 |
case None => |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
203 |
x match { |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
204 |
case Markup(name, props) => |
51663 | 205 |
store(Markup(cache_string(name), cache_props(props))) |
44704
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
206 |
} |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
207 |
} |
68265 | 208 |
} |
209 |
||
210 |
protected def cache_tree(x: XML.Tree): XML.Tree = |
|
211 |
{ |
|
44704
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
212 |
lookup(x) match { |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
213 |
case Some(y) => y |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
214 |
case None => |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
215 |
x match { |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
216 |
case XML.Elem(markup, body) => |
51663 | 217 |
store(XML.Elem(cache_markup(markup), cache_body(body))) |
218 |
case XML.Text(text) => store(XML.Text(cache_string(text))) |
|
44704
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
219 |
} |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
220 |
} |
68265 | 221 |
} |
222 |
||
223 |
protected def cache_body(x: XML.Body): XML.Body = |
|
224 |
{ |
|
44704
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
225 |
if (x.isEmpty) x |
528d635ef6f0
synchronous XML.Cache without actor -- potentially more efficient on machines with few cores;
wenzelm
parents:
44698
diff
changeset
|
226 |
else |
34133 | 227 |
lookup(x) match { |
228 |
case Some(y) => y |
|
71601 | 229 |
case None => x.map(cache_tree) |
34133 | 230 |
} |
68265 | 231 |
} |
38446
9d59dab38fef
XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents:
38268
diff
changeset
|
232 |
|
73030 | 233 |
// support hash-consing |
234 |
def tree0(x: XML.Tree): XML.Tree = |
|
235 |
if (no_cache) x else synchronized { lookup(x) getOrElse store(x) } |
|
236 |
||
38446
9d59dab38fef
XML.Cache: pipe-lined (thread-safe) version using actor;
wenzelm
parents:
38268
diff
changeset
|
237 |
// main methods |
73024 | 238 |
def props(x: Properties.T): Properties.T = |
239 |
if (no_cache) x else synchronized { cache_props(x) } |
|
240 |
def markup(x: Markup): Markup = |
|
241 |
if (no_cache) x else synchronized { cache_markup(x) } |
|
242 |
def tree(x: XML.Tree): XML.Tree = |
|
243 |
if (no_cache) x else synchronized { cache_tree(x) } |
|
244 |
def body(x: XML.Body): XML.Body = |
|
245 |
if (no_cache) x else synchronized { cache_body(x) } |
|
246 |
def elem(x: XML.Elem): XML.Elem = |
|
247 |
if (no_cache) x else synchronized { cache_tree(x).asInstanceOf[XML.Elem] } |
|
34108 | 248 |
} |
249 |
||
250 |
||
43767 | 251 |
|
252 |
/** XML as data representation language **/ |
|
253 |
||
51987 | 254 |
abstract class Error(s: String) extends Exception(s) |
255 |
class XML_Atom(s: String) extends Error(s) |
|
256 |
class XML_Body(body: XML.Body) extends Error("") |
|
43767 | 257 |
|
258 |
object Encode |
|
259 |
{ |
|
260 |
type T[A] = A => XML.Body |
|
65334 | 261 |
type V[A] = PartialFunction[A, (List[String], XML.Body)] |
70828 | 262 |
type P[A] = PartialFunction[A, List[String]] |
43767 | 263 |
|
264 |
||
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
265 |
/* atomic values */ |
43767 | 266 |
|
57909
0fb331032f02
more compact representation of special string values;
wenzelm
parents:
55618
diff
changeset
|
267 |
def long_atom(i: Long): String = Library.signed_string_of_long(i) |
43767 | 268 |
|
57909
0fb331032f02
more compact representation of special string values;
wenzelm
parents:
55618
diff
changeset
|
269 |
def int_atom(i: Int): String = Library.signed_string_of_int(i) |
43767 | 270 |
|
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
271 |
def bool_atom(b: Boolean): String = if (b) "1" else "0" |
43767 | 272 |
|
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
273 |
def unit_atom(u: Unit) = "" |
43767 | 274 |
|
275 |
||
276 |
/* structural nodes */ |
|
277 |
||
278 |
private def node(ts: XML.Body): XML.Tree = XML.Elem(Markup(":", Nil), ts) |
|
279 |
||
43781 | 280 |
private def vector(xs: List[String]): XML.Attributes = |
46839
f7232c078fa5
simplified -- plain map_index is sufficient (pointed out by Enrico Tassi);
wenzelm
parents:
45673
diff
changeset
|
281 |
xs.zipWithIndex.map({ case (x, i) => (int_atom(i), x) }) |
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
282 |
|
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
283 |
private def tagged(tag: Int, data: (List[String], XML.Body)): XML.Tree = |
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
284 |
XML.Elem(Markup(int_atom(tag), vector(data._1)), data._2) |
43767 | 285 |
|
286 |
||
287 |
/* representation of standard types */ |
|
288 |
||
65333 | 289 |
val tree: T[XML.Tree] = (t => List(t)) |
290 |
||
43780 | 291 |
val properties: T[Properties.T] = |
43767 | 292 |
(props => List(XML.Elem(Markup(":", props), Nil))) |
293 |
||
294 |
val string: T[String] = (s => if (s.isEmpty) Nil else List(XML.Text(s))) |
|
295 |
||
296 |
val long: T[Long] = (x => string(long_atom(x))) |
|
297 |
||
298 |
val int: T[Int] = (x => string(int_atom(x))) |
|
299 |
||
300 |
val bool: T[Boolean] = (x => string(bool_atom(x))) |
|
301 |
||
302 |
val unit: T[Unit] = (x => string(unit_atom(x))) |
|
303 |
||
304 |
def pair[A, B](f: T[A], g: T[B]): T[(A, B)] = |
|
305 |
(x => List(node(f(x._1)), node(g(x._2)))) |
|
306 |
||
307 |
def triple[A, B, C](f: T[A], g: T[B], h: T[C]): T[(A, B, C)] = |
|
308 |
(x => List(node(f(x._1)), node(g(x._2)), node(h(x._3)))) |
|
309 |
||
310 |
def list[A](f: T[A]): T[List[A]] = |
|
311 |
(xs => xs.map((x: A) => node(f(x)))) |
|
312 |
||
313 |
def option[A](f: T[A]): T[Option[A]] = |
|
314 |
{ |
|
315 |
case None => Nil |
|
316 |
case Some(x) => List(node(f(x))) |
|
317 |
} |
|
318 |
||
65334 | 319 |
def variant[A](fs: List[V[A]]): T[A] = |
43767 | 320 |
{ |
321 |
case x => |
|
322 |
val (f, tag) = fs.iterator.zipWithIndex.find(p => p._1.isDefinedAt(x)).get |
|
323 |
List(tagged(tag, f(x))) |
|
324 |
} |
|
325 |
} |
|
326 |
||
327 |
object Decode |
|
328 |
{ |
|
329 |
type T[A] = XML.Body => A |
|
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
330 |
type V[A] = (List[String], XML.Body) => A |
70828 | 331 |
type P[A] = PartialFunction[List[String], A] |
43767 | 332 |
|
333 |
||
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
334 |
/* atomic values */ |
43767 | 335 |
|
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
336 |
def long_atom(s: String): Long = |
43767 | 337 |
try { java.lang.Long.parseLong(s) } |
338 |
catch { case e: NumberFormatException => throw new XML_Atom(s) } |
|
339 |
||
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
340 |
def int_atom(s: String): Int = |
43767 | 341 |
try { Integer.parseInt(s) } |
342 |
catch { case e: NumberFormatException => throw new XML_Atom(s) } |
|
343 |
||
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
344 |
def bool_atom(s: String): Boolean = |
43767 | 345 |
if (s == "1") true |
346 |
else if (s == "0") false |
|
347 |
else throw new XML_Atom(s) |
|
348 |
||
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
349 |
def unit_atom(s: String): Unit = |
43767 | 350 |
if (s == "") () else throw new XML_Atom(s) |
351 |
||
352 |
||
353 |
/* structural nodes */ |
|
354 |
||
355 |
private def node(t: XML.Tree): XML.Body = |
|
356 |
t match { |
|
357 |
case XML.Elem(Markup(":", Nil), ts) => ts |
|
358 |
case _ => throw new XML_Body(List(t)) |
|
359 |
} |
|
360 |
||
43781 | 361 |
private def vector(atts: XML.Attributes): List[String] = |
46839
f7232c078fa5
simplified -- plain map_index is sufficient (pointed out by Enrico Tassi);
wenzelm
parents:
45673
diff
changeset
|
362 |
atts.iterator.zipWithIndex.map( |
f7232c078fa5
simplified -- plain map_index is sufficient (pointed out by Enrico Tassi);
wenzelm
parents:
45673
diff
changeset
|
363 |
{ case ((a, x), i) => if (int_atom(a) == i) x else throw new XML_Atom(a) }).toList |
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
364 |
|
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
365 |
private def tagged(t: XML.Tree): (Int, (List[String], XML.Body)) = |
43767 | 366 |
t match { |
43781 | 367 |
case XML.Elem(Markup(name, atts), ts) => (int_atom(name), (vector(atts), ts)) |
43767 | 368 |
case _ => throw new XML_Body(List(t)) |
369 |
} |
|
370 |
||
371 |
||
372 |
/* representation of standard types */ |
|
373 |
||
65333 | 374 |
val tree: T[XML.Tree] = |
375 |
{ |
|
376 |
case List(t) => t |
|
377 |
case ts => throw new XML_Body(ts) |
|
378 |
} |
|
379 |
||
43780 | 380 |
val properties: T[Properties.T] = |
43767 | 381 |
{ |
382 |
case List(XML.Elem(Markup(":", props), Nil)) => props |
|
383 |
case ts => throw new XML_Body(ts) |
|
384 |
} |
|
385 |
||
386 |
val string: T[String] = |
|
387 |
{ |
|
388 |
case Nil => "" |
|
389 |
case List(XML.Text(s)) => s |
|
390 |
case ts => throw new XML_Body(ts) |
|
391 |
} |
|
392 |
||
393 |
val long: T[Long] = (x => long_atom(string(x))) |
|
394 |
||
395 |
val int: T[Int] = (x => int_atom(string(x))) |
|
396 |
||
397 |
val bool: T[Boolean] = (x => bool_atom(string(x))) |
|
398 |
||
399 |
val unit: T[Unit] = (x => unit_atom(string(x))) |
|
400 |
||
401 |
def pair[A, B](f: T[A], g: T[B]): T[(A, B)] = |
|
402 |
{ |
|
403 |
case List(t1, t2) => (f(node(t1)), g(node(t2))) |
|
404 |
case ts => throw new XML_Body(ts) |
|
405 |
} |
|
406 |
||
407 |
def triple[A, B, C](f: T[A], g: T[B], h: T[C]): T[(A, B, C)] = |
|
408 |
{ |
|
409 |
case List(t1, t2, t3) => (f(node(t1)), g(node(t2)), h(node(t3))) |
|
410 |
case ts => throw new XML_Body(ts) |
|
411 |
} |
|
412 |
||
413 |
def list[A](f: T[A]): T[List[A]] = |
|
414 |
(ts => ts.map(t => f(node(t)))) |
|
415 |
||
416 |
def option[A](f: T[A]): T[Option[A]] = |
|
417 |
{ |
|
418 |
case Nil => None |
|
419 |
case List(t) => Some(f(node(t))) |
|
420 |
case ts => throw new XML_Body(ts) |
|
421 |
} |
|
422 |
||
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
423 |
def variant[A](fs: List[V[A]]): T[A] = |
43767 | 424 |
{ |
425 |
case List(t) => |
|
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
426 |
val (tag, (xs, ts)) = tagged(t) |
43768 | 427 |
val f = |
428 |
try { fs(tag) } |
|
429 |
catch { case _: IndexOutOfBoundsException => throw new XML_Body(List(t)) } |
|
43778
ce9189450447
more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents:
43768
diff
changeset
|
430 |
f(xs, ts) |
43767 | 431 |
case ts => throw new XML_Body(ts) |
432 |
} |
|
433 |
} |
|
27931 | 434 |
} |