| author | wenzelm | 
| Thu, 16 Apr 2015 15:22:44 +0200 | |
| changeset 60097 | d20ca79d50e4 | 
| parent 60074 | 38a64cc17403 | 
| child 60610 | f52b4b0c10c4 | 
| 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 _ =  | 
| 56387 | 11  | 
Isabelle_Process.protocol_command "Prover.echo"  | 
| 52786 | 12  | 
(fn args => List.app writeln args);  | 
13  | 
||
14  | 
val _ =  | 
|
| 56387 | 15  | 
Isabelle_Process.protocol_command "Prover.options"  | 
| 52786 | 16  | 
(fn [options_yxml] =>  | 
17  | 
let val options = Options.decode (YXML.parse_body options_yxml) in  | 
|
18  | 
Options.set_default options;  | 
|
| 
60074
 
38a64cc17403
GUI controls for ML_statistics, for more digestible protocol dump;
 
wenzelm 
parents: 
59934 
diff
changeset
 | 
19  | 
Future.ML_statistics := Options.bool options "ML_statistics";  | 
| 52786 | 20  | 
Multithreading.trace := Options.int options "threads_trace";  | 
| 54717 | 21  | 
Multithreading.max_threads_update (Options.int options "threads");  | 
| 52786 | 22  | 
Goal.parallel_proofs := (if Options.int options "parallel_proofs" > 0 then 3 else 0)  | 
23  | 
end);  | 
|
| 52585 | 24  | 
|
25  | 
val _ =  | 
|
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
26  | 
Isabelle_Process.protocol_command "Document.define_blob"  | 
| 
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
27  | 
(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
 | 
28  | 
|
| 
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
29  | 
val _ =  | 
| 
46119
 
0d7172a7672c
tuned signature -- emphasize special nature of protocol commands;
 
wenzelm 
parents: 
45709 
diff
changeset
 | 
30  | 
Isabelle_Process.protocol_command "Document.define_command"  | 
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
31  | 
(fn id :: name :: blobs_yxml :: toks_yxml :: sources =>  | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
32  | 
let  | 
| 
59685
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
33  | 
val (blobs, blobs_index) =  | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
34  | 
YXML.parse_body blobs_yxml |>  | 
| 
59714
 
ae322325adbb
tuned protocol -- resolve command positions in ML;
 
wenzelm 
parents: 
59685 
diff
changeset
 | 
35  | 
let  | 
| 
 
ae322325adbb
tuned protocol -- resolve command positions in ML;
 
wenzelm 
parents: 
59685 
diff
changeset
 | 
36  | 
val message =  | 
| 
 
ae322325adbb
tuned protocol -- resolve command positions in ML;
 
wenzelm 
parents: 
59685 
diff
changeset
 | 
37  | 
YXML.string_of_body o Protocol_Message.command_positions id;  | 
| 
 
ae322325adbb
tuned protocol -- resolve command positions in ML;
 
wenzelm 
parents: 
59685 
diff
changeset
 | 
38  | 
open XML.Decode;  | 
| 
 
ae322325adbb
tuned protocol -- resolve command positions in ML;
 
wenzelm 
parents: 
59685 
diff
changeset
 | 
39  | 
in  | 
| 
59685
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
40  | 
pair  | 
| 
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
41  | 
(list (variant  | 
| 
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
42  | 
[fn ([], a) => Exn.Res (pair string (option string) a),  | 
| 
59714
 
ae322325adbb
tuned protocol -- resolve command positions in ML;
 
wenzelm 
parents: 
59685 
diff
changeset
 | 
43  | 
fn ([], a) => Exn.Exn (ERROR (message a))]))  | 
| 
59685
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
44  | 
int  | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
45  | 
end;  | 
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
46  | 
val toks =  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
47  | 
(YXML.parse_body toks_yxml |> let open XML.Decode in list (pair int int) end) ~~ sources;  | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
48  | 
in  | 
| 
59685
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
49  | 
Document.change_state  | 
| 
 
c043306d2598
clarified markup for embedded files, early before execution;
 
wenzelm 
parents: 
59671 
diff
changeset
 | 
50  | 
(Document.define_command (Document_ID.parse id) name blobs blobs_index toks)  | 
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
53192 
diff
changeset
 | 
51  | 
end);  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
|
| 
38418
 
9a7af64d71bb
more explicit / functional ML version of document model;
 
wenzelm 
parents: 
38417 
diff
changeset
 | 
53  | 
val _ =  | 
| 
47343
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
54  | 
Isabelle_Process.protocol_command "Document.discontinue_execution"  | 
| 52606 | 55  | 
(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
 | 
56  | 
|
| 
 
b8aeab386414
less aggressive discontinue_execution before document update, to avoid unstable execs that need to be re-assigned;
 
wenzelm 
parents: 
46938 
diff
changeset
 | 
57  | 
val _ =  | 
| 
52931
 
ac6648c0c0fb
cancel_query via direct access to the exec_id of the running query process;
 
wenzelm 
parents: 
52862 
diff
changeset
 | 
58  | 
Isabelle_Process.protocol_command "Document.cancel_exec"  | 
| 
 
ac6648c0c0fb
cancel_query via direct access to the exec_id of the running query process;
 
wenzelm 
parents: 
52862 
diff
changeset
 | 
59  | 
(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
 | 
60  | 
|
| 
 
ac6648c0c0fb
cancel_query via direct access to the exec_id of the running query process;
 
wenzelm 
parents: 
52862 
diff
changeset
 | 
61  | 
val _ =  | 
| 
46119
 
0d7172a7672c
tuned signature -- emphasize special nature of protocol commands;
 
wenzelm 
parents: 
45709 
diff
changeset
 | 
62  | 
Isabelle_Process.protocol_command "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
 | 
63  | 
(Future.task_context "Document.update" (Future.new_group NONE)  | 
| 
 
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
 | 
64  | 
(fn [old_id_string, new_id_string, edits_yxml] => Document.change_state (fn state =>  | 
| 
 
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
 | 
65  | 
let  | 
| 
 
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
 | 
66  | 
val _ = Execution.discontinue ();  | 
| 
47404
 
e6e5750f1311
simplified Future.cancel/cancel_group (again) -- running threads only;
 
wenzelm 
parents: 
47388 
diff
changeset
 | 
67  | 
|
| 
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
 | 
68  | 
val old_id = Document_ID.parse old_id_string;  | 
| 
 
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
 | 
69  | 
val new_id = Document_ID.parse new_id_string;  | 
| 
 
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
 | 
70  | 
val edits =  | 
| 
 
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
 | 
71  | 
YXML.parse_body edits_yxml |>  | 
| 
 
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
 | 
72  | 
let open XML.Decode in  | 
| 
 
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
 | 
73  | 
list (pair string  | 
| 
 
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
 | 
74  | 
(variant  | 
| 
 
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
 | 
75  | 
[fn ([], a) => Document.Edits (list (pair (option int) (option int)) a),  | 
| 
 
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
 | 
76  | 
fn ([], a) =>  | 
| 
 
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
 | 
77  | 
let  | 
| 
 
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
 | 
78  | 
val (master, (name, (imports, (keywords, errors)))) =  | 
| 
 
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
 | 
79  | 
pair string (pair string (pair (list string)  | 
| 
 
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
 | 
80  | 
(pair (list (pair string  | 
| 
 
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
 | 
81  | 
(option (pair (pair string (list string)) (list string)))))  | 
| 
59671
 
9715eb8e9408
more precise position information in Isabelle/Scala, with YXML markup as in Isabelle/ML;
 
wenzelm 
parents: 
59463 
diff
changeset
 | 
82  | 
(list YXML.string_of_body)))) a;  | 
| 
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
 | 
83  | 
val imports' = map (rpair Position.none) imports;  | 
| 
59934
 
b65c4370f831
more position information and PIDE markup for command keywords;
 
wenzelm 
parents: 
59715 
diff
changeset
 | 
84  | 
val keywords' = map (fn (x, y) => ((x, Position.none), y)) keywords;  | 
| 
 
b65c4370f831
more position information and PIDE markup for command keywords;
 
wenzelm 
parents: 
59715 
diff
changeset
 | 
85  | 
val header = Thy_Header.make (name, Position.none) imports' keywords';  | 
| 
59715
 
4f0d0e4ad68d
avoid duplicate header errors, more precise positions;
 
wenzelm 
parents: 
59714 
diff
changeset
 | 
86  | 
                      in Document.Deps {master = master, header = header, errors = errors} end,
 | 
| 
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
 | 
87  | 
fn (a :: b, c) =>  | 
| 
 
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
 | 
88  | 
Document.Perspective (bool_atom a, map int_atom b,  | 
| 
 
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
 | 
89  | 
list (pair int (pair string (list string))) c)]))  | 
| 
 
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
 | 
90  | 
end;  | 
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
|
| 
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
 | 
92  | 
val (removed, assign_update, state') = Document.update old_id new_id edits state;  | 
| 
59463
 
b91dc7ab3464
more robust protocol command: purge removed execs asynchronously, to remain reactive despite problems to cancel "Command.run_process" in a situation of overrunning non-terminating tasks (see also 59f1591a11cb);
 
wenzelm 
parents: 
59370 
diff
changeset
 | 
93  | 
val _ =  | 
| 
 
b91dc7ab3464
more robust protocol command: purge removed execs asynchronously, to remain reactive despite problems to cancel "Command.run_process" in a situation of overrunning non-terminating tasks (see also 59f1591a11cb);
 
wenzelm 
parents: 
59370 
diff
changeset
 | 
94  | 
(singleton o Future.forks)  | 
| 
 
b91dc7ab3464
more robust protocol command: purge removed execs asynchronously, to remain reactive despite problems to cancel "Command.run_process" in a situation of overrunning non-terminating tasks (see also 59f1591a11cb);
 
wenzelm 
parents: 
59370 
diff
changeset
 | 
95  | 
             {name = "Document.update/remove", group = NONE,
 | 
| 
 
b91dc7ab3464
more robust protocol command: purge removed execs asynchronously, to remain reactive despite problems to cancel "Command.run_process" in a situation of overrunning non-terminating tasks (see also 59f1591a11cb);
 
wenzelm 
parents: 
59370 
diff
changeset
 | 
96  | 
deps = maps Future.group_snapshot (maps Execution.peek removed),  | 
| 
 
b91dc7ab3464
more robust protocol command: purge removed execs asynchronously, to remain reactive despite problems to cancel "Command.run_process" in a situation of overrunning non-terminating tasks (see also 59f1591a11cb);
 
wenzelm 
parents: 
59370 
diff
changeset
 | 
97  | 
pri = 1, interrupts = false}  | 
| 
 
b91dc7ab3464
more robust protocol command: purge removed execs asynchronously, to remain reactive despite problems to cancel "Command.run_process" in a situation of overrunning non-terminating tasks (see also 59f1591a11cb);
 
wenzelm 
parents: 
59370 
diff
changeset
 | 
98  | 
(fn () => (Execution.purge removed; List.app Isabelle_Process.reset_tracing removed));  | 
| 52601 | 99  | 
|
| 
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
 | 
100  | 
val _ =  | 
| 
 
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  | 
Output.protocol_message Markup.assign_update  | 
| 
 
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
 | 
102  | 
[(new_id, assign_update) |>  | 
| 
 
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
 | 
103  | 
let open XML.Encode  | 
| 
 
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
 | 
104  | 
in pair int (list (pair int (list int))) end  | 
| 
 
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
 | 
105  | 
|> YXML.string_of_body];  | 
| 
 
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
 | 
106  | 
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
 | 
107  | 
|
| 43748 | 108  | 
val _ =  | 
| 
46119
 
0d7172a7672c
tuned signature -- emphasize special nature of protocol commands;
 
wenzelm 
parents: 
45709 
diff
changeset
 | 
109  | 
Isabelle_Process.protocol_command "Document.remove_versions"  | 
| 44673 | 110  | 
(fn [versions_yxml] => Document.change_state (fn state =>  | 
111  | 
let  | 
|
112  | 
val versions =  | 
|
113  | 
YXML.parse_body versions_yxml |>  | 
|
114  | 
let open XML.Decode in list int end;  | 
|
| 44676 | 115  | 
val state1 = Document.remove_versions versions state;  | 
| 
56333
 
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
 
wenzelm 
parents: 
54717 
diff
changeset
 | 
116  | 
val _ = Output.protocol_message Markup.removed_versions [versions_yxml];  | 
| 44676 | 117  | 
in state1 end));  | 
| 44673 | 118  | 
|
119  | 
val _ =  | 
|
| 50498 | 120  | 
Isabelle_Process.protocol_command "Document.dialog_result"  | 
| 
50500
 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 
wenzelm 
parents: 
50498 
diff
changeset
 | 
121  | 
(fn [serial, result] =>  | 
| 
 
c94bba7906d2
identify dialogs via official serial and maintain as result message;
 
wenzelm 
parents: 
50498 
diff
changeset
 | 
122  | 
Active.dialog_result (Markup.parse_int serial) result  | 
| 
52775
 
e0169f13bd37
keep memo_exec execution running, which is important to cancel goal forks eventually;
 
wenzelm 
parents: 
52774 
diff
changeset
 | 
123  | 
handle exn => if Exn.is_interrupt exn then () (*sic!*) else reraise exn);  | 
| 50498 | 124  | 
|
| 
56616
 
abc2da18d08d
added protocol command "use_theories", with core functionality of batch build;
 
wenzelm 
parents: 
56458 
diff
changeset
 | 
125  | 
val _ =  | 
| 
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
 | 
126  | 
Isabelle_Process.protocol_command "ML_System.share_common_data"  | 
| 
 
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
 | 
127  | 
(fn [] => ML_System.share_common_data ());  | 
| 
 
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
 | 
128  | 
|
| 
38412
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
end;  | 
| 
 
c23f3abbf42d
moved isar_document.ML/scala to Pure/System/ -- side-by-side with isar.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
130  |