| author | wenzelm | 
| Sun, 13 Aug 2023 19:27:58 +0200 | |
| changeset 78528 | 3d6dbf215559 | 
| parent 76652 | 90abc28523d7 | 
| permissions | -rw-r--r-- | 
| 61617 | 1  | 
(* Title: Pure/Tools/jedit.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 71513 | 4  | 
Support for Isabelle/jEdit.  | 
| 61617 | 5  | 
*)  | 
6  | 
||
7  | 
signature JEDIT =  | 
|
8  | 
sig  | 
|
| 71513 | 9  | 
val get_actions: unit -> string list  | 
| 61620 | 10  | 
val check_action: string * Position.T -> string  | 
| 61617 | 11  | 
end;  | 
12  | 
||
13  | 
structure JEdit: JEDIT =  | 
|
14  | 
struct  | 
|
15  | 
||
| 71513 | 16  | 
(* parse XML *)  | 
| 61617 | 17  | 
|
18  | 
fun parse_named a (XML.Elem ((b, props), _)) =  | 
|
19  | 
(case Properties.get props "NAME" of  | 
|
20  | 
SOME name => if a = b then [name] else []  | 
|
21  | 
| NONE => [])  | 
|
22  | 
| parse_named _ _ = [];  | 
|
23  | 
||
| 71513 | 24  | 
fun parse_actions (XML.Elem (("ACTIONS", _), body)) = maps (parse_named "ACTION") body
 | 
25  | 
| parse_actions _ = [];  | 
|
26  | 
||
27  | 
fun parse_dockables (XML.Elem (("DOCKABLES", _), body)) =
 | 
|
28  | 
maps (parse_named "DOCKABLE") body  | 
|
29  | 
|> maps (fn a => [a, a ^ "-toggle", a ^ "-float"])  | 
|
30  | 
| parse_dockables _ = [];  | 
|
31  | 
||
32  | 
||
| 
76552
 
13fde66c7cf6
more direct access to jEdit jar resources, without unzip;
 
wenzelm 
parents: 
74563 
diff
changeset
 | 
33  | 
(* actions *)  | 
| 61617 | 34  | 
|
| 
76552
 
13fde66c7cf6
more direct access to jEdit jar resources, without unzip;
 
wenzelm 
parents: 
74563 
diff
changeset
 | 
35  | 
fun parse_resources [actions, dockables] =  | 
| 
 
13fde66c7cf6
more direct access to jEdit jar resources, without unzip;
 
wenzelm 
parents: 
74563 
diff
changeset
 | 
36  | 
parse_actions (XML.parse actions) @ parse_dockables (XML.parse dockables);  | 
| 71513 | 37  | 
|
38  | 
val lazy_actions =  | 
|
| 61617 | 39  | 
Lazy.lazy (fn () =>  | 
| 
76552
 
13fde66c7cf6
more direct access to jEdit jar resources, without unzip;
 
wenzelm 
parents: 
74563 
diff
changeset
 | 
40  | 
(parse_resources o map File.read)  | 
| 
 
13fde66c7cf6
more direct access to jEdit jar resources, without unzip;
 
wenzelm 
parents: 
74563 
diff
changeset
 | 
41  | 
[\<^file>\<open>~~/src/Tools/jEdit/jedit_main/actions.xml\<close>,  | 
| 
 
13fde66c7cf6
more direct access to jEdit jar resources, without unzip;
 
wenzelm 
parents: 
74563 
diff
changeset
 | 
42  | 
\<^file>\<open>~~/src/Tools/jEdit/jedit_main/dockables.xml\<close>] @  | 
| 
76652
 
90abc28523d7
tuned names: avoid overlap with instances of class Resources;
 
wenzelm 
parents: 
76552 
diff
changeset
 | 
43  | 
(parse_resources o \<^scala>\<open>jEdit.resource\<close>) ["actions.xml", "dockables.xml"]  | 
| 71513 | 44  | 
|> sort_strings);  | 
| 61617 | 45  | 
|
| 71513 | 46  | 
fun get_actions () = Lazy.force lazy_actions;  | 
| 61620 | 47  | 
|
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
48  | 
fun check_action (name, pos) =  | 
| 71513 | 49  | 
if member (op =) (get_actions ()) name then  | 
| 63681 | 50  | 
let  | 
51  | 
val ((bg1, bg2), en) =  | 
|
52  | 
YXML.output_markup_elem  | 
|
53  | 
          (Active.make_markup Markup.jedit_actionN {implicit = false, properties = []});
 | 
|
54  | 
val msg = "Invoke " ^ bg1 ^ name ^ bg2 ^ name ^ en ^ " jEdit action";  | 
|
55  | 
in writeln (msg ^ Position.here pos); name end  | 
|
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
56  | 
else  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
57  | 
let  | 
| 69289 | 58  | 
val completion_report =  | 
59  | 
Completion.make_report (name, pos)  | 
|
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
60  | 
(fn completed =>  | 
| 71513 | 61  | 
get_actions ()  | 
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
62  | 
|> filter completed  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
63  | 
|> sort_strings  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
64  | 
            |> map (fn a => (a, ("action", a))));
 | 
| 71513 | 65  | 
    in error ("Bad jEdit action " ^ quote name ^ Position.here pos ^ completion_report) end;
 | 
| 61620 | 66  | 
|
| 63681 | 67  | 
val _ =  | 
68  | 
Theory.setup  | 
|
| 74563 | 69  | 
(Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>action\<close>  | 
70  | 
(Scan.lift Parse.embedded_position)  | 
|
| 67463 | 71  | 
(fn ctxt => fn (name, pos) =>  | 
72  | 
let  | 
|
73  | 
val _ =  | 
|
74  | 
if Context_Position.is_reported ctxt pos  | 
|
75  | 
then ignore (check_action (name, pos))  | 
|
76  | 
else ();  | 
|
77  | 
in name end));  | 
|
| 63681 | 78  | 
|
| 61617 | 79  | 
end;  |