src/Pure/System/isabelle_tool.ML
author wenzelm
Sat, 28 Nov 2020 21:56:24 +0100
changeset 72763 3cc73d00553c
child 73761 ef1a18e20ace
permissions -rw-r--r--
added document antiquotation @{tool}; formal check of isabelle tools via Isabelle/Scala;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
72763
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/System/isabelle_tool.ML
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     3
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     4
Support for Isabelle system tools.
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     5
*)
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     6
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     7
signature ISABELLE_TOOL =
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     8
sig
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
     9
  val isabelle_tools: unit -> (string * Position.T) list
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    10
  val check: Proof.context -> string * Position.T -> string
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    11
end;
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    12
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    13
structure Isabelle_Tool: ISABELLE_TOOL =
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    14
struct
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    15
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    16
(* list tools *)
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    17
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    18
fun symbolic_file (a, b) =
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    19
  if a = Markup.fileN
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    20
  then (a, Path.explode b |> Path.implode_symbolic)
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    21
  else (a, b);
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    22
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    23
fun isabelle_tools () =
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    24
  \<^scala>\<open>isabelle_tools\<close> ""
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    25
  |> YXML.parse_body
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    26
  |> let open XML.Decode in list (pair string properties) end
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    27
  |> map (apsnd (map symbolic_file #> Position.of_properties));
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    28
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    29
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    30
(* check *)
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    31
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    32
fun check ctxt arg =
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    33
  Completion.check_item Markup.toolN
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    34
    (fn (name, pos) =>
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    35
      Markup.entity Markup.toolN name
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    36
      |> Markup.properties (Position.def_properties_of pos))
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    37
    (isabelle_tools ()) ctxt arg;
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    38
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    39
val _ =
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    40
  Theory.setup
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    41
   (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>tool\<close>
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    42
      (Scan.lift Parse.embedded_position) check);
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    43
3cc73d00553c added document antiquotation @{tool};
wenzelm
parents:
diff changeset
    44
end;