ML context and antiquotations (material from context.ML);
added thm/thms (from pure_thy.ML);
added support for ML antiquotations;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/ml_context.ML Fri Jan 19 22:08:13 2007 +0100
@@ -0,0 +1,195 @@
+(* Title: Pure/Thy/ml_context.ML
+ ID: $Id$
+ Author: Makarius
+
+ML context and antiquotations.
+*)
+
+signature BASIC_ML_CONTEXT =
+sig
+ val the_context: unit -> theory
+ val thm: xstring -> thm
+ val thms: xstring -> thm list
+end;
+
+signature ML_CONTEXT =
+sig
+ include BASIC_ML_CONTEXT
+ val get_context: unit -> Context.generic option
+ val set_context: Context.generic option -> unit
+ val setmp: Context.generic option -> ('a -> 'b) -> 'a -> 'b
+ val the_generic_context: unit -> Context.generic
+ val the_local_context: unit -> Proof.context
+ val pass: Context.generic option -> ('a -> 'b) -> 'a -> 'b * Context.generic option
+ val pass_context: Context.generic -> ('a -> 'b) -> 'a -> 'b * Context.generic
+ val save: ('a -> 'b) -> 'a -> 'b
+ val >> : (theory -> theory) -> unit
+ val add_keywords: string list -> unit
+ val inline_antiq: string -> (Args.T list -> string * Args.T list) -> unit
+ val value_antiq: string -> (Args.T list -> (string * string) * Args.T list) -> unit
+ val use_mltext: string -> bool -> Context.generic option -> unit
+ val use_mltext_update: string -> bool -> Context.generic -> Context.generic
+ val use_let: string -> string -> string -> Context.generic -> Context.generic
+end;
+
+structure ML_Context: ML_CONTEXT =
+struct
+
+(** Implicit ML context **)
+
+local
+ val current_context = ref (NONE: Context.generic option);
+in
+ fun get_context () = ! current_context;
+ fun set_context opt_context = current_context := opt_context;
+ fun setmp opt_context f x = Library.setmp current_context opt_context f x;
+end;
+
+fun the_generic_context () =
+ (case get_context () of
+ SOME context => context
+ | _ => error "Unknown context");
+
+val the_context = Context.theory_of o the_generic_context;
+val the_local_context = Context.proof_of o the_generic_context;
+
+fun pass opt_context f x =
+ setmp opt_context (fn x => let val y = f x in (y, get_context ()) end) x;
+
+fun pass_context context f x =
+ (case pass (SOME context) f x of
+ (y, SOME context') => (y, context')
+ | (_, NONE) => error "Lost context in ML");
+
+fun save f x = setmp (get_context ()) f x;
+
+fun change_theory f =
+ set_context (SOME (Context.map_theory f (the_generic_context ())));
+
+
+
+(** ML antiquotations **)
+
+(* maintain keywords *)
+
+val global_lexicon = ref Scan.empty_lexicon;
+
+fun add_keywords keywords =
+ change global_lexicon (Scan.extend_lexicon (map Symbol.explode keywords));
+
+
+(* maintain commands *)
+
+datatype antiq = Inline of string | Value of string * string;
+val global_parsers = ref (Symtab.empty: (Args.T list -> antiq * Args.T list) Symtab.table);
+
+local
+
+fun add_item kind name scan = change global_parsers (fn tab =>
+ (if not (Symtab.defined tab name) then ()
+ else warning ("Redefined ML antiquotation: " ^ quote name);
+ Symtab.update (name, scan >> kind) tab));
+
+in
+
+val inline_antiq = add_item Inline;
+val value_antiq = add_item Value;
+
+end;
+
+
+(* command syntax *)
+
+fun syntax scan src =
+ #1 (Args.context_syntax "ML antiquotation" (Scan.lift scan) src (the_local_context ()));
+
+fun command src =
+ let val ((name, _), pos) = Args.dest_src src in
+ (case Symtab.lookup (! global_parsers) name of
+ NONE => error ("Unknown ML antiquotation command: " ^ quote name ^ Position.str_of pos)
+ | SOME scan => transform_failure (curry Antiquote.ANTIQUOTE_FAIL (name, pos))
+ (fn () => syntax scan src) ())
+ end;
+
+
+(* outer syntax *)
+
+structure T = OuterLex;
+structure P = OuterParse;
+
+val antiq =
+ P.!!! (P.position P.xname -- P.arguments --| Scan.ahead P.eof)
+ >> (fn ((x, pos), y) => Args.src ((x, y), pos));
+
+
+(* input/output wrappers *)
+
+local
+
+val isabelle_struct = enclose "structure Isabelle =\nstruct\n" "end;";
+val isabelle_struct0 = isabelle_struct "";
+
+fun eval_antiquotes txt =
+ let
+ val ants = Antiquote.scan_antiquotes (txt, Position.line 1);
+
+ fun expand (Antiquote.Text s) names = (("", Symbol.escape s), names)
+ | expand (Antiquote.Antiq x) names =
+ (case command (Antiquote.scan_arguments (! global_lexicon) antiq x) of
+ Inline s => (("", s), names)
+ | Value (a, s) =>
+ let
+ val a' = if ML_Syntax.is_identifier a then a else "val";
+ val ([b], names') = Name.variants [a'] names;
+ in (("val " ^ b ^ " = " ^ s ^ ";\n", "Isabelle." ^ b), names') end);
+
+ val (decls, body) =
+ fold_map expand ants ML_Syntax.reserved
+ |> #1 |> split_list |> pairself implode;
+ in (isabelle_struct decls, body) end;
+
+fun eval verbose txt =
+ Output.ML_errors (use_text Output.ml_output verbose) txt;
+
+in
+
+fun eval_wrapper verbose txt =
+ let val (txt1, txt2) = eval_antiquotes txt
+ in eval false txt1; eval verbose txt2; eval false isabelle_struct0 end;
+
+end;
+
+
+
+(** ML evaluation **)
+
+fun use_mltext txt verbose opt_context = setmp opt_context (fn () => eval_wrapper verbose txt) ();
+fun use_mltext_update txt verbose context = #2 (pass_context context (eval_wrapper verbose) txt);
+
+fun use_context txt = use_mltext_update
+ ("ML_Context.set_context (SOME ((" ^ txt ^ ") (ML_Context.the_generic_context ())));") false;
+
+fun use_let bind body txt =
+ use_context ("let " ^ bind ^ " = " ^ txt ^ " in\n" ^ body ^ " end");
+
+
+
+(** implicit fact references **)
+
+fun thm name = ProofContext.get_thm (the_local_context ()) (Name name);
+fun thms name = ProofContext.get_thms (the_local_context ()) (Name name);
+
+val _ = value_antiq "thm"
+ (Args.name >> (fn s => (s, "ML_Context.thm " ^ ML_Syntax.print_string s)));
+val _ = value_antiq "thms"
+ (Args.name >> (fn s => (s, "ML_Context.thms " ^ ML_Syntax.print_string s)));
+
+
+(*final declarations of this structure!*)
+nonfix >>;
+fun >> f = change_theory f;
+
+end;
+
+structure Basic_ML_Context: BASIC_ML_CONTEXT = ML_Context;
+open Basic_ML_Context;