| author | haftmann |
| Wed, 11 Aug 2010 17:19:27 +0200 | |
| changeset 38381 | 7d1e2a6831ec |
| parent 38271 | 36187e8443dd |
| child 38355 | 8cb265fb12fe |
| permissions | -rw-r--r-- |
| 29553 | 1 |
/* Title: Pure/Isar/isar_document.scala |
2 |
Author: Makarius |
|
3 |
||
4 |
Interactive Isar documents. |
|
5 |
*/ |
|
6 |
||
7 |
package isabelle |
|
8 |
||
| 32450 | 9 |
|
| 32467 | 10 |
object Isar_Document |
| 32450 | 11 |
{
|
|
38236
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
wenzelm
parents:
38231
diff
changeset
|
12 |
/* protocol messages */ |
| 36682 | 13 |
|
14 |
object Assign {
|
|
15 |
def unapply(msg: XML.Tree): Option[List[XML.Tree]] = |
|
16 |
msg match {
|
|
| 38231 | 17 |
case XML.Elem(Markup.Assign, edits) => Some(edits) |
| 36682 | 18 |
case _ => None |
19 |
} |
|
20 |
} |
|
21 |
||
22 |
object Edit {
|
|
|
38150
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
wenzelm
parents:
36682
diff
changeset
|
23 |
def unapply(msg: XML.Tree): Option[(Document.Command_ID, Document.State_ID)] = |
| 36682 | 24 |
msg match {
|
|
38230
ed147003de4b
simplified type XML.Tree: embed Markup directly, avoid slightly odd triple;
wenzelm
parents:
38150
diff
changeset
|
25 |
case XML.Elem(Markup(Markup.EDIT, List((Markup.ID, cmd_id), (Markup.STATE, state_id))), Nil) => |
| 36682 | 26 |
Some(cmd_id, state_id) |
27 |
case _ => None |
|
28 |
} |
|
29 |
} |
|
| 29645 | 30 |
} |
31 |
||
| 32450 | 32 |
|
| 32467 | 33 |
trait Isar_Document extends Isabelle_Process |
| 29645 | 34 |
{
|
| 32467 | 35 |
import Isar_Document._ |
| 29553 | 36 |
|
37 |
||
38 |
/* commands */ |
|
39 |
||
|
38271
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
40 |
def define_command(id: Document.Command_ID, text: String): Unit = |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
41 |
input("Isar_Document.define_command", id, text)
|
| 29553 | 42 |
|
43 |
||
44 |
/* documents */ |
|
45 |
||
|
38150
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
wenzelm
parents:
36682
diff
changeset
|
46 |
def edit_document(old_id: Document.Version_ID, new_id: Document.Version_ID, |
|
67fc24df3721
simplified/refined document model: collection of named nodes, without proper dependencies yet;
wenzelm
parents:
36682
diff
changeset
|
47 |
edits: List[Document.Edit[Document.Command_ID]]) |
| 29553 | 48 |
{
|
|
38271
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
49 |
def make_id1(id1: Option[Document.Command_ID]): XML.Body = |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
50 |
XML_Data.make_string(id1 getOrElse Document.NO_ID) |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
51 |
val arg = |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
52 |
XML_Data.make_list( |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
53 |
XML_Data.make_pair(XML_Data.make_string)( |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
54 |
XML_Data.make_option(XML_Data.make_list( |
|
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
55 |
XML_Data.make_pair(make_id1)(XML_Data.make_option(XML_Data.make_string))))))(edits) |
| 29553 | 56 |
|
|
38271
36187e8443dd
Isar_Document command input via native Isabelle_Process commands, using YXML and XML_Data representation;
wenzelm
parents:
38259
diff
changeset
|
57 |
input("Isar_Document.edit_document", old_id, new_id, YXML.string_of_body(arg))
|
| 29553 | 58 |
} |
59 |
} |