| 72763 |      1 | (*  Title:      Pure/System/isabelle_tool.ML
 | 
|  |      2 |     Author:     Makarius
 | 
|  |      3 | 
 | 
|  |      4 | Support for Isabelle system tools.
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | signature ISABELLE_TOOL =
 | 
|  |      8 | sig
 | 
|  |      9 |   val isabelle_tools: unit -> (string * Position.T) list
 | 
|  |     10 |   val check: Proof.context -> string * Position.T -> string
 | 
|  |     11 | end;
 | 
|  |     12 | 
 | 
|  |     13 | structure Isabelle_Tool: ISABELLE_TOOL =
 | 
|  |     14 | struct
 | 
|  |     15 | 
 | 
|  |     16 | (* list tools *)
 | 
|  |     17 | 
 | 
|  |     18 | fun symbolic_file (a, b) =
 | 
|  |     19 |   if a = Markup.fileN
 | 
|  |     20 |   then (a, Path.explode b |> Path.implode_symbolic)
 | 
|  |     21 |   else (a, b);
 | 
|  |     22 | 
 | 
|  |     23 | fun isabelle_tools () =
 | 
|  |     24 |   \<^scala>\<open>isabelle_tools\<close> ""
 | 
|  |     25 |   |> YXML.parse_body
 | 
|  |     26 |   |> let open XML.Decode in list (pair string properties) end
 | 
|  |     27 |   |> map (apsnd (map symbolic_file #> Position.of_properties));
 | 
|  |     28 | 
 | 
|  |     29 | 
 | 
|  |     30 | (* check *)
 | 
|  |     31 | 
 | 
|  |     32 | fun check ctxt arg =
 | 
|  |     33 |   Completion.check_item Markup.toolN
 | 
|  |     34 |     (fn (name, pos) =>
 | 
|  |     35 |       Markup.entity Markup.toolN name
 | 
|  |     36 |       |> Markup.properties (Position.def_properties_of pos))
 | 
|  |     37 |     (isabelle_tools ()) ctxt arg;
 | 
|  |     38 | 
 | 
|  |     39 | val _ =
 | 
|  |     40 |   Theory.setup
 | 
| 73761 |     41 |    (Document_Output.antiquotation_verbatim_embedded \<^binding>\<open>tool\<close>
 | 
| 72763 |     42 |       (Scan.lift Parse.embedded_position) check);
 | 
|  |     43 | 
 | 
|  |     44 | end;
 |