author | wenzelm |
Thu, 06 Jun 2024 11:53:52 +0200 | |
changeset 80266 | d52be75ae60b |
parent 75393 | 87ebf5a50283 |
child 80455 | 99e276c44121 |
permissions | -rw-r--r-- |
59713 | 1 |
/* Title: Pure/PIDE/protocol_message.scala |
2 |
Author: Makarius |
|
3 |
||
4 |
Auxiliary operations on protocol messages. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
9 |
||
75393 | 10 |
object Protocol_Message { |
71620 | 11 |
/* message markers */ |
12 |
||
75393 | 13 |
object Marker { |
71620 | 14 |
def apply(a: String): Marker = |
15 |
new Marker { override def name: String = a } |
|
71623
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71620
diff
changeset
|
16 |
|
b3bddebe44ca
clarified signature: more explicit type Protocol_Message.Marker;
wenzelm
parents:
71620
diff
changeset
|
17 |
def test(line: String): Boolean = line.startsWith("\f") |
71620 | 18 |
} |
19 |
||
75393 | 20 |
abstract class Marker private { |
71620 | 21 |
def name: String |
22 |
val prefix: String = "\f" + name + " = " |
|
23 |
||
24 |
def apply(text: String): String = prefix + Library.encode_lines(text) |
|
25 |
def apply(props: Properties.T): String = apply(YXML.string_of_body(XML.Encode.properties(props))) |
|
26 |
||
27 |
def test(line: String): Boolean = line.startsWith(prefix) |
|
28 |
def test_yxml(line: String): Boolean = test(line) && YXML.detect(line) |
|
29 |
||
30 |
def unapply(line: String): Option[String] = |
|
31 |
Library.try_unprefix(prefix, line).map(Library.decode_lines) |
|
32 |
||
33 |
override def toString: String = "Marker(" + quote(name) + ")" |
|
34 |
} |
|
35 |
||
36 |
||
59713 | 37 |
/* inlined reports */ |
38 |
||
39 |
private val report_elements = |
|
40 |
Markup.Elements(Markup.REPORT, Markup.NO_REPORT) |
|
41 |
||
42 |
def clean_reports(body: XML.Body): XML.Body = |
|
43 |
body filter { |
|
44 |
case XML.Wrapped_Elem(Markup(name, _), _, _) => !report_elements(name) |
|
45 |
case XML.Elem(Markup(name, _), _) => !report_elements(name) |
|
46 |
case _ => true |
|
47 |
} map { |
|
48 |
case XML.Wrapped_Elem(markup, body, ts) => XML.Wrapped_Elem(markup, body, clean_reports(ts)) |
|
49 |
case XML.Elem(markup, ts) => XML.Elem(markup, clean_reports(ts)) |
|
50 |
case t => t |
|
51 |
} |
|
52 |
||
71631 | 53 |
def expose_no_reports(body: XML.Body): XML.Body = |
54 |
body flatMap { |
|
55 |
case XML.Wrapped_Elem(markup, body, ts) => List(XML.Wrapped_Elem(markup, body, expose_no_reports(ts))) |
|
56 |
case XML.Elem(Markup(Markup.NO_REPORT, _), ts) => ts |
|
57 |
case XML.Elem(markup, ts) => List(XML.Elem(markup, expose_no_reports(ts))) |
|
58 |
case t => List(t) |
|
59 |
} |
|
60 |
||
59713 | 61 |
def reports(props: Properties.T, body: XML.Body): List[XML.Elem] = |
62 |
body flatMap { |
|
63 |
case XML.Wrapped_Elem(Markup(Markup.REPORT, ps), body, ts) => |
|
64 |
List(XML.Wrapped_Elem(Markup(Markup.REPORT, props ::: ps), body, ts)) |
|
65 |
case XML.Elem(Markup(Markup.REPORT, ps), ts) => |
|
66 |
List(XML.Elem(Markup(Markup.REPORT, props ::: ps), ts)) |
|
67 |
case XML.Wrapped_Elem(_, _, ts) => reports(props, ts) |
|
68 |
case XML.Elem(_, ts) => reports(props, ts) |
|
69 |
case XML.Text(_) => Nil |
|
70 |
} |
|
71 |
} |