src/Pure/Tools/doc.ML
author wenzelm
Fri, 08 Jan 2021 22:30:32 +0100
changeset 73110 c87ca43ebd3b
parent 72760 042180540068
child 73761 ef1a18e20ace
permissions -rw-r--r--
avoid rescaled fonts, e.g. dockable buttons on Windows L&F after opening a new view;

(*  Title:      Pure/Tools/doc.ML
    Author:     Makarius

Access to Isabelle documentation.
*)

signature DOC =
sig
  val check: Proof.context -> string * Position.T -> string
end;

structure Doc: DOC =
struct

fun check ctxt arg =
  Completion.check_item "documentation" (Markup.doc o #1)
    (\<^scala>\<open>doc_names\<close> "" |> split_lines |> map (rpair ())) ctxt arg;

val _ =
  Theory.setup
   (Thy_Output.antiquotation_verbatim_embedded \<^binding>\<open>doc\<close>
      (Scan.lift Parse.embedded_position) check);

end;