| author | wenzelm | 
| Tue, 02 May 2017 18:27:51 +0200 | |
| changeset 65686 | 4a762cad298f | 
| parent 63681 | d2448471ffba | 
| child 66670 | e5188cb1c3d8 | 
| permissions | -rw-r--r-- | 
| 61617 | 1  | 
(* Title: Pure/Tools/jedit.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
jEdit support.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature JEDIT =  | 
|
8  | 
sig  | 
|
| 61620 | 9  | 
val check_action: string * Position.T -> string  | 
| 61617 | 10  | 
end;  | 
11  | 
||
12  | 
structure JEdit: JEDIT =  | 
|
13  | 
struct  | 
|
14  | 
||
15  | 
(* jEdit actions *)  | 
|
16  | 
||
17  | 
local  | 
|
18  | 
||
19  | 
fun parse_named a (XML.Elem ((b, props), _)) =  | 
|
20  | 
(case Properties.get props "NAME" of  | 
|
21  | 
SOME name => if a = b then [name] else []  | 
|
22  | 
| NONE => [])  | 
|
23  | 
| parse_named _ _ = [];  | 
|
24  | 
||
25  | 
val isabelle_jedit_actions =  | 
|
26  | 
Lazy.lazy (fn () =>  | 
|
| 63680 | 27  | 
(case XML.parse (File.read \<^file>\<open>~~/src/Tools/jEdit/src/actions.xml\<close>) of  | 
| 61617 | 28  | 
      XML.Elem (("ACTIONS", _), body) => maps (parse_named "ACTION") body
 | 
29  | 
| _ => []));  | 
|
30  | 
||
31  | 
val isabelle_jedit_dockables =  | 
|
32  | 
Lazy.lazy (fn () =>  | 
|
| 63680 | 33  | 
(case XML.parse (File.read \<^file>\<open>~~/src/Tools/jEdit/src/dockables.xml\<close>) of  | 
| 61617 | 34  | 
      XML.Elem (("DOCKABLES", _), body) => maps (parse_named "DOCKABLE") body
 | 
35  | 
| _ => []));  | 
|
36  | 
||
37  | 
val jedit_actions =  | 
|
38  | 
Lazy.lazy (fn () =>  | 
|
39  | 
(case Isabelle_System.bash_output  | 
|
40  | 
"unzip -p \"$JEDIT_HOME/dist/jedit.jar\" org/gjt/sp/jedit/actions.xml" of  | 
|
41  | 
(txt, 0) =>  | 
|
42  | 
(case XML.parse txt of  | 
|
43  | 
          XML.Elem (("ACTIONS", _), body) => maps (parse_named "ACTION") body
 | 
|
44  | 
| _ => [])  | 
|
45  | 
    | (_, rc) => error ("Cannot unzip jedit.jar\nreturn code = " ^ string_of_int rc)));
 | 
|
46  | 
||
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
47  | 
val all_actions =  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
48  | 
Lazy.lazy (fn () =>  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
49  | 
Lazy.force isabelle_jedit_actions @  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
50  | 
Lazy.force isabelle_jedit_dockables @  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
51  | 
Lazy.force jedit_actions);  | 
| 61617 | 52  | 
|
| 61620 | 53  | 
in  | 
54  | 
||
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
55  | 
fun check_action (name, pos) =  | 
| 63681 | 56  | 
if member (op =) (Lazy.force all_actions) name then  | 
57  | 
let  | 
|
58  | 
val ((bg1, bg2), en) =  | 
|
59  | 
YXML.output_markup_elem  | 
|
60  | 
          (Active.make_markup Markup.jedit_actionN {implicit = false, properties = []});
 | 
|
61  | 
val msg = "Invoke " ^ bg1 ^ name ^ bg2 ^ name ^ en ^ " jEdit action";  | 
|
62  | 
in writeln (msg ^ Position.here pos); name end  | 
|
| 
61621
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
63  | 
else  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
64  | 
let  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
65  | 
val completion =  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
66  | 
Completion.make (name, pos)  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
67  | 
(fn completed =>  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
68  | 
Lazy.force all_actions  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
69  | 
|> filter completed  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
70  | 
|> sort_strings  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
71  | 
            |> map (fn a => (a, ("action", a))));
 | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
72  | 
val report = Markup.markup_report (Completion.reported_text completion);  | 
| 
 
70b8085f51b4
more thorough check_action, including completion;
 
wenzelm 
parents: 
61620 
diff
changeset
 | 
73  | 
    in error ("Bad jEdit action " ^ quote name ^ Position.here pos ^ report) end
 | 
| 61620 | 74  | 
|
| 63681 | 75  | 
val _ =  | 
76  | 
Theory.setup  | 
|
77  | 
    (Thy_Output.antiquotation @{binding action} (Scan.lift (Parse.position Parse.embedded))
 | 
|
78  | 
      (fn {context = ctxt, ...} => fn (name, pos) =>
 | 
|
79  | 
(if Context_Position.is_reported ctxt pos then ignore (check_action (name, pos)) else ();  | 
|
80  | 
Thy_Output.verbatim_text ctxt name)));  | 
|
81  | 
||
| 61617 | 82  | 
end;  | 
83  | 
||
84  | 
end;  |