--- a/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy Wed Oct 01 00:09:51 2008 +0200
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Setup.thy Wed Oct 01 08:42:42 2008 +0200
@@ -1,61 +1,10 @@
theory Setup
imports Complex_Main
-uses "../../../antiquote_setup.ML"
+uses "../../../antiquote_setup.ML" "../../../more_antiquote.ML"
begin
ML {* no_document use_thys ["Efficient_Nat", "Code_Char_chr"] *}
ML_val {* Code_Target.code_width := 74 *}
-ML {*
-fun pr_class ctxt = enclose "\\isa{" "}" o Sign.extern_class (ProofContext.theory_of ctxt)
- o Sign.read_class (ProofContext.theory_of ctxt);
-*}
-
-ML {*
-val _ = ThyOutput.add_commands
- [("class", ThyOutput.args (Scan.lift Args.name) (K pr_class))]
-*}
-
-ML {*
-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"
-*}
-
-ML {*
-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 => map (Code_Name.const thy) (mk_cs thy));
- val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
- >> (fn tycos => fn thy => map (Code_Name.tyco thy o Sign.intern_type thy) tycos);
- val parse_classes = Scan.lift (Args.parens (Args.$$$ "classes") |-- Scan.repeat1 Args.name)
- >> (fn classes => fn thy => map (Code_Name.class thy 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 => map (Code_Name.instance thy 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))
- (maps (fn f => f (ProofContext.theory_of ctxt)) mk_stmtss)
- |> split_lines
- |> verbatim_lines;
-
-in
-
-val _ = ThyOutput.add_commands
- [("code_stmts", ThyOutput.args
- (parse_const_terms -- Scan.repeat parse_names -- Scan.lift (Args.parens Args.name))
- code_stmts)]
-
end
-*}
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/more_antiquote.ML Wed Oct 01 08:42:42 2008 +0200
@@ -0,0 +1,75 @@
+(* 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
+
+fun pr_class ctxt = enclose "\\isa{" "}" o Sign.extern_class (ProofContext.theory_of ctxt)
+ o Sign.read_class (ProofContext.theory_of ctxt);
+
+in
+
+val _ = O.add_commands
+ [("class", ThyOutput.args (Scan.lift Args.name) (K pr_class))]
+
+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 => map (Code_Name.const thy) (mk_cs thy));
+ val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
+ >> (fn tycos => fn thy => map (Code_Name.tyco thy o Sign.intern_type thy) tycos);
+ val parse_classes = Scan.lift (Args.parens (Args.$$$ "classes") |-- Scan.repeat1 Args.name)
+ >> (fn classes => fn thy => map (Code_Name.class thy 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 => map (Code_Name.instance thy 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))
+ (maps (fn f => f (ProofContext.theory_of ctxt)) 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;