# HG changeset patch # User wenzelm # Date 1160353195 -7200 # Node ID 905effde63d90a4a357f4772bb5653129c8dd9b4 # Parent a0034e545c13f3c87c7cba8f4914e7cd112f25bc attribute: Context.mapping; Proof.theorem_i; diff -r a0034e545c13 -r 905effde63d9 src/HOL/Tools/specification_package.ML --- a/src/HOL/Tools/specification_package.ML Mon Oct 09 02:19:54 2006 +0200 +++ b/src/HOL/Tools/specification_package.ML Mon Oct 09 02:19:55 2006 +0200 @@ -87,9 +87,6 @@ |> proc_exprop axiomatic cos |> apsnd standard -fun add_spec x y (context, thm) = - (Context.the_theory context, thm) |> add_specification x y |>> Context.Theory; - (* Collect all intances of constants in term *) @@ -225,13 +222,15 @@ arg |> apsnd Thm.freezeT |> process_all (zip3 alt_names rew_imps frees) end - fun post_proc (context, th) = - post_process (Context.theory_of context, th) |>> Context.Theory; + + fun after_qed [[thm]] = ProofContext.theory (fn thy => + #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm)))); in - IsarThy.theorem_i PureThy.internalK - ("", [add_spec axiomatic (zip3 names cnames overloaded), post_proc]) - (HOLogic.mk_Trueprop ex_prop, []) thy - end + thy + |> ProofContext.init + |> Proof.theorem_i PureThy.internalK NONE after_qed + (SOME "") ("", []) [(("", []), [(HOLogic.mk_Trueprop ex_prop, [])])] + end; (* outer syntax *)