doc-src/more_antiquote.ML
author ballarin
Tue, 28 Oct 2008 17:53:46 +0100
changeset 28707 548703affff5
parent 28679 d7384e8e99b3
child 28714 1992553cccfe
permissions -rw-r--r--
Revoked workaround (incompatible with HOL/ex/LocaleTest2.thy).

(*  Title:      Doc/more_antiquote.ML
    ID:         $Id$
    Author:     Florian Haftmann, TU Muenchen

More antiquotations.
*)

signature MORE_ANTIQUOTE =
sig
  val verbatim_lines: string list -> string
end;

structure More_Antiquote : MORE_ANTIQUOTE =
struct

structure O = ThyOutput;

(* printing verbatim lines *)

val verbatim_line = space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|";
val verbatim_lines = rev
  #> dropwhile (fn s => s = "")
  #> rev
  #> map verbatim_line #> space_implode "\\newline%\n"
  #> prefix "\\isaverbatim%\n\\noindent%\n"


(* class antiquotation *)

local

val pr_text = enclose "\\isa{" "}" o Pretty.output o Pretty.str;

fun pr_class ctxt = Sign.read_class (ProofContext.theory_of ctxt)
  #> Sign.extern_class (ProofContext.theory_of ctxt)
  #> pr_text;

fun pr_type ctxt = Sign.intern_type (ProofContext.theory_of ctxt)
  #> tap (Sign.arity_number (ProofContext.theory_of ctxt))
  #> Sign.extern_type (ProofContext.theory_of ctxt)
  #> pr_text;

in

val _ = O.add_commands
  [("class", ThyOutput.args (Scan.lift Args.name) (K pr_class)),
    ("type", ThyOutput.args (Scan.lift Args.name) (K pr_type))]

end;


(* code antiquotation *)

local

  val parse_const_terms = Scan.repeat1 Args.term
    >> (fn ts => fn thy => map (Code_Unit.check_const thy) ts);
  val parse_consts = Scan.lift (Args.parens (Args.$$$ "consts")) |-- parse_const_terms
    >> (fn mk_cs => fn thy => fn naming => map (the o Code_Thingol.lookup_const naming) (mk_cs thy));
  val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
    >> (fn tycos => fn thy => fn naming => map (the o Code_Thingol.lookup_tyco naming o Sign.intern_type thy) tycos);
  val parse_classes = Scan.lift (Args.parens (Args.$$$ "classes") |-- Scan.repeat1 Args.name)
    >> (fn classes => fn thy => fn naming => map (the o Code_Thingol.lookup_class naming o Sign.intern_class thy) classes);
  val parse_instances = Scan.lift (Args.parens (Args.$$$ "instances") |-- Scan.repeat1 (Args.name --| Args.$$$ "::" -- Args.name))
    >> (fn insts => fn thy => fn naming => map (the o Code_Thingol.lookup_instance naming o apsnd (Sign.intern_type thy) o apfst (Sign.intern_class thy) o swap) insts);
  val parse_names = parse_consts || parse_types || parse_classes || parse_instances; 

  fun code_stmts src ctxt ((mk_cs, mk_stmtss), target) =
    Code_Target.code_of (ProofContext.theory_of ctxt)
      target "Example" (mk_cs (ProofContext.theory_of ctxt))
        (fn naming => maps (fn f => f (ProofContext.theory_of ctxt) naming) mk_stmtss)
    |> split_lines
    |> verbatim_lines;

in

val _ = O.add_commands
  [("code_stmts", O.args
      (parse_const_terms -- Scan.repeat parse_names -- Scan.lift (Args.parens Args.name))
        code_stmts)]

end

end;