common antiquote_setup.ML;
authorwenzelm
Wed, 15 Nov 2006 15:36:09 +0100
changeset 21374 27ae6bc4102a
parent 21373 18f519614978
child 21375 ae8a112b62d7
common antiquote_setup.ML;
doc-src/IsarImplementation/IsaMakefile
doc-src/IsarImplementation/Thy/base.thy
doc-src/IsarImplementation/Thy/setup.ML
--- a/doc-src/IsarImplementation/IsaMakefile	Wed Nov 15 11:33:59 2006 +0100
+++ b/doc-src/IsarImplementation/IsaMakefile	Wed Nov 15 15:36:09 2006 +0100
@@ -23,7 +23,7 @@
 
 $(LOG)/Pure-Thy.gz: Thy/ROOT.ML Thy/base.thy Thy/integration.thy Thy/isar.thy \
   Thy/locale.thy Thy/logic.thy Thy/prelim.thy Thy/proof.thy Thy/tactic.thy \
-  Thy/ML.thy Thy/setup.ML
+  Thy/ML.thy ../antiquote_setup.ML
 	@$(USEDIR) Pure Thy
 
 
--- a/doc-src/IsarImplementation/Thy/base.thy	Wed Nov 15 11:33:59 2006 +0100
+++ b/doc-src/IsarImplementation/Thy/base.thy	Wed Nov 15 15:36:09 2006 +0100
@@ -3,7 +3,7 @@
 
 theory base
 imports CPure
-uses "setup.ML"
+uses "../../antiquote_setup.ML"
 begin
 
 end
--- a/doc-src/IsarImplementation/Thy/setup.ML	Wed Nov 15 11:33:59 2006 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,47 +0,0 @@
-
-local structure O = IsarOutput
-
-val str_of_source = space_implode " " o map Args.string_of o #2 o #1 o Args.dest_src;
-
-fun ml_val (txt1, "") = "fn _ => (" ^ txt1 ^ ");"
-  | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ ");";
-
-fun ml_type (txt1, "") = "val _ = NONE : (" ^ txt1 ^ ") option;"
-  | ml_type (txt1, txt2) = "val _ = [NONE : (" ^ txt1 ^ ") option, NONE : (" ^ txt2 ^ ") option];";
-  
-fun ml_structure (txt, _) = "functor DUMMY_FUNCTOR() = struct structure DUMMY = " ^ txt ^ " end;"
-
-fun ml_functor _ = "val _ = ();";  (*no check!*)
-
-val verbatim = space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|";
-
-fun index_ml kind ml src ctxt (txt1, txt2) =
-  let
-    val txt = if txt2 = "" then txt1 else
-      if kind = "type"
-        then txt1 ^ " = " ^ txt2
-        else txt1 ^ ": " ^ txt2;
-    val txt' = if kind = "" then txt else kind ^ " " ^ txt;
-    val _ = Context.use_mltext (ml (txt1, txt2)) false (SOME (ProofContext.theory_of ctxt));
-  in
-    "\\indexml" ^ kind ^ enclose "{" "}"
-      (translate_string (fn "_" => "-" | ">" => "$>$" | c => c) txt1) ^
-    ((if ! O.source then str_of_source src else txt')
-    |> (if ! O.quotes then quote else I)
-    |> (if ! O.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
-    else split_lines #> map verbatim #> space_implode "\\isasep\\isanewline%\n"))
-  end;
-
-fun output_verbatim _ _ = split_lines #> map verbatim #> space_implode "\\isasep\\isanewline%\n";
-
-fun arguments x = O.args (Scan.lift (Args.name -- Scan.optional (Args.colon |-- Args.name) "")) x;
-
-val _ = O.add_commands
- [("index_ML", arguments (index_ml "" ml_val)),
-  ("index_ML_type", arguments (index_ml "type" ml_type)),
-  ("index_ML_structure", arguments (index_ml "structure" ml_structure)),
-  ("index_ML_functor", arguments (index_ml "functor" ml_functor)),
-  ("ML_functor", O.args (Scan.lift Args.name) output_verbatim),
-  ("verbatim", O.args (Scan.lift Args.name) output_verbatim)];
-
-in val _ = () end;