| author | wenzelm | 
| Fri, 07 May 2021 21:03:20 +0200 | |
| changeset 73649 | 029de1598940 | 
| parent 73559 | 22b5ecb53dd9 | 
| child 74671 | df12779c3ce8 | 
| permissions | -rw-r--r-- | 
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
1  | 
(* Title: Pure/PIDE/protocol.ML  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
4  | 
Protocol message formats for interactive proof documents.  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
45709
 
87017fcbad83
clarified modules (again) -- NB: both Document and Protocol are specific to this particular prover;
 
wenzelm 
parents: 
45672 
diff
changeset
 | 
7  | 
structure Protocol: sig end =  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
struct  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
|
| 
38418
 
9a7af64d71bb
more explicit / functional ML version of document model;
 
wenzelm 
parents: 
38417 
diff
changeset
 | 
10  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
11  | 
Protocol_Command.define "Prover.echo"  | 
| 52786 | 12  | 
(fn args => List.app writeln args);  | 
13  | 
||
14  | 
val _ =  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
15  | 
Protocol_Command.define "Prover.stop"  | 
| 72217 | 16  | 
(fn rc :: msgs =>  | 
17  | 
(List.app Output.system_message msgs;  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
18  | 
raise Protocol_Command.STOP (Value.parse_int rc)));  | 
| 71879 | 19  | 
|
20  | 
val _ =  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
21  | 
Protocol_Command.define "Prover.options"  | 
| 52786 | 22  | 
(fn [options_yxml] =>  | 
| 65300 | 23  | 
(Options.set_default (Options.decode (YXML.parse_body options_yxml));  | 
24  | 
Isabelle_Process.init_options_interactive ()));  | 
|
| 52585 | 25  | 
|
26  | 
val _ =  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
27  | 
Protocol_Command.define "Prover.init_session"  | 
| 72637 | 28  | 
(fn [yxml] => Resources.init_session_yxml yxml);  | 
| 65470 | 29  | 
|
30  | 
val _ =  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
31  | 
Protocol_Command.define "Document.define_blob"  | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
32  | 
(fn [digest, content] => Document.change_state (Document.define_blob digest content));  | 
| 
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
33  | 
|
| 72946 | 34  | 
fun decode_command id name parents_xml blobs_xml toks_xml sources : Document.command =  | 
| 70663 | 35  | 
let  | 
36  | 
open XML.Decode;  | 
|
| 72946 | 37  | 
val parents = list string parents_xml;  | 
| 70663 | 38  | 
val (blobs_digests, blobs_index) =  | 
| 
70664
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
39  | 
blobs_xml |>  | 
| 70663 | 40  | 
let  | 
41  | 
val message = YXML.string_of_body o Protocol_Message.command_positions id;  | 
|
| 
72747
 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 
wenzelm 
parents: 
72637 
diff
changeset
 | 
42  | 
val blob =  | 
| 
 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 
wenzelm 
parents: 
72637 
diff
changeset
 | 
43  | 
triple string string (option string)  | 
| 
 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 
wenzelm 
parents: 
72637 
diff
changeset
 | 
44  | 
            #> (fn (a, b, c) => {file_node = a, src_path = Path.explode b, digest = c});
 | 
| 70663 | 45  | 
in  | 
46  | 
pair  | 
|
47  | 
(list (variant  | 
|
| 
72747
 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 
wenzelm 
parents: 
72637 
diff
changeset
 | 
48  | 
[fn ([], a) => Exn.Res (blob a),  | 
| 70663 | 49  | 
fn ([], a) => Exn.Exn (ERROR (message a))]))  | 
50  | 
int  | 
|
51  | 
end;  | 
|
| 
70664
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
52  | 
val toks = list (pair int int) toks_xml;  | 
| 70663 | 53  | 
in  | 
54  | 
   {command_id = Document_ID.parse id,
 | 
|
55  | 
name = name,  | 
|
| 72946 | 56  | 
parents = parents,  | 
| 70663 | 57  | 
blobs_digests = blobs_digests,  | 
58  | 
blobs_index = blobs_index,  | 
|
59  | 
tokens = toks ~~ sources}  | 
|
60  | 
end;  | 
|
61  | 
||
| 71024 | 62  | 
fun commands_accepted ids =  | 
| 73559 | 63  | 
Output.protocol_message Markup.commands_accepted [[XML.Text (space_implode "," ids)]];  | 
| 
70665
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
64  | 
|
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
65  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
66  | 
Protocol_Command.define "Document.define_command"  | 
| 72946 | 67  | 
(fn id :: name :: parents :: blobs :: toks :: sources =>  | 
| 
70664
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
68  | 
let  | 
| 72946 | 69  | 
val command =  | 
70  | 
decode_command id name (YXML.parse_body parents) (YXML.parse_body blobs)  | 
|
71  | 
(YXML.parse_body toks) sources;  | 
|
| 
70665
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
72  | 
val _ = Document.change_state (Document.define_command command);  | 
| 
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
73  | 
in commands_accepted [id] end);  | 
| 
70664
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
74  | 
|
| 
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
75  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
76  | 
Protocol_Command.define "Document.define_commands"  | 
| 
70664
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
77  | 
(fn args =>  | 
| 
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
78  | 
let  | 
| 
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
79  | 
fun decode arg =  | 
| 
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
80  | 
let  | 
| 
 
2bd9e30183b1
prefer define_commands_bulk: fewer protocol messages;
 
wenzelm 
parents: 
70663 
diff
changeset
 | 
81  | 
open XML.Decode;  | 
| 72946 | 82  | 
val (id, (name, (parents_xml, (blobs_xml, (toks_xml, sources))))) =  | 
83  | 
pair string (pair string (pair I (pair I (pair I (list string)))))  | 
|
84  | 
(YXML.parse_body arg);  | 
|
85  | 
in decode_command id name parents_xml blobs_xml toks_xml sources end;  | 
|
| 
70665
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
86  | 
|
| 
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
87  | 
val commands = map decode args;  | 
| 
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
88  | 
val _ = Document.change_state (fold Document.define_command commands);  | 
| 
 
94442fce40a5
prefer commands_accepted: fewer protocol messages;
 
wenzelm 
parents: 
70664 
diff
changeset
 | 
89  | 
in commands_accepted (map (Value.print_int o #command_id) commands) end);  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
90  | 
|
| 
38418
 
9a7af64d71bb
more explicit / functional ML version of document model;
 
wenzelm 
parents: 
38417 
diff
changeset
 | 
91  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
92  | 
Protocol_Command.define "Document.discontinue_execution"  | 
| 52606 | 93  | 
(fn [] => Execution.discontinue ());  | 
| 
47343
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
94  | 
|
| 
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
95  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
96  | 
Protocol_Command.define "Document.cancel_exec"  | 
| 
52931
 
ac6648c0c0fb
cancel_query via direct access to the exec_id of the running query process;
 
wenzelm 
parents: 
52862 
diff
changeset
 | 
97  | 
(fn [exec_id] => Execution.cancel (Document_ID.parse exec_id));  | 
| 
 
ac6648c0c0fb
cancel_query via direct access to the exec_id of the running query process;
 
wenzelm 
parents: 
52862 
diff
changeset
 | 
98  | 
|
| 
 
ac6648c0c0fb
cancel_query via direct access to the exec_id of the running query process;
 
wenzelm 
parents: 
52862 
diff
changeset
 | 
99  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
100  | 
Protocol_Command.define "Document.update"  | 
| 
59370
 
b13ff987c559
refrain from default task_context for all protocol commands, e.g. relevant for "build_theories" to admit Session.shutdown;
 
wenzelm 
parents: 
59366 
diff
changeset
 | 
101  | 
(Future.task_context "Document.update" (Future.new_group NONE)  | 
| 69849 | 102  | 
(fn old_id_string :: new_id_string :: consolidate_yxml :: edits_yxml =>  | 
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
103  | 
Document.change_state (fn state =>  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
104  | 
let  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
105  | 
val old_id = Document_ID.parse old_id_string;  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
106  | 
val new_id = Document_ID.parse new_id_string;  | 
| 69849 | 107  | 
val consolidate =  | 
108  | 
YXML.parse_body consolidate_yxml |>  | 
|
109  | 
let open XML.Decode in list string end;  | 
|
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
110  | 
val edits =  | 
| 69849 | 111  | 
edits_yxml |> map (YXML.parse_body #>  | 
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
112  | 
let open XML.Decode in  | 
| 69849 | 113  | 
pair string  | 
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
114  | 
(variant  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
115  | 
[fn ([], a) => Document.Edits (list (pair (option int) (option int)) a),  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
116  | 
fn ([], a) =>  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
117  | 
let  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
118  | 
val (master, (name, (imports, (keywords, errors)))) =  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
119  | 
pair string (pair string (pair (list string)  | 
| 
72747
 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 
wenzelm 
parents: 
72637 
diff
changeset
 | 
120  | 
(pair (list (pair string (pair string (list string))))  | 
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
121  | 
(list YXML.string_of_body)))) a;  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
122  | 
val imports' = map (rpair Position.none) imports;  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
123  | 
val keywords' = map (fn (x, y) => ((x, Position.none), y)) keywords;  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
124  | 
val header = Thy_Header.make (name, Position.none) imports' keywords';  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
125  | 
                        in Document.Deps {master = master, header = header, errors = errors} end,
 | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
126  | 
fn (a :: b, c) =>  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
127  | 
Document.Perspective (bool_atom a, map int_atom b,  | 
| 69849 | 128  | 
list (pair int (pair string (list string))) c)])  | 
129  | 
end);  | 
|
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
130  | 
|
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
131  | 
val _ = Execution.discontinue ();  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
|
| 
70284
 
3e17c3a5fd39
more thorough assignment, e.g. when "purge" removes commands that were not assigned;
 
wenzelm 
parents: 
69849 
diff
changeset
 | 
133  | 
val (edited, removed, assign_update, state') =  | 
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
134  | 
Document.update old_id new_id edits consolidate state;  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
135  | 
val _ =  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
136  | 
(singleton o Future.forks)  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
137  | 
               {name = "Document.update/remove", group = NONE,
 | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
138  | 
deps = Execution.snapshot removed,  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
139  | 
pri = Task_Queue.urgent_pri + 2, interrupts = false}  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
140  | 
(fn () => (Execution.purge removed; List.app Isabelle_Process.reset_tracing removed));  | 
| 52601 | 141  | 
|
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
142  | 
val _ =  | 
| 
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
143  | 
Output.protocol_message Markup.assign_update  | 
| 73559 | 144  | 
[(new_id, edited, assign_update) |>  | 
| 69846 | 145  | 
let  | 
146  | 
open XML.Encode;  | 
|
147  | 
fun encode_upd (a, bs) =  | 
|
148  | 
string (space_implode "," (map Value.print_int (a :: bs)));  | 
|
| 73559 | 149  | 
in triple int (list string) (list encode_upd) end];  | 
| 
68336
 
09ac56914b29
Document.update includes node consolidation / presentation as regular print operation: avoid user operations on protocol thread;
 
wenzelm 
parents: 
67493 
diff
changeset
 | 
150  | 
in Document.start_execution state' end)));  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
|
| 43748 | 152  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
153  | 
Protocol_Command.define "Document.remove_versions"  | 
| 44673 | 154  | 
(fn [versions_yxml] => Document.change_state (fn state =>  | 
155  | 
let  | 
|
156  | 
val versions =  | 
|
157  | 
YXML.parse_body versions_yxml |>  | 
|
158  | 
let open XML.Decode in list int end;  | 
|
| 44676 | 159  | 
val state1 = Document.remove_versions versions state;  | 
| 73559 | 160  | 
val _ = Output.protocol_message Markup.removed_versions [[XML.Text (versions_yxml)]];  | 
| 44676 | 161  | 
in state1 end));  | 
| 44673 | 162  | 
|
163  | 
val _ =  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
164  | 
Protocol_Command.define "Document.dialog_result"  | 
| 
50500
 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 
wenzelm 
parents: 
50498 
diff
changeset
 | 
165  | 
(fn [serial, result] =>  | 
| 63806 | 166  | 
Active.dialog_result (Value.parse_int serial) result  | 
| 62505 | 167  | 
handle exn => if Exn.is_interrupt exn then () (*sic!*) else Exn.reraise exn);  | 
| 50498 | 168  | 
|
| 
56616
 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 
wenzelm 
parents: 
56458 
diff
changeset
 | 
169  | 
val _ =  | 
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
170  | 
Protocol_Command.define "ML_Heap.full_gc"  | 
| 72146 | 171  | 
(fn [] => ML_Heap.full_gc ());  | 
172  | 
||
173  | 
val _ =  | 
|
| 
73225
 
3ab0cedaccad
clarified modules: allow early definition of protocol commands;
 
wenzelm 
parents: 
72946 
diff
changeset
 | 
174  | 
Protocol_Command.define "ML_Heap.share_common_data"  | 
| 62467 | 175  | 
(fn [] => ML_Heap.share_common_data ());  | 
| 
57868
 
0b954ac94827
protocol command for heap management, e.g. in Isabelle/jEdit/Scala console: PIDE.session.protocol_command("ML_System.share_common_data");
 
wenzelm 
parents: 
56616 
diff
changeset
 | 
176  | 
|
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
end;  |