# HG changeset patch # User wenzelm # Date 1221679658 -7200 # Node ID 17f6aa02ded3fb3c46156bdcba555a5818451017 # Parent ed959a0f650bbef9c8c59591229b26cd5b101219 simplified ML_Context.eval_in -- expect immutable Proof.context value; diff -r ed959a0f650b -r 17f6aa02ded3 doc-src/antiquote_setup.ML --- a/doc-src/antiquote_setup.ML Wed Sep 17 21:27:36 2008 +0200 +++ b/doc-src/antiquote_setup.ML Wed Sep 17 21:27:38 2008 +0200 @@ -75,7 +75,7 @@ else txt1 ^ ": " ^ txt2; val txt' = if kind = "" then txt else kind ^ " " ^ txt; val _ = writeln (ml (txt1, txt2)); - val _ = ML_Context.eval_in (SOME (Context.Proof ctxt)) false Position.none (ml (txt1, txt2)); + val _ = ML_Context.eval_in (SOME ctxt) false Position.none (ml (txt1, txt2)); in "\\indexml" ^ kind ^ enclose "{" "}" (clean_string txt1) ^ ((if ! O.source then str_of_source src else txt') diff -r ed959a0f650b -r 17f6aa02ded3 src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML Wed Sep 17 21:27:36 2008 +0200 +++ b/src/Pure/Isar/isar_cmd.ML Wed Sep 17 21:27:38 2008 +0200 @@ -303,7 +303,8 @@ (* diagnostic ML evaluation *) fun ml_diag verbose (txt, pos) = Toplevel.keep (fn state => - (ML_Context.eval_in (try Toplevel.generic_theory_of state) verbose pos txt)); + (ML_Context.eval_in + (Option.map Context.proof_of (try Toplevel.generic_theory_of state)) verbose pos txt)); (* current working directory *) diff -r ed959a0f650b -r 17f6aa02ded3 src/Pure/Thy/thy_output.ML --- a/src/Pure/Thy/thy_output.ML Wed Sep 17 21:27:36 2008 +0200 +++ b/src/Pure/Thy/thy_output.ML Wed Sep 17 21:27:38 2008 +0200 @@ -507,7 +507,7 @@ fun ml_struct txt = "functor DUMMY_FUNCTOR() = struct structure DUMMY = " ^ txt ^ " end;" fun output_ml ml src ctxt (txt, pos) = - (ML_Context.eval_in (SOME (Context.Proof ctxt)) false pos (ml txt); + (ML_Context.eval_in (SOME ctxt) false pos (ml txt); (if ! source then str_of_source src else SymbolPos.content (SymbolPos.explode (txt, pos))) |> (if ! quotes then quote else I) |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"