| author | wenzelm | 
| Fri, 17 May 2013 19:11:03 +0200 | |
| changeset 52057 | 69137d20ab0b | 
| parent 51605 | eca8acb42e4a | 
| child 52509 | 2193d2c7f586 | 
| permissions | -rw-r--r-- | 
| 47336 | 1  | 
(* Title: Pure/PIDE/command.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Prover command execution.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature COMMAND =  | 
|
8  | 
sig  | 
|
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
9  | 
val range: Token.T list -> Position.range  | 
| 
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
10  | 
val proper_range: Token.T list -> Position.range  | 
| 
47341
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
11  | 
type 'a memo  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
12  | 
val memo: (unit -> 'a) -> 'a memo  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
13  | 
val memo_value: 'a -> 'a memo  | 
| 
47342
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
14  | 
val memo_eval: 'a memo -> 'a  | 
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
15  | 
val memo_result: 'a memo -> 'a  | 
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
16  | 
val run_command: Toplevel.transition * Token.T list ->  | 
| 48772 | 17  | 
Toplevel.state * bool -> (Toplevel.state * bool) * unit lazy  | 
| 47336 | 18  | 
end;  | 
19  | 
||
20  | 
structure Command: COMMAND =  | 
|
21  | 
struct  | 
|
22  | 
||
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
23  | 
(* span range *)  | 
| 
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
24  | 
|
| 49866 | 25  | 
val range = Token.position_range_of;  | 
| 
51266
 
3007d0bc9cb1
unified Command.is_proper in ML with Scala (see also 123be08eed88);
 
wenzelm 
parents: 
50914 
diff
changeset
 | 
26  | 
val proper_range = Token.position_range_of o #1 o take_suffix Token.is_improper;  | 
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
27  | 
|
| 
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
47395 
diff
changeset
 | 
28  | 
|
| 
47341
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
29  | 
(* memo results *)  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
30  | 
|
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
31  | 
datatype 'a expr =  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
32  | 
Expr of unit -> 'a |  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
33  | 
Result of 'a Exn.result;  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
34  | 
|
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
35  | 
abstype 'a memo = Memo of 'a expr Synchronized.var  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
36  | 
with  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
37  | 
|
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
38  | 
fun memo e = Memo (Synchronized.var "Command.memo" (Expr e));  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
39  | 
fun memo_value a = Memo (Synchronized.var "Command.memo" (Result (Exn.Res a)));  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
40  | 
|
| 
47342
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
41  | 
fun memo_eval (Memo v) =  | 
| 
47341
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
42  | 
(case Synchronized.value v of  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
43  | 
Result res => res  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
44  | 
| _ =>  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
45  | 
Synchronized.guarded_access v  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
46  | 
(fn Result res => SOME (res, Result res)  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
47  | 
| Expr e =>  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
48  | 
let val res = Exn.capture e (); (*memoing of physical interrupts!*)  | 
| 
47342
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
49  | 
in SOME (res, Result res) end))  | 
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
50  | 
|> Exn.release;  | 
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
51  | 
|
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
52  | 
fun memo_result (Memo v) =  | 
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
53  | 
(case Synchronized.value v of  | 
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
54  | 
Result res => Exn.release res  | 
| 
 
7828c7b3c143
more explicit memo_eval vs. memo_result, to enforce bottom-up execution;
 
wenzelm 
parents: 
47341 
diff
changeset
 | 
55  | 
| _ => raise Fail "Unfinished memo result");  | 
| 
47341
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
56  | 
|
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
57  | 
end;  | 
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
58  | 
|
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
59  | 
|
| 
 
00f6279bb67a
Command.memo including physical interrupts (unlike Lazy.lazy);
 
wenzelm 
parents: 
47336 
diff
changeset
 | 
60  | 
(* run command *)  | 
| 47336 | 61  | 
|
62  | 
local  | 
|
63  | 
||
64  | 
fun run int tr st =  | 
|
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51266 
diff
changeset
 | 
65  | 
if Goal.future_enabled () andalso Keyword.is_diag (Toplevel.name_of tr) then  | 
| 
51605
 
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
 
wenzelm 
parents: 
51603 
diff
changeset
 | 
66  | 
    (Goal.fork_params {name = "Toplevel.diag", pos = Toplevel.pos_of tr, pri = ~1}
 | 
| 
 
eca8acb42e4a
more explicit Goal.fork_params -- avoid implicit arguments via thread data;
 
wenzelm 
parents: 
51603 
diff
changeset
 | 
67  | 
(fn () => Toplevel.command_exception int tr st); ([], SOME st))  | 
| 
51284
 
59a03019f3bf
fork diagnostic commands (theory loader and PIDE interaction);
 
wenzelm 
parents: 
51266 
diff
changeset
 | 
68  | 
else Toplevel.command_errors int tr st;  | 
| 47336 | 69  | 
|
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
70  | 
fun check_cmts tr cmts st =  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
71  | 
Toplevel.setmp_thread_position tr  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
72  | 
(fn () => cmts  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
73  | 
|> maps (fn cmt =>  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
74  | 
(Thy_Output.check_text (Token.source_position_of cmt) st; [])  | 
| 
50914
 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 
wenzelm 
parents: 
50911 
diff
changeset
 | 
75  | 
handle exn => ML_Compiler.exn_messages_ids exn)) ();  | 
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
76  | 
|
| 47336 | 77  | 
fun proof_status tr st =  | 
78  | 
(case try Toplevel.proof_of st of  | 
|
79  | 
SOME prf => Toplevel.status tr (Proof.status_markup prf)  | 
|
80  | 
| NONE => ());  | 
|
81  | 
||
82  | 
val no_print = Lazy.value ();  | 
|
83  | 
||
84  | 
fun print_state tr st =  | 
|
85  | 
(Lazy.lazy o Toplevel.setmp_thread_position tr)  | 
|
86  | 
(fn () => Toplevel.print_state false st);  | 
|
87  | 
||
88  | 
in  | 
|
89  | 
||
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
90  | 
fun run_command (tr, cmts) (st, malformed) =  | 
| 48772 | 91  | 
if malformed then ((Toplevel.toplevel, malformed), no_print)  | 
92  | 
else  | 
|
93  | 
let  | 
|
94  | 
val malformed' = Toplevel.is_malformed tr;  | 
|
95  | 
val is_init = Toplevel.is_init tr;  | 
|
96  | 
val is_proof = Keyword.is_proof (Toplevel.name_of tr);  | 
|
| 47336 | 97  | 
|
| 48772 | 98  | 
val _ = Multithreading.interrupted ();  | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49866 
diff
changeset
 | 
99  | 
val _ = Toplevel.status tr Markup.running;  | 
| 
48918
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
100  | 
val (errs1, result) = run (is_init orelse is_proof) (Toplevel.set_print false tr) st;  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
101  | 
val errs2 = (case result of NONE => [] | SOME st' => check_cmts tr cmts st');  | 
| 
 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 
wenzelm 
parents: 
48772 
diff
changeset
 | 
102  | 
val errs = errs1 @ errs2;  | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49866 
diff
changeset
 | 
103  | 
val _ = Toplevel.status tr Markup.finished;  | 
| 
50914
 
fe4714886d92
identify future results more carefully, to avoid odd duplication of error messages, notably from forked goals;
 
wenzelm 
parents: 
50911 
diff
changeset
 | 
104  | 
val _ = List.app (Future.error_msg (Toplevel.pos_of tr)) errs;  | 
| 48772 | 105  | 
in  | 
106  | 
(case result of  | 
|
107  | 
NONE =>  | 
|
108  | 
let  | 
|
109  | 
val _ = if null errs then Exn.interrupt () else ();  | 
|
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49866 
diff
changeset
 | 
110  | 
val _ = Toplevel.status tr Markup.failed;  | 
| 48772 | 111  | 
in ((st, malformed'), no_print) end  | 
112  | 
| SOME st' =>  | 
|
113  | 
let  | 
|
114  | 
val _ = proof_status tr st';  | 
|
115  | 
val do_print =  | 
|
116  | 
not is_init andalso  | 
|
117  | 
(Toplevel.print_of tr orelse (is_proof andalso Toplevel.is_proof st'));  | 
|
118  | 
in ((st', malformed'), if do_print then print_state tr st' else no_print) end)  | 
|
119  | 
end;  | 
|
| 47336 | 120  | 
|
121  | 
end;  | 
|
122  | 
||
123  | 
end;  | 
|
124  |