| author | wenzelm | 
| Tue, 22 Jul 2014 13:36:51 +0200 | |
| changeset 57604 | 30885e25c6de | 
| parent 57415 | e721124f1b1e | 
| child 57605 | 8e0a7eaffe47 | 
| permissions | -rw-r--r-- | 
| 56864 | 1  | 
(* Title: Pure/Tools/print_operation.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Print operations as asynchronous query.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
||
8  | 
signature PRINT_OPERATION =  | 
|
9  | 
sig  | 
|
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
10  | 
val register: string -> string -> (Toplevel.state -> Pretty.T list) -> unit  | 
| 56864 | 11  | 
end;  | 
12  | 
||
13  | 
structure Print_Operation: PRINT_OPERATION =  | 
|
14  | 
struct  | 
|
15  | 
||
16  | 
(* maintain print operations *)  | 
|
17  | 
||
18  | 
local  | 
|
19  | 
||
20  | 
val print_operations =  | 
|
21  | 
Synchronized.var "print_operations"  | 
|
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
22  | 
([]: (string * (string * (Toplevel.state -> Pretty.T list))) list);  | 
| 56864 | 23  | 
|
24  | 
fun report () =  | 
|
25  | 
Output.try_protocol_message Markup.print_operations  | 
|
26  | 
let  | 
|
27  | 
val yxml =  | 
|
| 56867 | 28  | 
Synchronized.value print_operations  | 
29  | 
|> map (fn (x, (y, _)) => (x, y)) |> rev  | 
|
| 56864 | 30  | 
|> let open XML.Encode in list (pair string string) end  | 
31  | 
|> YXML.string_of_body;  | 
|
32  | 
in [yxml] end;  | 
|
33  | 
||
34  | 
val _ = Isabelle_Process.protocol_command "print_operations" (fn [] => report ());  | 
|
35  | 
||
36  | 
val _ = Session.protocol_handler "isabelle.Print_Operation$Handler";  | 
|
37  | 
||
38  | 
in  | 
|
39  | 
||
40  | 
fun register name description pr =  | 
|
41  | 
(Synchronized.change print_operations (fn tab =>  | 
|
| 56867 | 42  | 
(if not (AList.defined (op =) tab name) then ()  | 
| 56864 | 43  | 
    else warning ("Redefining print operation: " ^ quote name);
 | 
| 56867 | 44  | 
AList.update (op =) (name, (description, pr)) tab));  | 
| 56864 | 45  | 
report ());  | 
46  | 
||
47  | 
val _ =  | 
|
48  | 
  Query_Operation.register "print_operation" (fn {state, args, output_result} =>
 | 
|
49  | 
let  | 
|
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
50  | 
val _ = Toplevel.context_of state handle Toplevel.UNDEF => error "Unknown context";  | 
| 
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
51  | 
fun err s = Pretty.mark_str (Markup.bad, s);  | 
| 
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
52  | 
fun print name =  | 
| 56867 | 53  | 
(case AList.lookup (op =) (Synchronized.value print_operations) name of  | 
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
54  | 
SOME (_, pr) => (pr state handle Toplevel.UNDEF => [err "Unknown context"])  | 
| 
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
55  | 
        | NONE => [err ("Unknown print operation: " ^ quote name)]);
 | 
| 
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
56  | 
in output_result (Pretty.string_of (Pretty.chunks (maps print args))) end);  | 
| 56864 | 57  | 
|
58  | 
end;  | 
|
59  | 
||
60  | 
||
61  | 
(* common print operations *)  | 
|
62  | 
||
63  | 
val _ =  | 
|
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
64  | 
register "context" "context of local theory target" Toplevel.pretty_context;  | 
| 56867 | 65  | 
|
66  | 
val _ =  | 
|
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
67  | 
register "cases" "cases of proof context" (Proof_Context.pretty_cases o Toplevel.context_of);  | 
| 56867 | 68  | 
|
69  | 
val _ =  | 
|
| 
56868
 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
70  | 
register "terms" "term bindings of proof context"  | 
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
71  | 
(Proof_Context.pretty_term_bindings o Toplevel.context_of);  | 
| 56867 | 72  | 
|
73  | 
val _ =  | 
|
| 
56868
 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
74  | 
register "theorems" "theorems of local theory or proof context"  | 
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
75  | 
(single o Isar_Cmd.pretty_theorems false);  | 
| 56864 | 76  | 
|
| 56867 | 77  | 
val _ =  | 
| 
57604
 
30885e25c6de
support multiple selected print operations instead of slightly odd "menu";
 
wenzelm 
parents: 
57415 
diff
changeset
 | 
78  | 
register "state" "proof state" Toplevel.pretty_state;  | 
| 56867 | 79  | 
|
| 56864 | 80  | 
end;  | 
81  |