# HG changeset patch # User wenzelm # Date 1366297621 -7200 # Node ID 9e7d1c13956963ee38b8c20e6f678b3b3e538447 # Parent 19b47bfac6ef4d278868aaaa441a2b61bc9a880f simplifier uses proper Proof.context instead of historic type simpset; diff -r 19b47bfac6ef -r 9e7d1c139569 NEWS --- a/NEWS Tue Apr 16 17:54:14 2013 +0200 +++ b/NEWS Thu Apr 18 17:07:01 2013 +0200 @@ -134,6 +134,17 @@ addEs, addDs etc. Note that claset_of and put_claset allow to manage clasets separately from the context. +* Simplifier tactics and tools use proper Proof.context instead of +historic type simpset. Old-style declarations like addsimps, +addsimprocs etc. operate directly on Proof.context. Raw type simpset +retains its use as snapshot of the main Simplifier context, using +simpset_of and put_simpset on Proof.context. INCOMPATIBILITY -- port +old tools by making them depend on (ctxt : Proof.context) instead of +(ss : simpset), then turn (simpset_of ctxt) into ctxt. + +* Discontinued obsolete ML antiquotations @{claset} and @{simpset}. +INCOMPATIBILITY, use @{context} instead. + *** System *** diff -r 19b47bfac6ef -r 9e7d1c139569 src/CCL/CCL.thy --- a/src/CCL/CCL.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/CCL/CCL.thy Thu Apr 18 17:07:01 2013 +0200 @@ -206,7 +206,7 @@ in CHANGED (REPEAT (ares_tac [@{thm iffI}, @{thm allI}, @{thm conjI}] i ORELSE eresolve_tac inj_lemmas i ORELSE - asm_simp_tac (simpset_of ctxt addsimps rews) i)) + asm_simp_tac (ctxt addsimps rews) i)) end; *} @@ -281,7 +281,7 @@ Goal.prove_global thy [] [] (Syntax.read_prop ctxt s) (fn _ => rewrite_goals_tac defs THEN - simp_tac (simpset_of ctxt addsimps (rls @ inj_rls)) 1) + simp_tac (ctxt addsimps (rls @ inj_rls)) 1) in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end fun mkall_dstnct_thms ctxt defs i_rls xss = maps (mk_dstnct_thms ctxt defs i_rls) xss @@ -422,7 +422,7 @@ REPEAT (rtac @{thm notI} 1 THEN dtac @{thm case_pocong} 1 THEN etac @{thm rev_mp} 5 THEN - ALLGOALS (simp_tac @{simpset}) THEN + ALLGOALS (simp_tac @{context}) THEN REPEAT (resolve_tac [@{thm po_refl}, @{thm npo_lam_bot}] 1)) *}) lemmas npo_rls = npo_pair_lam npo_lam_pair npo_rls1 diff -r 19b47bfac6ef -r 9e7d1c139569 src/CCL/Term.thy --- a/src/CCL/Term.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/CCL/Term.thy Thu Apr 18 17:07:01 2013 +0200 @@ -204,7 +204,7 @@ method_setup beta_rl = {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (CHANGED o - simp_tac (simpset_of ctxt addsimps @{thms rawBs} setloop (stac @{thm letrecB})))) + simp_tac (ctxt addsimps @{thms rawBs} setloop (stac @{thm letrecB})))) *} lemma ifBtrue: "if true then t else u = t" diff -r 19b47bfac6ef -r 9e7d1c139569 src/CCL/Type.thy --- a/src/CCL/Type.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/CCL/Type.thy Thu Apr 18 17:07:01 2013 +0200 @@ -130,7 +130,7 @@ SUBPROOF (fn {context = ctxt, prems = major :: prems, ...} => resolve_tac ([major] RL top_crls) 1 THEN REPEAT_SOME (eresolve_tac (crls @ [@{thm exE}, @{thm bexE}, @{thm conjE}, @{thm disjE}])) THEN - ALLGOALS (asm_simp_tac (simpset_of ctxt)) THEN + ALLGOALS (asm_simp_tac ctxt) THEN ALLGOALS (ares_tac (prems RL [@{thm lem}]) ORELSE' etac @{thm bspec}) THEN safe_tac (ctxt addSIs prems)) *} @@ -415,7 +415,7 @@ ML {* fun genIs_tac ctxt genXH gen_mono = rtac (genXH RS @{thm iffD2}) THEN' - simp_tac (simpset_of ctxt) THEN' + simp_tac ctxt THEN' TRY o fast_tac (ctxt addIs [genXH RS @{thm iffD2}, gen_mono RS @{thm coinduct3_mono_lemma} RS @{thm lfpI}]) *} @@ -498,7 +498,7 @@ SELECT_GOAL (TRY (safe_tac ctxt) THEN resolve_tac ((rews @ [@{thm refl}]) RL ((rews @ [@{thm refl}]) RL [@{thm ssubst_pair}])) i THEN - ALLGOALS (simp_tac (simpset_of ctxt)) THEN + ALLGOALS (simp_tac ctxt) THEN ALLGOALS EQgen_raw_tac) i *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/Doc/Codegen/Further.thy --- a/src/Doc/Codegen/Further.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Doc/Codegen/Further.thy Thu Apr 18 17:07:01 2013 +0200 @@ -255,8 +255,8 @@ @{index_ML Code.read_const: "theory -> string -> string"} \\ @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\ @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\ - @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\ - @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\ + @{index_ML Code_Preproc.map_pre: "(Proof.context -> Proof.context) -> theory -> theory"} \\ + @{index_ML Code_Preproc.map_post: "(Proof.context -> Proof.context) -> theory -> theory"} \\ @{index_ML Code_Preproc.add_functrans: " string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory"} \\ diff -r 19b47bfac6ef -r 9e7d1c139569 src/Doc/IsarImplementation/Isar.thy --- a/src/Doc/IsarImplementation/Isar.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Doc/IsarImplementation/Isar.thy Thu Apr 18 17:07:01 2013 +0200 @@ -385,7 +385,7 @@ Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (fn i => CHANGED (asm_full_simp_tac - (HOL_basic_ss addsimps thms) i))) + (put_simpset HOL_basic_ss ctxt addsimps thms) i))) *} "rewrite subgoal by given rules" text {* The concrete syntax wrapping of @{command method_setup} always @@ -424,7 +424,7 @@ SIMPLE_METHOD (CHANGED (ALLGOALS (asm_full_simp_tac - (HOL_basic_ss addsimps thms))))) + (put_simpset HOL_basic_ss ctxt addsimps thms))))) *} "rewrite all subgoals by given rules" notepad @@ -458,7 +458,8 @@ Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (fn i => CHANGED (asm_full_simp_tac - (HOL_basic_ss addsimps (thms @ My_Simps.get ctxt)) i))) + (put_simpset HOL_basic_ss ctxt + addsimps (thms @ My_Simps.get ctxt)) i))) *} "rewrite subgoal by given rules and my_simp rules from the context" text {* diff -r 19b47bfac6ef -r 9e7d1c139569 src/Doc/IsarRef/Generic.thy --- a/src/Doc/IsarRef/Generic.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Doc/IsarRef/Generic.thy Thu Apr 18 17:07:01 2013 +0200 @@ -996,9 +996,9 @@ text {* \begin{mldecls} - @{index_ML Simplifier.set_subgoaler: "(simpset -> int -> tactic) -> - simpset -> simpset"} \\ - @{index_ML Simplifier.prems_of: "simpset -> thm list"} \\ + @{index_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) -> + Proof.context -> Proof.context"} \\ + @{index_ML Simplifier.prems_of: "Proof.context -> thm list"} \\ \end{mldecls} The subgoaler is the tactic used to solve subgoals arising out of @@ -1010,14 +1010,12 @@ \begin{description} - \item @{ML Simplifier.set_subgoaler}~@{text "ss tac"} sets the - subgoaler of simpset @{text "ss"} to @{text "tac"}. The tactic will - be applied to the context of the running Simplifier instance, - expressed as a simpset. + \item @{ML Simplifier.set_subgoaler}~@{text "tac ctxt"} sets the + subgoaler of the context to @{text "tac"}. The tactic will + be applied to the context of the running Simplifier instance. - \item @{ML Simplifier.prems_of}~@{text "ss"} retrieves the current - set of premises from simpset @{text "ss"} that represents the - context of the running Simplifier. This may be non-empty only if + \item @{ML Simplifier.prems_of}~@{text "ctxt"} retrieves the current + set of premises from the context. This may be non-empty only if the Simplifier has been told to utilize local assumptions in the first place (cf.\ the options in \secref{sec:simp-meth}). @@ -1027,10 +1025,10 @@ *} ML {* - fun subgoaler_tac ss = + fun subgoaler_tac ctxt = assume_tac ORELSE' - resolve_tac (Simplifier.prems_of ss) ORELSE' - asm_simp_tac ss + resolve_tac (Simplifier.prems_of ctxt) ORELSE' + asm_simp_tac ctxt *} text {* This tactic first tries to solve the subgoal by assumption or @@ -1043,12 +1041,12 @@ text {* \begin{mldecls} @{index_ML_type solver} \\ - @{index_ML Simplifier.mk_solver: "string -> (simpset -> int -> tactic) -> - solver"} \\ - @{index_ML_op setSolver: "simpset * solver -> simpset"} \\ - @{index_ML_op addSolver: "simpset * solver -> simpset"} \\ - @{index_ML_op setSSolver: "simpset * solver -> simpset"} \\ - @{index_ML_op addSSolver: "simpset * solver -> simpset"} \\ + @{index_ML Simplifier.mk_solver: "string -> + (Proof.context -> int -> tactic) -> solver"} \\ + @{index_ML_op setSolver: "Proof.context * solver -> Proof.context"} \\ + @{index_ML_op addSolver: "Proof.context * solver -> Proof.context"} \\ + @{index_ML_op setSSolver: "Proof.context * solver -> Proof.context"} \\ + @{index_ML_op addSSolver: "Proof.context * solver -> Proof.context"} \\ \end{mldecls} A solver is a tactic that attempts to solve a subgoal after @@ -1085,24 +1083,24 @@ "tac"} into a solver; the @{text "name"} is only attached as a comment and has no further significance. - \item @{text "ss setSSolver solver"} installs @{text "solver"} as - the safe solver of @{text "ss"}. + \item @{text "ctxt setSSolver solver"} installs @{text "solver"} as + the safe solver of @{text "ctxt"}. - \item @{text "ss addSSolver solver"} adds @{text "solver"} as an + \item @{text "ctxt addSSolver solver"} adds @{text "solver"} as an additional safe solver; it will be tried after the solvers which had - already been present in @{text "ss"}. + already been present in @{text "ctxt"}. - \item @{text "ss setSolver solver"} installs @{text "solver"} as the - unsafe solver of @{text "ss"}. + \item @{text "ctxt setSolver solver"} installs @{text "solver"} as the + unsafe solver of @{text "ctxt"}. - \item @{text "ss addSolver solver"} adds @{text "solver"} as an + \item @{text "ctxt addSolver solver"} adds @{text "solver"} as an additional unsafe solver; it will be tried after the solvers which - had already been present in @{text "ss"}. + had already been present in @{text "ctxt"}. \end{description} - \medskip The solver tactic is invoked with a simpset that represents - the context of the running Simplifier. Further simpset operations + \medskip The solver tactic is invoked with the context of the + running Simplifier. Further operations may be used to retrieve relevant information, such as the list of local Simplifier premises via @{ML Simplifier.prems_of} --- this list may be non-empty only if the Simplifier runs in a mode that @@ -1144,14 +1142,18 @@ text {* \begin{mldecls} - @{index_ML_op setloop: "simpset * (int -> tactic) -> simpset"} \\ - @{index_ML_op setloop': "simpset * (simpset -> int -> tactic) -> simpset"} \\ - @{index_ML_op addloop: "simpset * (string * (int -> tactic)) -> simpset"} \\ - @{index_ML_op addloop': "simpset * (string * (simpset -> int -> tactic)) - -> simpset"} \\ - @{index_ML_op delloop: "simpset * string -> simpset"} \\ - @{index_ML_op Splitter.add_split: "thm -> simpset -> simpset"} \\ - @{index_ML_op Splitter.del_split: "thm -> simpset -> simpset"} \\ + @{index_ML_op setloop: "Proof.context * + (int -> tactic) -> Proof.context"} \\ + @{index_ML_op setloop': "Proof.context * + (Proof.context -> int -> tactic) -> Proof.context"} \\ + @{index_ML_op addloop: "Proof.context * + (string * (int -> tactic)) -> Proof.context"} \\ + @{index_ML_op addloop': "Proof.context * + (string * (Proof.context -> int -> tactic)) + -> Proof.context"} \\ + @{index_ML_op delloop: "Proof.context * string -> Proof.context"} \\ + @{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ + @{index_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\ \end{mldecls} The looper is a list of tactics that are applied after @@ -1166,28 +1168,26 @@ \begin{description} - \item @{text "ss setloop tac"} installs @{text "tac"} as the only - looper tactic of @{text "ss"}. The variant @{text "setloop'"} - allows the tactic to depend on the running Simplifier context, which - is represented as simpset. + \item @{text "ctxt setloop tac"} installs @{text "tac"} as the only + looper tactic of @{text "ctxt"}. The variant @{text "setloop'"} + allows the tactic to depend on the running Simplifier context. - \item @{text "ss addloop (name, tac)"} adds @{text "tac"} as an + \item @{text "ctxt addloop (name, tac)"} adds @{text "tac"} as an additional looper tactic with name @{text "name"}, which is significant for managing the collection of loopers. The tactic will be tried after the looper tactics that had already been present in - @{text "ss"}. The variant @{text "addloop'"} allows the tactic to - depend on the running Simplifier context, which is represented as - simpset. + @{text "ctxt"}. The variant @{text "addloop'"} allows the tactic to + depend on the running Simplifier context. - \item @{text "ss delloop name"} deletes the looper tactic that was - associated with @{text "name"} from @{text "ss"}. + \item @{text "ctxt delloop name"} deletes the looper tactic that was + associated with @{text "name"} from @{text "ctxt"}. - \item @{ML Splitter.add_split}~@{text "thm ss"} adds split tactics - for @{text "thm"} as additional looper tactics of @{text "ss"}. + \item @{ML Splitter.add_split}~@{text "thm ctxt"} adds split tactics + for @{text "thm"} as additional looper tactics of @{text "ctxt"}. - \item @{ML Splitter.del_split}~@{text "thm ss"} deletes the split + \item @{ML Splitter.del_split}~@{text "thm ctxt"} deletes the split tactic corresponding to @{text thm} from the looper tactics of - @{text "ss"}. + @{text "ctxt"}. \end{description} @@ -1817,16 +1817,16 @@ @{index_ML_op addSWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{index_ML_op addSbefore: "Proof.context * - (string * (int -> tactic)) -> Proof.context"} \\ + (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{index_ML_op addSafter: "Proof.context * - (string * (int -> tactic)) -> Proof.context"} \\ + (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{index_ML_op delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{index_ML_op addWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{index_ML_op addbefore: "Proof.context * - (string * (int -> tactic)) -> Proof.context"} \\ + (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{index_ML_op addafter: "Proof.context * - (string * (int -> tactic)) -> Proof.context"} \\ + (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{index_ML_op delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{index_ML addSss: "Proof.context -> Proof.context"} \\ @{index_ML addss: "Proof.context -> Proof.context"} \\ diff -r 19b47bfac6ef -r 9e7d1c139569 src/Doc/IsarRef/ML_Tactic.thy --- a/src/Doc/IsarRef/ML_Tactic.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Doc/IsarRef/ML_Tactic.thy Thu Apr 18 17:07:01 2013 +0200 @@ -88,12 +88,12 @@ \medskip \begin{tabular}{lll} - @{ML asm_full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp} \\ - @{ML ALLGOALS}~(@{ML asm_full_simp_tac}~@{text "@{simpset}"}) & & @{method simp_all} \\[0.5ex] - @{ML simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm)"} \\ - @{ML asm_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_simp)"} \\ - @{ML full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_use)"} \\ - @{ML asm_lr_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(asm_lr)"} \\ + @{ML asm_full_simp_tac}~@{text "@{context} 1"} & & @{method simp} \\ + @{ML ALLGOALS}~(@{ML asm_full_simp_tac}~@{text "@{context}"}) & & @{method simp_all} \\[0.5ex] + @{ML simp_tac}~@{text "@{context} 1"} & & @{method simp}~@{text "(no_asm)"} \\ + @{ML asm_simp_tac}~@{text "@{context} 1"} & & @{method simp}~@{text "(no_asm_simp)"} \\ + @{ML full_simp_tac}~@{text "@{context} 1"} & & @{method simp}~@{text "(no_asm_use)"} \\ + @{ML asm_lr_simp_tac}~@{text "@{context} 1"} & & @{method simp}~@{text "(asm_lr)"} \\ \end{tabular} \medskip *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Doc/Tutorial/Protocol/Message.thy Thu Apr 18 17:07:01 2013 +0200 @@ -840,12 +840,12 @@ impOfSubs Fake_parts_insert] THEN' eresolve_tac [asm_rl, @{thm synth.Inj}]; -fun Fake_insert_simp_tac ss i = - REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i; +fun Fake_insert_simp_tac ctxt i = + REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ctxt i; fun atomic_spy_analz_tac ctxt = SELECT_GOAL - (Fake_insert_simp_tac (simpset_of ctxt) 1 THEN + (Fake_insert_simp_tac ctxt 1 THEN IF_UNSOLVED (Blast.depth_tac (ctxt addIs [analz_insertI, impOfSubs analz_subset_parts]) 4 1)); fun spy_analz_tac ctxt i = @@ -856,7 +856,7 @@ (REPEAT o CHANGED) (res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1), (*...allowing further simplifications*) - simp_tac (simpset_of ctxt) 1, + simp_tac ctxt 1, REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])), DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i); *} @@ -914,7 +914,7 @@ "for debugging spy_analz" method_setup Fake_insert_simp = {* - Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o simpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac) *} "for debugging spy_analz" diff -r 19b47bfac6ef -r 9e7d1c139569 src/Doc/Tutorial/Protocol/Public.thy --- a/src/Doc/Tutorial/Protocol/Public.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Doc/Tutorial/Protocol/Public.thy Thu Apr 18 17:07:01 2013 +0200 @@ -159,7 +159,7 @@ ML {* fun possibility_tac ctxt = REPEAT (*omit used_Says so that Nonces start from different traces!*) - (ALLGOALS (simp_tac (simpset_of ctxt delsimps [used_Says])) + (ALLGOALS (simp_tac (ctxt delsimps [used_Says])) THEN REPEAT_FIRST (eq_assume_tac ORELSE' resolve_tac [refl, conjI, @{thm Nonce_supply}])); diff -r 19b47bfac6ef -r 9e7d1c139569 src/FOL/FOL.thy --- a/src/FOL/FOL.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/FOL/FOL.thy Thu Apr 18 17:07:01 2013 +0200 @@ -331,16 +331,20 @@ ML {* (*intuitionistic simprules only*) val IFOL_ss = - FOL_basic_ss + put_simpset FOL_basic_ss @{context} addsimps @{thms meta_simps IFOL_simps int_ex_simps int_all_simps} addsimprocs [@{simproc defined_All}, @{simproc defined_Ex}] - |> Simplifier.add_cong @{thm imp_cong}; + |> Simplifier.add_cong @{thm imp_cong} + |> simpset_of; (*classical simprules too*) -val FOL_ss = IFOL_ss addsimps @{thms cla_simps cla_ex_simps cla_all_simps}; +val FOL_ss = + put_simpset IFOL_ss @{context} + addsimps @{thms cla_simps cla_ex_simps cla_all_simps} + |> simpset_of; *} -setup {* Simplifier.map_simpset_global (K FOL_ss) *} +setup {* map_theory_simpset (put_simpset FOL_ss) *} setup "Simplifier.method_setup Splitter.split_modifiers" setup Splitter.setup diff -r 19b47bfac6ef -r 9e7d1c139569 src/FOL/ex/Classical.thy --- a/src/FOL/ex/Classical.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/FOL/ex/Classical.thy Thu Apr 18 17:07:01 2013 +0200 @@ -300,7 +300,7 @@ (*Other proofs: Can use auto, which cheats by using rewriting! Deepen_tac alone requires 253 secs. Or - by (mini_tac 1 THEN Deepen_tac 5 1) *) + by (mini_tac @{context} 1 THEN Deepen_tac 5 1) *) text{*44*} lemma "(\x. f(x) --> (\y. g(y) & h(x,y) & (\y. g(y) & ~ h(x,y)))) & diff -r 19b47bfac6ef -r 9e7d1c139569 src/FOL/ex/Miniscope.thy --- a/src/FOL/ex/Miniscope.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/FOL/ex/Miniscope.thy Thu Apr 18 17:07:01 2013 +0200 @@ -64,8 +64,8 @@ lemmas mini_simps = demorgans nnf_simps ex_simps all_simps ML {* -val mini_ss = @{simpset} addsimps @{thms mini_simps}; -val mini_tac = rtac @{thm ccontr} THEN' asm_full_simp_tac mini_ss; +val mini_ss = simpset_of (@{context} addsimps @{thms mini_simps}); +fun mini_tac ctxt = rtac @{thm ccontr} THEN' asm_full_simp_tac (put_simpset mini_ss ctxt); *} end diff -r 19b47bfac6ef -r 9e7d1c139569 src/FOL/simpdata.ML --- a/src/FOL/simpdata.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/FOL/simpdata.ML Thu Apr 18 17:07:01 2013 +0200 @@ -26,8 +26,8 @@ (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}])); (*Congruence rules for = or <-> (instead of ==)*) -fun mk_meta_cong ss rl = - Drule.zero_var_indexes (mk_meta_eq (mk_meta_prems (Simplifier.the_context ss) rl)) +fun mk_meta_cong ctxt rl = + Drule.zero_var_indexes (mk_meta_eq (mk_meta_prems ctxt rl)) handle THM _ => error("Premises and conclusion of congruence rules must use =-equality or <->"); @@ -48,7 +48,7 @@ | _ => [th]) in atoms end; -fun mksimps pairs (_: simpset) = map mk_eq o mk_atomize pairs o gen_all; +fun mksimps pairs (_: Proof.context) = map mk_eq o mk_atomize pairs o gen_all; (** make simplification procedures for quantifier elimination **) @@ -106,25 +106,25 @@ val triv_rls = [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl}, @{thm notFalseI}]; -fun unsafe_solver ss = - FIRST' [resolve_tac (triv_rls @ Simplifier.prems_of ss), atac, etac @{thm FalseE}]; +fun unsafe_solver ctxt = + FIRST' [resolve_tac (triv_rls @ Simplifier.prems_of ctxt), atac, etac @{thm FalseE}]; (*No premature instantiation of variables during simplification*) -fun safe_solver ss = - FIRST' [match_tac (triv_rls @ Simplifier.prems_of ss), eq_assume_tac, ematch_tac @{thms FalseE}]; +fun safe_solver ctxt = + FIRST' [match_tac (triv_rls @ Simplifier.prems_of ctxt), eq_assume_tac, ematch_tac @{thms FalseE}]; (*No simprules, but basic infastructure for simplification*) val FOL_basic_ss = - Simplifier.global_context @{theory} empty_ss + empty_simpset @{context} setSSolver (mk_solver "FOL safe" safe_solver) setSolver (mk_solver "FOL unsafe" unsafe_solver) |> Simplifier.set_subgoaler asm_simp_tac |> Simplifier.set_mksimps (mksimps mksimps_pairs) - |> Simplifier.set_mkcong mk_meta_cong; + |> Simplifier.set_mkcong mk_meta_cong + |> simpset_of; -fun unfold_tac ths = - let val ss0 = Simplifier.clear_ss FOL_basic_ss addsimps ths - in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; +fun unfold_tac ths ctxt = + ALLGOALS (full_simp_tac (clear_simpset (put_simpset FOL_basic_ss ctxt) addsimps ths)); (*** integration of simplifier with classical reasoner ***) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Algebra/ringsimp.ML --- a/src/HOL/Algebra/ringsimp.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Algebra/ringsimp.ML Thu Apr 18 17:07:01 2013 +0200 @@ -42,16 +42,16 @@ (** Method **) -fun struct_tac ((s, ts), simps) = +fun struct_tac ctxt ((s, ts), simps) = let val ops = map (fst o Term.strip_comb) ts; fun ord (Const (a, _)) = find_index (fn (Const (b, _)) => a=b | _ => false) ops | ord (Free (a, _)) = find_index (fn (Free (b, _)) => a=b | _ => false) ops; fun less (a, b) = (Term_Ord.term_lpo ord (a, b) = LESS); - in asm_full_simp_tac (HOL_ss addsimps simps |> Simplifier.set_termless less) end; + in asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps simps |> Simplifier.set_termless less) end; fun algebra_tac ctxt = - EVERY' (map (fn s => TRY o struct_tac s) (Data.get (Context.Proof ctxt))); + EVERY' (map (fn s => TRY o struct_tac ctxt s) (Data.get (Context.Proof ctxt))); (** Attribute **) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/Message.thy Thu Apr 18 17:07:01 2013 +0200 @@ -863,12 +863,12 @@ impOfSubs @{thm Fake_parts_insert}] THEN' eresolve_tac [asm_rl, @{thm synth.Inj}]; -fun Fake_insert_simp_tac ss i = - REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i; +fun Fake_insert_simp_tac ctxt i = + REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ctxt i; fun atomic_spy_analz_tac ctxt = SELECT_GOAL - (Fake_insert_simp_tac (simpset_of ctxt) 1 THEN + (Fake_insert_simp_tac ctxt 1 THEN IF_UNSOLVED (Blast.depth_tac (ctxt addIs [@{thm analz_insertI}, impOfSubs @{thm analz_subset_parts}]) 4 1)); @@ -881,7 +881,7 @@ (REPEAT o CHANGED) (res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1), (*...allowing further simplifications*) - simp_tac (simpset_of ctxt) 1, + simp_tac ctxt 1, REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])), DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i); *} @@ -933,7 +933,7 @@ "for debugging spy_analz" method_setup Fake_insert_simp = {* - Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o simpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac) *} "for debugging spy_analz" end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/OtwayReesBella.thy --- a/src/HOL/Auth/OtwayReesBella.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/OtwayReesBella.thy Thu Apr 18 17:07:01 2013 +0200 @@ -237,10 +237,11 @@ structure OtwayReesBella = struct -val analz_image_freshK_ss = - @{simpset} delsimps [image_insert, image_Un] +val analz_image_freshK_ss = + simpset_of + (@{context} delsimps [image_insert, image_Un] delsimps [@{thm imp_disjL}] (*reduces blow-up*) - addsimps @{thms analz_image_freshK_simps} + addsimps @{thms analz_image_freshK_simps}) end *} @@ -251,7 +252,7 @@ (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac @{thm analz_image_freshCryptK_lemma}), ALLGOALS (asm_simp_tac - (Simplifier.context ctxt OtwayReesBella.analz_image_freshK_ss))]))) *} + (put_simpset OtwayReesBella.analz_image_freshK_ss ctxt))]))) *} "for proving useful rewrite rule" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/Public.thy Thu Apr 18 17:07:01 2013 +0200 @@ -405,14 +405,16 @@ structure Public = struct -val analz_image_freshK_ss = @{simpset} delsimps [image_insert, image_Un] - delsimps [@{thm imp_disjL}] (*reduces blow-up*) - addsimps @{thms analz_image_freshK_simps} +val analz_image_freshK_ss = + simpset_of + (@{context} delsimps [image_insert, image_Un] + delsimps [@{thm imp_disjL}] (*reduces blow-up*) + addsimps @{thms analz_image_freshK_simps}) (*Tactic for possibility theorems*) fun possibility_tac ctxt = REPEAT (*omit used_Says so that Nonces start from different traces!*) - (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}])) + (ALLGOALS (simp_tac (ctxt delsimps [@{thm used_Says}])) THEN REPEAT_FIRST (eq_assume_tac ORELSE' resolve_tac [refl, conjI, @{thm Nonce_supply}])) @@ -421,7 +423,7 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) @@ -433,7 +435,7 @@ (SIMPLE_METHOD (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}), - ALLGOALS (asm_simp_tac (Simplifier.context ctxt Public.analz_image_freshK_ss))]))) *} + ALLGOALS (asm_simp_tac (put_simpset Public.analz_image_freshK_ss ctxt))]))) *} "for proving the Session Key Compromise theorem" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/Shared.thy --- a/src/HOL/Auth/Shared.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/Shared.thy Thu Apr 18 17:07:01 2013 +0200 @@ -200,7 +200,7 @@ such as Nonce ?N \ used evs that match Nonce_supply*) fun possibility_tac ctxt = (REPEAT - (ALLGOALS (simp_tac (simpset_of ctxt + (ALLGOALS (simp_tac (ctxt delsimps [@{thm used_Says}, @{thm used_Notes}, @{thm used_Gets}] setSolver safe_solver)) THEN @@ -211,15 +211,16 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) val analz_image_freshK_ss = - @{simpset} delsimps [image_insert, image_Un] + simpset_of + (@{context} delsimps [image_insert, image_Un] delsimps [@{thm imp_disjL}] (*reduces blow-up*) - addsimps @{thms analz_image_freshK_simps} + addsimps @{thms analz_image_freshK_simps}) end *} @@ -238,7 +239,7 @@ (SIMPLE_METHOD (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}), - ALLGOALS (asm_simp_tac (Simplifier.context ctxt Shared.analz_image_freshK_ss))]))) *} + ALLGOALS (asm_simp_tac (put_simpset Shared.analz_image_freshK_ss ctxt))]))) *} "for proving the Session Key Compromise theorem" method_setup possibility = {* diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/Smartcard/ShoupRubin.thy --- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Thu Apr 18 17:07:01 2013 +0200 @@ -819,7 +819,7 @@ (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}), - ALLGOALS (asm_simp_tac (Simplifier.context ctxt Smartcard.analz_image_freshK_ss + ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt addsimps [@{thm knows_Spy_Inputs_secureM_sr_Spy}, @{thm knows_Spy_Outpts_secureM_sr_Spy}, @{thm shouprubin_assumes_securemeans}, diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/Smartcard/ShoupRubinBella.thy --- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Thu Apr 18 17:07:01 2013 +0200 @@ -828,7 +828,7 @@ (SIMPLE_METHOD (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}), - ALLGOALS (asm_simp_tac (Simplifier.context ctxt Smartcard.analz_image_freshK_ss + ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt addsimps [@{thm knows_Spy_Inputs_secureM_srb_Spy}, @{thm knows_Spy_Outpts_secureM_srb_Spy}, @{thm shouprubin_assumes_securemeans}, diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Thu Apr 18 17:07:01 2013 +0200 @@ -369,9 +369,9 @@ such as Nonce ?N \ used evs that match Nonce_supply*) fun possibility_tac ctxt = (REPEAT - (ALLGOALS (simp_tac (simpset_of ctxt + (ALLGOALS (simp_tac (ctxt delsimps [@{thm used_Says}, @{thm used_Notes}, @{thm used_Gets}, - @{thm used_Inputs}, @{thm used_C_Gets}, @{thm used_Outpts}, @{thm used_A_Gets}] + @{thm used_Inputs}, @{thm used_C_Gets}, @{thm used_Outpts}, @{thm used_A_Gets}] setSolver safe_solver)) THEN REPEAT_FIRST (eq_assume_tac ORELSE' @@ -381,14 +381,15 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) val analz_image_freshK_ss = - @{simpset} delsimps [image_insert, image_Un] + simpset_of + (@{context} delsimps [image_insert, image_Un] delsimps [@{thm imp_disjL}] (*reduces blow-up*) - addsimps @{thms analz_image_freshK_simps} + addsimps @{thms analz_image_freshK_simps}) end *} @@ -405,7 +406,7 @@ (SIMPLE_METHOD (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]), REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}), - ALLGOALS (asm_simp_tac (Simplifier.context ctxt Smartcard.analz_image_freshK_ss))]))) *} + ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt))]))) *} "for proving the Session Key Compromise theorem" method_setup possibility = {* diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML --- a/src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML Thu Apr 18 17:07:01 2013 +0200 @@ -46,7 +46,7 @@ val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0; val sum_prod_thms_rel = @{thms prod_rel_simp sum_rel_simps}; -val ss_if_True_False = ss_only @{thms if_True if_False}; +val ss_if_True_False = simpset_of (ss_only @{thms if_True if_False} @{context}); fun mk_proj T k = let val binders = binder_types T in @@ -115,7 +115,8 @@ fun mk_corec_like_tac corec_like_defs map_comps'' map_comp's map_ids'' map_if_distribs ctor_dtor_corec_like pre_map_def ctr_def ctxt = unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN - (rtac (ctor_dtor_corec_like RS trans) THEN' asm_simp_tac ss_if_True_False) 1 THEN_MAYBE + (rtac (ctor_dtor_corec_like RS trans) THEN' + asm_simp_tac (put_simpset ss_if_True_False ctxt)) 1 THEN_MAYBE (unfold_thms_tac ctxt (pre_map_def :: map_comp's @ map_comps'' @ map_ids'' @ map_if_distribs @ corec_like_unfold_thms) THEN (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)) 1); @@ -123,7 +124,7 @@ fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt = EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc => case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN - asm_simp_tac (ss_only basic_simp_thms) 1 THEN + asm_simp_tac (ss_only basic_simp_thms ctxt) 1 THEN (if is_refl disc then all_tac else rtac disc 1)) (map rtac case_splits' @ [K all_tac]) corec_likes discs); @@ -162,12 +163,12 @@ SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels @ sum_prod_thms_rel)) THEN' (atac ORELSE' REPEAT o etac conjE THEN' full_simp_tac - (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms)) THEN_MAYBE' + (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN_MAYBE' REPEAT o hyp_subst_tac THEN' REPEAT o rtac conjI THEN' REPEAT o rtac refl); -fun mk_coinduct_distinct_ctrs discs discs' = +fun mk_coinduct_distinct_ctrs ctxt discs discs' = hyp_subst_tac THEN' REPEAT o etac conjE THEN' - full_simp_tac (ss_only (refl :: no_refl (discs @ discs') @ basic_simp_thms)); + full_simp_tac (ss_only (refl :: no_refl (discs @ discs') @ basic_simp_thms) ctxt); fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs discss selss = @@ -180,7 +181,7 @@ if k' = k then mk_coinduct_same_ctr ctxt rel_eqs' pre_rel_def dtor_ctor ctr_def discs sels else - mk_coinduct_distinct_ctrs discs discs') ks discss)) ks ctr_defs discss selss) + mk_coinduct_distinct_ctrs ctxt discs discs') ks discss)) ks ctr_defs discss selss) end; fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs dtor_ctors exhausts ctr_defss diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/BNF/Tools/bnf_tactics.ML --- a/src/HOL/BNF/Tools/bnf_tactics.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_tactics.ML Thu Apr 18 17:07:01 2013 +0200 @@ -8,7 +8,7 @@ signature BNF_TACTICS = sig - val ss_only: thm list -> simpset + val ss_only : thm list -> Proof.context -> Proof.context val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic val fo_rtac: thm -> Proof.context -> int -> tactic @@ -36,7 +36,7 @@ open BNF_Util -fun ss_only thms = Simplifier.clear_ss HOL_basic_ss addsimps thms; +fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl, tac, REPEAT_DETERM_N (n - k) o etac thin_rl]); diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/BNF/Tools/bnf_wrap.ML --- a/src/HOL/BNF/Tools/bnf_wrap.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_wrap.ML Thu Apr 18 17:07:01 2013 +0200 @@ -540,7 +540,7 @@ map (fn NONE => Drule.dummy_thm | SOME thm => thm RS sym) collapse_thm_opts; in [Goal.prove_sorry lthy [] [] goal (fn _ => - mk_expand_tac n ms (inst_thm u disc_exhaust_thm) + mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm) (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss disc_exclude_thmsss')] |> map Thm.close_derivation @@ -573,7 +573,7 @@ mk_Trueprop_eq (ufcase, vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); in - (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac uexhaust_thm case_thms), + (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms), Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1))) |> pairself (Thm.close_derivation #> singleton (Proof_Context.export names_lthy lthy)) end; @@ -596,7 +596,7 @@ val split_thm = Goal.prove_sorry lthy [] [] goal - (fn _ => mk_split_tac uexhaust_thm case_thms inject_thmss distinct_thmsss) + (fn _ => mk_split_tac lthy uexhaust_thm case_thms inject_thmss distinct_thmsss) |> Thm.close_derivation |> singleton (Proof_Context.export names_lthy lthy); val split_asm_thm = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/BNF/Tools/bnf_wrap_tactics.ML --- a/src/HOL/BNF/Tools/bnf_wrap_tactics.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_wrap_tactics.ML Thu Apr 18 17:07:01 2013 +0200 @@ -8,17 +8,18 @@ signature BNF_WRAP_TACTICS = sig val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic - val mk_case_cong_tac: thm -> thm list -> tactic + val mk_case_cong_tac: Proof.context -> thm -> thm list -> tactic val mk_case_conv_tac: Proof.context -> int -> thm -> thm list -> thm list list -> thm list list -> tactic val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic - val mk_expand_tac: int -> int list -> thm -> thm -> thm list -> thm list list list -> - thm list list list -> tactic + val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list -> + thm list list list -> thm list list list -> tactic val mk_half_disc_exclude_tac: int -> thm -> thm -> tactic val mk_nchotomy_tac: int -> thm -> tactic val mk_other_half_disc_exclude_tac: thm -> tactic - val mk_split_tac: thm -> thm list -> thm list list -> thm list list list -> tactic + val mk_split_tac: Proof.context -> + thm -> thm list -> thm list list -> thm list list list -> tactic val mk_split_asm_tac: Proof.context -> thm -> tactic val mk_unique_disc_def_tac: int -> thm -> tactic end; @@ -66,7 +67,8 @@ REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN' SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl)) 1; -fun mk_expand_tac n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss disc_excludesss' = +fun mk_expand_tac ctxt + n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss disc_excludesss' = if ms = [0] then rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses RS sym))) 1 else @@ -86,7 +88,8 @@ else [rtac (vuncollapse RS trans), maybe_atac, if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac], - REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, asm_simp_tac (ss_only [])] + REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, + asm_simp_tac (ss_only [] ctxt)] else [dtac (the_single (if k = n then disc_excludes else disc_excludes')), etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}), @@ -101,18 +104,18 @@ EVERY' [hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)), rtac casex]) cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss)) 1; -fun mk_case_cong_tac uexhaust cases = +fun mk_case_cong_tac ctxt uexhaust cases = (rtac uexhaust THEN' - EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex])]) cases)) 1; + EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases)) 1; val naked_ctxt = @{theory_context HOL}; (* TODO: More precise "simp_thms"; get rid of "blast_tac" *) -fun mk_split_tac uexhaust cases injectss distinctsss = +fun mk_split_tac ctxt uexhaust cases injectss distinctsss = rtac uexhaust 1 THEN ALLGOALS (fn k => (hyp_subst_tac THEN' simp_tac (ss_only (@{thms simp_thms} @ cases @ nth injectss (k - 1) @ - flat (nth distinctsss (k - 1))))) k) THEN + flat (nth distinctsss (k - 1))) ctxt)) k) THEN ALLGOALS (blast_tac naked_ctxt); val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex}; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/AxExample.thy --- a/src/HOL/Bali/AxExample.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/AxExample.thy Thu Apr 18 17:07:01 2013 +0200 @@ -126,7 +126,7 @@ apply (rule ax_subst_Var_allI) apply (tactic {* inst1_tac @{context} "P'" "\a vs l vf. ?PP a vs l vf\?x \. ?p" *}) apply (rule allI) -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac" delsimps [@{thm peek_and_def2}, @{thm heap_def2}, @{thm subst_res_def2}, @{thm normal_def2}]) 1 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac" delsimps [@{thm peek_and_def2}, @{thm heap_def2}, @{thm subst_res_def2}, @{thm normal_def2}]) 1 *}) apply (rule ax_derivs.Abrupt) apply (simp (no_asm)) apply (tactic "ax_tac 1" (* FVar *)) @@ -176,26 +176,26 @@ apply (rule ax_InitS) apply force apply (simp (no_asm)) -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac") 1 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *}) apply (rule ax_Init_Skip_lemma) -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac") 1 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *}) apply (rule ax_InitS [THEN conseq1] (* init Base *)) apply force apply (simp (no_asm)) apply (unfold arr_viewed_from_def) apply (rule allI) apply (rule_tac P' = "Normal ?P" in conseq1) -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac") 1 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *}) apply (tactic "ax_tac 1") apply (tactic "ax_tac 1") apply (rule_tac [2] ax_subst_Var_allI) apply (tactic {* inst1_tac @{context} "P'" "\vf l vfa. Normal (?P vf l vfa)" *}) -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac" delsimps [@{thm split_paired_All}, @{thm peek_and_def2}, @{thm heap_free_def2}, @{thm initd_def2}, @{thm normal_def2}, @{thm supd_lupd}]) 2 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac" delsimps [@{thm split_paired_All}, @{thm peek_and_def2}, @{thm heap_free_def2}, @{thm initd_def2}, @{thm normal_def2}, @{thm supd_lupd}]) 2 *}) apply (tactic "ax_tac 2" (* NewA *)) apply (tactic "ax_tac 3" (* ax_Alloc_Arr *)) apply (tactic "ax_tac 3") apply (tactic {* inst1_tac @{context} "P" "\vf l vfa. Normal (?P vf l vfa\\)" *}) -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac") 2 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 2 *}) apply (tactic "ax_tac 2") apply (tactic "ax_tac 1" (* FVar *)) apply (tactic "ax_tac 2" (* StatRef *)) @@ -206,7 +206,7 @@ apply (drule initedD) apply (clarsimp elim!: atleast_free_SucD simp add: arr_inv_def) apply force -apply (tactic {* simp_tac (@{simpset} delloop "split_all_tac") 1 *}) +apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *}) apply (rule ax_triv_Init_Object [THEN peek_and_forget2, THEN conseq1]) apply (rule wf_tprg) apply clarsimp diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/AxSem.thy --- a/src/HOL/Bali/AxSem.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/AxSem.thy Thu Apr 18 17:07:01 2013 +0200 @@ -464,7 +464,7 @@ declare split_paired_All [simp del] split_paired_Ex [simp del] declare split_if [split del] split_if_asm [split del] option.split [split del] option.split_asm [split del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *} inductive diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/Basis.thy Thu Apr 18 17:07:01 2013 +0200 @@ -12,7 +12,7 @@ ML {* fun strip_tac i = REPEAT (resolve_tac [impI, allI] i) *} declare split_if_asm [split] option.split [split] option.split_asm [split] -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} declare if_weak_cong [cong del] option.weak_case_cong [cong del] declare length_Suc_conv [iff] @@ -180,7 +180,7 @@ ML {* fun sum3_instantiate ctxt thm = map (fn s => - simplify (simpset_of ctxt delsimps [@{thm not_None_eq}]) + simplify (ctxt delsimps [@{thm not_None_eq}]) (read_instantiate ctxt [(("t", 0), "In" ^ s ^ " ?x")] thm)) ["1l","2","3","1r"] *} (* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/DefiniteAssignment.thy --- a/src/HOL/Bali/DefiniteAssignment.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/DefiniteAssignment.thy Thu Apr 18 17:07:01 2013 +0200 @@ -818,7 +818,7 @@ declare inj_term_sym_simps [simp] declare assigns_if.simps [simp del] declare split_paired_All [simp del] split_paired_Ex [simp del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} inductive_cases da_elim_cases [cases set]: "Env\ B \\Skip\\ A" @@ -884,7 +884,7 @@ declare inj_term_sym_simps [simp del] declare assigns_if.simps [simp] declare split_paired_All [simp] split_paired_Ex [simp] -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} (* To be able to eliminate both the versions with the overloaded brackets: (B \\Skip\\ A) and with the explicit constructor (B \In1r Skip\ A), diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/Eval.thy --- a/src/HOL/Bali/Eval.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/Eval.thy Thu Apr 18 17:07:01 2013 +0200 @@ -780,7 +780,7 @@ declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *) declare split_paired_All [simp del] split_paired_Ex [simp del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} inductive_cases eval_cases: "G\s \t\\ (v, s')" @@ -818,7 +818,7 @@ "G\Norm s \In1r (Init C) \\ (x, s')" declare not_None_eq [simp] (* IntDef.Zero_def [simp] *) declare split_paired_All [simp] split_paired_Ex [simp] -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +declaration {* K (Simplifier.map_ss (fn ss => ss addloop' ("split_all_tac", split_all_tac))) *} declare split_if [split] split_if_asm [split] option.split [split] option.split_asm [split] diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/Evaln.thy --- a/src/HOL/Bali/Evaln.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/Evaln.thy Thu Apr 18 17:07:01 2013 +0200 @@ -197,7 +197,7 @@ option.split [split del] option.split_asm [split del] not_None_eq [simp del] split_paired_All [simp del] split_paired_Ex [simp del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} inductive_cases evaln_cases: "G\s \t\\n\ (v, s')" @@ -238,7 +238,7 @@ option.split [split] option.split_asm [split] not_None_eq [simp] split_paired_All [simp] split_paired_Ex [simp] -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +declaration {* K (Simplifier.map_ss (fn ss => ss addloop' ("split_all_tac", split_all_tac))) *} lemma evaln_Inj_elim: "G\s \t\\n\ (w,s') \ case t of In1 ec \ (case ec of Inl e \ (\v. w = In1 v) | Inr c \ w = \) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/Example.thy --- a/src/HOL/Bali/Example.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/Example.thy Thu Apr 18 17:07:01 2013 +0200 @@ -1188,8 +1188,7 @@ Base_foo_defs [simp] ML {* bind_thms ("eval_intros", map - (simplify (@{simpset} delsimps @{thms Skip_eq} - addsimps @{thms lvar_def}) o + (simplify (@{context} delsimps @{thms Skip_eq} addsimps @{thms lvar_def}) o rewrite_rule [@{thm assign_def}, @{thm Let_def}]) @{thms eval.intros}) *} lemmas eval_Is = eval_Init eval_StatRef AbruptIs eval_intros diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/TypeSafe.thy --- a/src/HOL/Bali/TypeSafe.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/TypeSafe.thy Thu Apr 18 17:07:01 2013 +0200 @@ -726,7 +726,7 @@ declare split_paired_All [simp del] split_paired_Ex [simp del] declare split_if [split del] split_if_asm [split del] option.split [split del] option.split_asm [split del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *} lemma FVar_lemma: @@ -756,7 +756,7 @@ declare split_if [split] split_if_asm [split] option.split [split] option.split_asm [split] setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *} -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} lemma AVar_lemma1: "\globs s (Inl a) = Some obj;tag obj=Arr ty i; @@ -871,7 +871,7 @@ declare split_paired_All [simp del] split_paired_Ex [simp del] declare split_if [split del] split_if_asm [split del] option.split [split del] option.split_asm [split del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *} lemma conforms_init_lvars: @@ -925,7 +925,7 @@ declare split_if [split] split_if_asm [split] option.split [split] option.split_asm [split] setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *} -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} subsection "accessibility" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/WellForm.thy --- a/src/HOL/Bali/WellForm.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/WellForm.thy Thu Apr 18 17:07:01 2013 +0200 @@ -2127,7 +2127,7 @@ qed declare split_paired_All [simp del] split_paired_Ex [simp del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *} lemma dynamic_mheadsD: @@ -2258,7 +2258,7 @@ qed declare split_paired_All [simp] split_paired_Ex [simp] setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *} -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} (* Tactical version *) (* @@ -2401,7 +2401,7 @@ declare split_paired_All [simp del] split_paired_Ex [simp del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} setup {* map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac") *} lemma wt_is_type: "E,dt\v\T \ wf_prog (prg E) \ @@ -2427,7 +2427,7 @@ done declare split_paired_All [simp] split_paired_Ex [simp] setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *} -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} lemma ty_expr_is_type: "\E\e\-T; wf_prog (prg E)\ \ is_type (prg E) T" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Bali/WellType.thy --- a/src/HOL/Bali/WellType.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Bali/WellType.thy Thu Apr 18 17:07:01 2013 +0200 @@ -458,7 +458,7 @@ declare not_None_eq [simp del] declare split_if [split del] split_if_asm [split del] declare split_paired_All [simp del] split_paired_Ex [simp del] -declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *} +setup {* map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac") *} inductive_cases wt_elim_cases [cases set]: "E,dt\In2 (LVar vn) \T" @@ -494,7 +494,7 @@ declare not_None_eq [simp] declare split_if [split] split_if_asm [split] declare split_paired_All [simp] split_paired_Ex [simp] -declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *} +setup {* map_theory_simpset (fn ctxt => ctxt addloop' ("split_all_tac", split_all_tac)) *} lemma is_acc_class_is_accessible: "is_acc_class G P C \ G\(Class C) accessible_in P" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/Approximation.thy Thu Apr 18 17:07:01 2013 +0200 @@ -3487,7 +3487,7 @@ (@{cpat "?prec::nat"}, p), (@{cpat "?ss::nat list"}, s)]) @{thm "approx_form"}) i - THEN simp_tac @{simpset} i) st + THEN simp_tac @{context} i) st end | SOME t => if length vs <> 1 then raise (TERM ("More than one variable used for taylor series expansion", [prop_of st])) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/Dense_Linear_Order.thy --- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Thu Apr 18 17:07:01 2013 +0200 @@ -579,7 +579,8 @@ else Ferrante_Rackoff_Data.Nox | _ => Ferrante_Rackoff_Data.Nox in h end - fun ss phi = HOL_ss addsimps (simps phi) + fun ss phi = + simpset_of (put_simpset HOL_ss @{context} addsimps (simps phi)) in Ferrante_Rackoff_Data.funs @{thm "ferrack_axiom"} {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss} @@ -749,7 +750,7 @@ val clt = Thm.dest_fun2 ct val cz = Thm.dest_arg ct val neg = cr Ferrante_Rackoff_Data.Nox in h end; fun class_field_ss phi = - HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}]) - |> fold Splitter.add_split [@{thm "abs_split"}, @{thm "split_max"}, @{thm "split_min"}] + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}]) + |> fold Splitter.add_split [@{thm "abs_split"}, @{thm "split_max"}, @{thm "split_min"}]) in Ferrante_Rackoff_Data.funs @{thm "class_dense_linordered_field.ferrack_axiom"} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/commutative_ring_tac.ML --- a/src/HOL/Decision_Procs/commutative_ring_tac.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML Thu Apr 18 17:07:01 2013 +0200 @@ -86,15 +86,14 @@ fun tac ctxt = SUBGOAL (fn (g, i) => let val thy = Proof_Context.theory_of ctxt; - val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) - addsimps cring_simps; + val cring_ctxt = ctxt addsimps cring_simps; (*FIXME really the full simpset!?*) val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) val norm_eq_th = - simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) + simplify cring_ctxt (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) in cut_tac norm_eq_th i - THEN (simp_tac cring_ss i) - THEN (simp_tac cring_ss i) + THEN (simp_tac cring_ctxt i) + THEN (simp_tac cring_ctxt i) end); end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/cooper_tac.ML --- a/src/HOL/Decision_Procs/cooper_tac.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/cooper_tac.ML Thu Apr 18 17:07:01 2013 +0200 @@ -14,7 +14,7 @@ val trace = Unsynchronized.ref false; fun trace_msg s = if !trace then tracing s else (); -val cooper_ss = @{simpset}; +val cooper_ss = simpset_of @{context}; val nT = HOLogic.natT; val comp_arith = @{thms simp_thms} @@ -68,7 +68,8 @@ (* Transform the term*) val (t,np,nh) = prepare_for_linz q g (* Some simpsets for dealing with mod div abs and nat*) - val mod_div_simpset = HOL_basic_ss + val mod_div_simpset = + put_simpset HOL_basic_ss ctxt addsimps [refl,mod_add_eq, mod_add_left_eq, mod_add_right_eq, nat_div_add_eq, int_div_add_eq, @@ -78,29 +79,32 @@ Suc_eq_plus1] addsimps @{thms add_ac} addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}] - val simpset0 = HOL_basic_ss + val simpset0 = + put_simpset HOL_basic_ss ctxt addsimps [mod_div_equality', Suc_eq_plus1] addsimps comp_arith |> fold Splitter.add_split [split_zdiv, split_zmod, split_div', @{thm "split_min"}, @{thm "split_max"}] (* Simp rules for changing (n::int) to int n *) - val simpset1 = HOL_basic_ss + val simpset1 = + put_simpset HOL_basic_ss ctxt addsimps [zdvd_int] @ map (fn r => r RS sym) [@{thm int_numeral}, @{thm int_int_eq}, @{thm zle_int}, @{thm zless_int}, @{thm zadd_int}, @{thm zmult_int}] |> Splitter.add_split zdiff_int_split (*simp rules for elimination of int n*) - val simpset2 = HOL_basic_ss + val simpset2 = + put_simpset HOL_basic_ss ctxt addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, @{thm zero_le_numeral}, @{thm order_refl}(* FIXME: necessary? *), @{thm int_0}, @{thm int_1}] |> fold Simplifier.add_cong [@{thm conj_le_cong}, @{thm imp_le_cong}] (* simp rules for elimination of abs *) - val simpset3 = HOL_basic_ss |> Splitter.add_split @{thm abs_split} + val simpset3 = put_simpset HOL_basic_ss ctxt |> Splitter.add_split @{thm abs_split} val ct = cterm_of thy (HOLogic.mk_Trueprop t) (* Theorem for the nat --> int transformation *) val pre_thm = Seq.hd (EVERY [simp_tac mod_div_simpset 1, simp_tac simpset0 1, TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), - TRY (simp_tac simpset3 1), TRY (simp_tac cooper_ss 1)] + TRY (simp_tac simpset3 1), TRY (simp_tac (put_simpset cooper_ss ctxt) 1)] (Thm.trivial ct)) fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) (* The result of the quantifier elimination *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/ferrack_tac.ML --- a/src/HOL/Decision_Procs/ferrack_tac.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/ferrack_tac.ML Thu Apr 18 17:07:01 2013 +0200 @@ -16,8 +16,8 @@ val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, @{thm real_of_int_le_iff}] - in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) - end; + in @{context} delsimps ths addsimps (map (fn th => th RS sym) ths) + end |> simpset_of; val binarith = @{thms arith_simps} val comp_arith = binarith @ @{thms simp_thms} @@ -74,12 +74,12 @@ (* Transform the term*) val (t,np,nh) = prepare_for_linr thy q g (* Some simpsets for dealing with mod div abs and nat*) - val simpset0 = Simplifier.context ctxt HOL_basic_ss addsimps comp_arith + val simpset0 = put_simpset HOL_basic_ss ctxt addsimps comp_arith val ct = cterm_of thy (HOLogic.mk_Trueprop t) (* Theorem for the nat --> int transformation *) val pre_thm = Seq.hd (EVERY [simp_tac simpset0 1, - TRY (simp_tac (Simplifier.context ctxt ferrack_ss) 1)] + TRY (simp_tac (put_simpset ferrack_ss ctxt) 1)] (Thm.trivial ct)) fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) (* The result of the quantifier elimination *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/ferrante_rackoff.ML --- a/src/HOL/Decision_Procs/ferrante_rackoff.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML Thu Apr 18 17:07:01 2013 +0200 @@ -27,7 +27,7 @@ funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) (funpow 2 Thm.dest_arg (cprop_of th)) |> Thm.dest_arg -fun ferrack_conv +fun ferrack_conv ctxt (entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = ld, qe = qe, atoms = atoms}, {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) = @@ -163,7 +163,7 @@ qe)) [fU, ld_th, nmi_th, npi_th, minf_th, pinf_th] val bex_conv = - Simplifier.rewrite (HOL_basic_ss addsimps @{thms simp_thms bex_simps(1-5)}) + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms bex_simps(1-5)}) val result_th = fconv_rule (arg_conv bex_conv) (Thm.transitive enth qe_th) in result_th end @@ -196,22 +196,21 @@ in h (bounds + 1) b' end; in h end; -fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset}) tm = +fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset = ss}) tm = let - val ss = simpset val ss' = - merge_ss (HOL_basic_ss addsimps @{thms simp_thms ex_simps all_simps - not_all all_not_ex ex_disj_distrib}, ss) - |> Simplifier.inherit_context ss - val pcv = Simplifier.rewrite ss' - val postcv = Simplifier.rewrite ss - val nnf = K (nnf_conv then_conv postcv) + merge_ss (simpset_of + (put_simpset HOL_basic_ss ctxt addsimps + @{thms simp_thms ex_simps all_simps not_all all_not_ex ex_disj_distrib}), ss); + val pcv = Simplifier.rewrite (put_simpset ss' ctxt); + val postcv = Simplifier.rewrite (put_simpset ss ctxt); + val nnf = K (nnf_conv ctxt then_conv postcv) val qe_conv = Qelim.gen_qelim_conv pcv postcv pcv cons (Thm.add_cterm_frees tm []) (isolate_conv ctxt) nnf - (fn vs => ferrack_conv (thy,{isolate_conv = isolate_conv ctxt, - whatis = whatis, simpset = simpset}) vs + (fn vs => ferrack_conv ctxt (thy,{isolate_conv = isolate_conv ctxt, + whatis = whatis, simpset = ss}) vs then_conv postcv) - in (Simplifier.rewrite ss then_conv qe_conv) tm end; + in (Simplifier.rewrite (put_simpset ss ctxt) then_conv qe_conv) tm end; fun dlo_instance ctxt tm = Ferrante_Rackoff_Data.match ctxt (grab_atom_bop 0 tm); @@ -226,8 +225,8 @@ NONE => no_tac | SOME instance => Object_Logic.full_atomize_tac i THEN - simp_tac (#simpset (snd instance)) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) + simp_tac (put_simpset (#simpset (snd instance)) ctxt) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) CONVERSION (Object_Logic.judgment_conv (raw_ferrack_qe_conv ctxt instance)) i THEN - simp_tac (simpset_of ctxt) i)); (* FIXME really? *) + simp_tac ctxt i)); (* FIXME really? *) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/langford.ML --- a/src/HOL/Decision_Procs/langford.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/langford.ML Thu Apr 18 17:07:01 2013 +0200 @@ -26,16 +26,18 @@ (Thm.cprop_of th), SOME x] th1) th in fold ins u th0 end; -val simp_rule = +fun simp_rule ctxt = Conv.fconv_rule - (Conv.arg_conv (Simplifier.rewrite (HOL_basic_ss addsimps @{thms ball_simps simp_thms}))); + (Conv.arg_conv + (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms ball_simps simp_thms}))); fun basic_dloqe ctxt stupid dlo_qeth dlo_qeth_nolb dlo_qeth_noub gather ep = case term_of ep of Const(@{const_name Ex},_)$_ => let val p = Thm.dest_arg ep - val ths = simplify (HOL_basic_ss addsimps gather) (instantiate' [] [SOME p] stupid) + val ths = + simplify (put_simpset HOL_basic_ss ctxt addsimps gather) (instantiate' [] [SOME p] stupid) val (L,U) = let val (x,q) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of ths)) @@ -53,17 +55,17 @@ (Const (@{const_name Orderings.bot}, _),_) => let val (neU,fU) = proveneF U - in simp_rule (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end + in simp_rule ctxt (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end | (_,Const (@{const_name Orderings.bot}, _)) => let val (neL,fL) = proveneF L - in simp_rule (Thm.transitive ths (dlo_qeth_noub OF [neL, fL])) end + in simp_rule ctxt (Thm.transitive ths (dlo_qeth_noub OF [neL, fL])) end | (_,_) => let val (neL,fL) = proveneF L val (neU,fU) = proveneF U - in simp_rule (Thm.transitive ths (dlo_qeth OF [neL, neU, fL, fU])) + in simp_rule ctxt (Thm.transitive ths (dlo_qeth OF [neL, neU, fL, fU])) end in qe end | _ => error "dlo_qe : Not an existential formula"; @@ -122,7 +124,7 @@ | _ => false ; local -fun proc ct = +fun proc ctxt ct = case term_of ct of Const(@{const_name Ex},_)$Abs (xn,_,_) => let @@ -140,35 +142,36 @@ (Thm.apply @{cterm Trueprop} (list_conj (ndx @dx)))) |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e |> Conv.fconv_rule (Conv.arg_conv - (Simplifier.rewrite (HOL_basic_ss addsimps @{thms simp_thms ex_simps}))) + (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps}))) |> SOME end | _ => conj_aci_rule (Thm.mk_binop @{cterm "op == :: prop => _"} Pp (Thm.apply @{cterm Trueprop} (list_conj (eqs@neqs)))) |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e |> Conv.fconv_rule (Conv.arg_conv - (Simplifier.rewrite (HOL_basic_ss addsimps @{thms simp_thms ex_simps}))) + (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps}))) |> SOME end | _ => NONE; in val reduce_ex_simproc = Simplifier.make_simproc {lhss = [@{cpat "EX x. ?P x"}] , name = "reduce_ex_simproc", - proc = K (K proc) , identifier = []} + proc = K proc, identifier = []} end; -fun raw_dlo_conv dlo_ss +fun raw_dlo_conv ctxt dlo_ss ({qe_bnds, qe_nolb, qe_noub, gst, gs, atoms}:Langford_Data.entry) = let - val ss = dlo_ss addsimps @{thms "dnf_simps"} addsimprocs [reduce_ex_simproc] - val dnfex_conv = Simplifier.rewrite ss + val ctxt' = put_simpset dlo_ss ctxt addsimps @{thms "dnf_simps"} addsimprocs [reduce_ex_simproc] + val dnfex_conv = Simplifier.rewrite ctxt' val pcv = Simplifier.rewrite - (dlo_ss addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}) + (put_simpset dlo_ss ctxt + addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}) in fn p => Qelim.gen_qelim_conv pcv pcv dnfex_conv cons (Thm.add_cterm_frees p []) (K Thm.reflexive) (K Thm.reflexive) - (K (basic_dloqe () gst qe_bnds qe_nolb qe_noub gs)) p + (K (basic_dloqe ctxt gst qe_bnds qe_nolb qe_noub gs)) p end; @@ -204,7 +207,7 @@ fun dlo_conv ctxt tm = (case dlo_instance ctxt tm of (_, NONE) => raise CTERM ("dlo_conv (langford): no corresponding instance in context!", [tm]) - | (ss, SOME instance) => raw_dlo_conv ss instance tm); + | (ss, SOME instance) => raw_dlo_conv ctxt ss instance tm); fun generalize_tac f = CSUBGOAL (fn (p, i) => PRIMITIVE (fn st => let @@ -232,13 +235,13 @@ fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => (case dlo_instance ctxt p of - (ss, NONE) => simp_tac ss i - | (ss, SOME instance) => + (ss, NONE) => simp_tac (put_simpset ss ctxt) i + | (ss, SOME instance) => Object_Logic.full_atomize_tac i THEN - simp_tac ss i + simp_tac (put_simpset ss ctxt) i THEN (CONVERSION Thm.eta_long_conversion) i THEN (TRY o generalize_tac (cfrees (#atoms instance))) i THEN Object_Logic.full_atomize_tac i - THEN CONVERSION (Object_Logic.judgment_conv (raw_dlo_conv ss instance)) i - THEN (simp_tac ss i))); + THEN CONVERSION (Object_Logic.judgment_conv (raw_dlo_conv ctxt ss instance)) i + THEN (simp_tac (put_simpset ss ctxt) i))); end; \ No newline at end of file diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/langford_data.ML --- a/src/HOL/Decision_Procs/langford_data.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/langford_data.ML Thu Apr 18 17:07:01 2013 +0200 @@ -36,9 +36,11 @@ Thm.declaration_attribute (fn key => fn context => context |> Data.map (del_data key #> apsnd (cons (key, entry)))); -val add_simp = Thm.declaration_attribute (Data.map o apfst o Simplifier.add_simp); +val add_simp = Thm.declaration_attribute (fn th => fn context => + (Data.map o apfst) (simpset_map (Context.proof_of context) (Simplifier.add_simp th)) context); -val del_simp = Thm.declaration_attribute (Data.map o apfst o Simplifier.del_simp); +val del_simp = Thm.declaration_attribute (fn th => fn context => + (Data.map o apfst) (simpset_map (Context.proof_of context) (Simplifier.del_simp th)) context); fun match ctxt tm = let diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Decision_Procs/mir_tac.ML Thu Apr 18 17:07:01 2013 +0200 @@ -16,7 +16,7 @@ val mir_ss = let val ths = [@{thm "real_of_int_inject"}, @{thm "real_of_int_less_iff"}, @{thm "real_of_int_le_iff"}] -in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) +in simpset_of (@{context} delsimps ths addsimps (map (fn th => th RS sym) ths)) end; val nT = HOLogic.natT; @@ -83,7 +83,8 @@ fun mir_tac ctxt q = Object_Logic.atomize_prems_tac - THEN' simp_tac (HOL_basic_ss addsimps [@{thm "abs_ge_zero"}] addsimps @{thms simp_thms}) + THEN' simp_tac (put_simpset HOL_basic_ss ctxt + addsimps [@{thm "abs_ge_zero"}] addsimps @{thms simp_thms}) THEN' (REPEAT_DETERM o split_tac [@{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}]) THEN' SUBGOAL (fn (g, i) => let @@ -91,7 +92,7 @@ (* Transform the term*) val (t,np,nh) = prepare_for_mir q g (* Some simpsets for dealing with mod div abs and nat*) - val mod_div_simpset = HOL_basic_ss + val mod_div_simpset = put_simpset HOL_basic_ss ctxt addsimps [refl, mod_add_eq, @{thm mod_self}, @{thm div_0}, @{thm mod_0}, @@ -99,21 +100,21 @@ @{thm "Suc_eq_plus1"}] addsimps @{thms add_ac} addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}] - val simpset0 = HOL_basic_ss + val simpset0 = put_simpset HOL_basic_ss ctxt addsimps [mod_div_equality', @{thm Suc_eq_plus1}] addsimps comp_ths |> fold Splitter.add_split [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}] (* Simp rules for changing (n::int) to int n *) - val simpset1 = HOL_basic_ss + val simpset1 = put_simpset HOL_basic_ss ctxt addsimps [@{thm "zdvd_int"}] @ map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, @{thm nat_numeral}, @{thm "zmult_int"}] |> Splitter.add_split @{thm "zdiff_int_split"} (*simp rules for elimination of int n*) - val simpset2 = HOL_basic_ss + val simpset2 = put_simpset HOL_basic_ss ctxt addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm zero_le_numeral}, @{thm "int_0"}, @{thm "int_1"}] |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] @@ -122,7 +123,8 @@ (* Theorem for the nat --> int transformation *) val pre_thm = Seq.hd (EVERY [simp_tac mod_div_simpset 1, simp_tac simpset0 1, - TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), TRY (simp_tac mir_ss 1)] + TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), + TRY (simp_tac (put_simpset mir_ss ctxt) 1)] (Thm.trivial ct)) fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) (* The result of the quantifier elimination *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Divides.thy Thu Apr 18 17:07:01 2013 +0200 @@ -1643,12 +1643,12 @@ val simps = @{thms arith_simps} @ @{thms rel_simps} @ map (fn th => th RS sym) [@{thm numeral_1_eq_1}] fun prove ctxt goal = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop goal) - (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps simps)))); - fun binary_proc proc ss ct = + (K (ALLGOALS (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps)))); + fun binary_proc proc ctxt ct = (case Thm.term_of ct of _ $ t $ u => (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of - SOME args => proc (Simplifier.the_context ss) args + SOME args => proc ctxt args | NONE => NONE) | _ => NONE); in diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Fun.thy Thu Apr 18 17:07:01 2013 +0200 @@ -795,9 +795,10 @@ | find t = NONE in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end - fun proc ss ct = + val ss = simpset_of @{context} + + fun proc ctxt ct = let - val ctxt = Simplifier.the_context ss val t = Thm.term_of ct in case find_double t of @@ -807,7 +808,7 @@ (fn _ => rtac eq_reflection 1 THEN rtac ext 1 THEN - simp_tac (Simplifier.inherit_context ss @{simpset}) 1)) + simp_tac (put_simpset ss ctxt) 1)) end in proc end *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOL.thy Thu Apr 18 17:07:01 2013 +0200 @@ -1189,7 +1189,7 @@ ML_file "Tools/simpdata.ML" ML {* open Simpdata *} -setup {* Simplifier.map_simpset_global (K HOL_basic_ss) *} +setup {* map_theory_simpset (put_simpset HOL_basic_ss) *} simproc_setup defined_Ex ("EX x. P x") = {* fn _ => Quantifier1.rearrange_ex *} simproc_setup defined_All ("ALL x. P x") = {* fn _ => Quantifier1.rearrange_all *} @@ -1241,10 +1241,9 @@ case t of Abs (_, _, t') => count_loose t' 0 <= 1 | _ => true; -in fn _ => fn ss => fn ct => if is_trivial_let (Thm.term_of ct) +in fn _ => fn ctxt => fn ct => if is_trivial_let (Thm.term_of ct) then SOME @{thm Let_def} (*no or one ocurrence of bound variable*) else let (*Norbert Schirmer's case*) - val ctxt = Simplifier.the_context ss; val thy = Proof_Context.theory_of ctxt; val t = Thm.term_of ct; val ([t'], ctxt') = Variable.import_terms false [t] ctxt; @@ -1258,7 +1257,7 @@ val cx = cterm_of thy x; val {T = xT, ...} = rep_cterm cx; val cf = cterm_of thy f; - val fx_g = Simplifier.rewrite ss (Thm.apply cf cx); + val fx_g = Simplifier.rewrite ctxt (Thm.apply cf cx); val (_ $ _ $ g) = prop_of fx_g; val g' = abstract_over (x,g); val abs_g'= Abs (n,xT,g'); @@ -1345,7 +1344,7 @@ lemmas [cong] = imp_cong simp_implies_cong lemmas [split] = split_if -ML {* val HOL_ss = @{simpset} *} +ML {* val HOL_ss = simpset_of @{context} *} text {* Simplifies x assuming c and y assuming ~c *} lemma if_cong: @@ -1482,13 +1481,13 @@ addsimprocs [Simplifier.simproc_global @{theory} "swap_induct_false" ["induct_false ==> PROP P ==> PROP Q"] - (fn _ => fn _ => + (fn _ => (fn _ $ (P as _ $ @{const induct_false}) $ (_ $ Q $ _) => if P <> Q then SOME Drule.swap_prems_eq else NONE | _ => NONE)), Simplifier.simproc_global @{theory} "induct_equal_conj_curry" ["induct_conj P Q ==> PROP R"] - (fn _ => fn _ => + (fn _ => (fn _ $ (_ $ P) $ _ => let fun is_conj (@{const induct_conj} $ P $ Q) = @@ -1583,7 +1582,7 @@ signature REORIENT_PROC = sig val add : (term -> bool) -> theory -> theory - val proc : morphism -> simpset -> cterm -> thm option + val proc : morphism -> Proof.context -> cterm -> thm option end; structure Reorient_Proc : REORIENT_PROC = @@ -1599,9 +1598,8 @@ fun matches thy t = exists (fn (m, _) => m t) (Data.get thy); val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; - fun proc phi ss ct = + fun proc phi ctxt ct = let - val ctxt = Simplifier.the_context ss; val thy = Proof_Context.theory_of ctxt; in case Thm.term_of ct of @@ -1701,9 +1699,9 @@ subsubsection {* Generic code generator preprocessor setup *} setup {* - Code_Preproc.map_pre (K HOL_basic_ss) - #> Code_Preproc.map_post (K HOL_basic_ss) - #> Code_Simp.map_ss (K HOL_basic_ss) + Code_Preproc.map_pre (put_simpset HOL_basic_ss) + #> Code_Preproc.map_post (put_simpset HOL_basic_ss) + #> Code_Simp.map_ss (put_simpset HOL_basic_ss) *} subsubsection {* Equality *} @@ -1728,10 +1726,9 @@ declare eq_equal [code] setup {* - Code_Preproc.map_pre (fn simpset => - simpset addsimprocs [Simplifier.simproc_global_i @{theory} "equal" [@{term HOL.eq}] - (fn thy => fn _ => - fn Const (_, Type ("fun", [Type _, _])) => SOME @{thm eq_equal} | _ => NONE)]) + Code_Preproc.map_pre (fn ctxt => + ctxt addsimprocs [Simplifier.simproc_global_i @{theory} "equal" [@{term HOL.eq}] + (fn _ => fn Const (_, Type ("fun", [Type _, _])) => SOME @{thm eq_equal} | _ => NONE)]) *} @@ -1994,7 +1991,8 @@ fun smp_tac j = EVERY'[dresolve_tac (smp j), atac]; end; -val nnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps @{thms simp_thms nnf_simps}); +fun nnf_conv ctxt = + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms nnf_simps}); *} hide_const (open) eq equal diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Cfun.thy --- a/src/HOL/HOLCF/Cfun.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Cfun.thy Thu Apr 18 17:07:01 2013 +0200 @@ -140,14 +140,14 @@ *} simproc_setup beta_cfun_proc ("Rep_cfun (Abs_cfun f)") = {* - fn phi => fn ss => fn ct => + fn phi => fn ctxt => fn ct => let val dest = Thm.dest_comb; val f = (snd o dest o snd o dest) ct; val [T, U] = Thm.dest_ctyp (ctyp_of_term f); val tr = instantiate' [SOME T, SOME U] [SOME f] (mk_meta_eq @{thm Abs_cfun_inverse2}); - val rules = Cont2ContData.get (Simplifier.the_context ss); + val rules = Cont2ContData.get ctxt; val tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules)); in SOME (perhaps (SINGLE (tac 1)) tr) end *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/IOA/ABP/Correctness.thy --- a/src/HOL/HOLCF/IOA/ABP/Correctness.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/IOA/ABP/Correctness.thy Thu Apr 18 17:07:01 2013 +0200 @@ -84,10 +84,9 @@ lemma last_ind_on_first: "l ~= [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))" apply simp - apply (tactic {* auto_tac (map_simpset (fn _ => - HOL_ss + apply (tactic {* auto_tac (put_simpset HOL_ss @{context} addsimps (@{thms reverse.simps} @ [@{thm hd_append}, @{thm rev_red_not_nil}]) - |> Splitter.add_split @{thm list.split}) @{context}) *}) + |> Splitter.add_split @{thm list.split}) *}) done text {* Main Lemma 1 for @{text "S_pkt"} in showing that reduce is refinement. *} @@ -166,16 +165,18 @@ lemma sender_abstraction: "is_weak_ref_map reduce srch_ioa srch_fin_ioa" apply (tactic {* - simp_tac (HOL_ss addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, - @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, - @{thm channel_abstraction}]) 1 *}) + simp_tac (put_simpset HOL_ss @{context} + addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, + @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, + @{thm channel_abstraction}]) 1 *}) done lemma receiver_abstraction: "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa" apply (tactic {* - simp_tac (HOL_ss addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, - @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, - @{thm channel_abstraction}]) 1 *}) + simp_tac (put_simpset HOL_ss @{context} + addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, + @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, + @{thm channel_abstraction}]) 1 *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/IOA/NTP/Impl.thy --- a/src/HOL/HOLCF/IOA/NTP/Impl.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/IOA/NTP/Impl.thy Thu Apr 18 17:07:01 2013 +0200 @@ -102,13 +102,15 @@ 3) renname_ss unfolds transitions and the abstract channel *) ML {* -val ss = @{simpset} addsimps @{thms "transitions"}; -val rename_ss = ss addsimps @{thms unfold_renaming}; +val ss = simpset_of (@{context} addsimps @{thms "transitions"}); +val rename_ss = simpset_of (put_simpset ss @{context} addsimps @{thms unfold_renaming}); -val tac = - asm_simp_tac (ss |> Simplifier.add_cong @{thm conj_cong} |> Splitter.add_split @{thm split_if}) -val tac_ren = - asm_simp_tac (rename_ss |> Simplifier.add_cong @{thm conj_cong} |> Splitter.add_split @{thm split_if}) +fun tac ctxt = + asm_simp_tac (put_simpset ss ctxt + |> Simplifier.add_cong @{thm conj_cong} |> Splitter.add_split @{thm split_if}) +fun tac_ren ctxt = + asm_simp_tac (put_simpset rename_ss ctxt + |> Simplifier.add_cong @{thm conj_cong} |> Splitter.add_split @{thm split_if}) *} @@ -129,34 +131,34 @@ apply (simp add: Impl.inv1_def split del: split_if) apply (induct_tac a) -apply (tactic "EVERY1[tac, tac, tac, tac]") -apply (tactic "tac 1") -apply (tactic "tac_ren 1") +apply (tactic "EVERY1[tac @{context}, tac @{context}, tac @{context}, tac @{context}]") +apply (tactic "tac @{context} 1") +apply (tactic "tac_ren @{context} 1") txt {* 5 + 1 *} -apply (tactic "tac 1") -apply (tactic "tac_ren 1") +apply (tactic "tac @{context} 1") +apply (tactic "tac_ren @{context} 1") txt {* 4 + 1 *} -apply (tactic {* EVERY1[tac, tac, tac, tac] *}) +apply (tactic {* EVERY1[tac @{context}, tac @{context}, tac @{context}, tac @{context}] *}) txt {* Now the other half *} apply (simp add: Impl.inv1_def split del: split_if) apply (induct_tac a) -apply (tactic "EVERY1 [tac, tac]") +apply (tactic "EVERY1 [tac @{context}, tac @{context}]") txt {* detour 1 *} -apply (tactic "tac 1") -apply (tactic "tac_ren 1") +apply (tactic "tac @{context} 1") +apply (tactic "tac_ren @{context} 1") apply (rule impI) apply (erule conjE)+ apply (simp (no_asm_simp) add: hdr_sum_def Multiset.count_def Multiset.countm_nonempty_def split add: split_if) txt {* detour 2 *} -apply (tactic "tac 1") -apply (tactic "tac_ren 1") +apply (tactic "tac @{context} 1") +apply (tactic "tac_ren @{context} 1") apply (rule impI) apply (erule conjE)+ apply (simp add: Impl.hdr_sum_def Multiset.count_def Multiset.countm_nonempty_def @@ -181,7 +183,8 @@ apply (rule countm_spurious_delm) apply (simp (no_asm)) -apply (tactic "EVERY1 [tac, tac, tac, tac, tac, tac]") +apply (tactic "EVERY1 [tac @{context}, tac @{context}, tac @{context}, + tac @{context}, tac @{context}, tac @{context}]") done @@ -200,7 +203,7 @@ txt {* 10 cases. First 4 are simple, since state doesn't change *} - ML_prf {* val tac2 = asm_full_simp_tac (ss addsimps [@{thm inv2_def}]) *} + ML_prf {* val tac2 = asm_full_simp_tac (put_simpset ss @{context} addsimps [@{thm inv2_def}]) *} txt {* 10 - 7 *} apply (tactic "EVERY1 [tac2,tac2,tac2,tac2]") @@ -256,13 +259,13 @@ apply (simp (no_asm_simp) add: impl_ioas split del: split_if) apply (induct_tac "a") - ML_prf {* val tac3 = asm_full_simp_tac (ss addsimps [@{thm inv3_def}]) *} + ML_prf {* val tac3 = asm_full_simp_tac (put_simpset ss @{context} addsimps [@{thm inv3_def}]) *} txt {* 10 - 8 *} apply (tactic "EVERY1[tac3,tac3,tac3]") - apply (tactic "tac_ren 1") + apply (tactic "tac_ren @{context} 1") apply (intro strip, (erule conjE)+) apply hypsubst apply (erule exE) @@ -270,7 +273,7 @@ txt {* 7 *} apply (tactic "tac3 1") - apply (tactic "tac_ren 1") + apply (tactic "tac_ren @{context} 1") apply force txt {* 6 - 3 *} @@ -278,7 +281,7 @@ apply (tactic "EVERY1[tac3,tac3,tac3,tac3]") txt {* 2 *} - apply (tactic "asm_full_simp_tac ss 1") + apply (tactic "asm_full_simp_tac (put_simpset ss @{context}) 1") apply (simp (no_asm) add: inv3_def) apply (intro strip, (erule conjE)+) apply (rule imp_disjL [THEN iffD1]) @@ -321,7 +324,7 @@ apply (simp (no_asm_simp) add: impl_ioas split del: split_if) apply (induct_tac "a") - ML_prf {* val tac4 = asm_full_simp_tac (ss addsimps [@{thm inv4_def}]) *} + ML_prf {* val tac4 = asm_full_simp_tac (put_simpset ss @{context} addsimps [@{thm inv4_def}]) *} txt {* 10 - 2 *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy --- a/src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy Thu Apr 18 17:07:01 2013 +0200 @@ -606,8 +606,7 @@ fun abstraction_tac ctxt = SELECT_GOAL (auto_tac (ctxt addSIs @{thms weak_strength_lemmas} - |> map_simpset (fn ss => - ss addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}]))) + addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}])) *} method_setup abstraction = {* Scan.succeed (SIMPLE_METHOD' o abstraction_tac) *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy --- a/src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy Thu Apr 18 17:07:01 2013 +0200 @@ -295,20 +295,18 @@ in fun mkex_induct_tac ctxt sch exA exB = - let val ss = simpset_of ctxt in - EVERY'[Seq_induct_tac ctxt sch defs, - asm_full_simp_tac ss, - SELECT_GOAL (safe_tac @{theory_context Fun}), - Seq_case_simp_tac ctxt exA, - Seq_case_simp_tac ctxt exB, - asm_full_simp_tac ss, - Seq_case_simp_tac ctxt exA, - asm_full_simp_tac ss, - Seq_case_simp_tac ctxt exB, - asm_full_simp_tac ss, - asm_full_simp_tac (ss addsimps asigs) - ] - end + EVERY'[Seq_induct_tac ctxt sch defs, + asm_full_simp_tac ctxt, + SELECT_GOAL (safe_tac @{theory_context Fun}), + Seq_case_simp_tac ctxt exA, + Seq_case_simp_tac ctxt exB, + asm_full_simp_tac ctxt, + Seq_case_simp_tac ctxt exA, + asm_full_simp_tac ctxt, + Seq_case_simp_tac ctxt exB, + asm_full_simp_tac ctxt, + asm_full_simp_tac (ctxt addsimps asigs) + ] end *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/IOA/meta_theory/Sequence.thy --- a/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy Thu Apr 18 17:07:01 2013 +0200 @@ -1086,37 +1086,31 @@ (* on a>>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *) fun Seq_case_simp_tac ctxt s i = - let val ss = simpset_of ctxt in - Seq_case_tac ctxt s i - THEN asm_simp_tac ss (i+2) - THEN asm_full_simp_tac ss (i+1) - THEN asm_full_simp_tac ss i - end; + Seq_case_tac ctxt s i + THEN asm_simp_tac ctxt (i+2) + THEN asm_full_simp_tac ctxt (i+1) + THEN asm_full_simp_tac ctxt i; (* rws are definitions to be unfolded for admissibility check *) fun Seq_induct_tac ctxt s rws i = - let val ss = simpset_of ctxt in - res_inst_tac ctxt [(("x", 0), s)] @{thm Seq_induct} i - THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ss (i+1)))) - THEN simp_tac (ss addsimps rws) i - end; + res_inst_tac ctxt [(("x", 0), s)] @{thm Seq_induct} i + THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt (i+1)))) + THEN simp_tac (ctxt addsimps rws) i; fun Seq_Finite_induct_tac ctxt i = etac @{thm Seq_Finite_ind} i - THEN (REPEAT_DETERM (CHANGED (asm_simp_tac (simpset_of ctxt) i))); + THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt i))); fun pair_tac ctxt s = res_inst_tac ctxt [(("p", 0), s)] @{thm PairE} - THEN' hyp_subst_tac THEN' asm_full_simp_tac (simpset_of ctxt); + THEN' hyp_subst_tac THEN' asm_full_simp_tac ctxt; (* induction on a sequence of pairs with pairsplitting and simplification *) fun pair_induct_tac ctxt s rws i = - let val ss = simpset_of ctxt in - res_inst_tac ctxt [(("x", 0), s)] @{thm Seq_induct} i - THEN pair_tac ctxt "a" (i+3) - THEN (REPEAT_DETERM (CHANGED (simp_tac ss (i+1)))) - THEN simp_tac (ss addsimps rws) i - end; + res_inst_tac ctxt [(("x", 0), s)] @{thm Seq_induct} i + THEN pair_tac ctxt "a" (i+3) + THEN (REPEAT_DETERM (CHANGED (simp_tac ctxt (i+1)))) + THEN simp_tac (ctxt addsimps rws) i; *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Lift.thy --- a/src/HOL/HOLCF/Lift.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Lift.thy Thu Apr 18 17:07:01 2013 +0200 @@ -46,7 +46,7 @@ method_setup defined = {* Scan.succeed (fn ctxt => SIMPLE_METHOD' - (etac @{thm lift_definedE} THEN' asm_simp_tac (simpset_of ctxt))) + (etac @{thm lift_definedE} THEN' asm_simp_tac ctxt)) *} lemma DefE: "Def x = \ \ R" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tools/Domain/domain_constructors.ML --- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Thu Apr 18 17:07:01 2013 +0200 @@ -64,13 +64,15 @@ (************************** miscellaneous functions ***************************) -val simple_ss = HOL_basic_ss addsimps @{thms simp_thms} +val simple_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms simp_thms}) val beta_rules = @{thms beta_cfun cont_id cont_const cont2cont_APP cont2cont_LAM'} @ @{thms cont2cont_fst cont2cont_snd cont2cont_Pair} -val beta_ss = HOL_basic_ss addsimps (@{thms simp_thms} @ beta_rules) +val beta_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms simp_thms} @ beta_rules)) fun define_consts (specs : (binding * term * mixfix) list) @@ -268,7 +270,7 @@ val bottom = mk_bottom (fastype_of v') val vs' = map (fn v => if v = v' then bottom else v) vs val goal = mk_trp (mk_undef (list_ccomb (con, vs'))) - val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1] + val tacs = [simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps rules) 1] in prove thy con_betas goal (K tacs) end in map one_strict nonlazy end @@ -282,7 +284,7 @@ val goal = mk_trp (iff_disj (lhs, rhss)) val rule1 = iso_locale RS @{thm iso.abs_bottom_iff} val rules = rule1 :: @{thms con_bottom_iff_rules} - val tacs = [simp_tac (HOL_ss addsimps rules) 1] + val tacs = [simp_tac (Simplifier.global_context thy HOL_ss addsimps rules) 1] in prove thy con_betas goal (K tacs) end in val con_stricts = maps con_strict spec' @@ -313,7 +315,7 @@ val rules1 = abs_below :: @{thms sinl_below sinr_below spair_below up_below} val rules2 = @{thms up_defined spair_defined ONE_defined} val rules = rules1 @ rules2 - val tacs = [asm_simp_tac (simple_ss addsimps rules) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy simple_ss addsimps rules) 1] in map (fn c => pgterm mk_below c (K tacs)) cons' end val injects = let @@ -321,7 +323,7 @@ val rules1 = abs_eq :: @{thms sinl_eq sinr_eq spair_eq up_eq} val rules2 = @{thms up_defined spair_defined ONE_defined} val rules = rules1 @ rules2 - val tacs = [asm_simp_tac (simple_ss addsimps rules) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy simple_ss addsimps rules) 1] in map (fn c => pgterm mk_eq c (K tacs)) cons' end end @@ -346,7 +348,7 @@ val goal = mk_trp (iff_disj (lhs, rhss)) val rule1 = iso_locale RS @{thm iso.abs_below} val rules = rule1 :: @{thms con_below_iff_rules} - val tacs = [simp_tac (HOL_ss addsimps rules) 1] + val tacs = [simp_tac (Simplifier.global_context thy HOL_ss addsimps rules) 1] in prove thy con_betas goal (K tacs) end fun dist_eq (con1, args1) (con2, args2) = let @@ -358,7 +360,7 @@ val goal = mk_trp (iff_disj2 (lhs, rhss1, rhss2)) val rule1 = iso_locale RS @{thm iso.abs_eq} val rules = rule1 :: @{thms con_eq_iff_rules} - val tacs = [simp_tac (HOL_ss addsimps rules) 1] + val tacs = [simp_tac (Simplifier.global_context thy HOL_ss addsimps rules) 1] in prove thy con_betas goal (K tacs) end in val dist_les = map_dist dist_le spec' @@ -514,7 +516,7 @@ val rules2 = @{thms con_bottom_iff_rules} val rules3 = @{thms cfcomp2 one_case2} val rules = abs_inverse :: rules1 @ rules2 @ rules3 - val tacs = [asm_simp_tac (beta_ss addsimps rules) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy beta_ss addsimps rules) 1] in prove thy defs goal (K tacs) end in val case_apps = map2 one_case spec fs @@ -582,7 +584,7 @@ val sel_stricts : thm list = let val rules = rep_strict :: @{thms sel_strict_rules} - val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1] + val tacs = [simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps rules) 1] fun sel_strict sel = let val goal = mk_trp (mk_strict sel) @@ -598,7 +600,7 @@ let val defs = con_betas @ sel_defs val rules = abs_inv :: @{thms sel_app_rules} - val tacs = [asm_simp_tac (simple_ss addsimps rules) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy simple_ss addsimps rules) 1] fun sel_apps_of (i, (con, args: (bool * term option * typ) list)) = let val Ts : typ list = map #3 args @@ -643,7 +645,7 @@ val sel_defins : thm list = let val rules = rep_bottom_iff :: @{thms sel_bottom_iff_rules} - val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1] + val tacs = [simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps rules) 1] fun sel_defin sel = let val (T, U) = dest_cfunT (fastype_of sel) @@ -720,7 +722,7 @@ val assms = map (mk_trp o mk_defined) nonlazy val concl = mk_trp (mk_eq (lhs, rhs)) val goal = Logic.list_implies (assms, concl) - val tacs = [asm_simp_tac (beta_ss addsimps case_rews) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy beta_ss addsimps case_rews) 1] in prove thy dis_defs goal (K tacs) end fun one_dis (i, dis) = map_index (dis_app (i, dis)) spec @@ -736,9 +738,9 @@ val simps = dis_apps @ @{thms dist_eq_tr} val tacs = [rtac @{thm iffI} 1, - asm_simp_tac (HOL_basic_ss addsimps dis_stricts) 2, + asm_simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps dis_stricts) 2, rtac exhaust 1, atac 1, - ALLGOALS (asm_full_simp_tac (simple_ss addsimps simps))] + ALLGOALS (asm_full_simp_tac (Simplifier.global_context thy simple_ss addsimps simps))] val goal = mk_trp (mk_eq (mk_undef (dis ` x), mk_undef x)) in prove thy [] goal (K tacs) end in @@ -809,7 +811,7 @@ val (T, (U, V)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat)) val k = Free ("k", U) val goal = mk_trp (mk_eq (mat ` mk_bottom T ` k, mk_bottom V)) - val tacs = [asm_simp_tac (beta_ss addsimps case_rews) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy beta_ss addsimps case_rews) 1] in prove thy match_defs goal (K tacs) end in val match_stricts = map match_strict match_consts @@ -828,7 +830,7 @@ val assms = map (mk_trp o mk_defined) nonlazy val concl = mk_trp (mk_eq (lhs, rhs)) val goal = Logic.list_implies (assms, concl) - val tacs = [asm_simp_tac (beta_ss addsimps case_rews) 1] + val tacs = [asm_simp_tac (Simplifier.global_context thy beta_ss addsimps case_rews) 1] in prove thy match_defs goal (K tacs) end fun one_match (i, mat) = map_index (match_app (i, mat)) spec diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tools/Domain/domain_induction.ML --- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Thu Apr 18 17:07:01 2013 +0200 @@ -71,7 +71,7 @@ val rules = [abs_inverse] @ con_betas @ @{thms take_con_rules} @ take_Suc_thms @ deflation_thms @ deflation_take_thms - val tac = simp_tac (HOL_basic_ss addsimps rules) 1 + val tac = simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps rules) 1 in Goal.prove_global thy [] [] goal (K tac) end @@ -132,7 +132,8 @@ mk_trp (p $ HOLCF_Library.mk_bottom T) :: map (con_assm true p) cons val assms = maps eq_assms (Ps ~~ newTs ~~ map #con_specs constr_infos) - val take_ss = HOL_ss addsimps (@{thm Rep_cfun_strict1} :: take_rews) + val take_ss = + simpset_of (put_simpset HOL_ss @{context} addsimps (@{thm Rep_cfun_strict1} :: take_rews)) fun quant_tac ctxt i = EVERY (map (fn name => res_inst_tac ctxt [(("x", 0), name)] spec i) x_names) @@ -157,7 +158,7 @@ let val subgoal = con_assm false p (con, args) val rules = prems @ con_rews @ @{thms simp_thms} - val simplify = asm_simp_tac (HOL_basic_ss addsimps rules) + val simplify = asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) fun arg_tac (lazy, _) = rtac (if lazy then allI else case_UU_allI) 1 val tacs = @@ -173,16 +174,16 @@ val tacs1 = [ quant_tac ctxt 1, - simp_tac HOL_ss 1, + simp_tac (put_simpset HOL_ss ctxt) 1, Induct_Tacs.induct_tac ctxt [[SOME "n"]] 1, - simp_tac (take_ss addsimps prems) 1, + simp_tac (put_simpset take_ss ctxt addsimps prems) 1, TRY (safe_tac (put_claset HOL_cs ctxt))] fun con_tac _ = - asm_simp_tac take_ss 1 THEN + asm_simp_tac (put_simpset take_ss ctxt) 1 THEN (resolve_tac prems' THEN_ALL_NEW etac spec) 1 fun cases_tacs (cons, exhaust) = res_inst_tac ctxt [(("y", 0), "x")] exhaust 1 :: - asm_simp_tac (take_ss addsimps prems) 1 :: + asm_simp_tac (put_simpset take_ss ctxt addsimps prems) 1 :: map con_tac cons val tacs = tacs1 @ maps cases_tacs (conss ~~ exhausts) in @@ -325,10 +326,10 @@ val dests = map (fn th => th RS spec RS spec RS mp) prems' fun one_tac (dest, rews) = dtac dest 1 THEN safe_tac (put_claset HOL_cs ctxt) THEN - ALLGOALS (asm_simp_tac (HOL_basic_ss addsimps rews)) + ALLGOALS (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps rews)) in rtac @{thm nat.induct} 1 THEN - simp_tac (HOL_ss addsimps rules) 1 THEN + simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1 THEN safe_tac (put_claset HOL_cs ctxt) THEN EVERY (map one_tac (dests ~~ take_rews)) end @@ -344,12 +345,12 @@ val assm1 = mk_trp (list_comb (bisim_const, Rs)) val assm2 = mk_trp (R $ x $ y) val goal = mk_trp (mk_eq (x, y)) - fun tacf {prems, context = _} = + fun tacf {prems, context = ctxt} = let val rule = hd prems RS coind_lemma in rtac take_lemma 1 THEN - asm_simp_tac (HOL_basic_ss addsimps (rule :: prems)) 1 + asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (rule :: prems)) 1 end in Goal.prove_global thy [] [assm1, assm2] goal tacf diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML --- a/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML Thu Apr 18 17:07:01 2013 +0200 @@ -36,7 +36,8 @@ struct val beta_ss = - HOL_basic_ss addsimps @{thms simp_thms} addsimprocs [@{simproc beta_cfun_proc}] + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps @{thms simp_thms} addsimprocs [@{simproc beta_cfun_proc}]) fun is_cpo thy T = Sign.of_sort thy (T, @{sort cpo}) @@ -156,7 +157,8 @@ (* prove applied version of definitions *) fun prove_proj (lhs, rhs) = let - val tac = rewrite_goals_tac fixdef_thms THEN (simp_tac beta_ss) 1 + val tac = rewrite_goals_tac fixdef_thms THEN + (simp_tac (Simplifier.global_context thy beta_ss)) 1 val goal = Logic.mk_equals (lhs, rhs) in Goal.prove_global thy [] [] goal (K tac) end val proj_thms = map prove_proj projs @@ -324,13 +326,13 @@ @ deflation_abs_rep_thms @ Domain_Take_Proofs.get_deflation_thms thy in - Goal.prove_global thy [] assms goal (fn {prems, ...} => + Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} => EVERY [rewrite_goals_tac map_apply_thms, rtac (map_cont_thm RS @{thm cont_fix_ind}) 1, REPEAT (resolve_tac adm_rules 1), - simp_tac (HOL_basic_ss addsimps bottom_rules) 1, - simp_tac (HOL_basic_ss addsimps tuple_rules) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps bottom_rules) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps tuple_rules) 1, REPEAT (etac @{thm conjE} 1), REPEAT (resolve_tac (deflation_rules @ prems) 1 ORELSE atac 1)]) end @@ -638,15 +640,15 @@ @ isodefl_abs_rep_thms @ IsodeflData.get (Proof_Context.init_global thy) in - Goal.prove_global thy [] assms goal (fn {prems, ...} => + Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} => EVERY [rewrite_goals_tac (defl_apply_thms @ map_apply_thms), rtac (@{thm cont_parallel_fix_ind} OF [defl_cont_thm, map_cont_thm]) 1, REPEAT (resolve_tac adm_rules 1), - simp_tac (HOL_basic_ss addsimps bottom_rules) 1, - simp_tac (HOL_basic_ss addsimps tuple_rules) 1, - simp_tac (HOL_basic_ss addsimps map_ID_simps) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps bottom_rules) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps tuple_rules) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps map_ID_simps) 1, REPEAT (etac @{thm conjE} 1), REPEAT (resolve_tac (isodefl_rules @ prems) 1 ORELSE atac 1)]) end @@ -712,16 +714,16 @@ val rules1 = @{thms iterate_Suc prod_eq_iff fst_conv snd_conv} @ take_Suc_thms - val tac = + fun tac ctxt = EVERY - [simp_tac (HOL_basic_ss addsimps start_rules) 1, - simp_tac (HOL_basic_ss addsimps @{thms fix_def2}) 1, + [simp_tac (put_simpset HOL_basic_ss ctxt addsimps start_rules) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms fix_def2}) 1, rtac @{thm lub_eq} 1, rtac @{thm nat.induct} 1, - simp_tac (HOL_basic_ss addsimps rules0) 1, - asm_full_simp_tac (beta_ss addsimps rules1) 1] + simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules0) 1, + asm_full_simp_tac (put_simpset beta_ss ctxt addsimps rules1) 1] in - Goal.prove_global thy [] [] goal (K tac) + Goal.prove_global thy [] [] goal (tac o #context) end (* prove lub of take equals ID *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML --- a/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML Thu Apr 18 17:07:01 2013 +0200 @@ -108,7 +108,8 @@ } val beta_ss = - HOL_basic_ss addsimps @{thms simp_thms} addsimprocs [@{simproc beta_cfun_proc}] + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps @{thms simp_thms} addsimprocs [@{simproc beta_cfun_proc}]) (******************************************************************************) (******************************** theory data *********************************) @@ -272,7 +273,7 @@ let val goal = mk_trp (mk_chain take_const) val rules = take_defs @ @{thms chain_iterate ch2ch_fst ch2ch_snd} - val tac = simp_tac (HOL_basic_ss addsimps rules) 1 + val tac = simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps rules) 1 val thm = Goal.prove_global thy [] [] goal (K tac) in add_qualified_simp_thm "chain_take" (dbind, thm) thy @@ -286,7 +287,7 @@ val lhs = take_const $ @{term "0::nat"} val goal = mk_eqs (lhs, mk_bottom (lhsT ->> lhsT)) val rules = take_defs @ @{thms iterate_0 fst_strict snd_strict} - val tac = simp_tac (HOL_basic_ss addsimps rules) 1 + val tac = simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps rules) 1 val take_0_thm = Goal.prove_global thy [] [] goal (K tac) in add_qualified_simp_thm "take_0" (dbind, take_0_thm) thy @@ -306,7 +307,7 @@ val goal = mk_eqs (lhs, rhs) val simps = @{thms iterate_Suc fst_conv snd_conv} val rules = take_defs @ simps - val tac = simp_tac (beta_ss addsimps rules) 1 + val tac = simp_tac (Simplifier.global_context thy beta_ss addsimps rules) 1 val take_Suc_thm = Goal.prove_global thy [] [] goal (K tac) in add_qualified_thm "take_Suc" (dbind, take_Suc_thm) thy @@ -332,8 +333,8 @@ Goal.prove_global thy [] [] goal (fn _ => EVERY [rtac @{thm nat.induct} 1, - simp_tac (HOL_basic_ss addsimps bottom_rules) 1, - asm_simp_tac (HOL_basic_ss addsimps take_Suc_thms) 1, + simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps bottom_rules) 1, + asm_simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps take_Suc_thms) 1, REPEAT (etac @{thm conjE} 1 ORELSE resolve_tac deflation_rules 1 ORELSE atac 1)]) @@ -456,8 +457,8 @@ @ @{thms decisive_ID decisive_ssum_map decisive_sprod_map} val tac = EVERY [ rtac @{thm nat.induct} 1, - simp_tac (HOL_ss addsimps rules0) 1, - asm_simp_tac (HOL_ss addsimps rules1) 1] + simp_tac (Simplifier.global_context thy HOL_ss addsimps rules0) 1, + asm_simp_tac (Simplifier.global_context thy HOL_ss addsimps rules1) 1] in Goal.prove_global thy [] [] goal (K tac) end fun conjuncts 1 thm = [thm] | conjuncts n thm = let diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tools/cont_proc.ML --- a/src/HOL/HOLCF/Tools/cont_proc.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tools/cont_proc.ML Thu Apr 18 17:07:01 2013 +0200 @@ -119,8 +119,9 @@ end local - fun solve_cont thy _ t = + fun solve_cont ctxt t = let + val thy = Proof_Context.theory_of ctxt val tr = instantiate' [] [SOME (cterm_of thy t)] @{thm Eq_TrueI} in Option.map fst (Seq.pull (cont_tac 1 tr)) end in @@ -128,6 +129,6 @@ Simplifier.simproc_global thy "cont_proc" ["cont f"] solve_cont end -fun setup thy = Simplifier.map_simpset_global (fn ss => ss addsimprocs [cont_proc thy]) thy +fun setup thy = map_theory_simpset (fn ctxt => ctxt addsimprocs [cont_proc thy]) thy end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tools/fixrec.ML --- a/src/HOL/HOLCF/Tools/fixrec.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tools/fixrec.ML Thu Apr 18 17:07:01 2013 +0200 @@ -132,7 +132,7 @@ Syntax.string_of_term lthy prop) val rules = Cont2ContData.get lthy val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules)) - val slow_tac = SOLVED' (simp_tac (simpset_of lthy)) + val slow_tac = SOLVED' (simp_tac lthy) val tac = fast_tac 1 ORELSE slow_tac 1 ORELSE err in Goal.prove lthy [] [] prop (K tac) @@ -293,7 +293,6 @@ fun fixrec_simp_tac ctxt = let val tab = FixrecUnfoldData.get (Context.Proof ctxt) - val ss = Simplifier.simpset_of ctxt val concl = HOLogic.dest_Trueprop o Logic.strip_imp_concl o strip_alls fun tac (t, i) = let @@ -302,7 +301,7 @@ val unfold_thm = the (Symtab.lookup tab c) val rule = unfold_thm RS @{thm ssubst_lhs} in - CHANGED (rtac rule i THEN eta_tac i THEN asm_simp_tac ss i) + CHANGED (rtac rule i THEN eta_tac i THEN asm_simp_tac ctxt i) end in SUBGOAL (fn ti => the_default no_tac (try tac ti)) @@ -311,9 +310,8 @@ (* proves a block of pattern matching equations as theorems, using unfold *) fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = let - val ss = Simplifier.simpset_of ctxt val rule = unfold_thm RS @{thm ssubst_lhs} - val tac = rtac rule 1 THEN eta_tac 1 THEN asm_simp_tac ss 1 + val tac = rtac rule 1 THEN eta_tac 1 THEN asm_simp_tac ctxt 1 fun prove_term t = Goal.prove ctxt [] [] t (K tac) fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t) in diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/Tr.thy --- a/src/HOL/HOLCF/Tr.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/Tr.thy Thu Apr 18 17:07:01 2013 +0200 @@ -150,9 +150,10 @@ apply (simp_all) done +(* FIXME unused!? *) ML {* -val split_If_tac = - simp_tac (HOL_basic_ss addsimps [@{thm If2_def} RS sym]) +fun split_If_tac ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm If2_def} RS sym]) THEN' (split_tac [@{thm split_If2}]) *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/ex/Focus_ex.thy --- a/src/HOL/HOLCF/ex/Focus_ex.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/ex/Focus_ex.thy Thu Apr 18 17:07:01 2013 +0200 @@ -180,7 +180,8 @@ done lemma lemma3: "def_g(g) --> is_g(g)" -apply (tactic {* simp_tac (HOL_ss addsimps [@{thm def_g_def}, @{thm lemma1}, @{thm lemma2}]) 1 *}) +apply (tactic {* simp_tac (put_simpset HOL_ss @{context} + addsimps [@{thm def_g_def}, @{thm lemma1}, @{thm lemma2}]) 1 *}) apply (rule impI) apply (erule exE) apply (rule_tac x = "f" in exI) @@ -204,7 +205,8 @@ done lemma lemma4: "is_g(g) --> def_g(g)" -apply (tactic {* simp_tac (HOL_ss delsimps (@{thms HOL.ex_simps} @ @{thms HOL.all_simps}) +apply (tactic {* simp_tac (put_simpset HOL_ss @{context} + delsimps (@{thms HOL.ex_simps} @ @{thms HOL.all_simps}) addsimps [@{thm lemma1}, @{thm lemma2}, @{thm def_g_def}]) 1 *}) apply (rule impI) apply (erule exE) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/HOLCF/ex/Pattern_Match.thy --- a/src/HOL/HOLCF/ex/Pattern_Match.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/HOLCF/ex/Pattern_Match.thy Thu Apr 18 17:07:01 2013 +0200 @@ -381,7 +381,8 @@ @{thms beta_cfun cont_id cont_const cont2cont_APP cont2cont_LAM'} @ @{thms cont2cont_fst cont2cont_snd cont2cont_Pair}; -val beta_ss = HOL_basic_ss addsimps (@{thms simp_thms} @ beta_rules); +val beta_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms simp_thms} @ beta_rules)); fun define_consts (specs : (binding * term * mixfix) list) @@ -557,7 +558,7 @@ val defs = @{thm branch_def} :: pat_defs; val goal = mk_trp (mk_strict fun1); val rules = @{thms match_bind_simps} @ case_rews; - val tacs = [simp_tac (beta_ss addsimps rules) 1]; + val tacs = [simp_tac (Simplifier.global_context thy beta_ss addsimps rules) 1]; in prove thy defs goal (K tacs) end; fun pat_apps (i, (pat, (con, args))) = let @@ -572,7 +573,7 @@ val goal = Logic.list_implies (assms, concl); val defs = @{thm branch_def} :: pat_defs; val rules = @{thms match_bind_simps} @ case_rews; - val tacs = [asm_simp_tac (beta_ss addsimps rules) 1]; + val tacs = [asm_simp_tac (Simplifier.global_context thy beta_ss addsimps rules) 1]; in prove thy defs goal (K tacs) end; in map_index pat_app spec end; in diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Hoare/Hoare_Logic.thy --- a/src/HOL/Hoare/Hoare_Logic.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Hoare/Hoare_Logic.thy Thu Apr 18 17:07:01 2013 +0200 @@ -101,7 +101,7 @@ method_setup vcg_simp = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *} + SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac ctxt))) *} "verification condition generator plus simplification" end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Hoare/Hoare_Logic_Abort.thy --- a/src/HOL/Hoare/Hoare_Logic_Abort.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy Thu Apr 18 17:07:01 2013 +0200 @@ -112,7 +112,7 @@ method_setup vcg_simp = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *} + SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac ctxt))) *} "verification condition generator plus simplification" (* Special syntax for guarded statements and guarded array updates: *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Hoare/hoare_tac.ML --- a/src/HOL/Hoare/hoare_tac.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Hoare/hoare_tac.ML Thu Apr 18 17:07:01 2013 +0200 @@ -97,10 +97,11 @@ (**Simp_tacs**) -val before_set2pred_simp_tac = - (simp_tac (HOL_basic_ss addsimps [Collect_conj_eq RS sym, @{thm Compl_Collect}])); +fun before_set2pred_simp_tac ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Collect_conj_eq RS sym, @{thm Compl_Collect}]); -val split_simp_tac = (simp_tac (HOL_basic_ss addsimps [@{thm split_conv}])); +fun split_simp_tac ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]); (*****************************************************************************) (** set2pred_tac transforms sets inclusion into predicates implication, **) @@ -114,14 +115,15 @@ (** simplification done by (split_all_tac) **) (*****************************************************************************) -fun set2pred_tac var_names = SUBGOAL (fn (goal, i) => - before_set2pred_simp_tac i THEN_MAYBE +fun set2pred_tac ctxt var_names = SUBGOAL (fn (goal, i) => + before_set2pred_simp_tac ctxt i THEN_MAYBE EVERY [ rtac subsetI i, rtac CollectI i, dtac CollectD i, - TRY (split_all_tac i) THEN_MAYBE - (rename_tac var_names i THEN full_simp_tac (HOL_basic_ss addsimps [@{thm split_conv}]) i)]); + TRY (split_all_tac ctxt i) THEN_MAYBE + (rename_tac var_names i THEN + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i)]); (*****************************************************************************) (** BasicSimpTac is called to simplify all verification conditions. It does **) @@ -131,17 +133,19 @@ (** the tactic chosen by the user, which may solve the subgoal completely. **) (*****************************************************************************) -fun MaxSimpTac var_names tac = FIRST'[rtac subset_refl, set2pred_tac var_names THEN_MAYBE' tac]; +fun MaxSimpTac ctxt var_names tac = + FIRST'[rtac subset_refl, set2pred_tac ctxt var_names THEN_MAYBE' tac]; -fun BasicSimpTac var_names tac = +fun BasicSimpTac ctxt var_names tac = simp_tac - (HOL_basic_ss addsimps [mem_Collect_eq, @{thm split_conv}] addsimprocs [Record.simproc]) - THEN_MAYBE' MaxSimpTac var_names tac; + (put_simpset HOL_basic_ss ctxt + addsimps [mem_Collect_eq, @{thm split_conv}] addsimprocs [Record.simproc]) + THEN_MAYBE' MaxSimpTac ctxt var_names tac; (** hoare_rule_tac **) -fun hoare_rule_tac (vars, Mlem) tac = +fun hoare_rule_tac ctxt (vars, Mlem) tac = let val var_names = map (fst o dest_Free) vars; fun wlp_tac i = @@ -155,16 +159,16 @@ EVERY [ rtac @{thm BasicRule} i, rtac Mlem i, - split_simp_tac i], + split_simp_tac ctxt i], EVERY [ rtac @{thm CondRule} i, rule_tac false (i + 2), rule_tac false (i + 1)], EVERY [ rtac @{thm WhileRule} i, - BasicSimpTac var_names tac (i + 2), + BasicSimpTac ctxt var_names tac (i + 2), rule_tac true (i + 1)]] - THEN (if pre_cond then BasicSimpTac var_names tac i else rtac subset_refl i))); + THEN (if pre_cond then BasicSimpTac ctxt var_names tac i else rtac subset_refl i))); in rule_tac end; @@ -172,5 +176,5 @@ (** the final verification conditions **) fun hoare_tac ctxt (tac: int -> tactic) = SUBGOAL (fn (goal, i) => - SELECT_GOAL (hoare_rule_tac (Mset ctxt goal) tac true 1) i); + SELECT_GOAL (hoare_rule_tac ctxt (Mset ctxt goal) tac true 1) i); diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Hoare_Parallel/Gar_Coll.thy --- a/src/HOL/Hoare_Parallel/Gar_Coll.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy Thu Apr 18 17:07:01 2013 +0200 @@ -769,7 +769,7 @@ apply interfree_aux apply(simp_all add:collector_mutator_interfree) apply(unfold modules collector_defs mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* TRYALL (interfree_aux_tac @{context}) *}) --{* 32 subgoals left *} apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) --{* 20 subgoals left *} @@ -790,7 +790,7 @@ apply interfree_aux apply(simp_all add:collector_mutator_interfree) apply(unfold modules collector_defs mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* TRYALL (interfree_aux_tac @{context}) *}) --{* 64 subgoals left *} apply(simp_all add:nth_list_update Invariants Append_to_free0)+ apply(tactic{* TRYALL (clarify_tac @{context}) *}) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy --- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Thu Apr 18 17:07:01 2013 +0200 @@ -131,7 +131,7 @@ apply(interfree_aux) apply(simp_all add:mul_mutator_interfree) apply(simp_all add: mul_mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* TRYALL (interfree_aux_tac @{context}) *}) apply(tactic {* ALLGOALS (clarify_tac @{context}) *}) apply (simp_all add:nth_list_update) done @@ -1171,7 +1171,7 @@ apply interfree_aux apply(simp_all add:mul_collector_mutator_interfree) apply(unfold mul_modules mul_collector_defs mul_mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* TRYALL (interfree_aux_tac @{context}) *}) --{* 42 subgoals left *} apply (clarify,simp add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)+ --{* 24 subgoals left *} @@ -1201,8 +1201,8 @@ apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac @{thm subset_psubset_trans}, etac @{thm Graph11},force_tac @{context}]) *}) --{* 41 subgoals left *} apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans}, - force_tac (map_simpset (fn ss => ss addsimps - [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}]) @{context})]) *}) + force_tac (@{context} addsimps + [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *}) --{* 35 subgoals left *} apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{context}]) *}) --{* 31 subgoals left *} @@ -1211,8 +1211,8 @@ apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac @{thm subset_psubset_trans},etac @{thm subset_psubset_trans},etac @{thm Graph11},force_tac @{context}]) *}) --{* 25 subgoals left *} apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans}, - force_tac (map_simpset (fn ss => ss addsimps - [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}]) @{context})]) *}) + force_tac (@{context} addsimps + [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *}) --{* 10 subgoals left *} apply(rule disjI2,rule disjI2,rule conjI,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update, rule disjI1, rule less_imp_le, erule less_le_trans, force simp add:Queue_def less_Suc_eq_le le_length_filter_update)+ done @@ -1225,7 +1225,7 @@ apply interfree_aux apply(simp_all add:mul_collector_mutator_interfree) apply(unfold mul_modules mul_collector_defs mul_mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* TRYALL (interfree_aux_tac @{context}) *}) --{* 76 subgoals left *} apply (clarsimp simp add: nth_list_update)+ --{* 56 subgoals left *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Hoare_Parallel/OG_Tactics.thy --- a/src/HOL/Hoare_Parallel/OG_Tactics.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Thu Apr 18 17:07:01 2013 +0200 @@ -273,11 +273,16 @@ lemmas ParallelConseq_list = INTER_eq Collect_conj_eq length_map length_upt length_append list_length ML {* -val before_interfree_simp_tac = simp_tac (HOL_basic_ss addsimps [@{thm com.simps}, @{thm post.simps}]) +fun before_interfree_simp_tac ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm com.simps}, @{thm post.simps}]) -val interfree_simp_tac = asm_simp_tac (HOL_ss addsimps [@{thm split}, @{thm ball_Un}, @{thm ball_empty}] @ @{thms my_simp_list}) +fun interfree_simp_tac ctxt = + asm_simp_tac (put_simpset HOL_ss ctxt + addsimps [@{thm split}, @{thm ball_Un}, @{thm ball_empty}] @ @{thms my_simp_list}) -val ParallelConseq = simp_tac (HOL_basic_ss addsimps @{thms ParallelConseq_list} @ @{thms my_simp_list}) +fun ParallelConseq ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt + addsimps @{thms ParallelConseq_list} @ @{thms my_simp_list}) *} text {* The following tactic applies @{text tac} to each conjunct in a @@ -320,120 +325,120 @@ ML {* - fun WlpTac i = (rtac (@{thm SeqRule}) i) THEN (HoareRuleTac false (i+1)) -and HoareRuleTac precond i st = st |> - ( (WlpTac i THEN HoareRuleTac precond i) +fun WlpTac ctxt i = (rtac (@{thm SeqRule}) i) THEN (HoareRuleTac ctxt false (i+1)) +and HoareRuleTac ctxt precond i st = st |> + ( (WlpTac ctxt i THEN HoareRuleTac ctxt precond i) ORELSE (FIRST[rtac (@{thm SkipRule}) i, rtac (@{thm BasicRule}) i, EVERY[rtac (@{thm ParallelConseqRule}) i, - ParallelConseq (i+2), - ParallelTac (i+1), - ParallelConseq i], + ParallelConseq ctxt (i+2), + ParallelTac ctxt (i+1), + ParallelConseq ctxt i], EVERY[rtac (@{thm CondRule}) i, - HoareRuleTac false (i+2), - HoareRuleTac false (i+1)], + HoareRuleTac ctxt false (i+2), + HoareRuleTac ctxt false (i+1)], EVERY[rtac (@{thm WhileRule}) i, - HoareRuleTac true (i+1)], + HoareRuleTac ctxt true (i+1)], K all_tac i ] THEN (if precond then (K all_tac i) else (rtac (@{thm subset_refl}) i)))) -and AnnWlpTac i = (rtac (@{thm AnnSeq}) i) THEN (AnnHoareRuleTac (i+1)) -and AnnHoareRuleTac i st = st |> - ( (AnnWlpTac i THEN AnnHoareRuleTac i ) +and AnnWlpTac ctxt i = (rtac (@{thm AnnSeq}) i) THEN (AnnHoareRuleTac ctxt (i+1)) +and AnnHoareRuleTac ctxt i st = st |> + ( (AnnWlpTac ctxt i THEN AnnHoareRuleTac ctxt i ) ORELSE (FIRST[(rtac (@{thm AnnskipRule}) i), EVERY[rtac (@{thm AnnatomRule}) i, - HoareRuleTac true (i+1)], + HoareRuleTac ctxt true (i+1)], (rtac (@{thm AnnwaitRule}) i), rtac (@{thm AnnBasic}) i, EVERY[rtac (@{thm AnnCond1}) i, - AnnHoareRuleTac (i+3), - AnnHoareRuleTac (i+1)], + AnnHoareRuleTac ctxt (i+3), + AnnHoareRuleTac ctxt (i+1)], EVERY[rtac (@{thm AnnCond2}) i, - AnnHoareRuleTac (i+1)], + AnnHoareRuleTac ctxt (i+1)], EVERY[rtac (@{thm AnnWhile}) i, - AnnHoareRuleTac (i+2)], + AnnHoareRuleTac ctxt (i+2)], EVERY[rtac (@{thm AnnAwait}) i, - HoareRuleTac true (i+1)], + HoareRuleTac ctxt true (i+1)], K all_tac i])) -and ParallelTac i = EVERY[rtac (@{thm ParallelRule}) i, - interfree_Tac (i+1), - MapAnn_Tac i] +and ParallelTac ctxt i = EVERY[rtac (@{thm ParallelRule}) i, + interfree_Tac ctxt (i+1), + MapAnn_Tac ctxt i] -and MapAnn_Tac i st = st |> +and MapAnn_Tac ctxt i st = st |> (FIRST[rtac (@{thm MapAnnEmpty}) i, EVERY[rtac (@{thm MapAnnList}) i, - MapAnn_Tac (i+1), - AnnHoareRuleTac i], + MapAnn_Tac ctxt (i+1), + AnnHoareRuleTac ctxt i], EVERY[rtac (@{thm MapAnnMap}) i, - rtac (@{thm allI}) i,rtac (@{thm impI}) i, - AnnHoareRuleTac i]]) + rtac (@{thm allI}) i, rtac (@{thm impI}) i, + AnnHoareRuleTac ctxt i]]) -and interfree_swap_Tac i st = st |> +and interfree_swap_Tac ctxt i st = st |> (FIRST[rtac (@{thm interfree_swap_Empty}) i, EVERY[rtac (@{thm interfree_swap_List}) i, - interfree_swap_Tac (i+2), - interfree_aux_Tac (i+1), - interfree_aux_Tac i ], + interfree_swap_Tac ctxt (i+2), + interfree_aux_Tac ctxt (i+1), + interfree_aux_Tac ctxt i ], EVERY[rtac (@{thm interfree_swap_Map}) i, rtac (@{thm allI}) i,rtac (@{thm impI}) i, - conjI_Tac (interfree_aux_Tac) i]]) + conjI_Tac (interfree_aux_Tac ctxt) i]]) -and interfree_Tac i st = st |> +and interfree_Tac ctxt i st = st |> (FIRST[rtac (@{thm interfree_Empty}) i, EVERY[rtac (@{thm interfree_List}) i, - interfree_Tac (i+1), - interfree_swap_Tac i], + interfree_Tac ctxt (i+1), + interfree_swap_Tac ctxt i], EVERY[rtac (@{thm interfree_Map}) i, rtac (@{thm allI}) i,rtac (@{thm allI}) i,rtac (@{thm impI}) i, - interfree_aux_Tac i ]]) + interfree_aux_Tac ctxt i ]]) -and interfree_aux_Tac i = (before_interfree_simp_tac i ) THEN +and interfree_aux_Tac ctxt i = (before_interfree_simp_tac ctxt i ) THEN (FIRST[rtac (@{thm interfree_aux_rule1}) i, - dest_assertions_Tac i]) + dest_assertions_Tac ctxt i]) -and dest_assertions_Tac i st = st |> +and dest_assertions_Tac ctxt i st = st |> (FIRST[EVERY[rtac (@{thm AnnBasic_assertions}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], + dest_atomics_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnSeq_assertions}) i, - dest_assertions_Tac (i+1), - dest_assertions_Tac i], + dest_assertions_Tac ctxt (i+1), + dest_assertions_Tac ctxt i], EVERY[rtac (@{thm AnnCond1_assertions}) i, - dest_assertions_Tac (i+2), - dest_assertions_Tac (i+1), - dest_atomics_Tac i], + dest_assertions_Tac ctxt (i+2), + dest_assertions_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnCond2_assertions}) i, - dest_assertions_Tac (i+1), - dest_atomics_Tac i], + dest_assertions_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnWhile_assertions}) i, - dest_assertions_Tac (i+2), - dest_atomics_Tac (i+1), - dest_atomics_Tac i], + dest_assertions_Tac ctxt (i+2), + dest_atomics_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnAwait_assertions}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], - dest_atomics_Tac i]) + dest_atomics_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], + dest_atomics_Tac ctxt i]) -and dest_atomics_Tac i st = st |> +and dest_atomics_Tac ctxt i st = st |> (FIRST[EVERY[rtac (@{thm AnnBasic_atomics}) i, - HoareRuleTac true i], + HoareRuleTac ctxt true i], EVERY[rtac (@{thm AnnSeq_atomics}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], + dest_atomics_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnCond1_atomics}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], + dest_atomics_Tac ctxt (i+1), + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnCond2_atomics}) i, - dest_atomics_Tac i], + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm AnnWhile_atomics}) i, - dest_atomics_Tac i], + dest_atomics_Tac ctxt i], EVERY[rtac (@{thm Annatom_atomics}) i, - HoareRuleTac true i], + HoareRuleTac ctxt true i], EVERY[rtac (@{thm AnnAwait_atomics}) i, - HoareRuleTac true i], + HoareRuleTac ctxt true i], K all_tac i]) *} @@ -441,8 +446,7 @@ text {* The final tactic is given the name @{text oghoare}: *} ML {* -val oghoare_tac = SUBGOAL (fn (_, i) => - (HoareRuleTac true i)) +fun oghoare_tac ctxt = SUBGOAL (fn (_, i) => HoareRuleTac ctxt true i) *} text {* Notice that the tactic for parallel programs @{text @@ -453,26 +457,25 @@ verification conditions for annotated sequential programs and to generate verification conditions out of interference freedom tests: *} -ML {* val annhoare_tac = SUBGOAL (fn (_, i) => - (AnnHoareRuleTac i)) +ML {* +fun annhoare_tac ctxt = SUBGOAL (fn (_, i) => AnnHoareRuleTac ctxt i) -val interfree_aux_tac = SUBGOAL (fn (_, i) => - (interfree_aux_Tac i)) +fun interfree_aux_tac ctxt = SUBGOAL (fn (_, i) => interfree_aux_Tac ctxt i) *} text {* The so defined ML tactics are then ``exported'' to be used in Isabelle proofs. *} method_setup oghoare = {* - Scan.succeed (K (SIMPLE_METHOD' oghoare_tac)) *} + Scan.succeed (SIMPLE_METHOD' o oghoare_tac) *} "verification condition generator for the oghoare logic" method_setup annhoare = {* - Scan.succeed (K (SIMPLE_METHOD' annhoare_tac)) *} + Scan.succeed (SIMPLE_METHOD' o annhoare_tac) *} "verification condition generator for the ann_hoare logic" method_setup interfree_aux = {* - Scan.succeed (K (SIMPLE_METHOD' interfree_aux_tac)) *} + Scan.succeed (SIMPLE_METHOD' o interfree_aux_tac) *} "verification condition generator for interference freedom tests" text {* Tactics useful for dealing with the generated verification conditions: *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/IMPP/Hoare.thy --- a/src/HOL/IMPP/Hoare.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/IMPP/Hoare.thy Thu Apr 18 17:07:01 2013 +0200 @@ -287,7 +287,7 @@ apply (blast) (* weaken *) apply (tactic {* ALLGOALS (EVERY' [REPEAT o thin_tac @{context} "hoare_derivs ?x ?y", - simp_tac @{simpset}, clarify_tac @{context}, REPEAT o smp_tac 1]) *}) + simp_tac @{context}, clarify_tac @{context}, REPEAT o smp_tac 1]) *}) apply (simp_all (no_asm_use) add: triple_valid_def2) apply (intro strip, tactic "smp_tac 2 1", blast) (* conseq *) apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *}) (* Skip, Ass, Local *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/IOA/Solve.thy --- a/src/HOL/IOA/Solve.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/IOA/Solve.thy Thu Apr 18 17:07:01 2013 +0200 @@ -146,7 +146,7 @@ apply (simp (no_asm) add: conj_disj_distribR cong add: conj_cong split add: split_if) apply (tactic {* REPEAT((resolve_tac [conjI,impI] 1 ORELSE etac conjE 1) THEN - asm_full_simp_tac(@{simpset} addsimps [@{thm comp1_reachable}, @{thm comp2_reachable}]) 1) *}) + asm_full_simp_tac(@{context} addsimps [@{thm comp1_reachable}, @{thm comp2_reachable}]) 1) *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Isar_Examples/Hoare.thy --- a/src/HOL/Isar_Examples/Hoare.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Isar_Examples/Hoare.thy Thu Apr 18 17:07:01 2013 +0200 @@ -406,7 +406,8 @@ method_setup hoare = {* Scan.succeed (fn ctxt => (SIMPLE_METHOD' - (hoare_tac ctxt (simp_tac (HOL_basic_ss addsimps [@{thm "Record.K_record_comp"}] ))))) *} + (hoare_tac ctxt + (simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm "Record.K_record_comp"}] ))))) *} "verification condition generator for Hoare logic" end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Library/Extended_Nat.thy --- a/src/HOL/Library/Extended_Nat.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Library/Extended_Nat.thy Thu Apr 18 17:07:01 2013 +0200 @@ -507,11 +507,12 @@ fun dest_sum t = dest_summing (t, []) val find_first = find_first_t [] val trans_tac = Numeral_Simprocs.trans_tac - val norm_ss = HOL_basic_ss addsimps - @{thms add_ac add_0_left add_0_right} - fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) - fun simplify_meta_eq ss cancel_th th = - Arith_Data.simplify_meta_eq [] ss + val norm_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps @{thms add_ac add_0_left add_0_right}) + fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) + fun simplify_meta_eq ctxt cancel_th th = + Arith_Data.simplify_meta_eq [] ctxt ([th, cancel_th] MRS trans) fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end @@ -540,15 +541,15 @@ simproc_setup enat_eq_cancel ("(l::enat) + m = n" | "(l::enat) = m + n") = - {* fn phi => fn ss => fn ct => Eq_Enat_Cancel.proc ss (term_of ct) *} + {* fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (term_of ct) *} simproc_setup enat_le_cancel ("(l::enat) + m \ n" | "(l::enat) \ m + n") = - {* fn phi => fn ss => fn ct => Le_Enat_Cancel.proc ss (term_of ct) *} + {* fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (term_of ct) *} simproc_setup enat_less_cancel ("(l::enat) + m < n" | "(l::enat) < m + n") = - {* fn phi => fn ss => fn ct => Less_Enat_Cancel.proc ss (term_of ct) *} + {* fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (term_of ct) *} text {* TODO: add regression tests for these simprocs *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Library/Sum_of_Squares/sum_of_squares.ML --- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Thu Apr 18 17:07:01 2013 +0200 @@ -723,7 +723,9 @@ in (let val th = tryfind trivial_axiom (keq @ klep @ kltp) in - (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv Numeral_Simprocs.field_comp_conv) th, RealArith.Trivial) + (fconv_rule (arg_conv (arg1_conv (real_poly_conv ctxt)) + then_conv Numeral_Simprocs.field_comp_conv ctxt) th, + RealArith.Trivial) end) handle Failure _ => (let val proof = @@ -820,8 +822,8 @@ let val (c,v) = substitutable_monomial [] (Thm.dest_arg1(concl th)) val th1 = Drule.arg_cong_rule (Thm.apply @{cterm "op * :: real => _"} (RealArith.cterm_of_rat (Rat.inv c))) (mk_meta_eq th) - val th2 = fconv_rule (binop_conv real_poly_mul_conv) th1 - in fconv_rule (arg_conv real_poly_conv) (isolate_variable v th2) + val th2 = fconv_rule (binop_conv (real_poly_mul_conv ctxt)) th1 + in fconv_rule (arg_conv (real_poly_conv ctxt)) (isolate_variable v th2) end fun oprconv cv ct = let val g = Thm.dest_fun2 ct @@ -834,7 +836,8 @@ fun substfirst(eqs,les,lts) = ((let val eth = tryfind make_substitution eqs - val modify = fconv_rule (arg_conv (oprconv(subst_conv [eth] then_conv real_poly_conv))) + val modify = + fconv_rule (arg_conv (oprconv(subst_conv [eth] then_conv (real_poly_conv ctxt)))) in substfirst (filter_out (fn t => (Thm.dest_arg1 o Thm.dest_arg o cprop_of) t aconvc @{cterm "0::real"}) (map modify eqs), @@ -922,12 +925,13 @@ NONE => no_tac | SOME (d,ord) => let - val ss = simpset_of ctxt addsimps @{thms field_simps} - addsimps [@{thm nonzero_power_divide}, @{thm power_divide}] + val simp_ctxt = + ctxt addsimps @{thms field_simps} + addsimps [@{thm nonzero_power_divide}, @{thm power_divide}] val th = instantiate' [] [SOME d, SOME (Thm.dest_arg P)] (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto} else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast}) - in rtac th i THEN Simplifier.asm_full_simp_tac ss i end); + in rtac th i THEN Simplifier.asm_full_simp_tac simp_ctxt i end); fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i); diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Library/positivstellensatz.ML --- a/src/HOL/Library/positivstellensatz.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Library/positivstellensatz.ML Thu Apr 18 17:07:01 2013 +0200 @@ -358,15 +358,15 @@ poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv, absconv1,absconv2,prover) = let - val pre_ss = HOL_basic_ss addsimps + val pre_ss = put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps all_simps not_all not_ex ex_disj_distrib all_conj_distrib if_bool_eq_disj} - val prenex_ss = HOL_basic_ss addsimps prenex_simps - val skolemize_ss = HOL_basic_ss addsimps [choice_iff] - val presimp_conv = Simplifier.rewrite (Simplifier.context ctxt pre_ss) - val prenex_conv = Simplifier.rewrite (Simplifier.context ctxt prenex_ss) - val skolemize_conv = Simplifier.rewrite (Simplifier.context ctxt skolemize_ss) - val weak_dnf_ss = HOL_basic_ss addsimps weak_dnf_simps - val weak_dnf_conv = Simplifier.rewrite (Simplifier.context ctxt weak_dnf_ss) + val prenex_ss = put_simpset HOL_basic_ss ctxt addsimps prenex_simps + val skolemize_ss = put_simpset HOL_basic_ss ctxt addsimps [choice_iff] + val presimp_conv = Simplifier.rewrite pre_ss + val prenex_conv = Simplifier.rewrite prenex_ss + val skolemize_conv = Simplifier.rewrite skolemize_ss + val weak_dnf_ss = put_simpset HOL_basic_ss ctxt addsimps weak_dnf_simps + val weak_dnf_conv = Simplifier.rewrite weak_dnf_ss fun eqT_elim th = Thm.equal_elim (Thm.symmetric th) @{thm TrueI} fun oprconv cv ct = let val g = Thm.dest_fun2 ct @@ -423,7 +423,7 @@ end val init_conv = presimp_conv then_conv - nnf_conv then_conv skolemize_conv then_conv prenex_conv then_conv + nnf_conv ctxt then_conv skolemize_conv then_conv prenex_conv then_conv weak_dnf_conv val concl = Thm.dest_arg o cprop_of @@ -540,7 +540,7 @@ fun f ct = let val nnf_norm_conv' = - nnf_conv then_conv + nnf_conv ctxt then_conv literals_conv [@{term HOL.conj}, @{term HOL.disj}] [] (Conv.cache_conv (first_conv [real_lt_conv, real_le_conv, @@ -701,9 +701,10 @@ (* A less general generic arithmetic prover dealing with abs,max and min*) local - val absmaxmin_elim_ss1 = HOL_basic_ss addsimps real_abs_thms1 + val absmaxmin_elim_ss1 = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps real_abs_thms1) fun absmaxmin_elim_conv1 ctxt = - Simplifier.rewrite (Simplifier.context ctxt absmaxmin_elim_ss1) + Simplifier.rewrite (put_simpset absmaxmin_elim_ss1 ctxt) val absmaxmin_elim_conv2 = let @@ -758,8 +759,11 @@ (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord in gen_real_arith ctxt - (cterm_of_rat, Numeral_Simprocs.field_comp_conv, Numeral_Simprocs.field_comp_conv, Numeral_Simprocs.field_comp_conv, - main,neg,add,mul, prover) + (cterm_of_rat, + Numeral_Simprocs.field_comp_conv ctxt, + Numeral_Simprocs.field_comp_conv ctxt, + Numeral_Simprocs.field_comp_conv ctxt, + main ctxt, neg ctxt, add ctxt, mul ctxt, prover) end; end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Library/reflection.ML --- a/src/HOL/Library/reflection.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Library/reflection.ML Thu Apr 18 17:07:01 2013 +0200 @@ -275,7 +275,7 @@ val th' = Drule.instantiate_normalize ([], cvs) th val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th' val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) - (fn _ => simp_tac (simpset_of ctxt) 1) + (fn _ => simp_tac ctxt 1) in FWD trans [th'',th'] end @@ -290,8 +290,9 @@ val ft = (Thm.dest_arg1 o Thm.dest_arg o Thm.dest_arg o cprop_of) th val rth = conv ft in - simplify (HOL_basic_ss addsimps raw_eqs addsimps @{thms nth_Cons_0 nth_Cons_Suc}) - (simplify (HOL_basic_ss addsimps [rth]) th) + simplify + (put_simpset HOL_basic_ss ctxt addsimps raw_eqs addsimps @{thms nth_Cons_0 nth_Cons_Suc}) + (simplify (put_simpset HOL_basic_ss ctxt addsimps [rth]) th) end fun genreify_tac ctxt eqs to = SUBGOAL (fn (goal, i) => diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/List.thy --- a/src/HOL/List.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/List.thy Thu Apr 18 17:07:01 2013 +0200 @@ -489,7 +489,7 @@ signature LIST_TO_SET_COMPREHENSION = sig - val simproc : simpset -> cterm -> thm option + val simproc : Proof.context -> cterm -> thm option end structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION = @@ -529,9 +529,8 @@ datatype termlets = If | Case of (typ * int) -fun simproc ss redex = +fun simproc ctxt redex = let - val ctxt = Simplifier.the_context ss val thy = Proof_Context.theory_of ctxt val set_Nil_I = @{thm trans} OF [@{thm set.simps(1)}, @{thm empty_def}] val set_singleton = @{lemma "set [a] = {x. x = a}" by simp} @@ -836,7 +835,9 @@ | len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc | len t (ts,n) = (t::ts,n); -fun list_neq _ ss ct = +val ss = simpset_of @{context}; + +fun list_neq ctxt ct = let val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct; val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0); @@ -846,15 +847,15 @@ val size = HOLogic.size_const listT; val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs); val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len); - val thm = Goal.prove (Simplifier.the_context ss) [] [] neq_len - (K (simp_tac (Simplifier.inherit_context ss @{simpset}) 1)); + val thm = Goal.prove ctxt [] [] neq_len + (K (simp_tac (put_simpset ss ctxt) 1)); in SOME (thm RS @{thm neq_if_length_neq}) end in if m < n andalso submultiset (op aconv) (ls,rs) orelse n < m andalso submultiset (op aconv) (rs,ls) then prove_neq() else NONE end; -in list_neq end; +in K list_neq end; *} @@ -972,9 +973,10 @@ | butlast xs = Const(@{const_name Nil}, fastype_of xs); val rearr_ss = - HOL_basic_ss addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]; + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]); - fun list_eq ss (F as (eq as Const(_,eqT)) $ lhs $ rhs) = + fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) = let val lastl = last lhs and lastr = last rhs; fun rearr conv = @@ -985,15 +987,15 @@ val app = Const(@{const_name append},appT) val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2)); - val thm = Goal.prove (Simplifier.the_context ss) [] [] eq - (K (simp_tac (Simplifier.inherit_context ss rearr_ss) 1)); + val thm = Goal.prove ctxt [] [] eq + (K (simp_tac (put_simpset rearr_ss ctxt) 1)); in SOME ((conv RS (thm RS trans)) RS eq_reflection) end; in if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv} else if lastl aconv lastr then rearr @{thm append_same_eq} else NONE end; - in fn _ => fn ss => fn ct => list_eq ss (term_of ct) end; + in fn _ => fn ctxt => fn ct => list_eq ctxt (term_of ct) end; *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/MicroJava/J/JTypeSafe.thy --- a/src/HOL/MicroJava/J/JTypeSafe.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/MicroJava/J/JTypeSafe.thy Thu Apr 18 17:07:01 2013 +0200 @@ -200,7 +200,7 @@ apply( simp_all) apply( tactic "ALLGOALS (REPEAT o resolve_tac [impI, allI])") apply( tactic {* ALLGOALS (eresolve_tac [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}] - THEN_ALL_NEW (full_simp_tac (simpset_of @{theory_context Conform}))) *}) + THEN_ALL_NEW (full_simp_tac (put_simpset (simpset_of @{theory_context Conform}) @{context}))) *}) apply(tactic "ALLGOALS (EVERY' [REPEAT o (etac conjE), REPEAT o hyp_subst_tac])") -- "Level 7" @@ -240,11 +240,11 @@ -- "for FAss" apply( tactic {* EVERY'[eresolve_tac [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}] - THEN_ALL_NEW (full_simp_tac @{simpset}), REPEAT o (etac conjE), hyp_subst_tac] 3*}) + THEN_ALL_NEW (full_simp_tac @{context}), REPEAT o (etac conjE), hyp_subst_tac] 3*}) -- "for if" apply( tactic {* (Induct_Tacs.case_tac @{context} "the_Bool v" THEN_ALL_NEW - (asm_full_simp_tac @{simpset})) 7*}) + (asm_full_simp_tac @{context})) 7*}) apply (tactic "forward_hyp_tac") @@ -276,7 +276,7 @@ -- "7 LAss" apply (fold fun_upd_def) apply( tactic {* (eresolve_tac [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}] - THEN_ALL_NEW (full_simp_tac @{simpset})) 1 *}) + THEN_ALL_NEW (full_simp_tac @{context})) 1 *}) apply (intro strip) apply (case_tac E) apply (simp) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/MicroJava/J/WellForm.thy --- a/src/HOL/MicroJava/J/WellForm.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/MicroJava/J/WellForm.thy Thu Apr 18 17:07:01 2013 +0200 @@ -491,7 +491,8 @@ apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl) prefer 2 apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl) -apply( tactic "asm_full_simp_tac (HOL_ss addsimps [@{thm not_None_eq} RS sym]) 1") +apply( tactic "asm_full_simp_tac + (put_simpset HOL_ss @{context} addsimps [@{thm not_None_eq} RS sym]) 1") apply( simp_all (no_asm_simp) del: split_paired_Ex) apply( frule (1) class_wf) apply( simp (no_asm_simp) only: split_tupled_all) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Apr 18 17:07:01 2013 +0200 @@ -113,24 +113,27 @@ method_setup vector = {* let - val ss1 = HOL_basic_ss addsimps [@{thm setsum_addf} RS sym, - @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, - @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym] - val ss2 = @{simpset} addsimps + val ss1 = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm setsum_addf} RS sym, + @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, + @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym]) + val ss2 = + simpset_of (@{context} addsimps [@{thm plus_vec_def}, @{thm times_vec_def}, @{thm minus_vec_def}, @{thm uminus_vec_def}, @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def}, @{thm scaleR_vec_def}, - @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}] - fun vector_arith_tac ths = - simp_tac ss1 + @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}]) + fun vector_arith_tac ctxt ths = + simp_tac (put_simpset ss1 ctxt) THEN' (fn i => rtac @{thm setsum_cong2} i ORELSE rtac @{thm setsum_0'} i - ORELSE simp_tac (HOL_basic_ss addsimps [@{thm vec_eq_iff}]) i) + ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i) (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) - THEN' asm_full_simp_tac (ss2 addsimps ths) + THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths) in - Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths))) + Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths)) end *} "lift trivial vector statements to real arith statements" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Multivariate_Analysis/normarith.ML --- a/src/HOL/Multivariate_Analysis/normarith.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Multivariate_Analysis/normarith.ML Thu Apr 18 17:07:01 2013 +0200 @@ -165,7 +165,9 @@ val real_poly_conv = Semiring_Normalizer.semiring_normalize_wrapper ctxt (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"})) - in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.field_comp_conv then_conv real_poly_conv))) + in + fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv + arg_conv (Numeral_Simprocs.field_comp_conv ctxt then_conv real_poly_conv))) end; val apply_pth1 = rewr_conv @{thm pth_1}; @@ -175,8 +177,13 @@ val apply_pth5 = rewr_conv @{thm pth_5}; val apply_pth6 = rewr_conv @{thm pth_6}; val apply_pth7 = rewrs_conv @{thms pth_7}; - val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv Numeral_Simprocs.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left}))); - val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Numeral_Simprocs.field_comp_conv); + fun apply_pth8 ctxt = + rewr_conv @{thm pth_8} then_conv + arg1_conv (Numeral_Simprocs.field_comp_conv ctxt) then_conv + (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left}))); + fun apply_pth9 ctxt = + rewrs_conv @{thms pth_9} then_conv + arg1_conv (arg1_conv (Numeral_Simprocs.field_comp_conv ctxt)); val apply_ptha = rewr_conv @{thm pth_a}; val apply_pthb = rewrs_conv @{thms pth_b}; val apply_pthc = rewrs_conv @{thms pth_c}; @@ -188,13 +195,13 @@ | Const(@{const_name scaleR}, _)$_$v => v | _ => error "headvector: non-canonical term" -fun vector_cmul_conv ct = - ((apply_pth5 then_conv arg1_conv Numeral_Simprocs.field_comp_conv) else_conv - (apply_pth6 then_conv binop_conv vector_cmul_conv)) ct +fun vector_cmul_conv ctxt ct = + ((apply_pth5 then_conv arg1_conv (Numeral_Simprocs.field_comp_conv ctxt)) else_conv + (apply_pth6 then_conv binop_conv (vector_cmul_conv ctxt))) ct -fun vector_add_conv ct = apply_pth7 ct +fun vector_add_conv ctxt ct = apply_pth7 ct handle CTERM _ => - (apply_pth8 ct + (apply_pth8 ctxt ct handle CTERM _ => (case term_of ct of Const(@{const_name plus},_)$lt$rt => @@ -202,35 +209,35 @@ val l = headvector lt val r = headvector rt in (case Term_Ord.fast_term_ord (l,r) of - LESS => (apply_pthb then_conv arg_conv vector_add_conv + LESS => (apply_pthb then_conv arg_conv (vector_add_conv ctxt) then_conv apply_pthd) ct - | GREATER => (apply_pthc then_conv arg_conv vector_add_conv + | GREATER => (apply_pthc then_conv arg_conv (vector_add_conv ctxt) then_conv apply_pthd) ct - | EQUAL => (apply_pth9 then_conv - ((apply_ptha then_conv vector_add_conv) else_conv - arg_conv vector_add_conv then_conv apply_pthd)) ct) + | EQUAL => (apply_pth9 ctxt then_conv + ((apply_ptha then_conv (vector_add_conv ctxt)) else_conv + arg_conv (vector_add_conv ctxt) then_conv apply_pthd)) ct) end | _ => Thm.reflexive ct)) -fun vector_canon_conv ct = case term_of ct of +fun vector_canon_conv ctxt ct = case term_of ct of Const(@{const_name plus},_)$_$_ => let val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb - val lth = vector_canon_conv l - val rth = vector_canon_conv r + val lth = vector_canon_conv ctxt l + val rth = vector_canon_conv ctxt r val th = Drule.binop_cong_rule p lth rth - in fconv_rule (arg_conv vector_add_conv) th end + in fconv_rule (arg_conv (vector_add_conv ctxt)) th end | Const(@{const_name scaleR}, _)$_$_ => let val (p,r) = Thm.dest_comb ct - val rth = Drule.arg_cong_rule p (vector_canon_conv r) - in fconv_rule (arg_conv (apply_pth4 else_conv vector_cmul_conv)) rth + val rth = Drule.arg_cong_rule p (vector_canon_conv ctxt r) + in fconv_rule (arg_conv (apply_pth4 else_conv (vector_cmul_conv ctxt))) rth end -| Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv vector_canon_conv) ct +| Const(@{const_name minus},_)$_$_ => (apply_pth2 then_conv (vector_canon_conv ctxt)) ct -| Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv vector_canon_conv) ct +| Const(@{const_name uminus},_)$_ => (apply_pth3 then_conv (vector_canon_conv ctxt)) ct (* FIXME | Const(@{const_name vec},_)$n => @@ -241,8 +248,8 @@ *) | _ => apply_pth1 ct -fun norm_canon_conv ct = case term_of ct of - Const(@{const_name norm},_)$_ => arg_conv vector_canon_conv ct +fun norm_canon_conv ctxt ct = case term_of ct of + Const(@{const_name norm},_)$_ => arg_conv (vector_canon_conv ctxt) ct | _ => raise CTERM ("norm_canon_conv", [ct]) fun int_flip v eq = @@ -314,9 +321,9 @@ in fst (RealArith.real_linear_prover translator (map (fn t => Drule.instantiate_normalize ([(tv_n, ctyp_of_term t)],[]) pth_zero) zerodests, - map (fconv_rule (try_conv (Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv + map (fconv_rule (try_conv (Conv.top_sweep_conv (K (norm_canon_conv ctxt)) ctxt) then_conv arg_conv (arg_conv real_poly_conv))) ges', - map (fconv_rule (try_conv (Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv + map (fconv_rule (try_conv (Conv.top_sweep_conv (K (norm_canon_conv ctxt)) ctxt) then_conv arg_conv (arg_conv real_poly_conv))) gts)) end in val real_vector_combo_prover = real_vector_combo_prover @@ -367,7 +374,7 @@ (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt (the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord) val (th1,th2) = conj_pair(rawrule th) - in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc + in th1::fconv_rule (arg_conv (arg_conv (real_poly_neg_conv ctxt))) th2::acc end in fun real_vector_prover ctxt _ translator (eqs,ges,gts) = (real_vector_ineq_prover ctxt translator @@ -375,10 +382,11 @@ end; fun init_conv ctxt = - Simplifier.rewrite (Simplifier.context ctxt - (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths}))) - then_conv Numeral_Simprocs.field_comp_conv - then_conv nnf_conv + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt + addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm right_minus}, + @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})) + then_conv Numeral_Simprocs.field_comp_conv ctxt + then_conv nnf_conv ctxt fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt); fun norm_arith ctxt ct = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/NanoJava/Equivalence.thy --- a/src/HOL/NanoJava/Equivalence.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/NanoJava/Equivalence.thy Thu Apr 18 17:07:01 2013 +0200 @@ -101,7 +101,7 @@ by(simp add: cnvalids_def nvalids_def nvalid_def2) lemma hoare_sound_main:"\t. (A |\ C \ A |\ C) \ (A |\\<^sub>e t \ A |\\<^sub>e t)" -apply (tactic "split_all_tac 1", rename_tac P e Q) +apply (tactic "split_all_tac @{context} 1", rename_tac P e Q) apply (rule hoare_ehoare.induct) (*18*) apply (tactic {* ALLGOALS (REPEAT o dresolve_tac [@{thm all_conjunct2}, @{thm all3_conjunct2}]) *}) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_atoms.ML --- a/src/HOL/Nominal/nominal_atoms.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_atoms.ML Thu Apr 18 17:07:01 2013 +0200 @@ -116,23 +116,23 @@ val simp1 = @{thm inj_on_def} :: injects; - val proof1 = fn _ => EVERY [simp_tac (HOL_basic_ss addsimps simp1) 1, + fun proof1 ctxt = EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1, rtac @{thm ballI} 1, rtac @{thm ballI} 1, rtac @{thm impI} 1, atac 1] val (inj_thm,thy2) = - add_thms_string [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 proof1), [])] thy1 + add_thms_string [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 (proof1 o #context)), [])] thy1 (* second statement *) val y = Free ("y",ak_type) val stmnt2 = HOLogic.mk_Trueprop (HOLogic.mk_exists ("x",@{typ nat},HOLogic.mk_eq (y,Const (ak_sign,inj_type) $ Bound 0))) - val proof2 = fn {prems, context} => - Induct_Tacs.case_tac context "y" 1 THEN - asm_simp_tac (HOL_basic_ss addsimps simp1) 1 THEN + val proof2 = fn {prems, context = ctxt} => + Induct_Tacs.case_tac ctxt "y" 1 THEN + asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1 THEN rtac @{thm exI} 1 THEN rtac @{thm refl} 1 @@ -148,13 +148,13 @@ val simp2 = [@{thm image_def},@{thm bex_UNIV}]@inject_thm val simp3 = [@{thm UNIV_def}] - val proof3 = fn _ => EVERY [cut_facts_tac inj_thm 1, + fun proof3 ctxt = EVERY [cut_facts_tac inj_thm 1, dtac @{thm range_inj_infinite} 1, - asm_full_simp_tac (HOL_basic_ss addsimps simp2) 1, - simp_tac (HOL_basic_ss addsimps simp3) 1] + asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp2) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp3) 1] val (inf_thm,thy4) = - add_thms_string [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 proof3), [])] thy3 + add_thms_string [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 (proof3 o #context)), [])] thy3 in ((inj_thm,inject_thm,inf_thm),thy4) end) ak_names thy @@ -267,21 +267,19 @@ val i_type = Type(ak_name_qu,[]); val cat = Const ("Nominal.at",(Term.itselfT i_type) --> HOLogic.boolT); val at_type = Logic.mk_type i_type; - val simp_s = HOL_ss addsimps maps (Global_Theory.get_thms thy5) + fun proof ctxt = + simp_tac (put_simpset HOL_ss ctxt + addsimps maps (Global_Theory.get_thms thy5) ["at_def", ak_name ^ "_prm_" ^ ak_name ^ "_def", ak_name ^ "_prm_" ^ ak_name ^ ".simps", "swap_" ^ ak_name ^ "_def", "swap_" ^ ak_name ^ ".simps", - ak_name ^ "_infinite"] - + ak_name ^ "_infinite"]) 1; val name = "at_"^ak_name^ "_inst"; val statement = HOLogic.mk_Trueprop (cat $ at_type); - - val proof = fn _ => simp_tac simp_s 1 - in - ((name, Goal.prove_global thy5 [] [] statement proof), []) + ((name, Goal.prove_global thy5 [] [] statement (proof o #context)), []) end) ak_names_types); (* declares a perm-axclass for every atom-kind *) @@ -331,18 +329,17 @@ val cpt = Const ("Nominal.pt",(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); val pt_type = Logic.mk_type i_type1; val at_type = Logic.mk_type i_type2; - val simp_s = HOL_ss addsimps maps (Global_Theory.get_thms thy7) + fun proof ctxt = + simp_tac (put_simpset HOL_ss ctxt addsimps maps (Global_Theory.get_thms thy7) ["pt_def", "pt_" ^ ak_name ^ "1", "pt_" ^ ak_name ^ "2", - "pt_" ^ ak_name ^ "3"]; + "pt_" ^ ak_name ^ "3"]) 1; val name = "pt_"^ak_name^ "_inst"; val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type); - - val proof = fn _ => simp_tac simp_s 1; in - ((name, Goal.prove_global thy7 [] [] statement proof), []) + ((name, Goal.prove_global thy7 [] [] statement (proof o #context)), []) end) ak_names_types); (* declares an fs-axclass for every atom-kind *) @@ -379,16 +376,15 @@ (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); val fs_type = Logic.mk_type i_type1; val at_type = Logic.mk_type i_type2; - val simp_s = HOL_ss addsimps maps (Global_Theory.get_thms thy11) + fun proof ctxt = + simp_tac (put_simpset HOL_ss ctxt addsimps maps (Global_Theory.get_thms thy11) ["fs_def", - "fs_" ^ ak_name ^ "1"]; + "fs_" ^ ak_name ^ "1"]) 1; val name = "fs_"^ak_name^ "_inst"; val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type); - - val proof = fn _ => simp_tac simp_s 1; in - ((name, Goal.prove_global thy11 [] [] statement proof), []) + ((name, Goal.prove_global thy11 [] [] statement (proof o #context)), []) end) ak_names_types); (* declares for every atom-kind combination an axclass *) @@ -432,18 +428,18 @@ val at_type = Logic.mk_type i_type1; val at_type' = Logic.mk_type i_type2; val cp_type = Logic.mk_type i_type0; - val simp_s = HOL_basic_ss addsimps maps (Global_Theory.get_thms thy') ["cp_def"]; val cp1 = Global_Theory.get_thm thy' ("cp_" ^ ak_name ^ "_" ^ ak_name' ^ "1"); val name = "cp_"^ak_name^ "_"^ak_name'^"_inst"; val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type'); - val proof = fn _ => EVERY [simp_tac simp_s 1, - rtac allI 1, rtac allI 1, rtac allI 1, - rtac cp1 1]; + fun proof ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt + addsimps maps (Global_Theory.get_thms thy') ["cp_def"]) 1 + THEN EVERY [rtac allI 1, rtac allI 1, rtac allI 1, rtac cp1 1]; in yield_singleton add_thms_string ((name, - Goal.prove_global thy' [] [] statement proof), []) thy' + Goal.prove_global thy' [] [] statement (proof o #context)), []) thy' end) ak_names_types thy) ak_names_types thy12b; @@ -464,17 +460,17 @@ (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); val at_type = Logic.mk_type i_type1; val at_type' = Logic.mk_type i_type2; - val simp_s = HOL_ss addsimps maps (Global_Theory.get_thms thy') + fun proof ctxt = + simp_tac (put_simpset HOL_ss ctxt + addsimps maps (Global_Theory.get_thms thy') ["disjoint_def", ak_name ^ "_prm_" ^ ak_name' ^ "_def", - ak_name' ^ "_prm_" ^ ak_name ^ "_def"]; + ak_name' ^ "_prm_" ^ ak_name ^ "_def"]) 1; val name = "dj_"^ak_name^"_"^ak_name'; val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type'); - - val proof = fn _ => simp_tac simp_s 1; in - add_thms_string [((name, Goal.prove_global thy' [] [] statement proof), [])] thy' + add_thms_string [((name, Goal.prove_global thy' [] [] statement (proof o #context)), [])] thy' end else ([],thy'))) (* do nothing branch, if ak_name = ak_name' *) @@ -511,14 +507,15 @@ rtac ((at_inst RS at_pt_inst) RS pt2) 1, rtac ((at_inst RS at_pt_inst) RS pt3) 1, atac 1]; - val simp_s = HOL_basic_ss addsimps - maps (Global_Theory.get_thms thy') [ak_name ^ "_prm_" ^ ak_name' ^ "_def"]; - val proof2 = EVERY [Class.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; - + fun proof2 ctxt = + Class.intro_classes_tac [] THEN + REPEAT (asm_simp_tac + (put_simpset HOL_basic_ss ctxt addsimps + maps (Global_Theory.get_thms thy') [ak_name ^ "_prm_" ^ ak_name' ^ "_def"]) 1); in thy' |> Axclass.prove_arity (qu_name,[],[cls_name]) - (fn _ => if ak_name = ak_name' then proof1 else proof2) + (fn ctxt => if ak_name = ak_name' then proof1 else proof2 ctxt) end) ak_names thy) ak_names thy12d; (* show that *) @@ -581,7 +578,7 @@ let val qu_name = Sign.full_bname thy' ak_name'; val qu_class = Sign.full_bname thy' ("fs_"^ak_name); - val proof = + fun proof ctxt = (if ak_name = ak_name' then let val at_thm = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst"); @@ -589,10 +586,11 @@ rtac ((at_thm RS fs_at_inst) RS fs1) 1] end else let val dj_inst = Global_Theory.get_thm thy' ("dj_"^ak_name'^"_"^ak_name); - val simp_s = HOL_basic_ss addsimps [dj_inst RS dj_supp, finite_emptyI]; + val simp_s = + put_simpset HOL_basic_ss ctxt addsimps [dj_inst RS dj_supp, finite_emptyI]; in EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1] end) in - Axclass.prove_arity (qu_name,[],[qu_class]) (fn _ => proof) thy' + Axclass.prove_arity (qu_name,[],[qu_class]) proof thy' end) ak_names thy) ak_names thy18; (* shows that *) @@ -648,7 +646,7 @@ let val name = Sign.full_bname thy'' ak_name; val cls_name = Sign.full_bname thy'' ("cp_"^ak_name'^"_"^ak_name''); - val proof = + fun proof ctxt = (if (ak_name'=ak_name'') then (let val pt_inst = Global_Theory.get_thm thy'' ("pt_"^ak_name''^"_inst"); @@ -660,7 +658,7 @@ else (let val dj_inst = Global_Theory.get_thm thy'' ("dj_"^ak_name''^"_"^ak_name'); - val simp_s = HOL_basic_ss addsimps + val simp_s = put_simpset HOL_basic_ss ctxt addsimps ((dj_inst RS dj_pp_forget):: (maps (Global_Theory.get_thms thy'') [ak_name' ^"_prm_"^ak_name^"_def", @@ -669,7 +667,7 @@ EVERY [Class.intro_classes_tac [], simp_tac simp_s 1] end)) in - Axclass.prove_arity (name,[],[cls_name]) (fn _ => proof) thy'' + Axclass.prove_arity (name,[],[cls_name]) proof thy'' end) ak_names thy') ak_names thy) ak_names thy24; (* shows that *) @@ -719,10 +717,11 @@ fold (fn ak_name => fn thy => let val qu_class = Sign.full_bname thy ("pt_"^ak_name); - val simp_s = HOL_basic_ss addsimps [Simpdata.mk_eq defn]; - val proof = EVERY [Class.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; + fun proof ctxt = + Class.intro_classes_tac [] THEN + REPEAT (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Simpdata.mk_eq defn]) 1); in - Axclass.prove_arity (discrete_ty, [], [qu_class]) (fn _ => proof) thy + Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy end) ak_names; fun discrete_fs_inst discrete_ty defn = @@ -730,10 +729,12 @@ let val qu_class = Sign.full_bname thy ("fs_"^ak_name); val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"}; - val simp_s = HOL_ss addsimps [supp_def, Collect_const, finite_emptyI, Simpdata.mk_eq defn]; - val proof = EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1]; + fun proof ctxt = + Class.intro_classes_tac [] THEN + asm_simp_tac (put_simpset HOL_ss ctxt + addsimps [supp_def, Collect_const, finite_emptyI, Simpdata.mk_eq defn]) 1; in - Axclass.prove_arity (discrete_ty, [], [qu_class]) (fn _ => proof) thy + Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy end) ak_names; fun discrete_cp_inst discrete_ty defn = @@ -741,10 +742,11 @@ let val qu_class = Sign.full_bname thy ("cp_"^ak_name^"_"^ak_name'); val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"}; - val simp_s = HOL_ss addsimps [Simpdata.mk_eq defn]; - val proof = EVERY [Class.intro_classes_tac [], asm_simp_tac simp_s 1]; + fun proof ctxt = + Class.intro_classes_tac [] THEN + asm_simp_tac (put_simpset HOL_ss ctxt addsimps [Simpdata.mk_eq defn]) 1; in - Axclass.prove_arity (discrete_ty, [], [qu_class]) (fn _ => proof) thy + Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy end) ak_names)) ak_names; in diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_datatype.ML Thu Apr 18 17:07:01 2013 +0200 @@ -96,10 +96,11 @@ fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u | permTs_of _ = []; -fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) = +fun perm_simproc' ctxt (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) = let + val thy = Proof_Context.theory_of ctxt; val (aT as Type (a, []), S) = dest_permT T; - val (bT as Type (b, []), _) = dest_permT U + val (bT as Type (b, []), _) = dest_permT U; in if member (op =) (permTs_of u) aT andalso aT <> bT then let val cp = cp_inst_of thy a b; @@ -112,7 +113,7 @@ end else NONE end - | perm_simproc' thy ss _ = NONE; + | perm_simproc' _ _ = NONE; val perm_simproc = Simplifier.simproc_global @{theory} "perm_simp" ["pi1 \ (pi2 \ x)"] perm_simproc'; @@ -279,8 +280,7 @@ end) (perm_names_types ~~ perm_indnames)))) (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1, - ALLGOALS (asm_full_simp_tac - (simpset_of ctxt addsimps [perm_fun_def]))])), + ALLGOALS (asm_full_simp_tac (ctxt addsimps [perm_fun_def]))])), length new_type_names)); (**** prove [] \ t = t ****) @@ -300,7 +300,7 @@ (perm_names ~~ map body_type perm_types ~~ perm_indnames))))) (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of ctxt))])), + ALLGOALS (asm_full_simp_tac ctxt)])), length new_type_names)) end) atoms; @@ -335,7 +335,7 @@ (perm_names ~~ map body_type perm_types ~~ perm_indnames))))) (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of ctxt addsimps [pt2', pt2_ax]))]))), + ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt2', pt2_ax]))]))), length new_type_names) end) atoms; @@ -371,7 +371,7 @@ (perm_names ~~ map body_type perm_types ~~ perm_indnames)))))) (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac induct perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of ctxt addsimps [pt3', pt3_rev', pt3_ax]))]))), + ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt3', pt3_rev', pt3_ax]))]))), length new_type_names) end) atoms; @@ -393,7 +393,7 @@ val permT2 = mk_permT (Type (name2, [])); val Ts = map body_type perm_types; val cp_inst = cp_inst_of thy name1 name2; - fun simps ctxt = simpset_of ctxt addsimps (perm_fun_def :: + fun simps ctxt = ctxt addsimps (perm_fun_def :: (if name1 <> name2 then let val dj = dj_thm_of thy name2 name1 in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end @@ -563,7 +563,7 @@ end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))) (fn {context = ctxt, ...} => EVERY [Datatype_Aux.ind_tac rep_induct [] 1, - ALLGOALS (simp_tac (simpset_of ctxt addsimps + ALLGOALS (simp_tac (ctxt addsimps (Thm.symmetric perm_fun_def :: abs_perm))), ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])), length new_type_names)); @@ -623,10 +623,10 @@ map (inter_sort thy sort o snd) tvs, [pt_class]) (fn ctxt => EVERY [Class.intro_classes_tac [], rewrite_goals_tac [perm_def], - asm_full_simp_tac (simpset_of ctxt addsimps [Rep_inverse]) 1, - asm_full_simp_tac (simpset_of ctxt addsimps + asm_full_simp_tac (ctxt addsimps [Rep_inverse]) 1, + asm_full_simp_tac (ctxt addsimps [Rep RS perm_closed RS Abs_inverse]) 1, - asm_full_simp_tac (HOL_basic_ss addsimps [Global_Theory.get_thm thy + asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy |> Theory.checkpoint end) @@ -653,7 +653,7 @@ map (inter_sort thy sort o snd) tvs, [cp_class]) (fn ctxt => EVERY [Class.intro_classes_tac [], rewrite_goals_tac [perm_def], - asm_full_simp_tac (simpset_of ctxt addsimps + asm_full_simp_tac (ctxt addsimps ((Rep RS perm_closed1 RS Abs_inverse) :: (if atom1 = atom2 then [] else [Rep RS perm_closed2 RS Abs_inverse]))) 1, @@ -825,7 +825,8 @@ (HOLogic.mk_Trueprop (HOLogic.mk_eq (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x), Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x))))) - (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @ + (fn {context = ctxt, ...} => + simp_tac (put_simpset HOL_basic_ss ctxt addsimps (perm_defs @ Abs_inverse_thms @ perm_closed_thms @ Rep_thms)) 1) end) Rep_thms; @@ -842,7 +843,7 @@ | prove_distinct_thms (p as (rep_thms, dist_lemma)) (t :: ts) = let val dist_thm = Goal.prove_global_future thy8 [] [] t (fn {context = ctxt, ...} => - simp_tac (simpset_of ctxt addsimps (dist_lemma :: rep_thms)) 1) + simp_tac (ctxt addsimps (dist_lemma :: rep_thms)) 1) in dist_thm :: Drule.export_without_context (dist_thm RS not_sym) :: prove_distinct_thms p ts @@ -890,12 +891,12 @@ (HOLogic.mk_Trueprop (HOLogic.mk_eq (perm (list_comb (c, l_args)), list_comb (c, r_args))))) (fn {context = ctxt, ...} => EVERY - [simp_tac (simpset_of ctxt addsimps (constr_rep_thm :: perm_defs)) 1, - simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ + [simp_tac (ctxt addsimps (constr_rep_thm :: perm_defs)) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Rep_thms @ Abs_inverse_thms @ constr_defs @ perm_closed_thms)) 1, - TRY (simp_tac (HOL_basic_ss addsimps + TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Thm.symmetric perm_fun_def :: abs_perm)) 1), - TRY (simp_tac (HOL_basic_ss addsimps + TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ perm_closed_thms)) 1)]) end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss) @@ -946,9 +947,10 @@ (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), foldr1 HOLogic.mk_conj eqs)))) (fn {context = ctxt, ...} => EVERY - [asm_full_simp_tac (simpset_of ctxt addsimps (constr_rep_thm :: + [asm_full_simp_tac (ctxt addsimps (constr_rep_thm :: rep_inject_thms')) 1, - TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: + TRY (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt + addsimps (fresh_def :: supp_def :: alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ perm_rep_perm_thms)) 1)]) end) (constrs ~~ constr_rep_thms) @@ -989,8 +991,8 @@ (supp c, if null dts then HOLogic.mk_set atomT [] else foldr1 (HOLogic.mk_binop @{const_abbrev union}) (map supp args2))))) - (fn _ => - simp_tac (HOL_basic_ss addsimps (supp_def :: + (fn {context = ctxt, ...} => + simp_tac (put_simpset HOL_basic_ss ctxt addsimps (supp_def :: Un_assoc :: @{thm de_Morgan_conj} :: Collect_disj_eq :: finite_Un :: Collect_False_empty :: finite_emptyI :: @{thms simp_thms} @ abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1) @@ -1001,8 +1003,8 @@ (fresh c, if null dts then @{term True} else foldr1 HOLogic.mk_conj (map fresh args2))))) - (fn _ => - simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1)) + (fn {context = ctxt, ...} => + simp_tac (put_simpset HOL_ss ctxt addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1)) end) atoms) constrs end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps'))); @@ -1028,10 +1030,12 @@ val indrule_lemma = Goal.prove_global_future thy8 [] [] (Logic.mk_implies (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_prems), - HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_concls))) (fn _ => EVERY + HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_concls))) + (fn {context = ctxt, ...} => EVERY [REPEAT (etac conjE 1), REPEAT (EVERY - [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1, + [TRY (rtac conjI 1), + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps Rep_inverse_thms) 1, etac mp 1, resolve_tac Rep_thms 1])]); val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); @@ -1045,12 +1049,12 @@ val dt_induct_prop = Datatype_Prop.make_ind descr'; val dt_induct = Goal.prove_global_future thy8 [] (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) - (fn {prems, ...} => EVERY + (fn {prems, context = ctxt} => EVERY [rtac indrule_lemma' 1, (Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1, EVERY (map (fn (prem, r) => (EVERY [REPEAT (eresolve_tac Abs_inverse_thms' 1), - simp_tac (HOL_basic_ss addsimps [Thm.symmetric r]) 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Thm.symmetric r]) 1, DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) (prems ~~ constr_defs))]); @@ -1076,7 +1080,7 @@ (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T))) (indnames ~~ recTs))))) (fn {context = ctxt, ...} => Datatype_Aux.ind_tac dt_induct indnames 1 THEN - ALLGOALS (asm_full_simp_tac (simpset_of ctxt addsimps + ALLGOALS (asm_full_simp_tac (ctxt addsimps (abs_supp @ supp_atm @ Global_Theory.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @ flat supp_thms))))), @@ -1236,12 +1240,12 @@ Bound 0 $ p))) (fn _ => EVERY [resolve_tac exists_fresh' 1, - simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @ + simp_tac (put_simpset HOL_ss ctxt addsimps (supp_prod :: finite_Un :: fs_atoms @ fin_set_supp @ ths)) 1]); val (([(_, cx)], ths), ctxt') = Obtain.result - (fn _ => EVERY + (fn ctxt' => EVERY [etac exE 1, - full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, + full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1, REPEAT (etac conjE 1)]) [ex] ctxt in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; @@ -1281,16 +1285,16 @@ (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} => let val (prems1, prems2) = chop (length dt_atomTs) prems; - val ind_ss2 = HOL_ss addsimps + val ind_ss2 = put_simpset HOL_ss context addsimps finite_Diff :: abs_fresh @ abs_supp @ fs_atoms; val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @ fresh_atm @ rev_simps @ app_simps; - val ind_ss3 = HOL_ss addsimps abs_fun_eq1 :: + val ind_ss3 = put_simpset HOL_ss context addsimps abs_fun_eq1 :: abs_perm @ calc_atm @ perm_swap; - val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @ + val ind_ss4 = put_simpset HOL_basic_ss context addsimps fresh_left @ prems1 @ fin_set_fresh @ calc_atm; - val ind_ss5 = HOL_basic_ss addsimps pt1_atoms; - val ind_ss6 = HOL_basic_ss addsimps flat perm_simps'; + val ind_ss5 = put_simpset HOL_basic_ss context addsimps pt1_atoms; + val ind_ss6 = put_simpset HOL_basic_ss context addsimps flat perm_simps'; val th = Goal.prove context [] [] (augment_sort thy9 fs_cp_sort aux_ind_concl) (fn {context = context1, ...} => @@ -1332,7 +1336,7 @@ cs @ [fold_rev (mk_perm []) (map perm_of_pair (bs ~~ cs)) t]) (xs'' ~~ freshs1'))))) (fn _ => EVERY - (simp_tac (HOL_ss addsimps flat inject_thms) 1 :: + (simp_tac (put_simpset HOL_ss context3 addsimps flat inject_thms) 1 :: REPEAT (FIRSTGOAL (rtac conjI)) :: maps (fn ((bs, t), cs) => if null bs then [] @@ -1352,7 +1356,7 @@ simp_tac ind_ss1' i | _ $ (Const (@{const_name Not}, _) $ _) => resolve_tac freshs2' i - | _ => asm_simp_tac (HOL_basic_ss addsimps + | _ => asm_simp_tac (put_simpset HOL_basic_ss context3 addsimps pt2_atoms addsimprocs [perm_simproc]) i)) 1]) val final = Proof_Context.export context3 context2 [th] in @@ -1380,11 +1384,12 @@ val induct = Goal.prove_global_future thy9 [] (map (augment_sort thy9 fs_cp_sort) ind_prems) (augment_sort thy9 fs_cp_sort ind_concl) - (fn {prems, ...} => EVERY + (fn {prems, context = ctxt} => EVERY [rtac induct_aux' 1, REPEAT (resolve_tac fs_atoms 1), REPEAT ((resolve_tac prems THEN_ALL_NEW - (etac @{thm meta_spec} ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)]) + (etac @{thm meta_spec} ORELSE' + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [fresh_def]))) 1)]) val (_, thy10) = thy9 |> Sign.add_path big_name |> @@ -1526,20 +1531,20 @@ (Goal.prove_global_future thy11 [] [] (augment_sort thy1 pt_cp_sort (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps)))) - (fn _ => rtac rec_induct 1 THEN REPEAT + (fn {context = ctxt, ...} => rtac rec_induct 1 THEN REPEAT (simp_tac (Simplifier.global_context thy11 HOL_basic_ss addsimps flat perm_simps' addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN (resolve_tac rec_intrs THEN_ALL_NEW - asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1)))) + asm_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_bij @ perm_bij))) 1)))) val ths' = map (fn ((P, Q), th) => Goal.prove_global_future thy11 [] [] (augment_sort thy1 pt_cp_sort (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P))) - (fn _ => dtac (Thm.instantiate ([], + (fn {context = ctxt, ...} => dtac (Thm.instantiate ([], [(cterm_of thy11 (Var (("pi", 0), permT)), cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN - NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths) + NominalPermeq.perm_simp_tac (put_simpset HOL_ss ctxt) 1)) (ps ~~ ths) in (ths, ths') end) dt_atomTs); (** finite support **) @@ -1568,9 +1573,9 @@ finite $ (Const ("Nominal.supp", U --> aset) $ y)) end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ (1 upto length recTs)))))) - (fn {prems = fins, ...} => + (fn {prems = fins, context = ctxt} => (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT - (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1)))) + (NominalPermeq.finite_guess_tac (put_simpset HOL_ss ctxt addsimps [fs_name]) 1)))) end) dt_atomTs; (** freshness **) @@ -1620,7 +1625,7 @@ cterm_of thy11 (Const ("Nominal.supp", fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))] supports_fresh) 1, - simp_tac (HOL_basic_ss addsimps + simp_tac (put_simpset HOL_basic_ss context addsimps [supports_def, Thm.symmetric fresh_def, fresh_prod]) 1, REPEAT_DETERM (resolve_tac [allI, impI] 1), REPEAT_DETERM (etac conjE 1), @@ -1630,12 +1635,12 @@ rtac (Thm.instantiate ([], [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)), cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1, - asm_simp_tac (HOL_ss addsimps + asm_simp_tac (put_simpset HOL_ss context addsimps (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1, rtac rec_prem 1, - simp_tac (HOL_ss addsimps (fs_name :: + simp_tac (put_simpset HOL_ss context addsimps (fs_name :: supp_prod :: finite_Un :: finite_prems)) 1, - simp_tac (HOL_ss addsimps (Thm.symmetric fresh_def :: + simp_tac (put_simpset HOL_ss context addsimps (Thm.symmetric fresh_def :: fresh_prod :: fresh_prems)) 1] end)) end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths) @@ -1677,11 +1682,11 @@ [cut_facts_tac ths 1, REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1), resolve_tac exists_fresh' 1, - asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); + asm_simp_tac (put_simpset HOL_ss ctxt addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); val (([(_, cx)], ths), ctxt') = Obtain.result (fn _ => EVERY [etac exE 1, - full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, + full_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_prod :: fresh_atm)) 1, REPEAT (etac conjE 1)]) [ex] ctxt in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; @@ -1723,16 +1728,16 @@ ([rtac induct_aux_rec 1] @ maps (fn ((_, finite_ths), finite_th) => [cut_facts_tac (finite_th :: finite_ths) 1, - asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1]) + asm_simp_tac (put_simpset HOL_ss context addsimps [supp_prod, finite_Un]) 1]) (finite_thss ~~ finite_ctxt_ths) @ maps (fn ((_, idxss), elim) => maps (fn idxs => - [full_simp_tac (HOL_ss addsimps [Thm.symmetric fresh_def, supp_prod, Un_iff]) 1, + [full_simp_tac (put_simpset HOL_ss context addsimps [Thm.symmetric fresh_def, supp_prod, Un_iff]) 1, REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1), rtac ex1I 1, (resolve_tac rec_intrs THEN_ALL_NEW atac) 1, rotate_tac ~1 1, ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac - (HOL_ss addsimps flat distinct_thms)) 1] @ + (put_simpset HOL_ss context addsimps flat distinct_thms)) 1] @ (if null idxs then [] else [hyp_subst_tac 1, SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} => let @@ -1777,14 +1782,14 @@ (** as, bs, cs # K as ts, K bs us **) val _ = warning "step 2: as, bs, cs # K as ts, K bs us"; - val prove_fresh_ss = HOL_ss addsimps + val prove_fresh_simpset = put_simpset HOL_ss context'' addsimps (finite_Diff :: flat fresh_thms @ fs_atoms @ abs_fresh @ abs_supp @ fresh_atm); (* FIXME: avoid asm_full_simp_tac ? *) fun prove_fresh ths y x = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const (fastype_of x) (fastype_of y) $ x $ y)) - (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1); + (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_simpset 1); val constr_fresh_thms = map (prove_fresh fresh_prems lhs) boundsl @ map (prove_fresh fresh_prems rhs) boundsr @ @@ -1798,7 +1803,7 @@ (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs))) (fn _ => EVERY [cut_facts_tac constr_fresh_thms 1, - asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1, + asm_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh) 1, rtac prem 1]); (** pi1 o ts = pi2 o us **) @@ -1809,7 +1814,7 @@ (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u))) (fn _ => EVERY [cut_facts_tac [pi1_pi2_eq] 1, - asm_full_simp_tac (HOL_ss addsimps + asm_full_simp_tac (put_simpset HOL_ss context'' addsimps (calc_atm @ flat perm_simps' @ fresh_prems' @ freshs2' @ abs_perm @ alpha @ flat inject_thms)) 1])) @@ -1821,7 +1826,7 @@ Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) (rpi1 @ pi2) u, t))) - (fn _ => simp_tac (HOL_ss addsimps + (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps ((eq RS sym) :: perm_swap)) 1)) (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs); @@ -1850,12 +1855,12 @@ (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y)) (fn _ => EVERY (map eqvt_tac pi @ - [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @ + [simp_tac (put_simpset HOL_ss context'' addsimps (fresh_prems' @ freshs2' @ perm_swap @ perm_fresh_fresh)) 1, rtac th 1])) in Simplifier.simplify - (HOL_basic_ss addsimps rpi1_pi2_eqs) th' + (put_simpset HOL_basic_ss context'' addsimps rpi1_pi2_eqs) th' end) rec_prems2; val ihs = filter (fn th => case prop_of th of @@ -1874,7 +1879,7 @@ (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 (strip_perm rhs)))) - (fn _ => simp_tac (HOL_basic_ss addsimps + (fn _ => simp_tac (put_simpset HOL_basic_ss context'' addsimps (th' :: perm_swap)) 1) end) (rec_prems' ~~ ihs); @@ -1924,17 +1929,17 @@ (HOLogic.mk_Trueprop (fresh_const aT rT $ fold_rev (mk_perm []) (rpi2 @ pi1) a $ fold_rev (mk_perm []) (rpi2 @ pi1) rhs')) - (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN + (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps fresh_bij) 1 THEN rtac th 1) in Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs')) (fn _ => EVERY [cut_facts_tac [th'] 1, - full_simp_tac (Simplifier.global_context thy11 HOL_ss + full_simp_tac (Simplifier.global_context thy11 HOL_ss (* FIXME context'' (!?) *) addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap addsimprocs [NominalPermeq.perm_simproc_app]) 1, - full_simp_tac (HOL_ss addsimps (calc_atm @ + full_simp_tac (put_simpset HOL_ss context'' addsimps (calc_atm @ fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1]) end; @@ -1951,7 +1956,7 @@ REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp_thms' aT)) 1), NominalPermeq.fresh_guess_tac - (HOL_ss addsimps (freshs2 @ + (put_simpset HOL_ss context'' addsimps (freshs2 @ fs_atoms @ fresh_atm @ maps snd finite_thss)) 1]); @@ -1964,16 +1969,16 @@ val pi1_pi2_result = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs'))) - (fn _ => simp_tac (Simplifier.context context'' HOL_ss + (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps pi1_pi2_eqs @ rec_eqns addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN - TRY (simp_tac (HOL_ss addsimps + TRY (simp_tac (put_simpset HOL_ss context'' addsimps (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1)); val _ = warning "final result"; val final = Goal.prove context'' [] [] (term_of concl) (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN - full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @ + full_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh @ fresh_results @ fresh_results') 1); val final' = Proof_Context.export context'' context' [final]; val _ = warning "finished!" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_fresh_fun.ML --- a/src/HOL/Nominal/nominal_fresh_fun.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_fresh_fun.ML Thu Apr 18 17:07:01 2013 +0200 @@ -129,9 +129,9 @@ val thy = theory_of_thm thm; val abs_fresh = Global_Theory.get_thms thy "abs_fresh"; val fresh_perm_app = Global_Theory.get_thms thy "fresh_perm_app"; - val ss = simpset_of ctxt; - val ss' = ss addsimps fresh_prod::abs_fresh; - val ss'' = ss' addsimps fresh_perm_app; + val simp_ctxt = + ctxt addsimps (fresh_prod :: abs_fresh) + addsimps fresh_perm_app; val x = hd (tl (Misc_Legacy.term_vars (prop_of exI))); val goal = nth (prems_of thm) (i-1); val atom_name_opt = get_inner_fresh_fun goal; @@ -164,10 +164,10 @@ val post_rewrite_tacs = [rtac pt_name_inst, rtac at_name_inst, - TRY o SOLVED' (NominalPermeq.finite_guess_tac ss''), + TRY o SOLVED' (NominalPermeq.finite_guess_tac simp_ctxt), inst_fresh vars params THEN' - (TRY o SOLVED' (NominalPermeq.fresh_guess_tac ss'')) THEN' - (TRY o SOLVED' (asm_full_simp_tac ss''))] + (TRY o SOLVED' (NominalPermeq.fresh_guess_tac simp_ctxt)) THEN' + (TRY o SOLVED' (asm_full_simp_tac simp_ctxt))] in ((if no_asm then no_tac else (subst_inner_asm_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_induct.ML --- a/src/HOL/Nominal/nominal_induct.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_induct.ML Thu Apr 18 17:07:01 2013 +0200 @@ -21,8 +21,8 @@ Library.funpow (length Ts) HOLogic.mk_split (Var (xi, (HOLogic.unitT :: Ts) ---> Term.range_type T)); -val split_all_tuples = - Simplifier.full_simplify (HOL_basic_ss addsimps +fun split_all_tuples ctxt = + Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}, @{thm split_paired_all}, @{thm unit_all_eq1}, @{thm fresh_unit_elim}, @{thm fresh_prod_elim}] @ @{thms fresh_star_unit_elim} @ @{thms fresh_star_prod_elim}); @@ -90,7 +90,7 @@ val atomized_defs = map (map (Conv.fconv_rule Induct.atomize_cterm)) defs; val finish_rule = - split_all_tuples + split_all_tuples defs_ctxt #> rename_params_rule true (map (Name.clean o Variable.revert_fixed defs_ctxt o fst) avoiding); diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_inductive.ML Thu Apr 18 17:07:01 2013 +0200 @@ -20,12 +20,12 @@ fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; -val atomize_conv = +fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) - (HOL_basic_ss addsimps inductive_atomize); -val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv); + (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); +fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 - (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt)); + (Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); @@ -40,7 +40,7 @@ [(perm_boolI_pi, pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.simproc_global_i - (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn thy => fn ss => + (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt => fn Const (@{const_name Nominal.perm}, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE @@ -103,10 +103,10 @@ else NONE | inst_conj_all _ _ _ _ _ = NONE; -fun inst_conj_all_tac k = EVERY +fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]), REPEAT_DETERM_N k (etac allE 1), - simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1]; + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) @@ -271,10 +271,10 @@ val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn aT => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "2")) atomTs; - val eqvt_ss = Simplifier.global_context thy HOL_basic_ss + val eqvt_ss = simpset_of (Simplifier.global_context thy HOL_basic_ss addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [@{const_name Fun.id}], - NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; + NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]); val fresh_bij = Global_Theory.get_thms thy "fresh_bij"; val perm_bij = Global_Theory.get_thms thy "perm_bij"; val fs_atoms = map (fn aT => Global_Theory.get_thm thy @@ -299,10 +299,10 @@ [resolve_tac exists_fresh' 1, resolve_tac fs_atoms 1]); val (([(_, cx)], ths), ctxt') = Obtain.result - (fn _ => EVERY + (fn ctxt' => EVERY [etac exE 1, - full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, - full_simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1, + full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1, + full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, REPEAT (etac conjE 1)]) [ex] ctxt in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; @@ -312,7 +312,7 @@ let val th = Goal.prove ctxt [] [] concl (fn {context, ...} => rtac raw_induct 1 THEN EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => - [REPEAT (rtac allI 1), simp_tac eqvt_ss 1, + [REPEAT (rtac allI 1), simp_tac (put_simpset eqvt_ss context) 1, SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => let val (params', (pis, z)) = @@ -343,9 +343,9 @@ map (fold_rev (NominalDatatype.mk_perm []) (rev pis' @ pis)) params' @ [z])) ihyp; fun mk_pi th = - Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}] + Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) - (Simplifier.simplify eqvt_ss + (Simplifier.simplify (put_simpset eqvt_ss ctxt) (fold_rev (mk_perm_bool o cterm_of thy) (rev pis' @ pis) th)); val (gprems1, gprems2) = split_list @@ -355,7 +355,7 @@ (map_thm ctxt (split_conj (K o I) names) (etac conjunct1 1) monos NONE th, mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis'')) - (inst_conj_all_tac (length pis'')) monos (SOME t) th)))) + (inst_conj_all_tac ctxt (length pis'')) monos (SOME t) th)))) (gprems ~~ oprems)) |>> map_filter I; val vc_compat_ths' = map (fn th => let @@ -368,29 +368,29 @@ val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) (fold_rev (NominalDatatype.mk_perm []) pis rhs))) - (fn _ => simp_tac (HOL_basic_ss addsimps + (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps (fresh_bij @ perm_bij)) 1 THEN rtac th' 1) - in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end) + in Simplifier.simplify (put_simpset eqvt_ss ctxt'' addsimps fresh_atm) th'' end) vc_compat_ths; val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) (** we have to pre-simplify the rewrite rules **) - val swap_simps_ss = HOL_ss addsimps swap_simps @ - map (Simplifier.simplify (HOL_ss addsimps swap_simps)) + val swap_simps_simpset = put_simpset HOL_ss ctxt'' addsimps swap_simps @ + map (Simplifier.simplify (put_simpset HOL_ss ctxt'' addsimps swap_simps)) (vc_compat_ths'' @ freshs2'); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) - (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1, + (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, rtac ihyp' 1, REPEAT_DETERM_N (nprems_of ihyp - length gprems) - (simp_tac swap_simps_ss 1), + (simp_tac swap_simps_simpset 1), REPEAT_DETERM_N (length gprems) - (simp_tac (HOL_basic_ss + (simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac gprems2 1)])); val final = Goal.prove ctxt'' [] [] (term_of concl) - (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss + (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt'' addsimps vc_compat_ths'' @ freshs2' @ perm_fresh_fresh @ fresh_atm) 1); val final' = Proof_Context.export ctxt'' ctxt' [final]; @@ -400,7 +400,7 @@ cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN - asm_full_simp_tac (simpset_of ctxt) 1) + asm_full_simp_tac ctxt 1) end) |> singleton (Proof_Context.export ctxt' ctxt); (** strong case analysis rule **) @@ -452,13 +452,13 @@ concl)) in map mk_prem prems end) cases_prems; - val cases_eqvt_ss = Simplifier.global_context thy HOL_ss + val cases_eqvt_simpset = Simplifier.global_context thy HOL_ss addsimps eqvt_thms @ swap_simps @ perm_pi_simp addsimprocs [NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val simp_fresh_atm = map - (Simplifier.simplify (HOL_basic_ss addsimps fresh_atm)); + (Simplifier.simplify (Simplifier.global_context thy HOL_basic_ss addsimps fresh_atm)); fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt'))), prems') = @@ -520,16 +520,16 @@ SUBPROOF (fn {prems = fresh_hyps, ...} => let val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; - val case_ss = cases_eqvt_ss addsimps freshs2' @ + val case_simpset = cases_eqvt_simpset addsimps freshs2' @ simp_fresh_atm (vc_compat_ths' @ fresh_hyps'); - val fresh_fresh_ss = case_ss addsimps perm_fresh_fresh; + val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh; val hyps1' = map - (mk_pis #> Simplifier.simplify fresh_fresh_ss) hyps1; + (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1; val hyps2' = map - (mk_pis #> Simplifier.simplify case_ss) hyps2; + (mk_pis #> Simplifier.simplify case_simpset) hyps2; val case_hyps' = hyps1' @ hyps2' in - simp_tac case_ss 1 THEN + simp_tac case_simpset 1 THEN REPEAT_DETERM (TRY (rtac conjI 1) THEN resolve_tac case_hyps' 1) end) ctxt4 1) @@ -547,11 +547,11 @@ val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); - val thss' = map (map atomize_intr) thss; + val thss' = map (map (atomize_intr ctxt)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = - mk_ind_proof ctxt thss' |> Inductive.rulify; - val strong_cases = map (mk_cases_proof ##> Inductive.rulify) + mk_ind_proof ctxt thss' |> Inductive.rulify ctxt; + val strong_cases = map (mk_cases_proof ##> Inductive.rulify ctxt) (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); val strong_induct_atts = map (Attrib.internal o K) @@ -586,7 +586,7 @@ Inductive.the_inductive ctxt (Sign.intern_const thy s); val raw_induct = atomize_induct ctxt raw_induct; val elims = map (atomize_induct ctxt) elims; - val intrs = map atomize_intr intrs; + val intrs = map (atomize_intr ctxt) intrs; val monos = Inductive.get_monos ctxt; val intrs' = Inductive.unpartition_rules intrs (map (fn (((s, ths), (_, k)), th) => @@ -608,7 +608,7 @@ atoms) end; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; - val eqvt_ss = Simplifier.global_context thy HOL_basic_ss addsimps + val eqvt_simpset = Simplifier.global_context thy HOL_basic_ss addsimps (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs [mk_perm_bool_simproc names, NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; @@ -628,7 +628,7 @@ let val prems' = map (fn th => the_default th (map_thm ctxt' (split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems; - val prems'' = map (fn th => Simplifier.simplify eqvt_ss + val prems'' = map (fn th => Simplifier.simplify eqvt_simpset (mk_perm_bool (cterm_of thy pi) th)) prems'; val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~ map (cterm_of thy o NominalDatatype.mk_perm [] pi o term_of o #2) params) @@ -654,7 +654,7 @@ map (NominalDatatype.mk_perm [] pi') ts2)) end) ps))) (fn {context, ...} => EVERY (rtac raw_induct 1 :: map (fn intr_vs => - full_simp_tac eqvt_ss 1 THEN + full_simp_tac eqvt_simpset 1 THEN eqvt_tac context pi' intr_vs) intrs')) |> singleton (Proof_Context.export ctxt' ctxt))) end) atoms diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_inductive2.ML Thu Apr 18 17:07:01 2013 +0200 @@ -21,15 +21,15 @@ fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; -val atomize_conv = +fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) - (HOL_basic_ss addsimps inductive_atomize); -val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv); + (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); +fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 - (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt)); + (Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)); -val fresh_postprocess = - Simplifier.full_simplify (HOL_basic_ss addsimps +fun fresh_postprocess ctxt = + Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fresh_star_set_eq}, @{thm fresh_star_Un_elim}, @{thm fresh_star_insert_elim}, @{thm fresh_star_empty_elim}]); @@ -44,7 +44,7 @@ [(perm_boolI_pi, pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.simproc_global_i - (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn thy => fn ss => + (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt => fn Const ("Nominal.perm", _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE @@ -108,10 +108,10 @@ else NONE | inst_conj_all _ _ _ _ _ = NONE; -fun inst_conj_all_tac k = EVERY +fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]), REPEAT_DETERM_N k (etac allE 1), - simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1]; + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) @@ -290,10 +290,10 @@ val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn a => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name a ^ "2")) atoms; - val eqvt_ss = Simplifier.global_context thy HOL_basic_ss + val eqvt_ss = simpset_of (Simplifier.global_context thy HOL_basic_ss addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc ["Fun.id"], - NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; + NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]); val fresh_star_bij = Global_Theory.get_thms thy "fresh_star_bij"; val pt_insts = map (NominalAtoms.pt_inst_of thy) atoms; val at_insts = map (NominalAtoms.at_inst_of thy) atoms; @@ -322,10 +322,10 @@ [SOME (ctyp_of thy (fastype_of p))] [SOME (cterm_of thy p)] ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding}); val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result - (fn _ => EVERY + (fn ctxt' => EVERY [rtac avoid_th 1, - full_simp_tac (HOL_ss addsimps [@{thm fresh_star_prod_set}]) 1, - full_simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1, + full_simp_tac (put_simpset HOL_ss ctxt' addsimps [@{thm fresh_star_prod_set}]) 1, + full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, rotate_tac 1 1, REPEAT (etac conjE 1)]) [] ctxt; @@ -340,8 +340,8 @@ (f $ fold_rev (NominalDatatype.mk_perm (rev pTs)) (pis1 @ pi :: pis2) l $ r))) (fn _ => cut_facts_tac [th2] 1 THEN - full_simp_tac (HOL_basic_ss addsimps perm_set_forget) 1) |> - Simplifier.simplify eqvt_ss + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps perm_set_forget) 1) |> + Simplifier.simplify (put_simpset eqvt_ss ctxt) in (freshs @ [term_of cx], ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt') @@ -353,7 +353,7 @@ rtac raw_induct 1 THEN EVERY (maps (fn (((((_, sets, oprems, _), vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) => - [REPEAT (rtac allI 1), simp_tac eqvt_ss 1, + [REPEAT (rtac allI 1), simp_tac (put_simpset eqvt_ss context) 1, SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => let val (cparams', (pis, z)) = @@ -379,14 +379,14 @@ Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (list_comb (h, map (fold_rev (NominalDatatype.mk_perm []) pis) ts))) - (fn _ => simp_tac (HOL_basic_ss addsimps + (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (fresh_star_bij @ finite_ineq)) 1 THEN rtac th' 1) end) vc_compat_ths vc_compat_vs; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length sets) vc_compat_ths'; val vc_compat_ths1' = map (Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv - (Simplifier.rewrite eqvt_ss)))) vc_compat_ths1; + (Simplifier.rewrite (put_simpset eqvt_ss ctxt'))))) vc_compat_ths1; val (pis', fresh_ths1, fresh_ths2, fresh_ths3, ctxt'') = fold (obtain_fresh_name ts sets) (map snd sets' ~~ vc_compat_ths2) ([], [], [], [], ctxt'); @@ -401,16 +401,16 @@ (map (fold_rev (NominalDatatype.mk_perm []) (pis' @ pis) #> cterm_of thy) params' @ [cterm_of thy z]); fun mk_pi th = - Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}] + Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) - (Simplifier.simplify eqvt_ss + (Simplifier.simplify (put_simpset eqvt_ss ctxt) (fold_rev (mk_perm_bool o cterm_of thy) (pis' @ pis) th)); val gprems2 = map (fn (th, t) => if null (preds_of ps t) then mk_pi th else mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis'')) - (inst_conj_all_tac (length pis'')) monos (SOME t) th))) + (inst_conj_all_tac ctxt (length pis'')) monos (SOME t) th))) (gprems ~~ oprems); val perm_freshs_freshs' = map (fn (th, (_, T)) => th RS the (AList.lookup op = perm_freshs_freshs T)) @@ -418,15 +418,15 @@ val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts)))) - (fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1] @ + (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, rtac ihyp' 1] @ map (fn th => rtac th 1) fresh_ths3 @ [REPEAT_DETERM_N (length gprems) - (simp_tac (HOL_basic_ss + (simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac gprems2 1)])); val final = Goal.prove ctxt'' [] [] (term_of concl) - (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss + (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt'' addsimps vc_compat_ths1' @ fresh_ths1 @ perm_freshs_freshs') 1); val final' = Proof_Context.export ctxt'' ctxt' [final]; @@ -436,9 +436,9 @@ cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN - asm_full_simp_tac (simpset_of ctxt) 1) + asm_full_simp_tac ctxt 1) end) |> - fresh_postprocess |> + fresh_postprocess ctxt' |> singleton (Proof_Context.export ctxt' ctxt); in @@ -450,10 +450,10 @@ val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); - val thss' = map (map atomize_intr) thss; + val thss' = map (map (atomize_intr ctxt)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = - mk_ind_proof ctxt thss' |> Inductive.rulify; + mk_ind_proof ctxt thss' |> Inductive.rulify ctxt; val strong_induct_atts = map (Attrib.internal o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_permeq.ML --- a/src/HOL/Nominal/nominal_permeq.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_permeq.ML Thu Apr 18 17:07:01 2013 +0200 @@ -30,11 +30,11 @@ val perm_simproc_fun : simproc val perm_simproc_app : simproc - val perm_simp_tac : simpset -> int -> tactic - val perm_extend_simp_tac : simpset -> int -> tactic - val supports_tac : simpset -> int -> tactic - val finite_guess_tac : simpset -> int -> tactic - val fresh_guess_tac : simpset -> int -> tactic + val perm_simp_tac : Proof.context -> int -> tactic + val perm_extend_simp_tac : Proof.context -> int -> tactic + val supports_tac : Proof.context -> int -> tactic + val finite_guess_tac : Proof.context -> int -> tactic + val fresh_guess_tac : Proof.context -> int -> tactic val perm_simp_meth : (Proof.context -> Proof.method) context_parser val perm_simp_meth_debug : (Proof.context -> Proof.method) context_parser @@ -90,8 +90,9 @@ (* of applications; just adding this rule to the simplifier *) (* would loop; it also needs careful tuning with the simproc *) (* for functions to avoid further possibilities for looping *) -fun perm_simproc_app' sg ss redex = +fun perm_simproc_app' ctxt redex = let + val thy = Proof_Context.theory_of ctxt; (* the "application" case is only applicable when the head of f is not a *) (* constant or when (f x) is a permuation with two or more arguments *) fun applicable_app t = @@ -107,8 +108,8 @@ (if (applicable_app f) then let val name = Long_Name.base_name n - val at_inst = Global_Theory.get_thm sg ("at_" ^ name ^ "_inst") - val pt_inst = Global_Theory.get_thm sg ("pt_" ^ name ^ "_inst") + val at_inst = Global_Theory.get_thm thy ("at_" ^ name ^ "_inst") + val pt_inst = Global_Theory.get_thm thy ("pt_" ^ name ^ "_inst") in SOME ((at_inst RS (pt_inst RS perm_eq_app)) RS eq_reflection) end else NONE) | _ => NONE @@ -118,7 +119,7 @@ ["Nominal.perm pi x"] perm_simproc_app'; (* a simproc that deals with permutation instances in front of functions *) -fun perm_simproc_fun' sg ss redex = +fun perm_simproc_fun' ctxt redex = let fun applicable_fun t = (case (strip_comb t) of @@ -140,36 +141,36 @@ (* function for simplyfying permutations *) (* stac contains the simplifiation tactic that is *) (* applied (see (no_asm) options below *) -fun perm_simp_gen stac dyn_thms eqvt_thms ss i = +fun perm_simp_gen stac dyn_thms eqvt_thms ctxt i = ("general simplification of permutations", fn st => let - val ss' = Simplifier.global_context (theory_of_thm st) ss + val ctxt' = ctxt addsimps (maps (dynamic_thms st) dyn_thms @ eqvt_thms) addsimprocs [perm_simproc_fun, perm_simproc_app] |> fold Simplifier.del_cong weak_congs |> fold Simplifier.add_cong strong_congs in - stac ss' i st + stac ctxt' i st end); (* general simplification of permutations and permutation that arose from eqvt-problems *) -fun perm_simp stac ss = +fun perm_simp stac ctxt = let val simps = ["perm_swap","perm_fresh_fresh","perm_bij","perm_pi_simp","swap_simps"] in - perm_simp_gen stac simps [] ss + perm_simp_gen stac simps [] ctxt end; -fun eqvt_simp stac ss = +fun eqvt_simp stac ctxt = let val simps = ["perm_swap","perm_fresh_fresh","perm_pi_simp"] - val eqvts_thms = NominalThmDecls.get_eqvt_thms (Simplifier.the_context ss); + val eqvts_thms = NominalThmDecls.get_eqvt_thms ctxt; in - perm_simp_gen stac simps eqvts_thms ss + perm_simp_gen stac simps eqvts_thms ctxt end; (* main simplification tactics for permutations *) -fun perm_simp_tac_gen_i stac tactical ss i = DETERM (tactical (perm_simp stac ss i)); -fun eqvt_simp_tac_gen_i stac tactical ss i = DETERM (tactical (eqvt_simp stac ss i)); +fun perm_simp_tac_gen_i stac tactical ctxt i = DETERM (tactical (perm_simp stac ctxt i)); +fun eqvt_simp_tac_gen_i stac tactical ctxt i = DETERM (tactical (eqvt_simp stac ctxt i)); val perm_simp_tac_i = perm_simp_tac_gen_i simp_tac val perm_asm_simp_tac_i = perm_simp_tac_gen_i asm_simp_tac @@ -187,28 +188,29 @@ (* generating perm_aux'es for the outermost permutation and then un- *) (* folding the definition *) -fun perm_compose_simproc' sg ss redex = +fun perm_compose_simproc' ctxt redex = (case redex of (Const ("Nominal.perm", Type ("fun", [Type ("List.list", [Type (@{type_name Product_Type.prod}, [T as Type (tname,_),_])]),_])) $ pi1 $ (Const ("Nominal.perm", Type ("fun", [Type ("List.list", [Type (@{type_name Product_Type.prod}, [U as Type (uname,_),_])]),_])) $ pi2 $ t)) => let + val thy = Proof_Context.theory_of ctxt val tname' = Long_Name.base_name tname val uname' = Long_Name.base_name uname in if pi1 <> pi2 then (* only apply the composition rule in this case *) if T = U then SOME (Drule.instantiate' - [SOME (ctyp_of sg (fastype_of t))] - [SOME (cterm_of sg pi1), SOME (cterm_of sg pi2), SOME (cterm_of sg t)] - (mk_meta_eq ([Global_Theory.get_thm sg ("pt_"^tname'^"_inst"), - Global_Theory.get_thm sg ("at_"^tname'^"_inst")] MRS pt_perm_compose_aux))) + [SOME (ctyp_of thy (fastype_of t))] + [SOME (cterm_of thy pi1), SOME (cterm_of thy pi2), SOME (cterm_of thy t)] + (mk_meta_eq ([Global_Theory.get_thm thy ("pt_"^tname'^"_inst"), + Global_Theory.get_thm thy ("at_"^tname'^"_inst")] MRS pt_perm_compose_aux))) else SOME (Drule.instantiate' - [SOME (ctyp_of sg (fastype_of t))] - [SOME (cterm_of sg pi1), SOME (cterm_of sg pi2), SOME (cterm_of sg t)] - (mk_meta_eq (Global_Theory.get_thm sg ("cp_"^tname'^"_"^uname'^"_inst") RS + [SOME (ctyp_of thy (fastype_of t))] + [SOME (cterm_of thy pi1), SOME (cterm_of thy pi2), SOME (cterm_of thy t)] + (mk_meta_eq (Global_Theory.get_thm thy ("cp_"^tname'^"_"^uname'^"_inst") RS cp1_aux))) else NONE end @@ -217,13 +219,12 @@ val perm_compose_simproc = Simplifier.simproc_global @{theory} "perm_compose" ["Nominal.perm pi1 (Nominal.perm pi2 t)"] perm_compose_simproc'; -fun perm_compose_tac ss i = +fun perm_compose_tac ctxt i = ("analysing permutation compositions on the lhs", fn st => EVERY [rtac trans i, - asm_full_simp_tac (Simplifier.global_context (theory_of_thm st) empty_ss - addsimprocs [perm_compose_simproc]) i, - asm_full_simp_tac (HOL_basic_ss addsimps [perm_aux_fold]) i] st); + asm_full_simp_tac (empty_simpset ctxt addsimprocs [perm_compose_simproc]) i, + asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [perm_aux_fold]) i] st); fun apply_cong_tac i = ("application of congruence", cong_tac i); @@ -247,32 +248,32 @@ (* to decide equation that come from support problems *) (* since it contains looping rules the "recursion" - depth is set *) (* to 10 - this seems to be sufficient in most cases *) -fun perm_extend_simp_tac_i tactical ss = - let fun perm_extend_simp_tac_aux tactical ss n = +fun perm_extend_simp_tac_i tactical ctxt = + let fun perm_extend_simp_tac_aux tactical ctxt n = if n=0 then K all_tac else DETERM o (FIRST'[fn i => tactical ("splitting conjunctions on the rhs", rtac conjI i), - fn i => tactical (perm_simp asm_full_simp_tac ss i), - fn i => tactical (perm_compose_tac ss i), + fn i => tactical (perm_simp asm_full_simp_tac ctxt i), + fn i => tactical (perm_compose_tac ctxt i), fn i => tactical (apply_cong_tac i), fn i => tactical (unfold_perm_fun_def_tac i), fn i => tactical (ext_fun_tac i)] - THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ss (n-1)))) - in perm_extend_simp_tac_aux tactical ss 10 end; + THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ctxt (n-1)))) + in perm_extend_simp_tac_aux tactical ctxt 10 end; (* tactic that tries to solve "supports"-goals; first it *) (* unfolds the support definition and strips off the *) (* intros, then applies eqvt_simp_tac *) -fun supports_tac_i tactical ss i = +fun supports_tac_i tactical ctxt i = let val simps = [supports_def, Thm.symmetric fresh_def, fresh_prod] in - EVERY [tactical ("unfolding of supports ", simp_tac (HOL_basic_ss addsimps simps) i), + EVERY [tactical ("unfolding of supports ", simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps) i), tactical ("stripping of foralls ", REPEAT_DETERM (rtac allI i)), tactical ("geting rid of the imps ", rtac impI i), tactical ("eliminating conjuncts ", REPEAT_DETERM (etac conjE i)), - tactical ("applying eqvt_simp ", eqvt_simp_tac_gen_i asm_full_simp_tac tactical ss i )] + tactical ("applying eqvt_simp ", eqvt_simp_tac_gen_i asm_full_simp_tac tactical ctxt i )] end; @@ -288,7 +289,7 @@ | collect_vars i (t $ u) vs = collect_vars i u (collect_vars i t vs); (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) -fun finite_guess_tac_i tactical ss i st = +fun finite_guess_tac_i tactical ctxt i st = let val goal = nth (cprems_of st) (i - 1) in case Envir.eta_contract (Logic.strip_assums_concl (term_of goal)) of @@ -310,12 +311,12 @@ val supports_rule'' = Drule.cterm_instantiate [(cert (head_of S), cert s')] supports_rule' val fin_supp = dynamic_thms st ("fin_supp") - val ss' = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp + val ctxt' = ctxt addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp in (tactical ("guessing of the right supports-set", EVERY [compose_tac (false, supports_rule'', 2) i, - asm_full_simp_tac ss' (i+1), - supports_tac_i tactical ss i])) st + asm_full_simp_tac ctxt' (i+1), + supports_tac_i tactical ctxt i])) st end | _ => Seq.empty end @@ -327,13 +328,13 @@ (* it first collects all free variables and tries to show that the *) (* support of these free variables (op supports) the goal *) (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) -fun fresh_guess_tac_i tactical ss i st = +fun fresh_guess_tac_i tactical ctxt i st = let val goal = nth (cprems_of st) (i - 1) val fin_supp = dynamic_thms st ("fin_supp") val fresh_atm = dynamic_thms st ("fresh_atm") - val ss1 = ss addsimps [Thm.symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm - val ss2 = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp + val ctxt1 = ctxt addsimps [Thm.symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm + val ctxt2 = ctxt addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp in case Logic.strip_assums_concl (term_of goal) of _ $ (Const ("Nominal.fresh", Type ("fun", [T, _])) $ _ $ t) => @@ -356,14 +357,14 @@ in (tactical ("guessing of the right set that supports the goal", (EVERY [compose_tac (false, supports_fresh_rule'', 3) i, - asm_full_simp_tac ss1 (i+2), - asm_full_simp_tac ss2 (i+1), - supports_tac_i tactical ss i]))) st + asm_full_simp_tac ctxt1 (i+2), + asm_full_simp_tac ctxt2 (i+1), + supports_tac_i tactical ctxt i]))) st end (* when a term-constructor contains more than one binder, it is useful *) (* in nominal_primrecs to try whether the goal can be solved by an hammer *) | _ => (tactical ("if it is not of the form _\_, then try the simplifier", - (asm_full_simp_tac (HOL_ss addsimps [fresh_prod]@fresh_atm) i))) st + (asm_full_simp_tac (put_simpset HOL_ss ctxt addsimps [fresh_prod]@fresh_atm) i))) st end handle General.Subscript => Seq.empty; (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) @@ -399,19 +400,19 @@ val perm_simp_meth = Scan.lift perm_simp_options --| Method.sections (Simplifier.simp_modifiers') >> - (fn tac => fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o tac (simpset_of ctxt))); + (fn tac => fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o tac ctxt)); (* setup so that the simpset is used which is active at the moment when the tactic is called *) fun local_simp_meth_setup tac = Method.sections (Simplifier.simp_modifiers' @ Splitter.split_modifiers) >> - (K (SIMPLE_METHOD' o tac o simpset_of)); + (K (SIMPLE_METHOD' o tac)); (* uses HOL_basic_ss only and fails if the tactic does not solve the subgoal *) fun basic_simp_meth_setup debug tac = - Scan.depend (fn ctxt => Scan.succeed (Simplifier.map_ss (fn _ => HOL_basic_ss) ctxt, ())) -- + Scan.depend (fn context => Scan.succeed (Simplifier.map_ss (put_simpset HOL_basic_ss) context, ())) -- Method.sections (Simplifier.simp_modifiers' @ Splitter.split_modifiers) >> - (K (SIMPLE_METHOD' o (if debug then tac else SOLVED' o tac) o simpset_of)); + (K (SIMPLE_METHOD' o (if debug then tac else SOLVED' o tac))); val perm_simp_meth_debug = local_simp_meth_setup dperm_simp_tac; val perm_extend_simp_meth = local_simp_meth_setup perm_extend_simp_tac; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Nominal/nominal_thmdecls.ML --- a/src/HOL/Nominal/nominal_thmdecls.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Nominal/nominal_thmdecls.ML Thu Apr 18 17:07:01 2013 +0200 @@ -64,7 +64,7 @@ dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)), tactic ctxt ("getting rid of the pi on the right", rtac @{thm perm_boolI}), tactic ctxt ("getting rid of all remaining perms", - full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))] + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps perm_pi_simp))] end; fun get_derived_thm ctxt hyp concl orig_thm pi typi = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Old_Number_Theory/Chinese.thy --- a/src/HOL/Old_Number_Theory/Chinese.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Old_Number_Theory/Chinese.thy Thu Apr 18 17:07:01 2013 +0200 @@ -243,7 +243,7 @@ prefer 6 apply (simp add: mult_ac) apply (unfold xilin_sol_def) - apply (tactic {* asm_simp_tac @{simpset} 6 *}) + apply (tactic {* asm_simp_tac @{context} 6 *}) apply (rule_tac [6] ex1_implies_ex [THEN someI_ex]) apply (rule_tac [6] unique_xi_sol) apply (rule_tac [3] funprod_zdvd) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Old_Number_Theory/WilsonBij.thy --- a/src/HOL/Old_Number_Theory/WilsonBij.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Old_Number_Theory/WilsonBij.thy Thu Apr 18 17:07:01 2013 +0200 @@ -143,7 +143,7 @@ apply (rule_tac [7] zcong_trans) apply (tactic {* stac @{thm zcong_sym} 8 *}) apply (erule_tac [7] inv_is_inv) - apply (tactic "asm_simp_tac @{simpset} 9") + apply (tactic "asm_simp_tac @{context} 9") apply (erule_tac [9] inv_is_inv) apply (rule_tac [6] zless_zprime_imp_zrelprime) apply (rule_tac [8] inv_less) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Orderings.thy Thu Apr 18 17:07:01 2013 +0200 @@ -597,8 +597,8 @@ fun prp t thm = Thm.prop_of thm = t; (* FIXME aconv!? *) -fun prove_antisym_le sg ss ((le as Const(_,T)) $ r $ s) = - let val prems = Simplifier.prems_of ss; +fun prove_antisym_le ctxt ((le as Const(_,T)) $ r $ s) = + let val prems = Simplifier.prems_of ctxt; val less = Const (@{const_name less}, T); val t = HOLogic.mk_Trueprop(le $ s $ r); in case find_first (prp t) prems of @@ -612,8 +612,8 @@ end handle THM _ => NONE; -fun prove_antisym_less sg ss (NotC $ ((less as Const(_,T)) $ r $ s)) = - let val prems = Simplifier.prems_of ss; +fun prove_antisym_less ctxt (NotC $ ((less as Const(_,T)) $ r $ s)) = + let val prems = Simplifier.prems_of ctxt; val le = Const (@{const_name less_eq}, T); val t = HOLogic.mk_Trueprop(le $ r $ s); in case find_first (prp t) prems of @@ -628,13 +628,13 @@ handle THM _ => NONE; fun add_simprocs procs thy = - Simplifier.map_simpset_global (fn ss => ss + map_theory_simpset (fn ctxt => ctxt addsimprocs (map (fn (name, raw_ts, proc) => Simplifier.simproc_global thy name raw_ts proc) procs)) thy; fun add_solver name tac = - Simplifier.map_simpset_global (fn ss => ss addSolver - mk_solver name (fn ss => tac (Simplifier.the_context ss) (Simplifier.prems_of ss))); + map_theory_simpset (fn ctxt0 => ctxt0 addSolver + mk_solver name (fn ctxt => tac ctxt (Simplifier.prems_of ctxt))); in add_simprocs [ diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Probability/measurable.ML --- a/src/HOL/Probability/measurable.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Probability/measurable.ML Thu Apr 18 17:07:01 2013 +0200 @@ -8,7 +8,7 @@ sig datatype level = Concrete | Generic - val simproc : simpset -> cterm -> thm option + val simproc : Proof.context -> cterm -> thm option val method : (Proof.context -> Method.method) context_parser val measurable_tac : Proof.context -> thm list -> tactic @@ -151,7 +151,8 @@ in if null cps then no_tac else debug_tac ctxt (K "split countable fun") (resolve_tac cps i) end handle TERM _ => no_tac) 1) -fun measurable_tac' ctxt ss facts = let +fun measurable_tac' ctxt facts = + let val imported_thms = (maps (import_theorem (Context.Proof ctxt) o Simplifier.norm_hhf) facts) @ get_all ctxt @@ -202,7 +203,7 @@ val depth_measurable_tac = REPEAT_cnt (fn n => (COND (is_cond_formula 1) - (debug_tac ctxt (K ("simp " ^ string_of_int n)) (SOLVED' (asm_full_simp_tac ss) 1)) + (debug_tac ctxt (K ("simp " ^ string_of_int n)) (SOLVED' (asm_full_simp_tac ctxt) 1)) ((debug_tac ctxt (K ("single " ^ string_of_int n)) (resolve_tac imported_thms 1)) APPEND (split_app_tac ctxt 1) APPEND (splitter 1)))) 0 @@ -210,7 +211,7 @@ in debug_tac ctxt (debug_facts "start") depth_measurable_tac end; fun measurable_tac ctxt facts = - TAKE (Config.get ctxt backtrack) (measurable_tac' ctxt (simpset_of ctxt) facts); + TAKE (Config.get ctxt backtrack) (measurable_tac' ctxt facts); val attr_add = Thm.declaration_attribute o add_thm; @@ -227,11 +228,11 @@ val method : (Proof.context -> Method.method) context_parser = Scan.lift (Scan.succeed (fn ctxt => METHOD (fn facts => measurable_tac ctxt facts))); -fun simproc ss redex = let - val ctxt = Simplifier.the_context ss; +fun simproc ctxt redex = + let val t = HOLogic.mk_Trueprop (term_of redex); fun tac {context = ctxt, prems = _ } = - SOLVE (measurable_tac' ctxt ss (Simplifier.prems_of ss)); + SOLVE (measurable_tac' ctxt (Simplifier.prems_of ctxt)); in try (fn () => Goal.prove ctxt [] [] t tac RS @{thm Eq_TrueI}) () end; end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Product_Type.thy Thu Apr 18 17:07:01 2013 +0200 @@ -415,16 +415,21 @@ | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u | exists_paired_all (Abs (_, _, t)) = exists_paired_all t | exists_paired_all _ = false; - val ss = HOL_basic_ss - addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}] - addsimprocs [@{simproc unit_eq}]; + val ss = + simpset_of + (put_simpset HOL_basic_ss @{context} + addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}] + addsimprocs [@{simproc unit_eq}]); in - val split_all_tac = SUBGOAL (fn (t, i) => - if exists_paired_all t then safe_full_simp_tac ss i else no_tac); - val unsafe_split_all_tac = SUBGOAL (fn (t, i) => - if exists_paired_all t then full_simp_tac ss i else no_tac); - fun split_all th = - if exists_paired_all (Thm.prop_of th) then full_simplify ss th else th; + fun split_all_tac ctxt = SUBGOAL (fn (t, i) => + if exists_paired_all t then safe_full_simp_tac (put_simpset ss ctxt) i else no_tac); + + fun unsafe_split_all_tac ctxt = SUBGOAL (fn (t, i) => + if exists_paired_all t then full_simp_tac (put_simpset ss ctxt) i else no_tac); + + fun split_all ctxt th = + if exists_paired_all (Thm.prop_of th) + then full_simplify (put_simpset ss ctxt) th else th; end; *} @@ -451,7 +456,8 @@ ML {* local - val cond_split_eta_ss = HOL_basic_ss addsimps @{thms cond_split_eta}; + val cond_split_eta_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms cond_split_eta}); fun Pair_pat k 0 (Bound m) = (m = k) | Pair_pat k i (Const (@{const_name Pair}, _) $ Bound m $ t) = i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t @@ -463,9 +469,9 @@ fun split_pat tp i (Abs (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE | split_pat tp i (Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t | split_pat tp i _ = NONE; - fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] [] + fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))) - (K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1))); + (K (simp_tac (put_simpset cond_split_eta_ss ctxt) 1))); fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t | beta_term_pat k i (t $ u) = @@ -479,20 +485,20 @@ else (subst arg k i t $ subst arg k i u) | subst arg k i t = t; in - fun beta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t) $ arg) = + fun beta_proc ctxt (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t) $ arg) = (case split_pat beta_term_pat 1 t of - SOME (i, f) => SOME (metaeq ss s (subst arg 0 i f)) + SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f)) | NONE => NONE) | beta_proc _ _ = NONE; - fun eta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = + fun eta_proc ctxt (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = (case split_pat eta_term_pat 1 t of - SOME (_, ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end)) + SOME (_, ft) => SOME (metaeq ctxt s (let val (f $ arg) = ft in f end)) | NONE => NONE) | eta_proc _ _ = NONE; end; *} -simproc_setup split_beta ("split f z") = {* fn _ => fn ss => fn ct => beta_proc ss (term_of ct) *} -simproc_setup split_eta ("split f") = {* fn _ => fn ss => fn ct => eta_proc ss (term_of ct) *} +simproc_setup split_beta ("split f z") = {* fn _ => fn ctxt => fn ct => beta_proc ctxt (term_of ct) *} +simproc_setup split_eta ("split f") = {* fn _ => fn ctxt => fn ct => eta_proc ctxt (term_of ct) *} lemma split_beta [mono]: "(%(x, y). P x y) z = P (fst z) (snd z)" by (subst surjective_pairing, rule split_conv) @@ -572,10 +578,11 @@ | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u | exists_p_split (Abs (_, _, t)) = exists_p_split t | exists_p_split _ = false; - val ss = HOL_basic_ss addsimps @{thms split_conv}; in -val split_conv_tac = SUBGOAL (fn (t, i) => - if exists_p_split t then safe_full_simp_tac ss i else no_tac); +fun split_conv_tac ctxt = SUBGOAL (fn (t, i) => + if exists_p_split t + then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_conv}) i + else no_tac); end; *} @@ -1154,7 +1161,7 @@ ML_file "Tools/set_comprehension_pointfree.ML" setup {* - Code_Preproc.map_pre (fn ss => ss addsimprocs + Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs [Raw_Simplifier.make_simproc {name = "set comprehension", lhss = [@{cpat "Collect ?P"}], proc = K Set_Comprehension_Pointfree.code_simproc, identifier = []}]) *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Record_Benchmark/Record_Benchmark.thy --- a/src/HOL/Record_Benchmark/Record_Benchmark.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Record_Benchmark/Record_Benchmark.thy Thu Apr 18 17:07:01 2013 +0200 @@ -355,46 +355,50 @@ by simp lemma "(r\A255:=x,A253:=y,A255:=z \) = r\A253:=y,A255:=z\" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.upd_simproc]) 1*}) + apply (tactic {* simp_tac + (put_simpset HOL_basic_ss @{context} addsimprocs [Record.upd_simproc]) 1*}) done lemma "(\r. P (A155 r)) \ (\x. P x)" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.split_simproc (K ~1)]) 1*}) + apply (tactic {* simp_tac + (put_simpset HOL_basic_ss @{context} addsimprocs [Record.split_simproc (K ~1)]) 1*}) apply simp done lemma "(\r. P (A155 r)) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply simp done lemma "(\r. P (A155 r)) \ (\x. P x)" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.split_simproc (K ~1)]) 1*}) + apply (tactic {* simp_tac + (put_simpset HOL_basic_ss @{context} addsimprocs [Record.split_simproc (K ~1)]) 1*}) apply simp done lemma "(\r. P (A155 r)) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply simp done lemma "\r. P (A155 r) \ (\x. P x)" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.split_simproc (K ~1)]) 1*}) + apply (tactic {* simp_tac + (put_simpset HOL_basic_ss @{context} addsimprocs [Record.split_simproc (K ~1)]) 1*}) apply auto done lemma "\r. P (A155 r) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply auto done lemma "P (A155 r) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply auto done lemma fixes r shows "P (A155 r) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply auto done @@ -405,14 +409,15 @@ assume "P (A155 r)" then have "\x. P x" apply - - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply auto done end lemma "\r. A155 r = x" - apply (tactic {*simp_tac (HOL_basic_ss addsimprocs [Record.ex_sel_eq_simproc]) 1*}) + apply (tactic {*simp_tac + (put_simpset HOL_basic_ss @{context} addsimprocs [Record.ex_sel_eq_simproc]) 1*}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/SET_Protocol/Message_SET.thy --- a/src/HOL/SET_Protocol/Message_SET.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/SET_Protocol/Message_SET.thy Thu Apr 18 17:07:01 2013 +0200 @@ -853,12 +853,12 @@ impOfSubs @{thm Fake_parts_insert}] THEN' eresolve_tac [asm_rl, @{thm synth.Inj}]; -fun Fake_insert_simp_tac ss i = - REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i; +fun Fake_insert_simp_tac ctxt i = + REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ctxt i; fun atomic_spy_analz_tac ctxt = SELECT_GOAL - (Fake_insert_simp_tac (simpset_of ctxt) 1 THEN + (Fake_insert_simp_tac ctxt 1 THEN IF_UNSOLVED (Blast.depth_tac (ctxt addIs [@{thm analz_insertI}, impOfSubs @{thm analz_subset_parts}]) 4 1)); @@ -871,7 +871,7 @@ (REPEAT o CHANGED) (res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1), (*...allowing further simplifications*) - simp_tac (simpset_of ctxt) 1, + simp_tac ctxt 1, REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])), DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i); *} @@ -932,7 +932,7 @@ "for debugging spy_analz" method_setup Fake_insert_simp = {* - Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o simpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac) *} "for debugging spy_analz" end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/SET_Protocol/Public_SET.thy --- a/src/HOL/SET_Protocol/Public_SET.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/SET_Protocol/Public_SET.thy Thu Apr 18 17:07:01 2013 +0200 @@ -344,7 +344,7 @@ (*Tactic for possibility theorems*) fun possibility_tac ctxt = REPEAT (*omit used_Says so that Nonces start from different traces!*) - (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}])) + (ALLGOALS (simp_tac (ctxt delsimps [@{thm used_Says}, @{thm used_Notes}])) THEN REPEAT_FIRST (eq_assume_tac ORELSE' resolve_tac [refl, conjI, @{thm Nonce_supply}])) @@ -353,7 +353,7 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/SPARK/Tools/spark_vcs.ML --- a/src/HOL/SPARK/Tools/spark_vcs.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/SPARK/Tools/spark_vcs.ML Thu Apr 18 17:07:01 2013 +0200 @@ -211,17 +211,17 @@ rtac @{thm subsetI} 1 THEN Datatype_Aux.exh_tac (K (#exhaust (Datatype.the_info (Proof_Context.theory_of lthy) tyname'))) 1 THEN - ALLGOALS (asm_full_simp_tac (simpset_of lthy))); + ALLGOALS (asm_full_simp_tac lthy)); val finite_UNIV = Goal.prove lthy [] [] (HOLogic.mk_Trueprop (Const (@{const_name finite}, HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T)) - (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1); + (fn _ => simp_tac (lthy addsimps [UNIV_eq]) 1); val card_UNIV = Goal.prove lthy [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (card, HOLogic.mk_number HOLogic.natT k))) - (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1); + (fn _ => simp_tac (lthy addsimps [UNIV_eq]) 1); val range_pos = Goal.prove lthy [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq @@ -233,12 +233,12 @@ HOLogic.mk_number HOLogic.intT 0 $ (@{term int} $ card)))) (fn _ => - simp_tac (simpset_of lthy addsimps [card_UNIV]) 1 THEN - simp_tac (simpset_of lthy addsimps [UNIV_eq, def1]) 1 THEN + simp_tac (lthy addsimps [card_UNIV]) 1 THEN + simp_tac (lthy addsimps [UNIV_eq, def1]) 1 THEN rtac @{thm subset_antisym} 1 THEN - simp_tac (simpset_of lthy) 1 THEN + simp_tac lthy 1 THEN rtac @{thm subsetI} 1 THEN - asm_full_simp_tac (simpset_of lthy addsimps @{thms interval_expand} + asm_full_simp_tac (lthy addsimps @{thms interval_expand} delsimps @{thms atLeastLessThan_iff}) 1); val lthy' = @@ -246,34 +246,31 @@ Class.intro_classes_tac [] THEN rtac finite_UNIV 1 THEN rtac range_pos 1 THEN - simp_tac (HOL_basic_ss addsimps [def3]) 1 THEN - simp_tac (HOL_basic_ss addsimps [def2]) 1) lthy; + simp_tac (put_simpset HOL_basic_ss lthy addsimps [def3]) 1 THEN + simp_tac (put_simpset HOL_basic_ss lthy addsimps [def2]) 1) lthy; val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) => let val n = HOLogic.mk_number HOLogic.intT i; val th = Goal.prove lthy' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n))) - (fn _ => simp_tac (simpset_of lthy' addsimps [def1]) 1); + (fn _ => simp_tac (lthy' addsimps [def1]) 1); val th' = Goal.prove lthy' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c))) (fn _ => rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN - simp_tac (simpset_of lthy' addsimps - [@{thm pos_val}, range_pos, card_UNIV, th]) 1) + simp_tac (lthy' addsimps [@{thm pos_val}, range_pos, card_UNIV, th]) 1) in (th, th') end) cs); val first_el = Goal.prove lthy' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (@{const_name first_el}, T), hd cs))) - (fn _ => simp_tac (simpset_of lthy' addsimps - [@{thm first_el_def}, hd val_eqs]) 1); + (fn _ => simp_tac (lthy' addsimps [@{thm first_el_def}, hd val_eqs]) 1); val last_el = Goal.prove lthy' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (@{const_name last_el}, T), List.last cs))) - (fn _ => simp_tac (simpset_of lthy' addsimps - [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1); + (fn _ => simp_tac (lthy' addsimps [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1); in lthy' |> Local_Theory.note diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Set.thy --- a/src/HOL/Set.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Set.thy Thu Apr 18 17:07:01 2013 +0200 @@ -380,7 +380,8 @@ *} setup {* - map_theory_claset (fn ctxt => ctxt addbefore ("bspec", dtac @{thm bspec} THEN' assume_tac)) + map_theory_claset (fn ctxt => + ctxt addbefore ("bspec", fn _ => dtac @{thm bspec} THEN' assume_tac)) *} ML {* diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Statespace/distinct_tree_prover.ML --- a/src/HOL/Statespace/distinct_tree_prover.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Statespace/distinct_tree_prover.ML Thu Apr 18 17:07:01 2013 +0200 @@ -356,16 +356,15 @@ | _ => no_tac)) fun distinctFieldSolver names = - mk_solver "distinctFieldSolver" (distinctTree_tac names o Simplifier.the_context); + mk_solver "distinctFieldSolver" (distinctTree_tac names); fun distinct_simproc names = Simplifier.simproc_global @{theory HOL} "DistinctTreeProver.distinct_simproc" ["x = y"] - (fn thy => fn ss => fn (Const (@{const_name HOL.eq}, _) $ x $ y) => - (case try Simplifier.the_context ss of - SOME ctxt => + (fn ctxt => + (fn Const (@{const_name HOL.eq}, _) $ x $ y => Option.map (fn neq => @{thm neq_to_eq_False} OF [neq]) (get_fst_success (neq_x_y ctxt x y) names) - | NONE => NONE)); + | _ => NONE)); end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Statespace/state_fun.ML --- a/src/HOL/Statespace/state_fun.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Statespace/state_fun.ML Thu Apr 18 17:07:01 2013 +0200 @@ -15,7 +15,7 @@ val ex_lookup_eq_simproc : simproc val ex_lookup_ss : simpset val lazy_conj_simproc : simproc - val string_eq_simp_tac : int -> tactic + val string_eq_simp_tac : Proof.context -> int -> tactic val setup : theory -> theory end; @@ -54,44 +54,49 @@ val lazy_conj_simproc = Simplifier.simproc_global @{theory HOL} "lazy_conj_simp" ["P & Q"] - (fn thy => fn ss => fn t => - (case t of (Const (@{const_name HOL.conj},_) $ P $ Q) => - let - val P_P' = Simplifier.rewrite ss (cterm_of thy P); - val P' = P_P' |> prop_of |> Logic.dest_equals |> #2; - in - if isFalse P' then SOME (conj1_False OF [P_P']) - else - let - val Q_Q' = Simplifier.rewrite ss (cterm_of thy Q); - val Q' = Q_Q' |> prop_of |> Logic.dest_equals |> #2; - in - if isFalse Q' then SOME (conj2_False OF [Q_Q']) - else if isTrue P' andalso isTrue Q' then SOME (conj_True OF [P_P', Q_Q']) - else if P aconv P' andalso Q aconv Q' then NONE - else SOME (conj_cong OF [P_P', Q_Q']) - end - end - | _ => NONE)); + (fn ctxt => fn t => + let val thy = Proof_Context.theory_of ctxt in + (case t of (Const (@{const_name HOL.conj},_) $ P $ Q) => + let + val P_P' = Simplifier.rewrite ctxt (cterm_of thy P); + val P' = P_P' |> prop_of |> Logic.dest_equals |> #2; + in + if isFalse P' then SOME (conj1_False OF [P_P']) + else + let + val Q_Q' = Simplifier.rewrite ctxt (cterm_of thy Q); + val Q' = Q_Q' |> prop_of |> Logic.dest_equals |> #2; + in + if isFalse Q' then SOME (conj2_False OF [Q_Q']) + else if isTrue P' andalso isTrue Q' then SOME (conj_True OF [P_P', Q_Q']) + else if P aconv P' andalso Q aconv Q' then NONE + else SOME (conj_cong OF [P_P', Q_Q']) + end + end + | _ => NONE) + end); -val string_eq_simp_tac = simp_tac (HOL_basic_ss - addsimps (@{thms list.inject} @ @{thms char.inject} - @ @{thms list.distinct} @ @{thms char.distinct} @ @{thms simp_thms}) - addsimprocs [lazy_conj_simproc] - |> Simplifier.add_cong @{thm block_conj_cong}); +fun string_eq_simp_tac ctxt = + simp_tac (put_simpset HOL_basic_ss ctxt + addsimps (@{thms list.inject} @ @{thms char.inject} + @ @{thms list.distinct} @ @{thms char.distinct} @ @{thms simp_thms}) + addsimprocs [lazy_conj_simproc] + |> Simplifier.add_cong @{thm block_conj_cong}); end; -val lookup_ss = (HOL_basic_ss - addsimps (@{thms list.inject} @ @{thms char.inject} - @ @{thms list.distinct} @ @{thms char.distinct} @ @{thms simp_thms} - @ [@{thm StateFun.lookup_update_id_same}, @{thm StateFun.id_id_cancel}, - @{thm StateFun.lookup_update_same}, @{thm StateFun.lookup_update_other}]) - addsimprocs [lazy_conj_simproc] - addSolver StateSpace.distinctNameSolver - |> fold Simplifier.add_cong @{thms block_conj_cong}); +val lookup_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps (@{thms list.inject} @ @{thms char.inject} + @ @{thms list.distinct} @ @{thms char.distinct} @ @{thms simp_thms} + @ [@{thm StateFun.lookup_update_id_same}, @{thm StateFun.id_id_cancel}, + @{thm StateFun.lookup_update_same}, @{thm StateFun.lookup_update_other}]) + addsimprocs [lazy_conj_simproc] + addSolver StateSpace.distinctNameSolver + |> fold Simplifier.add_cong @{thms block_conj_cong}); -val ex_lookup_ss = HOL_ss addsimps @{thms StateFun.ex_id}; +val ex_lookup_ss = + simpset_of (put_simpset HOL_ss @{context} addsimps @{thms StateFun.ex_id}); structure Data = Generic_Data @@ -108,10 +113,11 @@ val lookup_simproc = Simplifier.simproc_global @{theory} "lookup_simp" ["lookup d n (update d' c m v s)"] - (fn thy => fn ss => fn t => + (fn ctxt => fn t => (case t of (Const (@{const_name StateFun.lookup}, lT) $ destr $ n $ (s as Const (@{const_name StateFun.update}, uT) $ _ $ _ $ _ $ _ $ _)) => (let + val thy = Proof_Context.theory_of ctxt; val (_::_::_::_::sT::_) = binder_types uT; val mi = maxidx_of_term t; fun mk_upds (Const (@{const_name StateFun.update}, uT) $ d' $ c $ m $ v $ s) = @@ -140,10 +146,9 @@ val ct = cterm_of thy (Const (@{const_name StateFun.lookup}, lT) $ destr $ n $ fst (mk_upds s)); - val ctxt = Simplifier.the_context ss; val basic_ss = #1 (Data.get (Context.Proof ctxt)); - val ss' = Simplifier.context (Config.put simp_depth_limit 100 ctxt) basic_ss; - val thm = Simplifier.rewrite ss' ct; + val ctxt' = ctxt |> Config.put simp_depth_limit 100 |> put_simpset basic_ss; + val thm = Simplifier.rewrite ctxt' ct; in if (op aconv) (Logic.dest_equals (prop_of thm)) then NONE @@ -156,17 +161,18 @@ local val meta_ext = @{thm StateFun.meta_ext}; -val ss' = (HOL_ss addsimps - (@{thm StateFun.update_apply} :: @{thm Fun.o_apply} :: @{thms list.inject} @ @{thms char.inject} - @ @{thms list.distinct} @ @{thms char.distinct}) - addsimprocs [lazy_conj_simproc, StateSpace.distinct_simproc] - |> fold Simplifier.add_cong @{thms block_conj_cong}); +val ss' = + simpset_of (put_simpset HOL_ss @{context} addsimps + (@{thm StateFun.update_apply} :: @{thm Fun.o_apply} :: @{thms list.inject} @ @{thms char.inject} + @ @{thms list.distinct} @ @{thms char.distinct}) + addsimprocs [lazy_conj_simproc, StateSpace.distinct_simproc] + |> fold Simplifier.add_cong @{thms block_conj_cong}); in val update_simproc = Simplifier.simproc_global @{theory} "update_simp" ["update d c n v s"] - (fn thy => fn ss => fn t => + (fn ctxt => fn t => (case t of ((upd as Const (@{const_name StateFun.update}, uT)) $ d $ c $ n $ v $ s) => let @@ -237,18 +243,18 @@ end | mk_updterm _ t = init_seed t; - val ctxt = Simplifier.the_context ss |> Config.put simp_depth_limit 100; - val ss1 = Simplifier.context ctxt ss'; - val ss2 = Simplifier.context ctxt (#1 (Data.get (Context.Proof ctxt))); + val ctxt0 = Config.put simp_depth_limit 100 ctxt; + val ctxt1 = put_simpset ss' ctxt0; + val ctxt2 = put_simpset (#1 (Data.get (Context.Proof ctxt0))) ctxt0; in (case mk_updterm [] t of (trm, trm', vars, _, true) => let val eq1 = - Goal.prove ctxt [] [] + Goal.prove ctxt0 [] [] (Logic.list_all (vars, Logic.mk_equals (trm, trm'))) - (fn _ => rtac meta_ext 1 THEN simp_tac ss1 1); - val eq2 = Simplifier.asm_full_rewrite ss2 (Thm.dest_equals_rhs (cprop_of eq1)); + (fn _ => rtac meta_ext 1 THEN simp_tac ctxt1 1); + val eq2 = Simplifier.asm_full_rewrite ctxt2 (Thm.dest_equals_rhs (cprop_of eq1)); in SOME (Thm.transitive eq1 eq2) end | _ => NONE) end @@ -269,14 +275,15 @@ val ex_lookup_eq_simproc = Simplifier.simproc_global @{theory HOL} "ex_lookup_eq_simproc" ["Ex t"] - (fn thy => fn ss => fn t => + (fn ctxt => fn t => let - val ctxt = Simplifier.the_context ss |> Config.put simp_depth_limit 100; + val thy = Proof_Context.theory_of ctxt; + val ex_lookup_ss = #2 (Data.get (Context.Proof ctxt)); - val ss' = Simplifier.context ctxt ex_lookup_ss; + val ctxt' = ctxt |> Config.put simp_depth_limit 100 |> put_simpset ex_lookup_ss; fun prove prop = Goal.prove_global thy [] [] prop - (fn _ => Record.split_simp_tac [] (K ~1) 1 THEN simp_tac ss' 1); + (fn _ => Record.split_simp_tac ctxt [] (K ~1) 1 THEN simp_tac ctxt' 1); fun mkeq (swap, Teq, lT, lo, d, n, x, s) i = let @@ -364,18 +371,21 @@ val mk_destr = gen_constr_destr (fn a => fn b => Syntax.const @{const_name Fun.comp} $ b $ a) "the_"; -val statefun_simp_attr = Thm.declaration_attribute (fn thm => fn ctxt => +val statefun_simp_attr = Thm.declaration_attribute (fn thm => fn context => let - val (lookup_ss, ex_lookup_ss, simprocs_active) = Data.get ctxt; + val ctxt = Context.proof_of context; + val (lookup_ss, ex_lookup_ss, simprocs_active) = Data.get context; val (lookup_ss', ex_lookup_ss') = (case concl_of thm of - (_ $ ((Const (@{const_name Ex}, _) $ _))) => (lookup_ss, ex_lookup_ss addsimps [thm]) - | _ => (lookup_ss addsimps [thm], ex_lookup_ss)); - fun activate_simprocs ctxt = - if simprocs_active then ctxt - else Simplifier.map_ss (fn ss => ss addsimprocs [lookup_simproc, update_simproc]) ctxt; + (_ $ ((Const (@{const_name Ex}, _) $ _))) => + (lookup_ss, simpset_map ctxt (Simplifier.add_simp thm) ex_lookup_ss) + | _ => + (simpset_map ctxt (Simplifier.add_simp thm) lookup_ss, ex_lookup_ss)); + val activate_simprocs = + if simprocs_active then I + else Simplifier.map_ss (fn ctxt => ctxt addsimprocs [lookup_simproc, update_simproc]); in - ctxt + context |> activate_simprocs |> Data.put (lookup_ss', ex_lookup_ss', true) end); diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Statespace/state_space.ML --- a/src/HOL/Statespace/state_space.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Statespace/state_space.ML Thu Apr 18 17:07:01 2013 +0200 @@ -225,22 +225,14 @@ | NONE => no_tac) | _ => no_tac)); -val distinctNameSolver = - mk_solver "distinctNameSolver" (distinctTree_tac o Simplifier.the_context); +val distinctNameSolver = mk_solver "distinctNameSolver" distinctTree_tac; val distinct_simproc = Simplifier.simproc_global @{theory HOL} "StateSpace.distinct_simproc" ["x = y"] - (fn thy => fn ss => (fn (Const (@{const_name HOL.eq},_)$(x as Free _)$(y as Free _)) => - (case try Simplifier.the_context ss of - SOME ctxt => - Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) - (neq_x_y ctxt x y) - | NONE => NONE) + (fn ctxt => (fn (Const (@{const_name HOL.eq},_)$(x as Free _)$(y as Free _)) => + Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) (neq_x_y ctxt x y) | _ => NONE)); -local - val ss = HOL_basic_ss -in fun interprete_parent name dist_thm_name parent_expr thy = let fun solve_tac ctxt = CSUBGOAL (fn (goal, i) => @@ -256,8 +248,6 @@ thy |> prove_interpretation_in tac (name, parent_expr) end; -end; - fun namespace_definition name nameT parent_expr parent_comps new_comps thy = let val all_comps = parent_comps @ new_comps; @@ -283,14 +273,12 @@ | NONE => Symtab.update (name,thm) tt) val tt' = tt |> fold upd all_names; - val activate_simproc = - Simplifier.map_ss - (Simplifier.with_context (Context_Position.set_visible false ctxt) - (fn ss => ss addsimprocs [distinct_simproc])); val context' = - context - |> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent} - |> activate_simproc; + Context_Position.set_visible false ctxt + addsimprocs [distinct_simproc] + |> Context_Position.restore_visible ctxt + |> Context.Proof + |> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent}; in context' end)); val attr = Attrib.internal type_attr; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/String.thy --- a/src/HOL/String.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/String.thy Thu Apr 18 17:07:01 2013 +0200 @@ -252,8 +252,9 @@ setup {* let val nibbles = map_range (Thm.cterm_of @{theory} o HOLogic.mk_nibble) 16; - val simpset = HOL_ss addsimps - @{thms nat_of_nibble.simps mult_0 mult_1 add_0 add_0_right arith_simps numeral_plus_one}; + val simpset = + put_simpset HOL_ss @{context} + addsimps @{thms nat_of_nibble.simps mult_0 mult_1 add_0 add_0_right arith_simps numeral_plus_one}; fun mk_code_eqn x y = Drule.instantiate' [] [SOME x, SOME y] @{thm nat_of_char_Char} |> simplify simpset; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/Buffer/DBuffer.thy --- a/src/HOL/TLA/Buffer/DBuffer.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/Buffer/DBuffer.thy Thu Apr 18 17:07:01 2013 +0200 @@ -59,7 +59,7 @@ apply (rule square_simulation) apply clarsimp apply (tactic - {* action_simp_tac (@{simpset} addsimps (@{thm hd_append} :: @{thms db_defs})) [] [] 1 *}) + {* action_simp_tac (@{context} addsimps (@{thm hd_append} :: @{thms db_defs})) [] [] 1 *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/Inc/Inc.thy --- a/src/HOL/TLA/Inc/Inc.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/Inc/Inc.thy Thu Apr 18 17:07:01 2013 +0200 @@ -170,9 +170,9 @@ --> (pc1 = #g ~> pc1 = #a)" apply (rule SF1) apply (tactic - {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) + {* action_simp_tac (@{context} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) apply (tactic - {* action_simp_tac (@{simpset} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *}) + {* action_simp_tac (@{context} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *}) (* reduce |- []A --> <>Enabled B to |- A --> Enabled B *) apply (auto intro!: InitDmd_gen [temp_use] N1_enabled_at_g [temp_use] dest!: STL2_gen [temp_use] simp: Init_def) @@ -191,8 +191,8 @@ "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True --> (pc2 = #g ~> pc2 = #a)" apply (rule SF1) - apply (tactic {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) - apply (tactic {* action_simp_tac (@{simpset} addsimps @{thm angle_def} :: @{thms Psi_defs}) + apply (tactic {* action_simp_tac (@{context} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) + apply (tactic {* action_simp_tac (@{context} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *}) apply (auto intro!: InitDmd_gen [temp_use] N2_enabled_at_g [temp_use] dest!: STL2_gen [temp_use] simp add: Init_def) @@ -211,9 +211,9 @@ --> (pc2 = #b ~> pc2 = #g)" apply (rule SF1) apply (tactic - {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) + {* action_simp_tac (@{context} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) apply (tactic - {* action_simp_tac (@{simpset} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *}) + {* action_simp_tac (@{context} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *}) apply (auto intro!: InitDmd_gen [temp_use] N2_enabled_at_b [temp_use] dest!: STL2_gen [temp_use] simp: Init_def) done @@ -253,9 +253,9 @@ & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2) --> (pc1 = #a ~> pc1 = #b)" apply (rule SF1) - apply (tactic {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) + apply (tactic {* action_simp_tac (@{context} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *}) apply (tactic - {* action_simp_tac (@{simpset} addsimps (@{thm angle_def} :: @{thms Psi_defs})) [] [] 1 *}) + {* action_simp_tac (@{context} addsimps (@{thm angle_def} :: @{thms Psi_defs})) [] [] 1 *}) apply (clarsimp intro!: N1_enabled_at_both_a [THEN DmdImpl [temp_use]]) apply (auto intro!: BoxDmd2_simple [temp_use] N2_live [temp_use] simp: split_box_conj more_temp_simps) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/Memory/MemClerk.thy --- a/src/HOL/TLA/Memory/MemClerk.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/Memory/MemClerk.thy Thu Apr 18 17:07:01 2013 +0200 @@ -85,7 +85,7 @@ lemma MClkFwd_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> |- Calling send p & ~Calling rcv p & cst!p = #clkA --> Enabled (MClkFwd send rcv cst p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm MClkFwd_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm MClkFwd_def}, @{thm Call_def}, @{thm caller_def}, @{thm rtrner_def}]) [exI] [@{thm base_enabled}, @{thm Pair_inject}] 1 *}) @@ -100,9 +100,9 @@ lemma MClkReply_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> |- Calling send p & ~Calling rcv p & cst!p = #clkB --> Enabled (_(cst!p, rtrner send!p, caller rcv!p))" - apply (tactic {* action_simp_tac @{simpset} + apply (tactic {* action_simp_tac @{context} [@{thm MClkReply_change} RSN (2, @{thm enabled_mono})] [] 1 *}) - apply (tactic {* action_simp_tac (@{simpset} addsimps + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm MClkReply_def}, @{thm Return_def}, @{thm caller_def}, @{thm rtrner_def}]) [exI] [@{thm base_enabled}, @{thm Pair_inject}] 1 *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/Memory/Memory.thy --- a/src/HOL/TLA/Memory/Memory.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/Memory/Memory.thy Thu Apr 18 17:07:01 2013 +0200 @@ -176,9 +176,9 @@ |- Calling ch p & (rs!p ~= #NotAResult) --> Enabled (_(rtrner ch ! p, rs!p))" apply (tactic - {* action_simp_tac @{simpset} [@{thm MemReturn_change} RSN (2, @{thm enabled_mono}) ] [] 1 *}) + {* action_simp_tac @{context} [@{thm MemReturn_change} RSN (2, @{thm enabled_mono}) ] [] 1 *}) apply (tactic - {* action_simp_tac (@{simpset} addsimps [@{thm MemReturn_def}, @{thm Return_def}, + {* action_simp_tac (@{context} addsimps [@{thm MemReturn_def}, @{thm Return_def}, @{thm rtrner_def}]) [exI] [@{thm base_enabled}, @{thm Pair_inject}] 1 *}) done @@ -222,11 +222,11 @@ --> Enabled (_(rtrner ch ! p, rs!p))" apply (auto simp: enabled_disj [try_rewrite] intro!: RWRNext_enabled [temp_use]) apply (case_tac "arg (ch w p)") - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm Read_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm Read_def}, temp_rewrite @{thm enabled_ex}]) [@{thm ReadInner_enabled}, exI] [] 1 *}) apply (force dest: base_pair [temp_use]) apply (erule contrapos_np) - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm Write_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm Write_def}, temp_rewrite @{thm enabled_ex}]) [@{thm WriteInner_enabled}, exI] [] 1 *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/Memory/MemoryImplementation.thy --- a/src/HOL/TLA/Memory/MemoryImplementation.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Thu Apr 18 17:07:01 2013 +0200 @@ -225,13 +225,13 @@ *) ML {* val config_fast_solver = Attrib.setup_config_bool @{binding fast_solver} (K false); - val fast_solver = mk_solver "fast_solver" (fn ss => - if Config.get (Simplifier.the_context ss) config_fast_solver + val fast_solver = mk_solver "fast_solver" (fn ctxt => + if Config.get ctxt config_fast_solver then assume_tac ORELSE' (etac notE) else K no_tac); *} -declaration {* K (Simplifier.map_ss (fn ss => ss addSSolver fast_solver)) *} +setup {* map_theory_simpset (fn ctxt => ctxt addSSolver fast_solver) *} ML {* val temp_elim = make_elim o temp_use *} @@ -248,9 +248,9 @@ apply (rule historyI) apply assumption+ apply (rule MI_base) - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm HInit_def}]) [] [] 1 *}) + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm HInit_def}]) [] [] 1 *}) apply (erule fun_cong) - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm HNext_def}]) + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm HNext_def}]) [@{thm busy_squareI}] [] 1 *}) apply (erule fun_cong) done @@ -350,7 +350,7 @@ lemma S1Hist: "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S1 rmhist p) --> unchanged (rmhist!p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm HNext_def}, @{thm S_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm HNext_def}, @{thm S_def}, @{thm S1_def}, @{thm MemReturn_def}, @{thm RPCFail_def}, @{thm MClkReply_def}, @{thm Return_def}]) [] [temp_use @{thm squareE}] 1 *}) @@ -366,7 +366,7 @@ lemma S2Forward: "|- $(S2 rmhist p) & MClkFwd memCh crCh cst p & unchanged (e p, r p, m p, rmhist!p) --> (S3 rmhist p)$" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm MClkFwd_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm MClkFwd_def}, @{thm Call_def}, @{thm e_def}, @{thm r_def}, @{thm m_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm S_def}, @{thm S2_def}, @{thm S3_def}, @{thm Calling_def}]) [] [] 1 *}) @@ -403,7 +403,7 @@ lemma S3Forward: "|- RPCFwd crCh rmCh rst p & HNext rmhist p & $(S3 rmhist p) & unchanged (e p, c p, m p) --> (S4 rmhist p)$ & unchanged (rmhist!p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm RPCFwd_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm RPCFwd_def}, @{thm HNext_def}, @{thm MemReturn_def}, @{thm RPCFail_def}, @{thm MClkReply_def}, @{thm Return_def}, @{thm Call_def}, @{thm e_def}, @{thm c_def}, @{thm m_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm S_def}, @@ -412,7 +412,7 @@ lemma S3Fail: "|- RPCFail crCh rmCh rst p & $(S3 rmhist p) & HNext rmhist p & unchanged (e p, c p, m p) --> (S6 rmhist p)$" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm HNext_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm HNext_def}, @{thm RPCFail_def}, @{thm Return_def}, @{thm e_def}, @{thm c_def}, @{thm m_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm MVOKBARF_def}, @{thm S_def}, @{thm S3_def}, @{thm S6_def}, @{thm Calling_def}]) [] [] 1 *}) @@ -439,7 +439,7 @@ lemma S4ReadInner: "|- ReadInner rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p) & HNext rmhist p & $(MemInv mm l) --> (S4 rmhist p)$ & unchanged (rmhist!p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ReadInner_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm ReadInner_def}, @{thm GoodRead_def}, @{thm BadRead_def}, @{thm HNext_def}, @{thm MemReturn_def}, @{thm RPCFail_def}, @{thm MClkReply_def}, @{thm Return_def}, @{thm e_def}, @{thm c_def}, @{thm r_def}, @{thm rtrner_def}, @{thm caller_def}, @@ -453,7 +453,7 @@ lemma S4WriteInner: "|- WriteInner rmCh mm ires p l v & $(S4 rmhist p) & unchanged (e p, c p, r p) & HNext rmhist p --> (S4 rmhist p)$ & unchanged (rmhist!p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm WriteInner_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm WriteInner_def}, @{thm GoodWrite_def}, @{thm BadWrite_def}, @{thm HNext_def}, @{thm MemReturn_def}, @{thm RPCFail_def}, @{thm MClkReply_def}, @{thm Return_def}, @{thm e_def}, @{thm c_def}, @{thm r_def}, @{thm rtrner_def}, @{thm caller_def}, @{thm MVNROKBA_def}, @@ -492,14 +492,14 @@ lemma S5Reply: "|- RPCReply crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) --> (S6 rmhist p)$" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm RPCReply_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm RPCReply_def}, @{thm Return_def}, @{thm e_def}, @{thm c_def}, @{thm m_def}, @{thm MVOKBA_def}, @{thm MVOKBARF_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm S_def}, @{thm S5_def}, @{thm S6_def}, @{thm Calling_def}]) [] [] 1 *}) lemma S5Fail: "|- RPCFail crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) --> (S6 rmhist p)$" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm RPCFail_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm RPCFail_def}, @{thm Return_def}, @{thm e_def}, @{thm c_def}, @{thm m_def}, @{thm MVOKBARF_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm S_def}, @{thm S5_def}, @{thm S6_def}, @{thm Calling_def}]) [] [] 1 *}) @@ -525,7 +525,7 @@ lemma S6Retry: "|- MClkRetry memCh crCh cst p & HNext rmhist p & $S6 rmhist p & unchanged (e p,r p,m p) --> (S3 rmhist p)$ & unchanged (rmhist!p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm HNext_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm HNext_def}, @{thm MClkReply_def}, @{thm MClkRetry_def}, @{thm Call_def}, @{thm Return_def}, @{thm e_def}, @{thm r_def}, @{thm m_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm S_def}, @{thm S6_def}, @{thm S3_def}, @{thm Calling_def}]) [] [] 1 *}) @@ -533,7 +533,7 @@ lemma S6Reply: "|- MClkReply memCh crCh cst p & HNext rmhist p & $S6 rmhist p & unchanged (e p,r p,m p) --> (S1 rmhist p)$" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm HNext_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm HNext_def}, @{thm MemReturn_def}, @{thm RPCFail_def}, @{thm Return_def}, @{thm MClkReply_def}, @{thm e_def}, @{thm r_def}, @{thm m_def}, @{thm caller_def}, @{thm rtrner_def}, @{thm S_def}, @{thm S6_def}, @{thm S1_def}, @{thm Calling_def}]) [] [] 1 *}) @@ -565,7 +565,7 @@ lemma Step1_2_1: "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S1 rmhist p --> (S2 rmhist p)$ & ENext p & unchanged (c p, r p, m p)" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ImpNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ImpNext_def}]) [] (map temp_elim [@{thm S1ClerkUnch}, @{thm S1RPCUnch}, @{thm S1MemUnch}, @{thm S1Hist}]) 1 *}) using [[fast_solver]] apply (auto elim!: squareE [temp_use] intro!: S1Env [temp_use]) @@ -575,7 +575,7 @@ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S2 rmhist p --> (S3 rmhist p)$ & MClkFwd memCh crCh cst p & unchanged (e p, r p, m p, rmhist!p)" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ImpNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ImpNext_def}]) [] (map temp_elim [@{thm S2EnvUnch}, @{thm S2RPCUnch}, @{thm S2MemUnch}, @{thm S2Hist}]) 1 *}) using [[fast_solver]] apply (auto elim!: squareE [temp_use] intro!: S2Clerk [temp_use] S2Forward [temp_use]) @@ -585,9 +585,9 @@ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S3 rmhist p --> ((S4 rmhist p)$ & RPCFwd crCh rmCh rst p & unchanged (e p, c p, m p, rmhist!p)) | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ImpNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ImpNext_def}]) [] (map temp_elim [@{thm S3EnvUnch}, @{thm S3ClerkUnch}, @{thm S3MemUnch}]) 1 *}) - apply (tactic {* action_simp_tac @{simpset} [] + apply (tactic {* action_simp_tac @{context} [] (@{thm squareE} :: map temp_elim [@{thm S3RPC}, @{thm S3Forward}, @{thm S3Fail}]) 1 *}) apply (auto dest!: S3Hist [temp_use]) done @@ -598,9 +598,9 @@ --> ((S4 rmhist p)$ & Read rmCh mm ires p & unchanged (e p, c p, r p, rmhist!p)) | ((S4 rmhist p)$ & (? l. Write rmCh mm ires p l) & unchanged (e p, c p, r p, rmhist!p)) | ((S5 rmhist p)$ & MemReturn rmCh ires p & unchanged (e p, c p, r p))" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ImpNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ImpNext_def}]) [] (map temp_elim [@{thm S4EnvUnch}, @{thm S4ClerkUnch}, @{thm S4RPCUnch}]) 1 *}) - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm RNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm RNext_def}]) [] (@{thm squareE} :: map temp_elim [@{thm S4Read}, @{thm S4Write}, @{thm S4Return}]) 1 *}) apply (auto dest!: S4Hist [temp_use]) done @@ -609,9 +609,9 @@ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S5 rmhist p --> ((S6 rmhist p)$ & RPCReply crCh rmCh rst p & unchanged (e p, c p, m p)) | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ImpNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ImpNext_def}]) [] (map temp_elim [@{thm S5EnvUnch}, @{thm S5ClerkUnch}, @{thm S5MemUnch}, @{thm S5Hist}]) 1 *}) - apply (tactic {* action_simp_tac @{simpset} [] [@{thm squareE}, temp_elim @{thm S5RPC}] 1 *}) + apply (tactic {* action_simp_tac @{context} [] [@{thm squareE}, temp_elim @{thm S5RPC}] 1 *}) using [[fast_solver]] apply (auto elim!: squareE [temp_use] dest!: S5Reply [temp_use] S5Fail [temp_use]) done @@ -620,9 +620,9 @@ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S6 rmhist p --> ((S1 rmhist p)$ & MClkReply memCh crCh cst p & unchanged (e p, r p, m p)) | ((S3 rmhist p)$ & MClkRetry memCh crCh cst p & unchanged (e p,r p,m p,rmhist!p))" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ImpNext_def}]) [] + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ImpNext_def}]) [] (map temp_elim [@{thm S6EnvUnch}, @{thm S6RPCUnch}, @{thm S6MemUnch}]) 1 *}) - apply (tactic {* action_simp_tac @{simpset} [] + apply (tactic {* action_simp_tac @{context} [] (@{thm squareE} :: map temp_elim [@{thm S6Clerk}, @{thm S6Retry}, @{thm S6Reply}]) 1 *}) apply (auto dest: S6Hist [temp_use]) done @@ -634,7 +634,7 @@ section "Initialization (Step 1.3)" lemma Step1_3: "|- S1 rmhist p --> PInit (resbar rmhist) p" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm resbar_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm resbar_def}, @{thm PInit_def}, @{thm S_def}, @{thm S1_def}]) [] [] 1 *}) (* ---------------------------------------------------------------------- @@ -653,7 +653,7 @@ & unchanged (e p, r p, m p, rmhist!p) --> unchanged (rtrner memCh!p, resbar rmhist!p)" by (tactic {* action_simp_tac - (@{simpset} addsimps [@{thm MClkFwd_def}, @{thm e_def}, @{thm r_def}, @{thm m_def}, + (@{context} addsimps [@{thm MClkFwd_def}, @{thm e_def}, @{thm r_def}, @{thm m_def}, @{thm resbar_def}, @{thm S_def}, @{thm S2_def}, @{thm S3_def}]) [] [] 1 *}) lemma Step1_4_3a: "|- RPCFwd crCh rmCh rst p & $S3 rmhist p & (S4 rmhist p)$ @@ -661,7 +661,7 @@ --> unchanged (rtrner memCh!p, resbar rmhist!p)" apply clarsimp apply (drule S3_excl [temp_use] S4_excl [temp_use])+ - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm e_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm e_def}, @{thm c_def}, @{thm m_def}, @{thm resbar_def}, @{thm S_def}, @{thm S3_def}]) [] [] 1 *}) done @@ -680,11 +680,11 @@ --> ReadInner memCh mm (resbar rmhist) p l" apply clarsimp apply (drule S4_excl [temp_use])+ - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm ReadInner_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm ReadInner_def}, @{thm GoodRead_def}, @{thm BadRead_def}, @{thm e_def}, @{thm c_def}, @{thm m_def}]) [] [] 1 *}) apply (auto simp: resbar_def) apply (tactic {* ALLGOALS (action_simp_tac - (@{simpset} addsimps [@{thm RPCRelayArg_def}, @{thm MClkRelayArg_def}, + (@{context} addsimps [@{thm RPCRelayArg_def}, @{thm MClkRelayArg_def}, @{thm S_def}, @{thm S4_def}, @{thm RdRequest_def}, @{thm MemInv_def}]) [] [@{thm impE}, @{thm MemValNotAResultE}]) *}) done @@ -699,11 +699,11 @@ --> WriteInner memCh mm (resbar rmhist) p l v" apply clarsimp apply (drule S4_excl [temp_use])+ - apply (tactic {* action_simp_tac (@{simpset} addsimps + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm WriteInner_def}, @{thm GoodWrite_def}, @{thm BadWrite_def}, @{thm e_def}, @{thm c_def}, @{thm m_def}]) [] [] 1 *}) apply (auto simp: resbar_def) - apply (tactic {* ALLGOALS (action_simp_tac (@{simpset} addsimps + apply (tactic {* ALLGOALS (action_simp_tac (@{context} addsimps [@{thm RPCRelayArg_def}, @{thm MClkRelayArg_def}, @{thm S_def}, @{thm S4_def}, @{thm WrRequest_def}]) [] []) *}) done @@ -716,7 +716,7 @@ lemma Step1_4_4c: "|- MemReturn rmCh ires p & $S4 rmhist p & (S5 rmhist p)$ & unchanged (e p, c p, r p) --> unchanged (rtrner memCh!p, resbar rmhist!p)" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm e_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm e_def}, @{thm c_def}, @{thm r_def}, @{thm resbar_def}]) [] [] 1 *}) apply (drule S4_excl [temp_use] S5_excl [temp_use])+ using [[fast_solver]] @@ -746,11 +746,11 @@ --> MemReturn memCh (resbar rmhist) p" apply clarsimp apply (drule S6_excl [temp_use])+ - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm e_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm e_def}, @{thm r_def}, @{thm m_def}, @{thm MClkReply_def}, @{thm MemReturn_def}, @{thm Return_def}, @{thm resbar_def}]) [] [] 1 *}) apply simp_all (* simplify if-then-else *) - apply (tactic {* ALLGOALS (action_simp_tac (@{simpset} addsimps + apply (tactic {* ALLGOALS (action_simp_tac (@{context} addsimps [@{thm MClkReplyVal_def}, @{thm S6_def}, @{thm S_def}]) [] [@{thm MVOKBARFnotNR}]) *}) done @@ -759,7 +759,7 @@ --> MemFail memCh (resbar rmhist) p" apply clarsimp apply (drule S3_excl [temp_use])+ - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm e_def}, @{thm r_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm e_def}, @{thm r_def}, @{thm m_def}, @{thm MClkRetry_def}, @{thm MemFail_def}, @{thm resbar_def}]) [] [] 1 *}) apply (auto simp: S6_def S_def) done @@ -797,7 +797,7 @@ Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" 1 THEN rewrite_goals_tac @{thms action_rews} THEN forward_tac [temp_use @{thm Step1_4_7}] 1 THEN - asm_full_simp_tac (simpset_of ctxt) 1); + asm_full_simp_tac ctxt 1); *} method_setup split_idle = {* @@ -897,14 +897,14 @@ lemma S1_RNextdisabled: "|- S1 rmhist p --> ~Enabled (_(rtrner memCh!p, resbar rmhist!p))" - apply (tactic {* action_simp_tac (@{simpset} addsimps [@{thm angle_def}, + apply (tactic {* action_simp_tac (@{context} addsimps [@{thm angle_def}, @{thm S_def}, @{thm S1_def}]) [notI] [@{thm enabledE}, temp_elim @{thm Memoryidle}] 1 *}) apply force done lemma S1_Returndisabled: "|- S1 rmhist p --> ~Enabled (_(rtrner memCh!p, resbar rmhist!p))" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm angle_def}, @{thm MemReturn_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm angle_def}, @{thm MemReturn_def}, @{thm Return_def}, @{thm S_def}, @{thm S1_def}]) [notI] [@{thm enabledE}] 1 *}) lemma RNext_fair: "|- []<>S1 rmhist p @@ -1083,7 +1083,7 @@ lemma MClkReplyS6: "|- $ImpInv rmhist p & _(c p) --> $S6 rmhist p" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm angle_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm angle_def}, @{thm MClkReply_def}, @{thm Return_def}, @{thm ImpInv_def}, @{thm S_def}, @{thm S1_def}, @{thm S2_def}, @{thm S3_def}, @{thm S4_def}, @{thm S5_def}]) [] [] 1 *}) @@ -1091,7 +1091,7 @@ apply (auto simp: c_def intro!: MClkReply_enabled [temp_use]) apply (cut_tac MI_base) apply (blast dest: base_pair) - apply (tactic {* ALLGOALS (action_simp_tac (@{simpset} + apply (tactic {* ALLGOALS (action_simp_tac (@{context} addsimps [@{thm S_def}, @{thm S6_def}]) [] []) *}) done @@ -1102,7 +1102,7 @@ apply (subgoal_tac "sigma |= []<> (_ (c p))") apply (erule InfiniteEnsures) apply assumption - apply (tactic {* action_simp_tac @{simpset} [] + apply (tactic {* action_simp_tac @{context} [] (map temp_elim [@{thm MClkReplyS6}, @{thm S6MClkReply_successors}]) 1 *}) apply (auto simp: SF_def) apply (erule contrapos_np) @@ -1189,7 +1189,7 @@ sigma |= []<>S6 rmhist p --> []<>S1 rmhist p |] ==> sigma |= []<>S1 rmhist p" apply (rule classical) - apply (tactic {* asm_lr_simp_tac (@{simpset} addsimps + apply (tactic {* asm_lr_simp_tac (@{context} addsimps [temp_use @{thm NotBox}, temp_rewrite @{thm NotDmd}]) 1 *}) apply (auto elim!: leadsto_infinite [temp_use] mp dest!: DBImplBD [temp_use]) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/Memory/RPC.thy --- a/src/HOL/TLA/Memory/RPC.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/Memory/RPC.thy Thu Apr 18 17:07:01 2013 +0200 @@ -99,14 +99,14 @@ (* Enabledness of some actions *) lemma RPCFail_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm RPCFail_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm RPCFail_def}, @{thm Return_def}, @{thm caller_def}, @{thm rtrner_def}]) [exI] [@{thm base_enabled}, @{thm Pair_inject}] 1 *}) lemma RPCReply_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> |- ~Calling rcv p & Calling send p & rst!p = #rpcB --> Enabled (RPCReply send rcv rst p)" - by (tactic {* action_simp_tac (@{simpset} addsimps [@{thm RPCReply_def}, + by (tactic {* action_simp_tac (@{context} addsimps [@{thm RPCReply_def}, @{thm Return_def}, @{thm caller_def}, @{thm rtrner_def}]) [exI] [@{thm base_enabled}, @{thm Pair_inject}] 1 *}) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TLA/TLA.thy --- a/src/HOL/TLA/TLA.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TLA/TLA.thy Thu Apr 18 17:07:01 2013 +0200 @@ -597,7 +597,7 @@ SELECT_GOAL (inv_tac ctxt 1 THEN (TRYALL (action_simp_tac - (simpset_of ctxt addsimps [@{thm Init_stp}, @{thm Init_act}]) [] [@{thm squareE}]))); + (ctxt addsimps [@{thm Init_stp}, @{thm Init_act}]) [] [@{thm squareE}]))); *} method_setup invariant = {* diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/TPTP/atp_problem_import.ML --- a/src/HOL/TPTP/atp_problem_import.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/TPTP/atp_problem_import.ML Thu Apr 18 17:07:01 2013 +0200 @@ -234,7 +234,7 @@ SOLVE_TIMEOUT (timeout div 20) "nitpick" (nitpick_finite_oracle_tac ctxt (timeout div 20) i) ORELSE SOLVE_TIMEOUT (timeout div 10) "simp" - (asm_full_simp_tac (simpset_of ctxt) i) + (asm_full_simp_tac ctxt i) ORELSE SOLVE_TIMEOUT (timeout div 10) "blast" (blast_tac ctxt i) ORELSE SOLVE_TIMEOUT (timeout div 5) "auto+spass" (auto_tac ctxt diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/ATP/atp_problem_generate.ML --- a/src/HOL/Tools/ATP/atp_problem_generate.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/ATP/atp_problem_generate.ML Thu Apr 18 17:07:01 2013 +0200 @@ -1230,10 +1230,10 @@ | _ => do_term bs t in do_formula [] end -fun presimplify_term thy t = +fun presimplify_term ctxt t = if exists_Const (member (op =) Meson.presimplified_consts o fst) t then - t |> Skip_Proof.make_thm thy - |> Meson.presimplify + t |> Skip_Proof.make_thm (Proof_Context.theory_of ctxt) + |> Meson.presimplify ctxt |> prop_of else t @@ -1273,7 +1273,7 @@ t |> need_trueprop ? HOLogic.mk_Trueprop |> (if is_ho then unextensionalize_def else cong_extensionalize_term thy #> abs_extensionalize_term ctxt) - |> presimplify_term thy + |> presimplify_term ctxt |> HOLogic.dest_Trueprop end handle TERM _ => @{const True} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Datatype/datatype.ML --- a/src/HOL/Tools/Datatype/datatype.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Datatype/datatype.ML Thu Apr 18 17:07:01 2013 +0200 @@ -537,8 +537,8 @@ fun prove [] = [] | prove (t :: ts) = let - val dist_thm = Goal.prove_sorry_global thy5 [] [] t (fn _ => - EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1]) + val dist_thm = Goal.prove_sorry_global thy5 [] [] t (fn {context = ctxt, ...} => + EVERY [simp_tac (put_simpset HOL_ss ctxt addsimps dist_rewrites') 1]) in dist_thm :: Drule.zero_var_indexes (dist_thm RS not_sym) :: prove ts end; in prove end; @@ -632,13 +632,14 @@ Goal.prove_sorry_global thy6 [] (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) - (fn {prems, ...} => + (fn {context = ctxt, prems, ...} => EVERY [rtac indrule_lemma' 1, (Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1, EVERY (map (fn (prem, r) => (EVERY [REPEAT (eresolve_tac Abs_inverse_thms 1), - simp_tac (HOL_basic_ss addsimps (Thm.symmetric r :: Rep_inverse_thms')) 1, + simp_tac (put_simpset HOL_basic_ss ctxt + addsimps (Thm.symmetric r :: Rep_inverse_thms')) 1, DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) (prems ~~ (constr_defs @ map mk_meta_eq iso_char_thms)))]); diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Datatype/datatype_codegen.ML --- a/src/HOL/Tools/Datatype/datatype_codegen.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Apr 18 17:07:01 2013 +0200 @@ -81,12 +81,11 @@ [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)]; val distincts = maps prep_distinct (nth (Datatype_Prop.make_distincts [descr]) index); val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty))); - val simpset = - Simplifier.global_context thy - (HOL_basic_ss addsimps - (map Simpdata.mk_eq (@{thms equal eq_True} @ inject_thms @ distinct_thms))); + val simp_ctxt = + Simplifier.global_context thy HOL_basic_ss + addsimps (map Simpdata.mk_eq (@{thms equal eq_True} @ inject_thms @ distinct_thms)); fun prove prop = - Goal.prove_sorry_global thy [] [] prop (K (ALLGOALS (simp_tac simpset))) + Goal.prove_sorry_global thy [] [] prop (K (ALLGOALS (simp_tac simp_ctxt))) |> Simpdata.mk_eq; in (map prove (triv_injects @ injects @ distincts), prove refl) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Datatype/datatype_realizer.ML --- a/src/HOL/Tools/Datatype/datatype_realizer.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Thu Apr 18 17:07:01 2013 +0200 @@ -192,8 +192,8 @@ EVERY [ rtac (cterm_instantiate [(cert y, cert y')] exhaust) 1, ALLGOALS (EVERY' - [asm_simp_tac (HOL_basic_ss addsimps case_rewrites), - resolve_tac prems, asm_simp_tac HOL_basic_ss])]) + [asm_simp_tac (Simplifier.global_context thy HOL_basic_ss addsimps case_rewrites), + resolve_tac prems, asm_simp_tac (Simplifier.global_context thy HOL_basic_ss)])]) |> Drule.export_without_context; val exh_name = Thm.derivation_name exhaust; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Datatype/rep_datatype.ML --- a/src/HOL/Tools/Datatype/rep_datatype.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Datatype/rep_datatype.ML Thu Apr 18 17:07:01 2013 +0200 @@ -156,12 +156,12 @@ val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ..."; - fun mk_unique_tac ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) = + fun mk_unique_tac ctxt ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) = let val distinct_tac = if i < length newTs then - full_simp_tac (HOL_ss addsimps (nth dist_rewrites i)) 1 - else full_simp_tac (HOL_ss addsimps (flat other_dist_rewrites)) 1; + full_simp_tac (put_simpset HOL_ss ctxt addsimps (nth dist_rewrites i)) 1 + else full_simp_tac (put_simpset HOL_ss ctxt addsimps (flat other_dist_rewrites)) 1; val inject = map (fn r => r RS iffD1) @@ -203,13 +203,13 @@ map (fn ((i, T), t) => absfree ("x" ^ string_of_int i, T) t) ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts); val induct' = cterm_instantiate (map cert induct_Ps ~~ map cert insts) induct; - val (tac, _) = - fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts) - (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1 THEN - rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs)); in Datatype_Aux.split_conj_thm (Goal.prove_sorry_global thy1 [] [] - (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts)) (K tac)) + (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts)) + (fn {context = ctxt, ...} => + #1 (fold (mk_unique_tac ctxt) (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts) + (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1 THEN + rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs))))) end; val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms; @@ -359,12 +359,13 @@ val cert = cterm_of thy; val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion))); val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion; - val tac = + fun tac ctxt = EVERY [rtac exhaustion' 1, - ALLGOALS (asm_simp_tac (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))]; + ALLGOALS (asm_simp_tac + (put_simpset HOL_ss ctxt addsimps (dist_rewrites' @ inject @ case_thms')))]; in - (Goal.prove_sorry_global thy [] [] t1 (K tac), - Goal.prove_sorry_global thy [] [] t2 (K tac)) + (Goal.prove_sorry_global thy [] [] t1 (tac o #context), + Goal.prove_sorry_global thy [] [] t2 (tac o #context)) end; val split_thm_pairs = @@ -429,10 +430,12 @@ val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'; in Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) - (fn {prems, ...} => - let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites)) in + (fn {context = ctxt, prems, ...} => + let + val simplify = asm_simp_tac (put_simpset HOL_ss ctxt addsimps (prems @ case_rewrites)) + in EVERY [ - simp_tac (HOL_ss addsimps [hd prems]) 1, + simp_tac (put_simpset HOL_ss ctxt addsimps [hd prems]) 1, cut_tac nchotomy'' 1, REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1), REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)] diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/context_tree.ML --- a/src/HOL/Tools/Function/context_tree.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/context_tree.ML Thu Apr 18 17:07:01 2013 +0200 @@ -33,7 +33,7 @@ (ctxt * thm) list * 'b) -> ctx_tree -> 'b -> 'b - val rewrite_by_tree : theory -> term -> thm -> (thm * thm) list -> + val rewrite_by_tree : Proof.context -> term -> thm -> (thm * thm) list -> ctx_tree -> thm * (thm * thm) list end @@ -240,8 +240,9 @@ snd o traverse_help ([], []) tr [] end -fun rewrite_by_tree thy h ih x tr = +fun rewrite_by_tree ctxt h ih x tr = let + val thy = Proof_Context.theory_of ctxt fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (cterm_of thy t), x) | rewrite_help fix h_as x (RCall (_ $ arg, st)) = let @@ -268,7 +269,7 @@ |> map (fn u_eq => (u_eq RS sym) RS eq_reflection) |> filter_out Thm.is_reflexive - val assumes' = map (simplify (HOL_basic_ss addsimps used)) assumes + val assumes' = map (simplify (put_simpset HOL_basic_ss ctxt addsimps used)) assumes val (subeq, x') = rewrite_help (fix @ fixes) (h_as @ assumes') x st diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/function.ML Thu Apr 18 17:07:01 2013 +0200 @@ -187,7 +187,8 @@ let val totality = Thm.close_derivation totality val remove_domain_condition = - full_simplify (HOL_basic_ss addsimps [totality, @{thm True_implies_equals}]) + full_simplify (put_simpset HOL_basic_ss lthy + addsimps [totality, @{thm True_implies_equals}]) val tsimps = map remove_domain_condition psimps val tinduct = map remove_domain_condition pinducts diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/function_core.ML --- a/src/HOL/Tools/Function/function_core.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/function_core.ML Thu Apr 18 17:07:01 2013 +0200 @@ -260,8 +260,9 @@ end (* Generates the replacement lemma in fully quantified form. *) -fun mk_replacement_lemma thy h ih_elim clause = +fun mk_replacement_lemma ctxt h ih_elim clause = let + val thy = Proof_Context.theory_of ctxt val ClauseInfo {cdata=ClauseContext {qs, lhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause local open Conv in @@ -276,7 +277,7 @@ Thm.assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs val (eql, _) = - Function_Ctx_Tree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree + Function_Ctx_Tree.rewrite_by_tree ctxt h ih_elim_case (Ris ~~ h_assums) tree val replace_lemma = (eql RS meta_eq_to_obj_eq) |> Thm.implies_intr (cprop_of case_hyp) @@ -328,13 +329,14 @@ -fun mk_uniqueness_case thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei = +fun mk_uniqueness_case ctxt globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei = let + val thy = Proof_Context.theory_of ctxt val Globals {x, y, ranT, fvar, ...} = globals val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs - val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro + val ih_intro_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ih_intro fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case) |> fold_rev (Thm.implies_intr o cprop_of) CCas @@ -366,7 +368,7 @@ ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)] |> curry (op COMP) existence |> curry (op COMP) uniqueness - |> simplify (HOL_basic_ss addsimps [case_hyp RS sym]) + |> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym]) |> Thm.implies_intr (cprop_of case_hyp) |> fold_rev (Thm.implies_intr o cprop_of) ags |> fold_rev Thm.forall_intr cqs @@ -401,11 +403,14 @@ |> instantiate' [] [NONE, SOME (cterm_of thy h)] val _ = trace_msg (K "Proving Replacement lemmas...") - val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses + val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses val _ = trace_msg (K "Proving cases for unique existence...") val (ex1s, values) = - split_list (map (mk_uniqueness_case thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses) + split_list + (map + (mk_uniqueness_case ctxt globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) + clauses) val _ = trace_msg (K "Proving: Graph is a function") val graph_is_function = complete @@ -551,8 +556,9 @@ * PROVING THE RULES **********************************************************) -fun mk_psimps thy globals R clauses valthms f_iff graph_is_function = +fun mk_psimps ctxt globals R clauses valthms f_iff graph_is_function = let + val thy = Proof_Context.theory_of ctxt val Globals {domT, z, ...} = globals fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm = @@ -566,7 +572,7 @@ |> Thm.forall_intr (cterm_of thy z) |> (fn it => it COMP valthm) |> Thm.implies_intr lhs_acc - |> asm_simplify (HOL_basic_ss addsimps [f_iff]) + |> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff]) |> fold_rev (Thm.implies_intr o cprop_of) ags |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) end @@ -714,13 +720,14 @@ val wf_in_rel = @{thm FunDef.wf_in_rel} val in_rel_def = @{thm FunDef.in_rel_def} -fun mk_nest_term_case thy globals R' ihyp clause = +fun mk_nest_term_case ctxt globals R' ihyp clause = let + val thy = Proof_Context.theory_of ctxt val Globals {z, ...} = globals val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree, qglr=(oqs, _, _, _), ...} = clause - val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp + val ih_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ihyp fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) = let @@ -763,8 +770,9 @@ end -fun mk_nest_term_rule thy globals R R_cases clauses = +fun mk_nest_term_rule ctxt globals R R_cases clauses = let + val thy = Proof_Context.theory_of ctxt val Globals { domT, x, z, ... } = globals val acc_R = mk_acc domT R @@ -788,7 +796,7 @@ val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x)) - val (hyps, cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([], []) + val (hyps, cases) = fold (mk_nest_term_case ctxt globals R' ihyp_a) clauses ([], []) in R_cases |> Thm.forall_elim (cterm_of thy z) @@ -810,7 +818,7 @@ |> Thm.forall_intr (cterm_of thy R') |> Thm.forall_elim (cterm_of thy (inrel_R)) |> curry op RS wf_in_rel - |> full_simplify (HOL_basic_ss addsimps [in_rel_def]) + |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps [in_rel_def]) |> Thm.forall_intr (cterm_of thy Rrel) end @@ -882,6 +890,7 @@ fun mk_partial_rules provedgoal = let val newthy = theory_of_thm provedgoal (*FIXME*) + val newctxt = Proof_Context.init_global newthy (*FIXME*) val (graph_is_function, complete_thm) = provedgoal @@ -891,13 +900,13 @@ val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff) val psimps = PROFILE "Proving simplification rules" - (mk_psimps newthy globals R xclauses values f_iff) graph_is_function + (mk_psimps newctxt globals R xclauses values f_iff) graph_is_function val simple_pinduct = PROFILE "Proving partial induction rule" (mk_partial_induct_rule newthy globals R complete_thm) xclauses val total_intro = PROFILE "Proving nested termination rule" - (mk_nest_term_rule newthy globals R R_elim) xclauses + (mk_nest_term_rule newctxt globals R R_elim) xclauses val dom_intros = if domintros then SOME (PROFILE "Proving domain introduction rules" diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/induction_schema.ML --- a/src/HOL/Tools/Function/induction_schema.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/induction_schema.ML Thu Apr 18 17:07:01 2013 +0200 @@ -231,7 +231,7 @@ val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches val pats = map_index (uncurry inject) xss val sum_split_rule = - Pat_Completeness.prove_completeness thy [x] (P_comp $ x) xss (map single pats) + Pat_Completeness.prove_completeness ctxt [x] (P_comp $ x) xss (map single pats) fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) = let @@ -253,8 +253,9 @@ val cqs = map (cert o Free) qs val ags = map (Thm.assume o cert) gs - val replace_x_ss = HOL_basic_ss addsimps (branch_hyp :: case_hyps) - val sih = full_simplify replace_x_ss aihyp + val replace_x_simpset = + put_simpset HOL_basic_ss ctxt addsimps (branch_hyp :: case_hyps) + val sih = full_simplify replace_x_simpset aihyp fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) = let @@ -373,7 +374,7 @@ in indthm |> Drule.instantiate' [] [SOME inst] - |> simplify SumTree.sumcase_split_ss + |> simplify (put_simpset SumTree.sumcase_split_ss ctxt'') |> Conv.fconv_rule ind_rulify end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Thu Apr 18 17:07:01 2013 +0200 @@ -212,7 +212,7 @@ fun lexicographic_order_tac quiet ctxt = TRY (Function_Common.apply_termination_rule ctxt 1) THEN lex_order_tac quiet ctxt - (auto_tac (map_simpset (fn ss => ss addsimps Function_Common.Termination_Simps.get ctxt) ctxt)) + (auto_tac (ctxt addsimps Function_Common.Termination_Simps.get ctxt)) val setup = Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false)) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/mutual.ML Thu Apr 18 17:07:01 2013 +0200 @@ -193,7 +193,7 @@ (fn _ => Local_Defs.unfold_tac ctxt all_orig_fdefs THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1 - THEN (simp_tac (simpset_of ctxt)) 1) + THEN (simp_tac ctxt) 1) |> restore_cond |> export end @@ -209,9 +209,9 @@ |> Thm.forall_elim_vars 0 end -fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, parts, ...}) = +fun mutual_induct_rules ctxt induct all_f_defs (Mutual {n, ST, parts, ...}) = let - val cert = cterm_of (Proof_Context.theory_of lthy) + val cert = cterm_of (Proof_Context.theory_of ctxt) val newPs = map2 (fn Pname => fn MutualPart {cargTs, ...} => Free (Pname, cargTs ---> HOLogic.boolT)) @@ -230,8 +230,8 @@ val induct_inst = Thm.forall_elim (cert case_exp) induct - |> full_simplify SumTree.sumcase_split_ss - |> full_simplify (HOL_basic_ss addsimps all_f_defs) + |> full_simplify (put_simpset SumTree.sumcase_split_ss ctxt) + |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps all_f_defs) fun project rule (MutualPart {cargTs, i, ...}) k = let @@ -240,7 +240,7 @@ in (rule |> Thm.forall_elim (cert inj) - |> full_simplify SumTree.sumcase_split_ss + |> full_simplify (put_simpset SumTree.sumcase_split_ss ctxt) |> fold_rev (Thm.forall_intr o cert) (afs @ newPs), k + length cargTs) end @@ -266,11 +266,11 @@ fun mk_mpsimp fqgar sum_psimp = in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp - val rew_ss = HOL_basic_ss addsimps all_f_defs + val rew_simpset = put_simpset HOL_basic_ss lthy addsimps all_f_defs val mpsimps = map2 mk_mpsimp fqgars psimps val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m - val mtermination = full_simplify rew_ss termination - val mdomintros = Option.map (map (full_simplify rew_ss)) domintros + val mtermination = full_simplify rew_simpset termination + val mdomintros = Option.map (map (full_simplify rew_simpset)) domintros in FunctionResult { fs=fs, G=G, R=R, psimps=mpsimps, simple_pinducts=minducts, diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/partial_function.ML --- a/src/HOL/Tools/Function/partial_function.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/partial_function.ML Thu Apr 18 17:07:01 2013 +0200 @@ -157,11 +157,13 @@ fun curry_n arity = funpow (arity - 1) mk_curry; fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split; -val curry_uncurry_ss = HOL_basic_ss addsimps - [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}] +val curry_uncurry_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}]) -val split_conv_ss = HOL_basic_ss addsimps - [@{thm Product_Type.split_conv}]; +val split_conv_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm Product_Type.split_conv}]); fun mk_curried_induct args ctxt ccurry cuncurry rule = let @@ -187,12 +189,12 @@ val inst_rule' = inst_rule |> Tactic.rule_by_tactic ctxt - (Simplifier.simp_tac curry_uncurry_ss 4 - THEN Simplifier.simp_tac curry_uncurry_ss 3 + (Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 4 + THEN Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3 THEN CONVERSION (split_params_conv ctxt then_conv (Conv.forall_conv (K split_paired_all_conv) ctxt)) 3) |> Drule.instantiate' [] [NONE, NONE, SOME P_inst, SOME x_inst] - |> Simplifier.full_simplify split_conv_ss + |> Simplifier.full_simplify (put_simpset split_conv_ss ctxt) |> singleton (Variable.export ctxt' ctxt) in inst_rule' @@ -253,7 +255,7 @@ val unfold = (cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq OF [mono_thm, f_def]) - |> Tactic.rule_by_tactic lthy (Simplifier.simp_tac curry_uncurry_ss 1); + |> Tactic.rule_by_tactic lthy (Simplifier.simp_tac (put_simpset curry_uncurry_ss lthy) 1); val mk_raw_induct = mk_curried_induct args args_ctxt (cert curry) (cert uncurry) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/pat_completeness.ML --- a/src/HOL/Tools/Function/pat_completeness.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/pat_completeness.ML Thu Apr 18 17:07:01 2013 +0200 @@ -7,7 +7,7 @@ signature PAT_COMPLETENESS = sig val pat_completeness_tac: Proof.context -> int -> tactic - val prove_completeness : theory -> term list -> term -> term list list -> + val prove_completeness : Proof.context -> term list -> term -> term list list -> term list list -> thm end @@ -61,12 +61,13 @@ | inst_constrs_of thy _ = raise Match -fun transform_pat thy avars c_assum ([] , thm) = raise Match - | transform_pat thy avars c_assum (pat :: pats, thm) = +fun transform_pat _ avars c_assum ([] , thm) = raise Match + | transform_pat ctxt avars c_assum (pat :: pats, thm) = let + val thy = Proof_Context.theory_of ctxt val (_, subps) = strip_comb pat val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps) - val c_eq_pat = simplify (HOL_basic_ss addsimps (map Thm.assume eqs)) c_assum + val c_eq_pat = simplify (put_simpset HOL_basic_ss ctxt addsimps (map Thm.assume eqs)) c_assum in (subps @ pats, fold_rev Thm.implies_intr eqs (Thm.implies_elim thm c_eq_pat)) @@ -75,40 +76,45 @@ exception COMPLETENESS -fun constr_case thy P idx (v :: vs) pats cons = +fun constr_case ctxt P idx (v :: vs) pats cons = let + val thy = Proof_Context.theory_of ctxt val (avars, pvars, newidx) = invent_vars cons idx val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars)))) val c_assum = Thm.assume c_hyp - val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats) + val newpats = map (transform_pat ctxt avars c_assum) (filter_pats thy cons pvars pats) in - o_alg thy P newidx (avars @ vs) newpats + o_alg ctxt P newidx (avars @ vs) newpats |> Thm.implies_intr c_hyp |> fold_rev (Thm.forall_intr o cterm_of thy) avars end | constr_case _ _ _ _ _ _ = raise Match -and o_alg thy P idx [] (([], Pthm) :: _) = Pthm - | o_alg thy P idx (v :: vs) [] = raise COMPLETENESS - | o_alg thy P idx (v :: vs) pts = +and o_alg _ P idx [] (([], Pthm) :: _) = Pthm + | o_alg _ P idx (v :: vs) [] = raise COMPLETENESS + | o_alg ctxt P idx (v :: vs) pts = if forall (is_Free o hd o fst) pts (* Var case *) - then o_alg thy P idx vs + then o_alg ctxt P idx vs (map (fn (pv :: pats, thm) => - (pats, refl RS (inst_free (cterm_of thy pv) (cterm_of thy v) thm))) pts) + (pats, refl RS + (inst_free (cterm_of (Proof_Context.theory_of ctxt) pv) + (cterm_of (Proof_Context.theory_of ctxt) v) thm))) pts) else (* Cons case *) let + val thy = Proof_Context.theory_of ctxt val T = fastype_of v val (tname, _) = dest_Type T val {exhaust=case_thm, ...} = Datatype.the_info thy tname val constrs = inst_constrs_of thy T - val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs + val c_cases = map (constr_case ctxt P idx (v :: vs) pts) constrs in inst_case_thm thy v P case_thm |> fold (curry op COMP) c_cases end | o_alg _ _ _ _ _ = raise Match -fun prove_completeness thy xs P qss patss = +fun prove_completeness ctxt xs P qss patss = let + val thy = Proof_Context.theory_of ctxt fun mk_assum qs pats = HOLogic.mk_Trueprop P |> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats) @@ -119,7 +125,7 @@ fun inst_hyps hyp qs = fold (Thm.forall_elim o cterm_of thy) qs (Thm.assume hyp) val assums = map2 inst_hyps hyps qss in - o_alg thy P 2 xs (patss ~~ assums) + o_alg ctxt P 2 xs (patss ~~ assums) |> fold_rev Thm.implies_intr hyps end @@ -143,7 +149,7 @@ handle List.Empty => raise COMPLETENESS val patss = map (map snd) x_pats - val complete_thm = prove_completeness thy xs thesis qss patss + val complete_thm = prove_completeness ctxt xs thesis qss patss |> fold_rev (Thm.forall_intr o cterm_of thy) vs in PRIMITIVE (fn st => Drule.compose_single(complete_thm, i, st)) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/scnp_reconstruct.ML --- a/src/HOL/Tools/Function/scnp_reconstruct.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Thu Apr 18 17:07:01 2013 +0200 @@ -289,7 +289,7 @@ THEN (rtac @{thm rp_inv_image_rp} 1) THEN (rtac (order_rpair ms_rp label) 1) THEN PRIMITIVE (instantiate' [] [SOME level_mapping]) - THEN unfold_tac @{thms rp_inv_image_def} (simpset_of ctxt) + THEN unfold_tac @{thms rp_inv_image_def} ctxt THEN Local_Defs.unfold_tac ctxt (@{thms split_conv} @ @{thms fst_conv} @ @{thms snd_conv}) THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}])) @@ -338,7 +338,7 @@ fun decomp_scnp_tac orders ctxt = let val extra_simps = Function_Common.Termination_Simps.get ctxt - val autom_tac = auto_tac (map_simpset (fn ss => ss addsimps extra_simps) ctxt) + val autom_tac = auto_tac (ctxt addsimps extra_simps) in gen_sizechange_tac orders autom_tac ctxt end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/size.ML --- a/src/HOL/Tools/Function/size.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/size.ML Thu Apr 18 17:07:01 2013 +0200 @@ -149,8 +149,9 @@ val ctxt = Proof_Context.init_global thy'; - val simpset1 = HOL_basic_ss addsimps @{thm Nat.add_0} :: @{thm Nat.add_0_right} :: - size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites; + val simpset1 = + put_simpset HOL_basic_ss ctxt addsimps @{thm Nat.add_0} :: @{thm Nat.add_0_right} :: + size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites; val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2); fun mk_unfolded_size_eq tab size_ofp fs (p as (x, T), r) = @@ -186,10 +187,12 @@ else foldl1 plus (ts @ [HOLogic.Suc_zero]))) end; - val simpset2 = HOL_basic_ss addsimps - rec_rewrites @ size_def_thms @ unfolded_size_eqs1; - val simpset3 = HOL_basic_ss addsimps - rec_rewrites @ size_def_thms' @ unfolded_size_eqs2; + val simpset2 = + put_simpset HOL_basic_ss ctxt + addsimps (rec_rewrites @ size_def_thms @ unfolded_size_eqs1); + val simpset3 = + put_simpset HOL_basic_ss ctxt + addsimps (rec_rewrites @ size_def_thms' @ unfolded_size_eqs2); fun prove_size_eqs p size_fns size_ofp simpset = maps (fn (((_, (_, _, constrs)), size_const), T) => diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Function/sum_tree.ML --- a/src/HOL/Tools/Function/sum_tree.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Function/sum_tree.ML Thu Apr 18 17:07:01 2013 +0200 @@ -21,7 +21,8 @@ (* Theory dependencies *) val sumcase_split_ss = - HOL_basic_ss addsimps (@{thm Product_Type.split} :: @{thms sum.cases}) + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps (@{thm Product_Type.split} :: @{thms sum.cases})) (* top-down access in balanced tree *) fun access_top_down {left, right, init} len i = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Meson/meson.ML --- a/src/HOL/Tools/Meson/meson.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Meson/meson.ML Thu Apr 18 17:07:01 2013 +0200 @@ -19,7 +19,7 @@ -> Proof.context -> thm list * Proof.context val finish_cnf: thm list -> thm list val presimplified_consts : string list - val presimplify: thm -> thm + val presimplify: Proof.context -> thm -> thm val make_nnf: Proof.context -> thm -> thm val choice_theorems : theory -> thm list val skolemize_with_choice_theorems : Proof.context -> thm list -> thm -> thm @@ -541,22 +541,23 @@ (* FIXME: "let_simp" is probably redundant now that we also rewrite with "Let_def [abs_def]". *) val nnf_ss = - HOL_basic_ss addsimps nnf_extra_simps + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps nnf_extra_simps addsimprocs [@{simproc defined_All}, @{simproc defined_Ex}, @{simproc neq}, - @{simproc let_simp}] + @{simproc let_simp}]) val presimplified_consts = [@{const_name simp_implies}, @{const_name False}, @{const_name True}, @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}, @{const_name If}, @{const_name Let}] -val presimplify = +fun presimplify ctxt = rewrite_rule (map safe_mk_meta_eq nnf_simps) - #> simplify nnf_ss + #> simplify (put_simpset nnf_ss ctxt) #> Raw_Simplifier.rewrite_rule @{thms Let_def [abs_def]} fun make_nnf ctxt th = case prems_of th of - [] => th |> presimplify |> make_nnf1 ctxt + [] => th |> presimplify ctxt |> make_nnf1 ctxt | _ => raise THM ("make_nnf: premises in argument", 0, [th]); fun choice_theorems thy = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Meson/meson_clausify.ML --- a/src/HOL/Tools/Meson/meson_clausify.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Meson/meson_clausify.ML Thu Apr 18 17:07:01 2013 +0200 @@ -14,7 +14,7 @@ val introduce_combinators_in_cterm : cterm -> thm val introduce_combinators_in_theorem : thm -> thm val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool - val ss_only : thm list -> simpset + val ss_only : thm list -> Proof.context -> Proof.context val cnf_axiom : Proof.context -> bool -> bool -> int -> thm -> (thm * term) option * thm list @@ -292,7 +292,7 @@ else Conv.all_conv | _ => Conv.all_conv) -fun ss_only ths = Simplifier.clear_ss HOL_basic_ss addsimps ths +fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths val cheat_choice = @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"} @@ -317,11 +317,11 @@ let fun skolemize choice_ths = skolemize_with_choice_theorems ctxt choice_ths - #> simplify (ss_only @{thms all_simps[symmetric]}) + #> simplify (ss_only @{thms all_simps[symmetric]} ctxt) val no_choice = null choice_ths val pull_out = if no_choice then - simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]}) + simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]} ctxt) else skolemize choice_ths val discharger_th = th |> pull_out diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Metis/metis_tactic.ML --- a/src/HOL/Tools/Metis/metis_tactic.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Metis/metis_tactic.ML Thu Apr 18 17:07:01 2013 +0200 @@ -249,7 +249,7 @@ fun preskolem_tac ctxt st0 = (if exists (Meson.has_too_many_clauses ctxt) (Logic.prems_of_goal (prop_of st0) 1) then - Simplifier.full_simp_tac (Meson_Clausify.ss_only @{thms not_all not_ex}) 1 + Simplifier.full_simp_tac (Meson_Clausify.ss_only @{thms not_all not_ex} ctxt) 1 THEN cnf.cnfx_rewrite_tac ctxt 1 else all_tac) st0 diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Apr 18 17:07:01 2013 +0200 @@ -873,8 +873,10 @@ val t_insts' = map rewrite_pat t_insts val intro'' = Thm.instantiate (T_insts, t_insts') intro val [intro'''] = Variable.export ctxt3 ctxt [intro''] - val intro'''' = Simplifier.full_simplify - (HOL_basic_ss addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}]) + val intro'''' = + Simplifier.full_simplify + (put_simpset HOL_basic_ss ctxt + addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}]) intro''' (* splitting conjunctions introduced by Pair_eq*) val intro''''' = split_conjuncts_in_assms ctxt intro'''' diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Thu Apr 18 17:07:01 2013 +0200 @@ -1054,7 +1054,8 @@ val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI) val simprules = [defthm, @{thm eval_pred}, @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}] - val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1 + val unfolddef_tac = + Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simprules) 1 val introthm = Goal.prove ctxt (argnames @ hoarg_names' @ ["y"]) [] introtrm (fn _ => unfolddef_tac) val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)); @@ -1070,7 +1071,7 @@ (Predicate_Comp_Funs.mk_not (list_comb (funtrm, inargs)), HOLogic.unit)) val neg_introtrm = Logic.list_implies (neg_predpropI :: param_eqs, neg_funpropI) val tac = - Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps + Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm if_False} :: @{thm Predicate.not_pred_eq} :: simprules)) 1 THEN rtac @{thm Predicate.singleI} 1 in SOME (Goal.prove ctxt (argnames @ hoarg_names') [] @@ -1308,7 +1309,7 @@ (HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs)) val def = predfun_definition_of ctxt predname full_mode val tac = fn _ => Simplifier.simp_tac - (HOL_basic_ss addsimps [def, @{thm holds_eq}, @{thm eval_pred}]) 1 + (put_simpset HOL_basic_ss ctxt addsimps [def, @{thm holds_eq}, @{thm eval_pred}]) 1 val eq = Goal.prove ctxt arg_names [] eq_term tac in (pred, result_thms @ [eq]) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Thu Apr 18 17:07:01 2013 +0200 @@ -159,8 +159,9 @@ fun inline_equations thy th = let - val inline_defs = Predicate_Compile_Inline_Defs.get (Proof_Context.init_global thy) - val th' = (Simplifier.full_simplify (HOL_basic_ss addsimps inline_defs)) th + val ctxt = Proof_Context.init_global thy + val inline_defs = Predicate_Compile_Inline_Defs.get ctxt + val th' = (Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps inline_defs)) th (*val _ = print_step options ("Inlining " ^ (Syntax.string_of_term_global thy (prop_of th)) ^ "with " ^ (commas (map ((Syntax.string_of_term_global thy) o prop_of) inline_defs)) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Thu Apr 18 17:07:01 2013 +0200 @@ -165,7 +165,7 @@ Logic.dest_implies o prop_of) @{thm exI} fun prove_introrule (index, (ps, introrule)) = let - val tac = Simplifier.simp_tac (HOL_basic_ss addsimps [th]) 1 + val tac = Simplifier.simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [th]) 1 THEN EVERY1 (select_disj (length disjuncts) (index + 1)) THEN (EVERY (map (fn y => rtac (Drule.cterm_instantiate [(x, cterm_of thy (Free y))] @{thm exI}) 1) ps)) @@ -179,16 +179,17 @@ end in maps introrulify' ths' |> Variable.export ctxt' ctxt end -val rewrite = - Simplifier.simplify (HOL_basic_ss addsimps [@{thm Ball_def}, @{thm Bex_def}]) - #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}]) - #> Conv.fconv_rule nnf_conv - #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm ex_disj_distrib}]) +fun rewrite ctxt = + Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm Ball_def}, @{thm Bex_def}]) + #> Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm all_not_ex}]) + #> Conv.fconv_rule (nnf_conv ctxt) + #> Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm ex_disj_distrib}]) fun rewrite_intros thy = - Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm all_not_ex}]) + Simplifier.full_simplify (Simplifier.global_context thy HOL_basic_ss addsimps [@{thm all_not_ex}]) #> Simplifier.full_simplify - (HOL_basic_ss addsimps (tl @{thms bool_simps}) addsimps @{thms nnf_simps}) + (Simplifier.global_context thy HOL_basic_ss + addsimps (tl @{thms bool_simps}) addsimps @{thms nnf_simps}) #> split_conjuncts_in_assms (Proof_Context.init_global thy) fun print_specs options thy msg ths = @@ -205,7 +206,7 @@ val ctxt = Proof_Context.init_global thy val intros = if forall is_pred_equation specs then - map (split_conjuncts_in_assms ctxt) (introrulify thy (map rewrite specs)) + map (split_conjuncts_in_assms ctxt) (introrulify thy (map (rewrite ctxt) specs)) else if forall (is_intro constname) specs then map (rewrite_intros thy) specs else diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Thu Apr 18 17:07:01 2013 +0200 @@ -38,9 +38,11 @@ (** special setup for simpset **) -val HOL_basic_ss' = HOL_basic_ss addsimps @{thms simp_thms Pair_eq} - setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac)) - setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI})) +val HOL_basic_ss' = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps @{thms simp_thms Pair_eq} + setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac)) + setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))) (* auxillary functions *) @@ -72,7 +74,7 @@ val param_derivations = param_derivations_of deriv val ho_args = ho_args_of mode args val f_tac = case f of - Const (name, _) => simp_tac (HOL_basic_ss addsimps + Const (name, _) => simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm eval_pred}, predfun_definition_of ctxt name mode, @{thm split_eta}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1 @@ -154,15 +156,15 @@ (* replace TRY by determining if it necessary - are there equations when calling compile match? *) in (* make this simpset better! *) - asm_full_simp_tac (HOL_basic_ss' addsimps simprules) 1 + asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt addsimps simprules) 1 THEN print_tac options "after prove_match:" THEN (DETERM (TRY (rtac eval_if_P 1 THEN (SUBPROOF (fn {context, params, prems, asms, concl, schematics} => (REPEAT_DETERM (rtac @{thm conjI} 1 - THEN (SOLVED (asm_simp_tac HOL_basic_ss' 1)))) + THEN (SOLVED (asm_simp_tac (put_simpset HOL_basic_ss' ctxt) 1)))) THEN print_tac options "if condition to be solved:" - THEN asm_simp_tac HOL_basic_ss' 1 + THEN asm_simp_tac (put_simpset HOL_basic_ss' ctxt) 1 THEN TRY ( let val prems' = maps dest_conjunct_prem (take nargs prems) @@ -190,8 +192,8 @@ (all_input_of T)) preds in - simp_tac (HOL_basic_ss addsimps - (@{thms HOL.simp_thms eval_pred} @ defs)) 1 + simp_tac (put_simpset HOL_basic_ss ctxt + addsimps (@{thms HOL.simp_thms eval_pred} @ defs)) 1 (* need better control here! *) end @@ -201,7 +203,7 @@ fun prove_prems out_ts [] = (prove_match options ctxt nargs out_ts) THEN print_tac options "before simplifying assumptions" - THEN asm_full_simp_tac HOL_basic_ss' 1 + THEN asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1 THEN print_tac options "before single intro rule" THEN Subgoal.FOCUS_PREMS (fn {context, params, prems, asms, concl, schematics} => @@ -225,7 +227,7 @@ val rec_tac = prove_prems out_ts''' ps in print_tac options "before clause:" - (*THEN asm_simp_tac HOL_basic_ss 1*) + (*THEN asm_simp_tac (put_simpset HOL_basic_ss ctxt) 1*) THEN print_tac options "before prove_expr:" THEN prove_expr options ctxt nargs premposition (t, deriv) THEN print_tac options "after prove_expr:" @@ -244,7 +246,7 @@ val params = ho_args_of mode args in print_tac options "before prove_neg_expr:" - THEN full_simp_tac (HOL_basic_ss addsimps + THEN full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_eta}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1 THEN (if (is_some name) then @@ -260,8 +262,10 @@ rtac @{thm not_predI'} 1 (* test: *) THEN dtac @{thm sym} 1 - THEN asm_full_simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1) - THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 + THEN asm_full_simp_tac + (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1) + THEN simp_tac + (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1 THEN rec_tac end | Sidecond t => @@ -321,7 +325,7 @@ "splitting with rules \n" ^ Display.string_of_thm ctxt split_asm) THEN TRY (Splitter.split_asm_tac [split_asm] 1 THEN (print_tac options "after splitting with split_asm rules") - (* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1) + (* THEN (Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt) 1) THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*) THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2))) @@ -347,7 +351,7 @@ val param_derivations = param_derivations_of deriv val ho_args = ho_args_of mode args val f_tac = case f of - Const (name, _) => full_simp_tac (HOL_basic_ss addsimps + Const (name, _) => full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm eval_pred}::(predfun_definition_of ctxt name mode) :: @{thm "Product_Type.split_conv"}::[])) 1 | Free _ => all_tac @@ -390,7 +394,7 @@ preds in (* only simplify the one assumption *) - full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 + full_simp_tac (put_simpset HOL_basic_ss' ctxt addsimps @{thm eval_pred} :: defs) 1 (* need better control here! *) THEN print_tac options "after sidecond2 simplification" end @@ -399,24 +403,26 @@ let val pred_intro_rule = nth (intros_of ctxt pred) (i - 1) val (in_ts, _) = split_mode mode ts; - val split_ss = HOL_basic_ss' addsimps [@{thm split_eta}, @{thm split_beta}, - @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] + val split_simpset = + put_simpset HOL_basic_ss' ctxt + addsimps [@{thm split_eta}, @{thm split_beta}, + @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] fun prove_prems2 out_ts [] = print_tac options "before prove_match2 - last call:" THEN prove_match2 options ctxt out_ts THEN print_tac options "after prove_match2 - last call:" THEN (etac @{thm singleE} 1) THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1)) - THEN (asm_full_simp_tac HOL_basic_ss' 1) + THEN (asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1) THEN TRY ( (REPEAT_DETERM (etac @{thm Pair_inject} 1)) - THEN (asm_full_simp_tac HOL_basic_ss' 1) + THEN (asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1) THEN SOLVED (print_tac options "state before applying intro rule:" THEN (rtac pred_intro_rule (* How to handle equality correctly? *) THEN_ALL_NEW (K (print_tac options "state before assumption matching") - THEN' (atac ORELSE' ((CHANGED o asm_full_simp_tac split_ss) THEN' (TRY o atac))) + THEN' (atac ORELSE' ((CHANGED o asm_full_simp_tac split_simpset) THEN' (TRY o atac))) THEN' (K (print_tac options "state after pre-simplification:")) THEN' (K (print_tac options "state after assumption matching:")))) 1)) | prove_prems2 out_ts ((p, deriv) :: ps) = @@ -443,10 +449,10 @@ print_tac options "before neg prem 2" THEN etac @{thm bindE} 1 THEN (if is_some name then - full_simp_tac (HOL_basic_ss addsimps + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [predfun_definition_of ctxt (the name) mode]) 1 THEN etac @{thm not_predE} 1 - THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 + THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1 THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations)) else etac @{thm not_predE'} 1) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Qelim/cooper.ML Thu Apr 18 17:07:01 2013 +0200 @@ -58,23 +58,26 @@ val get = Data.get o Context.Proof; fun add ts = Thm.declaration_attribute (fn th => fn context => - context |> Data.map (fn (ss,ts') => - (ss addsimps [th], merge (op aconv) (ts',ts) ))) + context |> Data.map (fn (ss, ts') => + (simpset_map (Context.proof_of context) (fn ctxt => ctxt addsimps [th]) ss, + merge (op aconv) (ts', ts)))) fun del ts = Thm.declaration_attribute (fn th => fn context => - context |> Data.map (fn (ss,ts') => - (ss delsimps [th], subtract (op aconv) ts' ts ))) + context |> Data.map (fn (ss, ts') => + (simpset_map (Context.proof_of context) (fn ctxt => ctxt delsimps [th]) ss, + subtract (op aconv) ts' ts))) fun simp_thms_conv ctxt = - Simplifier.rewrite (Simplifier.context ctxt HOL_basic_ss addsimps @{thms simp_thms}); + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms}); val FWD = Drule.implies_elim_list; val true_tm = @{cterm "True"}; val false_tm = @{cterm "False"}; val zdvd1_eq = @{thm "zdvd1_eq"}; -val presburger_ss = @{simpset} addsimps [zdvd1_eq]; +val presburger_ss = simpset_of (@{context} addsimps [zdvd1_eq]); val lin_ss = - presburger_ss addsimps (@{thm dvd_eq_mod_eq_0} :: zdvd1_eq :: @{thms add_ac [where 'a=int]}); + simpset_of (put_simpset presburger_ss @{context} + addsimps (@{thm dvd_eq_mod_eq_0} :: zdvd1_eq :: @{thms add_ac [where 'a=int]})); val iT = HOLogic.intT val bT = HOLogic.boolT; @@ -109,8 +112,10 @@ val [dvd_uminus, dvd_uminus'] = @{thms "uminus_dvd_conv"}; -val eval_ss = presburger_ss addsimps [simp_from_to] delsimps [insert_iff,bex_triv]; -val eval_conv = Simplifier.rewrite eval_ss; +val eval_ss = + simpset_of (put_simpset presburger_ss @{context} + addsimps [simp_from_to] delsimps [insert_iff, bex_triv]); +fun eval_conv ctxt = Simplifier.rewrite (put_simpset eval_ss ctxt); (* recognising cterm without moving to terms *) @@ -221,7 +226,7 @@ (* Canonical linear form for terms, formulae etc.. *) fun provelin ctxt t = Goal.prove ctxt [] [] t - (fn _ => EVERY [simp_tac lin_ss 1, TRY (Lin_Arith.tac ctxt 1)]); + (fn _ => EVERY [simp_tac (put_simpset lin_ss ctxt) 1, TRY (Lin_Arith.tac ctxt 1)]); fun linear_cmul 0 tm = zero | linear_cmul n tm = case tm of Const (@{const_name Groups.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b @@ -309,7 +314,7 @@ val (d',t') = Thm.dest_binop (Thm.rhs_of th) val (dt',tt') = (term_of d', term_of t') in if is_number dt' andalso is_number tt' - then Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite presburger_ss)) th + then Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite (put_simpset presburger_ss ctxt))) th else let val dth = @@ -369,7 +374,7 @@ fun nzprop x = let val th = - Simplifier.rewrite lin_ss + Simplifier.rewrite (put_simpset lin_ss ctxt) (Thm.apply @{cterm Trueprop} (Thm.apply @{cterm "Not"} (Thm.apply (Thm.apply @{cterm "op = :: int => _"} (Numeral.mk_cnumber @{ctyp "int"} x)) @{cterm "0::int"}))) @@ -460,7 +465,7 @@ fun divprop x = let val th = - Simplifier.rewrite lin_ss + Simplifier.rewrite (put_simpset lin_ss ctxt) (Thm.apply @{cterm Trueprop} (Thm.apply (Thm.apply dvdc (Numeral.mk_cnumber @{ctyp "int"} x)) cd)) in Thm.equal_elim (Thm.symmetric th) TrueI end; @@ -472,7 +477,7 @@ Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option) end val dp = - let val th = Simplifier.rewrite lin_ss + let val th = Simplifier.rewrite (put_simpset lin_ss ctxt) (Thm.apply @{cterm Trueprop} (Thm.apply (Thm.apply @{cterm "op < :: int => _"} @{cterm "0::int"}) cd)) in Thm.equal_elim (Thm.symmetric th) TrueI end; @@ -544,7 +549,7 @@ in [dp, inf, nb, pd] MRS cpth end val cpth' = Thm.transitive uth (cpth RS eq_reflection) -in Thm.transitive cpth' ((simp_thms_conv ctxt then_conv eval_conv) (Thm.rhs_of cpth')) +in Thm.transitive cpth' ((simp_thms_conv ctxt then_conv eval_conv ctxt) (Thm.rhs_of cpth')) end; fun literals_conv bops uops env cv = @@ -556,14 +561,20 @@ in h end; fun integer_nnf_conv ctxt env = - nnf_conv then_conv literals_conv [HOLogic.conj, HOLogic.disj] [] env (linearize_conv ctxt); + nnf_conv ctxt then_conv literals_conv [HOLogic.conj, HOLogic.disj] [] env (linearize_conv ctxt); -val conv_ss = HOL_basic_ss addsimps - (@{thms simp_thms} @ take 4 @{thms ex_simps} @ [not_all, all_not_ex, @{thm ex_disj_distrib}]); +val conv_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps (@{thms simp_thms} @ take 4 @{thms ex_simps} @ + [not_all, all_not_ex, @{thm ex_disj_distrib}])); fun conv ctxt p = - Qelim.gen_qelim_conv (Simplifier.rewrite conv_ss) (Simplifier.rewrite presburger_ss) (Simplifier.rewrite conv_ss) - (cons o term_of) (Misc_Legacy.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) + Qelim.gen_qelim_conv + (Simplifier.rewrite (put_simpset conv_ss ctxt)) + (Simplifier.rewrite (put_simpset presburger_ss ctxt)) + (Simplifier.rewrite (put_simpset conv_ss ctxt)) + (cons o term_of) (Misc_Legacy.term_frees (term_of p)) + (linearize_conv ctxt) (integer_nnf_conv ctxt) (cooperex_conv ctxt) p handle CTERM _ => raise COOPER "bad cterm" | THM _ => raise COOPER "bad thm" @@ -690,7 +701,8 @@ (Thm.cterm_of (Proof_Context.theory_of ctxt) o Logic.mk_equals o pairself HOLogic.mk_Trueprop) (t, procedure t))))); -val comp_ss = HOL_ss addsimps @{thms semiring_norm}; +val comp_ss = + simpset_of (put_simpset HOL_ss @{context} addsimps @{thms semiring_norm}); fun strip_objimp ct = (case Thm.term_of ct of @@ -708,10 +720,12 @@ | _ => ([],ct); local - val all_maxscope_ss = - HOL_basic_ss addsimps map (fn th => th RS sym) @{thms "all_simps"} + val all_maxscope_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps map (fn th => th RS sym) @{thms "all_simps"}) in -fun thin_prems_tac ctxt P = simp_tac all_maxscope_ss THEN' +fun thin_prems_tac ctxt P = + simp_tac (put_simpset all_maxscope_ss ctxt) THEN' CSUBGOAL (fn (p', i) => let val (qvs, p) = strip_objall (Thm.dest_arg p') @@ -791,41 +805,46 @@ in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end)); local -val ss1 = comp_ss - addsimps @{thms simp_thms} @ [@{thm "nat_numeral"} RS sym, @{thm "zdvd_int"}] - @ map (fn r => r RS sym) - [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, - @{thm "zmult_int"}] - |> Splitter.add_split @{thm "zdiff_int_split"} +val ss1 = + simpset_of (put_simpset comp_ss @{context} + addsimps @{thms simp_thms} @ [@{thm "nat_numeral"} RS sym, @{thm "zdvd_int"}] + @ map (fn r => r RS sym) + [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, + @{thm "zmult_int"}] + |> Splitter.add_split @{thm "zdiff_int_split"}) -val ss2 = HOL_basic_ss - addsimps [@{thm "nat_0_le"}, @{thm "int_numeral"}, - @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "zero_le_numeral"}, - @{thm "le_numeral_extra"(3)}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}] - |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] -val div_mod_ss = HOL_basic_ss addsimps @{thms simp_thms} - @ map (Thm.symmetric o mk_meta_eq) - [@{thm "dvd_eq_mod_eq_0"}, - @{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"}, - @{thm "mod_add_eq"}, @{thm "div_add1_eq"}, @{thm "zdiv_zadd1_eq"}] - @ [@{thm "mod_self"}, @{thm "mod_by_0"}, @{thm "div_by_0"}, - @{thm "div_0"}, @{thm "mod_0"}, @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, - @{thm "mod_1"}, @{thm "Suc_eq_plus1"}] - @ @{thms add_ac} - addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}] +val ss2 = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm "nat_0_le"}, @{thm "int_numeral"}, + @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "zero_le_numeral"}, + @{thm "le_numeral_extra"(3)}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}] + |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]) +val div_mod_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps @{thms simp_thms} + @ map (Thm.symmetric o mk_meta_eq) + [@{thm "dvd_eq_mod_eq_0"}, + @{thm "mod_add_left_eq"}, @{thm "mod_add_right_eq"}, + @{thm "mod_add_eq"}, @{thm "div_add1_eq"}, @{thm "zdiv_zadd1_eq"}] + @ [@{thm "mod_self"}, @{thm "mod_by_0"}, @{thm "div_by_0"}, + @{thm "div_0"}, @{thm "mod_0"}, @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, + @{thm "mod_1"}, @{thm "Suc_eq_plus1"}] + @ @{thms add_ac} + addsimprocs [@{simproc cancel_div_mod_nat}, @{simproc cancel_div_mod_int}]) val splits_ss = - comp_ss addsimps [@{thm "mod_div_equality'"}] - |> fold Splitter.add_split - [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, - @{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}] + simpset_of (put_simpset comp_ss @{context} + addsimps [@{thm "mod_div_equality'"}] + |> fold Splitter.add_split + [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, + @{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}]) in fun nat_to_int_tac ctxt = - simp_tac (Simplifier.context ctxt ss1) THEN_ALL_NEW - simp_tac (Simplifier.context ctxt ss2) THEN_ALL_NEW - simp_tac (Simplifier.context ctxt comp_ss); + simp_tac (put_simpset ss1 ctxt) THEN_ALL_NEW + simp_tac (put_simpset ss2 ctxt) THEN_ALL_NEW + simp_tac (put_simpset comp_ss ctxt); -fun div_mod_tac ctxt i = simp_tac (Simplifier.context ctxt div_mod_ss) i; -fun splits_tac ctxt i = simp_tac (Simplifier.context ctxt splits_ss) i; +fun div_mod_tac ctxt = simp_tac (put_simpset div_mod_ss ctxt); +fun splits_tac ctxt = simp_tac (put_simpset splits_ss ctxt); end; fun core_tac ctxt = CSUBGOAL (fn (p, i) => @@ -844,20 +863,21 @@ fun tac elim add_ths del_ths = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} => let - val ss = Simplifier.context ctxt (fst (get ctxt)) delsimps del_ths addsimps add_ths + val simpset_ctxt = + put_simpset (fst (get ctxt)) ctxt delsimps del_ths addsimps add_ths val aprems = Arith_Data.get_arith_facts ctxt in Method.insert_tac aprems THEN_ALL_NEW Object_Logic.full_atomize_tac THEN_ALL_NEW CONVERSION Thm.eta_long_conversion - THEN_ALL_NEW simp_tac ss + THEN_ALL_NEW simp_tac simpset_ctxt THEN_ALL_NEW (TRY o generalize_tac (int_nat_terms ctxt)) THEN_ALL_NEW Object_Logic.full_atomize_tac THEN_ALL_NEW (thin_prems_tac ctxt (is_relevant ctxt)) THEN_ALL_NEW Object_Logic.full_atomize_tac THEN_ALL_NEW div_mod_tac ctxt THEN_ALL_NEW splits_tac ctxt - THEN_ALL_NEW simp_tac ss + THEN_ALL_NEW simp_tac simpset_ctxt THEN_ALL_NEW CONVERSION Thm.eta_long_conversion THEN_ALL_NEW nat_to_int_tac ctxt THEN_ALL_NEW (core_tac ctxt) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Qelim/qelim.ML --- a/src/HOL/Tools/Qelim/qelim.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Qelim/qelim.ML Thu Apr 18 17:07:01 2013 +0200 @@ -8,7 +8,8 @@ sig val gen_qelim_conv: conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a -> ('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv - val standard_qelim_conv: (cterm list -> conv) -> (cterm list -> conv) -> + val standard_qelim_conv: Proof.context -> + (cterm list -> conv) -> (cterm list -> conv) -> (cterm list -> conv) -> conv end; @@ -53,12 +54,19 @@ (* Instantiation of some parameter for most common cases *) local -val pcv = - Simplifier.rewrite - (HOL_basic_ss addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}); + +val ss = + simpset_of + (put_simpset HOL_basic_ss @{context} + addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}); +fun pcv ctxt = Simplifier.rewrite (put_simpset ss ctxt) -in fun standard_qelim_conv atcv ncv qcv p = - gen_qelim_conv pcv pcv pcv cons (Thm.add_cterm_frees p []) atcv ncv qcv p +in + +fun standard_qelim_conv ctxt atcv ncv qcv p = + let val pcv = pcv ctxt + in gen_qelim_conv pcv pcv pcv cons (Thm.add_cterm_frees p []) atcv ncv qcv p end + end; end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Quickcheck/random_generators.ML --- a/src/HOL/Tools/Quickcheck/random_generators.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Quickcheck/random_generators.ML Thu Apr 18 17:07:01 2013 +0200 @@ -79,7 +79,7 @@ val rew_thms = map mk_meta_eq [@{thm natural_zero_minus_one}, @{thm Suc_natural_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}]; val rew_ts = map (Logic.dest_equals o Thm.prop_of) rew_thms; -val rew_ss = HOL_ss addsimps rew_thms; +val rew_ss = simpset_of (put_simpset HOL_ss @{context} addsimps rew_thms); in @@ -104,10 +104,11 @@ val rule = @{thm random_aux_rec} |> Drule.instantiate_normalize ([(aT, icT)], [(cT_random_aux, cert t_random_aux), (cT_rhs, cert t_rhs)]); - val tac = ALLGOALS (rtac rule) - THEN ALLGOALS (simp_tac rew_ss) + fun tac ctxt = + ALLGOALS (rtac rule) + THEN ALLGOALS (simp_tac (put_simpset rew_ss ctxt)) THEN (ALLGOALS (Proof_Context.fact_tac eqs2)) - val simp = Goal.prove_sorry lthy' [v] [] eq (K tac); + val simp = Goal.prove_sorry lthy' [v] [] eq (tac o #context); in (simp, lthy') end; end; @@ -141,9 +142,10 @@ let val proj_simps = map (snd o snd) proj_defs; fun tac { context = ctxt, prems = _ } = - ALLGOALS (simp_tac (HOL_ss addsimps proj_simps)) + ALLGOALS (simp_tac (put_simpset HOL_ss ctxt addsimps proj_simps)) THEN ALLGOALS (EqSubst.eqsubst_tac ctxt [0] [aux_simp]) - THEN ALLGOALS (simp_tac (HOL_ss addsimps [@{thm fst_conv}, @{thm snd_conv}])); + THEN ALLGOALS (simp_tac + (put_simpset HOL_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}])); in (map (fn prop => Goal.prove_sorry lthy [v] [] prop tac) eqs, lthy) end; in lthy diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/Quotient/quotient_tacs.ML --- a/src/HOL/Tools/Quotient/quotient_tacs.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Thu Apr 18 17:07:01 2013 +0200 @@ -32,8 +32,8 @@ (* Since HOL_basic_ss is too "big" for us, we *) (* need to set up our own minimal simpset. *) -fun mk_minimal_ss ctxt = - Simplifier.context ctxt empty_ss +fun mk_minimal_simpset ctxt = + empty_simpset ctxt |> Simplifier.set_subgoaler asm_simp_tac |> Simplifier.set_mksimps (mksimps []) @@ -57,16 +57,14 @@ fun equiv_tac ctxt = REPEAT_ALL_NEW (resolve_tac (Quotient_Info.equiv_rules ctxt)) -fun equiv_solver_tac ss = equiv_tac (Simplifier.the_context ss) -val equiv_solver = mk_solver "Equivalence goal solver" equiv_solver_tac +val equiv_solver = mk_solver "Equivalence goal solver" equiv_tac fun quotient_tac ctxt = (REPEAT_ALL_NEW (FIRST' [rtac @{thm identity_quotient3}, resolve_tac (Quotient_Info.quotient_rules ctxt)])) -fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss) -val quotient_solver = mk_solver "Quotient goal solver" quotient_solver_tac +val quotient_solver = mk_solver "Quotient goal solver" quotient_tac fun solve_quotient_assm ctxt thm = case Seq.pull (quotient_tac ctxt 1 thm) of @@ -113,21 +111,17 @@ | SOME inst2 => try (Drule.instantiate_normalize inst2) thm')) end -fun ball_bex_range_simproc ss redex = - let - val ctxt = Simplifier.the_context ss - in - case redex of - (Const (@{const_name "Ball"}, _) $ (Const (@{const_name "Respects"}, _) $ - (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) => - calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2 +fun ball_bex_range_simproc ctxt redex = + case redex of + (Const (@{const_name "Ball"}, _) $ (Const (@{const_name "Respects"}, _) $ + (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) => + calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2 - | (Const (@{const_name "Bex"}, _) $ (Const (@{const_name "Respects"}, _) $ - (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) => - calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2 + | (Const (@{const_name "Bex"}, _) $ (Const (@{const_name "Respects"}, _) $ + (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) => + calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2 - | _ => NONE - end + | _ => NONE (* Regularize works as follows: @@ -162,9 +156,9 @@ val ball_pat = @{term "Ball (Respects (R1 ===> R2)) P"} val bex_pat = @{term "Bex (Respects (R1 ===> R2)) P"} val simproc = - Simplifier.simproc_global_i thy "" [ball_pat, bex_pat] (K (ball_bex_range_simproc)) + Simplifier.simproc_global_i thy "" [ball_pat, bex_pat] ball_bex_range_simproc val simpset = - mk_minimal_ss ctxt + mk_minimal_simpset ctxt addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp} addsimprocs [simproc] addSolver equiv_solver addSolver quotient_solver @@ -531,12 +525,12 @@ val thms = @{thms Quotient3_abs_rep Quotient3_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs - val ss = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver + val simpset = (mk_minimal_simpset lthy) addsimps thms addSolver quotient_solver in EVERY' [ map_fun_tac lthy, lambda_prs_tac lthy, - simp_tac ss, + simp_tac simpset, TRY o rtac refl] end @@ -638,9 +632,9 @@ fun descend_procedure_tac ctxt simps = let - val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds) + val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in - full_simp_tac ss + full_simp_tac simpset THEN' Object_Logic.full_atomize_tac THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => @@ -688,9 +682,9 @@ fun partiality_descend_procedure_tac ctxt simps = let - val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds) + val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in - full_simp_tac ss + full_simp_tac simpset THEN' Object_Logic.full_atomize_tac THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => @@ -723,9 +717,9 @@ (* the tactic leaves three subgoals to be proved *) fun lift_procedure_tac ctxt simps rthm = let - val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds) + val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in - full_simp_tac ss + full_simp_tac simpset THEN' Object_Logic.full_atomize_tac THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => @@ -733,7 +727,7 @@ (* full_atomize_tac contracts eta redexes, so we do it also in the original theorem *) val rthm' = - rthm |> full_simplify ss + rthm |> full_simplify simpset |> Drule.eta_contraction_rule |> Thm.forall_intr_frees |> atomize_thm diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/SMT/smt_normalize.ML --- a/src/HOL/Tools/SMT/smt_normalize.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/SMT/smt_normalize.ML Thu Apr 18 17:07:01 2013 +0200 @@ -479,10 +479,9 @@ fun mk_number_eq ctxt i lhs = let val eq = SMT_Utils.mk_cequals lhs (Numeral.mk_cnumber @{ctyp int} i) - val ss = HOL_ss - addsimps [@{thm Int.int_numeral}] - fun tac _ = Simplifier.simp_tac (Simplifier.context ctxt ss) 1 - in Goal.norm_result (Goal.prove_internal [] eq tac) end + val tac = + Simplifier.simp_tac (put_simpset HOL_ss ctxt addsimps [@{thm Int.int_numeral}]) 1 + in Goal.norm_result (Goal.prove_internal [] eq (K tac)) end fun ite_conv cv1 cv2 = Conv.combination_conv (Conv.combination_conv (Conv.arg_conv cv1) cv2) cv2 @@ -539,13 +538,14 @@ | NONE => false) | is_strange_number _ _ = false - val pos_num_ss = HOL_ss - addsimps [@{thm Num.numeral_One}] - addsimps [@{thm Num.neg_numeral_def}] + val pos_num_ss = + simpset_of (put_simpset HOL_ss @{context} + addsimps [@{thm Num.numeral_One}] + addsimps [@{thm Num.neg_numeral_def}]) fun norm_num_conv ctxt = SMT_Utils.if_conv (is_strange_number ctxt) - (Simplifier.rewrite (Simplifier.context ctxt pos_num_ss)) Conv.no_conv + (Simplifier.rewrite (put_simpset pos_num_ss ctxt)) Conv.no_conv in fun normalize_numerals_conv ctxt = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/SMT/smt_real.ML --- a/src/HOL/Tools/SMT/smt_real.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/SMT/smt_real.ML Thu Apr 18 17:07:01 2013 +0200 @@ -121,7 +121,7 @@ by auto} val real_linarith_proc = Simplifier.simproc_global @{theory} "fast_real_arith" [ - "(m::real) < n", "(m::real) <= n", "(m::real) = n"] (K Lin_Arith.simproc) + "(m::real) < n", "(m::real) <= n", "(m::real) = n"] Lin_Arith.simproc (* setup *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/SMT/z3_proof_methods.ML --- a/src/HOL/Tools/SMT/z3_proof_methods.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/SMT/z3_proof_methods.ML Thu Apr 18 17:07:01 2013 +0200 @@ -99,7 +99,7 @@ in Z3_Proof_Tools.by_tac ( CONVERSION (SMT_Utils.prop_conv conv) - THEN' Simplifier.simp_tac HOL_ss) + THEN' Simplifier.simp_tac (put_simpset HOL_ss ctxt)) end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/SMT/z3_proof_reconstruction.ML --- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Thu Apr 18 17:07:01 2013 +0200 @@ -109,7 +109,7 @@ Classical.fast_tac (put_claset HOL_cs ctxt) fun simp_fast_tac ctxt = - Simplifier.simp_tac (HOL_ss addsimps [rewr_if]) + Simplifier.simp_tac (put_simpset HOL_ss ctxt addsimps [rewr_if]) THEN_ALL_NEW HOL_fast_tac ctxt end @@ -148,8 +148,7 @@ val remove_fun_app = mk_meta_eq @{thm SMT.fun_app_def} fun rewrite_conv _ [] = Conv.all_conv - | rewrite_conv ctxt eqs = Simplifier.full_rewrite - (Simplifier.context ctxt Simplifier.empty_ss addsimps eqs) + | rewrite_conv ctxt eqs = Simplifier.full_rewrite (empty_simpset ctxt addsimps eqs) val prep_rules = [@{thm Let_def}, remove_trigger, remove_weight, remove_fun_app, Z3_Proof_Literals.rewrite_true] @@ -658,7 +657,7 @@ Z3_Proof_Tools.by_tac ( NAMED ctxt' "arith" (Arith_Data.arith_tac ctxt') ORELSE' NAMED ctxt' "simp+arith" ( - Simplifier.asm_full_simp_tac simpset + Simplifier.asm_full_simp_tac (put_simpset simpset ctxt') THEN_ALL_NEW Arith_Data.arith_tac ctxt')))] @@ -718,19 +717,19 @@ named ctxt "pull-ite" Z3_Proof_Methods.prove_ite, Z3_Proof_Tools.by_abstraction 0 (true, false) ctxt [] (fn ctxt' => Z3_Proof_Tools.by_tac ( - NAMED ctxt' "simp (logic)" (Simplifier.simp_tac simpset) + NAMED ctxt' "simp (logic)" (Simplifier.simp_tac (put_simpset simpset ctxt')) THEN_ALL_NEW NAMED ctxt' "fast (logic)" (fast_tac ctxt'))), Z3_Proof_Tools.by_abstraction 0 (false, true) ctxt [] (fn ctxt' => Z3_Proof_Tools.by_tac ( (Tactic.rtac @{thm iff_allI} ORELSE' K all_tac) - THEN' NAMED ctxt' "simp (theory)" (Simplifier.simp_tac simpset) + THEN' NAMED ctxt' "simp (theory)" (Simplifier.simp_tac (put_simpset simpset ctxt')) THEN_ALL_NEW ( NAMED ctxt' "fast (theory)" (HOL_fast_tac ctxt') ORELSE' NAMED ctxt' "arith (theory)" (Arith_Data.arith_tac ctxt')))), Z3_Proof_Tools.by_abstraction 0 (true, true) ctxt [] (fn ctxt' => Z3_Proof_Tools.by_tac ( (Tactic.rtac @{thm iff_allI} ORELSE' K all_tac) - THEN' NAMED ctxt' "simp (full)" (Simplifier.simp_tac simpset) + THEN' NAMED ctxt' "simp (full)" (Simplifier.simp_tac (put_simpset simpset ctxt')) THEN_ALL_NEW ( NAMED ctxt' "fast (full)" (HOL_fast_tac ctxt') ORELSE' NAMED ctxt' "arith (full)" (Arith_Data.arith_tac ctxt')))), @@ -739,7 +738,7 @@ (fn ctxt' => Z3_Proof_Tools.by_tac ( (Tactic.rtac @{thm iff_allI} ORELSE' K all_tac) - THEN' NAMED ctxt' "simp (deepen)" (Simplifier.simp_tac simpset) + THEN' NAMED ctxt' "simp (deepen)" (Simplifier.simp_tac (put_simpset simpset ctxt')) THEN_ALL_NEW ( NAMED ctxt' "fast (deepen)" (HOL_fast_tac ctxt') ORELSE' NAMED ctxt' "arith (deepen)" (Arith_Data.arith_tac diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/SMT/z3_proof_tools.ML --- a/src/HOL/Tools/SMT/z3_proof_tools.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/SMT/z3_proof_tools.ML Thu Apr 18 17:07:01 2013 +0200 @@ -304,11 +304,11 @@ fun dest_binop ((c as Const _) $ t $ u) = (c, t, u) | dest_binop t = raise TERM ("dest_binop", [t]) - fun prove_antisym_le ss t = + fun prove_antisym_le ctxt t = let val (le, r, s) = dest_binop t val less = Const (@{const_name less}, Term.fastype_of le) - val prems = Simplifier.prems_of ss + val prems = Simplifier.prems_of ctxt in (case find_first (eq_prop (le $ s $ r)) prems of NONE => @@ -318,11 +318,11 @@ end handle THM _ => NONE - fun prove_antisym_less ss t = + fun prove_antisym_less ctxt t = let val (less, r, s) = dest_binop (HOLogic.dest_not t) val le = Const (@{const_name less_eq}, Term.fastype_of less) - val prems = Simplifier.prems_of ss + val prems = Simplifier.prems_of ctxt in (case find_first (eq_prop (le $ r $ s)) prems of NONE => @@ -332,21 +332,23 @@ end handle THM _ => NONE - val basic_simpset = HOL_ss addsimps @{thms field_simps} - addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}] - addsimps @{thms arith_special} addsimps @{thms arith_simps} - addsimps @{thms rel_simps} - addsimps @{thms array_rules} - addsimps @{thms term_true_def} addsimps @{thms term_false_def} - addsimps @{thms z3div_def} addsimps @{thms z3mod_def} - addsimprocs [@{simproc binary_int_div}, @{simproc binary_int_mod}] - addsimprocs [ - Simplifier.simproc_global @{theory} "fast_int_arith" [ - "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc), - Simplifier.simproc_global @{theory} "antisym_le" ["(x::'a::order) <= y"] - (K prove_antisym_le), - Simplifier.simproc_global @{theory} "antisym_less" ["~ (x::'a::linorder) < y"] - (K prove_antisym_less)] + val basic_simpset = + simpset_of (put_simpset HOL_ss @{context} + addsimps @{thms field_simps} + addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}] + addsimps @{thms arith_special} addsimps @{thms arith_simps} + addsimps @{thms rel_simps} + addsimps @{thms array_rules} + addsimps @{thms term_true_def} addsimps @{thms term_false_def} + addsimps @{thms z3div_def} addsimps @{thms z3mod_def} + addsimprocs [@{simproc binary_int_div}, @{simproc binary_int_mod}] + addsimprocs [ + Simplifier.simproc_global @{theory} "fast_int_arith" [ + "(m::int) < n", "(m::int) <= n", "(m::int) = n"] Lin_Arith.simproc, + Simplifier.simproc_global @{theory} "antisym_le" ["(x::'a::order) <= y"] + prove_antisym_le, + Simplifier.simproc_global @{theory} "antisym_less" ["~ (x::'a::linorder) < y"] + prove_antisym_less]) structure Simpset = Generic_Data ( @@ -357,10 +359,12 @@ ) in -fun add_simproc simproc = Simpset.map (fn ss => ss addsimprocs [simproc]) +fun add_simproc simproc context = + Simpset.map (simpset_map (Context.proof_of context) + (fn ctxt => ctxt addsimprocs [simproc])) context fun make_simpset ctxt rules = - Simplifier.context ctxt (Simpset.get (Context.Proof ctxt)) addsimps rules + simpset_of (put_simpset (Simpset.get (Context.Proof ctxt)) ctxt addsimps rules) end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/TFL/post.ML --- a/src/HOL/Tools/TFL/post.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/TFL/post.ML Thu Apr 18 17:07:01 2013 +0200 @@ -38,10 +38,10 @@ Prim.postprocess strict {wf_tac = REPEAT (ares_tac wfs 1), terminator = - asm_simp_tac (simpset_of ctxt) 1 + asm_simp_tac ctxt 1 THEN TRY (Arith_Data.arith_tac ctxt 1 ORELSE fast_force_tac (ctxt addSDs [@{thm not0_implies_Suc}]) 1), - simplifier = Rules.simpl_conv (simpset_of ctxt) []}; + simplifier = Rules.simpl_conv ctxt []}; @@ -103,7 +103,7 @@ val simplified' = map join_assums simplified val dummy = (Prim.trace_thms "solved =" solved; Prim.trace_thms "simplified' =" simplified') - val rewr = full_simplify (simpset_of ctxt addsimps (solved @ simplified')); + val rewr = full_simplify (ctxt addsimps (solved @ simplified')); val dummy = Prim.trace_thms "Simplifying the induction rule..." [induction] val induction' = rewr induction @@ -121,12 +121,12 @@ (*lcp: curry the predicate of the induction rule*) -fun curry_rule rl = - Split_Rule.split_rule_var (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl; +fun curry_rule ctxt rl = + Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl; (*lcp: put a theorem into Isabelle form, using meta-level connectives*) fun meta_outer ctxt = - curry_rule o Drule.export_without_context o + curry_rule ctxt o Drule.export_without_context o rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE))); (*Strip off the outer !P*) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/TFL/rules.ML --- a/src/HOL/Tools/TFL/rules.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/TFL/rules.ML Thu Apr 18 17:07:01 2013 +0200 @@ -43,7 +43,7 @@ val list_beta_conv: cterm -> cterm list -> thm val SUBS: thm list -> thm -> thm - val simpl_conv: simpset -> thm list -> cterm -> thm + val simpl_conv: Proof.context -> thm list -> cterm -> thm val rbeta: thm -> thm val tracing: bool Unsynchronized.ref @@ -422,8 +422,8 @@ val rew_conv = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)); -fun simpl_conv ss thl ctm = - rew_conv (ss addsimps thl) ctm RS meta_eq_to_obj_eq; +fun simpl_conv ctxt thl ctm = + rew_conv (ctxt addsimps thl) ctm RS meta_eq_to_obj_eq; val RIGHT_ASSOC = rewrite_rule [Thms.disj_assoc]; @@ -648,14 +648,14 @@ fun CONTEXT_REWRITE_RULE (func, G, cut_lemma, congs) th = let val globals = func::G - val ss0 = Simplifier.global_context (Thm.theory_of_thm th) empty_ss - val pbeta_reduce = simpl_conv ss0 [@{thm split_conv} RS eq_reflection]; + val ctxt0 = Simplifier.global_context (Thm.theory_of_thm th) empty_ss + val pbeta_reduce = simpl_conv ctxt0 [@{thm split_conv} RS eq_reflection]; val tc_list = Unsynchronized.ref []: term list Unsynchronized.ref val cut_lemma' = cut_lemma RS eq_reflection - fun prover used ss thm = - let fun cong_prover ss thm = + fun prover used ctxt thm = + let fun cong_prover ctxt thm = let val dummy = say "cong_prover:" - val cntxt = Simplifier.prems_of ss + val cntxt = Simplifier.prems_of ctxt val dummy = print_thms "cntxt:" cntxt val dummy = say "cong rule:" val dummy = say (Display.string_of_thm_without_context thm) @@ -666,8 +666,8 @@ val ants = map tych (Logic.strip_imp_prems imp) val eq = Logic.strip_imp_concl imp val lhs = tych(get_lhs eq) - val ss' = Simplifier.add_prems (map ASSUME ants) ss - val lhs_eq_lhs1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used) ss' lhs + val ctxt' = Simplifier.add_prems (map ASSUME ants) ctxt + val lhs_eq_lhs1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used) ctxt' lhs handle Utils.ERR _ => Thm.reflexive lhs val dummy = print_thms "proven:" [lhs_eq_lhs1] val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1 @@ -687,8 +687,8 @@ val Q = get_lhs eq1 val QeqQ1 = pbeta_reduce (tych Q) val Q1 = #2(Dcterm.dest_eq(cconcl QeqQ1)) - val ss' = Simplifier.add_prems (map ASSUME ants1) ss - val Q1eeqQ2 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ss' Q1 + val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt + val Q1eeqQ2 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ctxt' Q1 handle Utils.ERR _ => Thm.reflexive Q1 val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2)) val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl)) @@ -712,9 +712,9 @@ else let val tych = cterm_of thy val ants1 = map tych ants - val ss' = Simplifier.add_prems (map ASSUME ants1) ss + val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt val Q_eeq_Q1 = Raw_Simplifier.rewrite_cterm - (false,true,false) (prover used') ss' (tych Q) + (false,true,false) (prover used') ctxt' (tych Q) handle Utils.ERR _ => Thm.reflexive (tych Q) val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1 val lhs_eq_lhs2 = lhs_eeq_lhs2 RS meta_eq_to_obj_eq @@ -735,9 +735,9 @@ in SOME(eliminate (rename thm)) end handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *) - fun restrict_prover ss thm = + fun restrict_prover ctxt thm = let val dummy = say "restrict_prover:" - val cntxt = rev (Simplifier.prems_of ss) + val cntxt = rev (Simplifier.prems_of ctxt) val dummy = print_thms "cntxt:" cntxt val thy = Thm.theory_of_thm thm val Const("==>",_) $ (Const(@{const_name Trueprop},_) $ A) $ _ = Thm.prop_of thm @@ -777,13 +777,13 @@ in SOME (th'') end handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *) in - (if (is_cong thm) then cong_prover else restrict_prover) ss thm + (if (is_cong thm) then cong_prover else restrict_prover) ctxt thm end val ctm = cprop_of th val names = Misc_Legacy.add_term_names (term_of ctm, []) val th1 = Raw_Simplifier.rewrite_cterm (false, true, false) - (prover names) (ss0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm + (prover names) (ctxt0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm val th2 = Thm.equal_elim th1 th in (th2, filter_out restricted (!tc_list)) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/TFL/tfl.ML --- a/src/HOL/Tools/TFL/tfl.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/TFL/tfl.ML Thu Apr 18 17:07:01 2013 +0200 @@ -431,14 +431,14 @@ (*case_ss causes minimal simplification: bodies of case expressions are not simplified. Otherwise large examples (Red-Black trees) are too slow.*) - val case_ss = - Simplifier.global_context theory - (HOL_basic_ss addsimps case_rewrites + val case_simpset = + Simplifier.global_context theory HOL_basic_ss + addsimps case_rewrites |> fold (Simplifier.add_cong o #weak_case_cong o snd) - (Symtab.dest (Datatype.get_all theory))) - val corollaries' = map (Simplifier.simplify case_ss) corollaries + (Symtab.dest (Datatype.get_all theory)) + val corollaries' = map (Simplifier.simplify case_simpset) corollaries val extract = Rules.CONTEXT_REWRITE_RULE - (f, [R], @{thm cut_apply}, meta_tflCongs@context_congs) + (f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs) val (rules, TCs) = ListPair.unzip (map extract corollaries') val rules0 = map (rewrite_rule [Thms.CUT_DEF]) rules val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/arith_data.ML --- a/src/HOL/Tools/arith_data.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/arith_data.ML Thu Apr 18 17:07:01 2013 +0200 @@ -18,9 +18,9 @@ val prove_conv_nohyps: tactic list -> Proof.context -> term * term -> thm option val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option - val prove_conv2: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm - val simp_all_tac: thm list -> simpset -> tactic - val simplify_meta_eq: thm list -> simpset -> thm -> thm + val prove_conv2: tactic -> (Proof.context -> tactic) -> Proof.context -> term * term -> thm + val simp_all_tac: thm list -> Proof.context -> tactic + val simplify_meta_eq: thm list -> Proof.context -> thm -> thm val setup: theory -> theory end; @@ -105,17 +105,16 @@ fun prove_conv tacs ctxt (_: thm list) = prove_conv_nohyps tacs ctxt; -fun prove_conv2 expand_tac norm_tac ss tu = (* FIXME avoid Drule.export_without_context *) - mk_meta_eq (Drule.export_without_context (Goal.prove (Simplifier.the_context ss) [] [] +fun prove_conv2 expand_tac norm_tac ctxt tu = (* FIXME avoid Drule.export_without_context *) + mk_meta_eq (Drule.export_without_context (Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq tu)) - (K (EVERY [expand_tac, norm_tac ss])))); + (K (EVERY [expand_tac, norm_tac ctxt])))); -fun simp_all_tac rules = - let val ss0 = HOL_ss addsimps rules - in fn ss => ALLGOALS (safe_simp_tac (Simplifier.inherit_context ss ss0)) end; +fun simp_all_tac rules ctxt = + ALLGOALS (safe_simp_tac (put_simpset HOL_ss ctxt addsimps rules)); -fun simplify_meta_eq rules = - let val ss0 = HOL_basic_ss addsimps rules |> Simplifier.add_eqcong @{thm eq_cong2} - in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end; +fun simplify_meta_eq rules ctxt = + simplify (put_simpset HOL_basic_ss ctxt addsimps rules |> Simplifier.add_eqcong @{thm eq_cong2}) + o mk_meta_eq; end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/enriched_type.ML --- a/src/HOL/Tools/enriched_type.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/enriched_type.ML Thu Apr 18 17:07:01 2013 +0200 @@ -79,7 +79,8 @@ (* mapper properties *) -val compositionality_ss = Simplifier.add_simp (Simpdata.mk_eq @{thm comp_def}) HOL_basic_ss; +val compositionality_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps [Simpdata.mk_eq @{thm comp_def}]); fun make_comp_prop ctxt variances (tyco, mapper) = let @@ -127,11 +128,12 @@ val compositionality_prop = fold_rev Logic.all (map Free (args21 @ args32) @ [x]) eq2; fun prove_compositionality ctxt comp_thm = Goal.prove_sorry ctxt [] [] compositionality_prop (K (ALLGOALS (Method.insert_tac [@{thm fun_cong} OF [comp_thm]] - THEN' Simplifier.asm_lr_simp_tac compositionality_ss + THEN' Simplifier.asm_lr_simp_tac (put_simpset compositionality_ss ctxt) THEN_ALL_NEW (Goal.assume_rule_tac ctxt)))); in (comp_prop, prove_compositionality) end; -val identity_ss = Simplifier.add_simp (Simpdata.mk_eq @{thm id_def}) HOL_basic_ss; +val identity_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps [Simpdata.mk_eq @{thm id_def}]); fun make_id_prop ctxt variances (tyco, mapper) = let @@ -149,7 +151,8 @@ val (id_prop, identity_prop) = pairself (HOLogic.mk_Trueprop o HOLogic.mk_eq o rpair rhs) (lhs1, lhs2); fun prove_identity ctxt id_thm = Goal.prove_sorry ctxt [] [] identity_prop - (K (ALLGOALS (Method.insert_tac [id_thm] THEN' Simplifier.asm_lr_simp_tac identity_ss))); + (K (ALLGOALS (Method.insert_tac [id_thm] THEN' + Simplifier.asm_lr_simp_tac (put_simpset identity_ss ctxt)))); in (id_prop, prove_identity) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/groebner.ML --- a/src/HOL/Tools/groebner.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/groebner.ML Thu Apr 18 17:07:01 2013 +0200 @@ -9,10 +9,10 @@ vars: cterm list, semiring: cterm list * thm list, ideal : thm list} -> (cterm -> Rat.rat) -> (Rat.rat -> cterm) -> conv -> conv -> - {ring_conv : conv, + {ring_conv: Proof.context -> conv, simple_ideal: (cterm list -> cterm -> (cterm * cterm -> order) -> cterm list), multi_ideal: cterm list -> cterm list -> cterm list -> (cterm * cterm) list, - poly_eq_ss: simpset, unwind_conv : conv} + poly_eq_ss: simpset, unwind_conv: Proof.context -> conv} val ring_tac: thm list -> thm list -> Proof.context -> int -> tactic val ideal_tac: thm list -> thm list -> Proof.context -> int -> tactic val algebra_tac: thm list -> thm list -> Proof.context -> int -> tactic @@ -437,13 +437,19 @@ val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq; val bool_simps = @{thms bool_simps}; val nnf_simps = @{thms nnf_simps}; -val nnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps bool_simps addsimps nnf_simps) -val weak_dnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps @{thms weak_dnf_simps}); -val initial_conv = - Simplifier.rewrite - (HOL_basic_ss addsimps nnf_simps - addsimps [not_all, not_ex] - addsimps map (fn th => th RS sym) (@{thms ex_simps} @ @{thms all_simps})); +fun nnf_conv ctxt = + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps bool_simps addsimps nnf_simps) + +fun weak_dnf_conv ctxt = + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms weak_dnf_simps}); + +val initial_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps nnf_simps + addsimps [not_all, not_ex] + addsimps map (fn th => th RS sym) (@{thms ex_simps} @ @{thms all_simps})); +fun initial_conv ctxt = + Simplifier.rewrite (put_simpset initial_ss ctxt); val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec)); @@ -564,8 +570,8 @@ fun mkeq s t = Thm.apply @{cterm Trueprop} (Thm.apply (Thm.apply @{cterm "op = :: bool => _"} s) t) fun mk_exists v th = Drule.arg_cong_rule (ext (ctyp_of_term v)) (Thm.abstract_rule (getname v) v th) - val simp_ex_conv = - Simplifier.rewrite (HOL_basic_ss addsimps @{thms simp_thms(39)}) + fun simp_ex_conv ctxt = + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)}) fun frees t = Thm.add_cterm_frees t []; fun free_in v t = member op aconvc (frees t) v; @@ -704,14 +710,15 @@ else end_itlist ring_mk_add (map (holify_monomial vars) p) in holify_polynomial end ; -val idom_rule = simplify (HOL_basic_ss addsimps [idom_thm]); + +fun idom_rule ctxt = simplify (put_simpset HOL_basic_ss ctxt addsimps [idom_thm]); fun prove_nz n = eqF_elim (ring_eq_conv(mk_binop eq_tm (mk_const n) (mk_const(rat_0)))); val neq_01 = prove_nz (rat_1); fun neq_rule n th = [prove_nz n, th] MRS neq_thm; fun mk_add th1 = Thm.combination (Drule.arg_cong_rule ring_add_tm th1); -fun refute tm = +fun refute ctxt tm = if tm aconvc false_tm then assume_Trueprop tm else ((let val (nths0,eths0) = List.partition (is_neg o concl) (HOLogic.conj_elims (assume_Trueprop tm)) @@ -720,7 +727,7 @@ in if null eths then let - val th1 = end_itlist (fn th1 => fn th2 => idom_rule(HOLogic.conj_intr th1 th2)) nths + val th1 = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr th1 th2)) nths val th2 = Conv.fconv_rule ((Conv.arg_conv #> Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) th1 @@ -740,13 +747,13 @@ end else let - val nth = end_itlist (fn th1 => fn th2 => idom_rule(HOLogic.conj_intr th1 th2)) nths + val nth = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr th1 th2)) nths val (vars,pol::pols) = grobify_equations(list_mk_conj(Thm.dest_arg(concl nth)::map concl eths)) val (deg,l,cert) = grobner_strong vars pols pol val th1 = Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) nth - val th2 = funpow deg (idom_rule o HOLogic.conj_intr th1) neq_01 + val th2 = funpow deg (idom_rule ctxt o HOLogic.conj_intr th1) neq_01 in (vars,l,cert,th2) end) val cert_pos = map (fn (i,p) => (i,filter (fn (c,m) => c >/ rat_0) p)) cert @@ -772,7 +779,7 @@ end end) handle ERROR _ => raise CTERM ("Groebner-refute: unable to refute",[tm])) -fun ring tm = +fun ring ctxt tm = let fun mk_forall x p = Thm.apply @@ -780,19 +787,20 @@ @{cpat "All:: (?'a => bool) => _"}) (Thm.lambda x p) val avs = Thm.add_cterm_frees tm [] val P' = fold mk_forall avs tm - val th1 = initial_conv(mk_neg P') + val th1 = initial_conv ctxt (mk_neg P') val (evs,bod) = strip_exists(concl th1) in if is_forall bod then raise CTERM("ring: non-universal formula",[tm]) else let - val th1a = weak_dnf_conv bod + val th1a = weak_dnf_conv ctxt bod val boda = concl th1a - val th2a = refute_disj refute boda + val th2a = refute_disj (refute ctxt) boda val th2b = [mk_object_eq th1a, (th2a COMP notI) COMP PFalse'] MRS trans val th2 = fold (fn v => fn th => (Thm.forall_intr v th) COMP allI) evs (th2b RS PFalse) val th3 = Thm.equal_elim - (Simplifier.rewrite (HOL_basic_ss addsimps [not_ex RS sym]) (th2 |> cprop_of)) th2 + (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [not_ex RS sym]) + (th2 |> cprop_of)) th2 in specl avs ([[[mk_object_eq th1, th3 RS PFalse'] MRS trans] MRS PFalse] MRS notnotD) end @@ -815,15 +823,17 @@ end val poly_eq_simproc = let - fun proc phi ss t = + fun proc phi ctxt t = let val th = poly_eq_conv t in if Thm.is_reflexive th then NONE else SOME th end in make_simproc {lhss = [Thm.lhs_of idl_sub], name = "poly_eq_simproc", proc = proc, identifier = []} end; - val poly_eq_ss = HOL_basic_ss addsimps @{thms simp_thms} - addsimprocs [poly_eq_simproc] + val poly_eq_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps @{thms simp_thms} + addsimprocs [poly_eq_simproc]) local fun is_defined v t = @@ -849,7 +859,7 @@ in Conv.fconv_rule(funpow 2 Conv.arg_conv ring_normalize_conv) th2 end in - fun unwind_polys_conv tm = + fun unwind_polys_conv ctxt tm = let val (vars,bod) = strip_exists tm val cjs = striplist (dest_binary @{cterm HOL.conj}) bod @@ -864,7 +874,7 @@ (Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2)))) val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3)) val vars' = (remove op aconvc v vars) @ [v] - val th4 = Conv.fconv_rule (Conv.arg_conv simp_ex_conv) (mk_exists v th3) + val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists v th3) val th5 = ex_eq_conv (mk_eq tm (fold mk_ex (remove op aconvc v vars) (Thm.lhs_of th4))) in Thm.transitive th5 (fold mk_exists (remove op aconvc v vars) th4) end; @@ -966,10 +976,12 @@ | SOME (res as (theory, {is_const, dest_const, mk_const, conv = ring_eq_conv})) => #ring_conv (ring_and_ideal_conv theory dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt) - (Semiring_Normalizer.semiring_normalize_wrapper ctxt res)) form)); + (Semiring_Normalizer.semiring_normalize_wrapper ctxt res)) ctxt form)); -fun presimplify ctxt add_thms del_thms = asm_full_simp_tac (Simplifier.context ctxt - (HOL_basic_ss addsimps (Algebra_Simplification.get ctxt) delsimps del_thms addsimps add_thms)); +fun presimplify ctxt add_thms del_thms = + asm_full_simp_tac (put_simpset HOL_basic_ss ctxt + addsimps (Algebra_Simplification.get ctxt) + delsimps del_thms addsimps add_thms); fun ring_tac add_ths del_ths ctxt = Object_Logic.full_atomize_tac @@ -978,7 +990,7 @@ rtac (let val form = Object_Logic.dest_judgment p in case get_ring_ideal_convs ctxt form of NONE => Thm.reflexive form - | SOME thy => #ring_conv thy form + | SOME thy => #ring_conv thy ctxt form end) i handle TERM _ => no_tac | CTERM _ => no_tac @@ -1013,9 +1025,9 @@ in clarify_tac @{context} i THEN Object_Logic.full_atomize_tac i - THEN asm_full_simp_tac (Simplifier.context ctxt (#poly_eq_ss thy)) i + THEN asm_full_simp_tac (put_simpset (#poly_eq_ss thy) ctxt) i THEN clarify_tac @{context} i - THEN (REPEAT (CONVERSION (#unwind_conv thy) i)) + THEN (REPEAT (CONVERSION (#unwind_conv thy ctxt) i)) THEN SUBPROOF poly_exists_tac ctxt i end handle TERM _ => no_tac diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/inductive.ML Thu Apr 18 17:07:01 2013 +0200 @@ -32,7 +32,7 @@ val mono_del: attribute val mk_cases: Proof.context -> term -> thm val inductive_forall_def: thm - val rulify: thm -> thm + val rulify: Proof.context -> thm -> thm val inductive_cases: (Attrib.binding * string list) list -> local_theory -> thm list list * local_theory val inductive_cases_i: (Attrib.binding * term list) list -> local_theory -> @@ -346,10 +346,10 @@ ((binding, att), arule) end; -val rulify = - hol_simplify inductive_conj - #> hol_simplify inductive_rulify - #> hol_simplify inductive_rulify_fallback +fun rulify ctxt = + hol_simplify ctxt inductive_conj + #> hol_simplify ctxt inductive_rulify + #> hol_simplify ctxt inductive_rulify_fallback #> Simplifier.norm_hhf; end; @@ -515,7 +515,7 @@ EVERY (map (fn ci => etac @{thm disjE} 1 THEN prove_intr2 ci) c_intrs') THEN prove_intr2 last_c_intr end)) - |> rulify + |> rulify ctxt' |> singleton (Proof_Context.export ctxt' ctxt'') end; in @@ -533,15 +533,14 @@ val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE]; val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls; -fun simp_case_tac ss i = - EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i; +fun simp_case_tac ctxt i = + EVERY' [elim_tac, asm_full_simp_tac ctxt, elim_tac, REPEAT o bound_hyp_subst_tac] i; in fun mk_cases ctxt prop = let val thy = Proof_Context.theory_of ctxt; - val ss = simpset_of ctxt; fun err msg = error (Pretty.string_of (Pretty.block @@ -550,7 +549,7 @@ val elims = Induct.find_casesP ctxt prop; val cprop = Thm.cterm_of thy prop; - val tac = ALLGOALS (simp_case_tac ss) THEN prune_params_tac; + val tac = ALLGOALS (simp_case_tac ctxt) THEN prune_params_tac; fun mk_elim rl = Thm.implies_intr cprop (Tactic.rule_by_tactic ctxt tac (Thm.assume cprop RS rl)) |> singleton (Variable.export (Variable.auto_fixes prop ctxt) ctxt); @@ -617,7 +616,7 @@ (Term.add_vars (lhs_of eq) []); in Drule.cterm_instantiate inst eq - |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.full_rewrite (simpset_of ctxt)))) + |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.full_rewrite ctxt))) |> singleton (Variable.export ctxt' ctxt) end @@ -822,7 +821,7 @@ (Const (fp_name, (predT --> predT) --> predT) $ fp_fun))) ||> Local_Theory.restore_naming lthy; val fp_def' = - Simplifier.rewrite (HOL_basic_ss addsimps [fp_def]) + Simplifier.rewrite (put_simpset HOL_basic_ss lthy' addsimps [fp_def]) (cterm_of (Proof_Context.theory_of lthy') (list_comb (rec_const, params))); val specs = if length cs < 2 then [] @@ -895,7 +894,7 @@ apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>> Local_Theory.note ((rec_qualified true (Binding.name (coind_prefix coind ^ "induct")), - map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]); + map (Attrib.internal o K) (#2 induct)), [rulify lthy1 (#1 induct)]); val (eqs', lthy3) = lthy2 |> fold_map (fn (name, eq) => Local_Theory.note @@ -963,8 +962,8 @@ val eqs = if no_elim then [] else prove_eqs quiet_mode cs params intr_ts intrs elims lthy2 lthy1; - val elims' = map (fn (th, ns, i) => (rulify th, ns, i)) elims; - val intrs' = map rulify intrs; + val elims' = map (fn (th, ns, i) => (rulify lthy1 th, ns, i)) elims; + val intrs' = map (rulify lthy1) intrs; val (intrs'', elims'', eqs', induct, inducts, lthy3) = declare_rules rec_name coind no_ind @@ -974,7 +973,7 @@ {preds = preds, intrs = intrs'', elims = elims'', - raw_induct = rulify raw_induct, + raw_induct = rulify lthy3 raw_induct, induct = induct, inducts = inducts, eqs = eqs'}; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/inductive_realizer.ML Thu Apr 18 17:07:01 2013 +0200 @@ -455,13 +455,14 @@ val rlz' = attach_typeS (strip_all (Logic.unvarify_global rlz)); val rews = map mk_meta_eq case_thms; val thm = Goal.prove_global thy [] - (Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz') (fn {prems, ...} => EVERY - [cut_tac (hd prems) 1, - etac elimR 1, - ALLGOALS (asm_simp_tac HOL_basic_ss), - rewrite_goals_tac rews, - REPEAT ((resolve_tac prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac THEN' - DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE])) 1)]); + (Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz') + (fn {context = ctxt, prems, ...} => EVERY + [cut_tac (hd prems) 1, + etac elimR 1, + ALLGOALS (asm_simp_tac (put_simpset HOL_basic_ss ctxt)), + rewrite_goals_tac rews, + REPEAT ((resolve_tac prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac THEN' + DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE])) 1)]); val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_" (name_of_thm elim :: vs @ Ps @ ["correctness"])), thm) thy in diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/inductive_set.ML Thu Apr 18 17:07:01 2013 +0200 @@ -34,7 +34,7 @@ (**** simplify {(x1, ..., xn). (x1, ..., xn) : S} to S ****) val collect_mem_simproc = - Simplifier.simproc_global @{theory Set} "Collect_mem" ["Collect t"] (fn thy => fn ss => + Simplifier.simproc_global @{theory Set} "Collect_mem" ["Collect t"] (fn ctxt => fn S as Const (@{const_name Collect}, Type ("fun", [_, T])) $ t => let val (u, _, ps) = HOLogic.strip_psplits t in case u of @@ -45,10 +45,11 @@ if not (Term.is_open S') andalso ts = map Bound (length ps downto 0) then - let val simp = full_simp_tac (Simplifier.inherit_context ss - (HOL_basic_ss addsimps [@{thm split_paired_all}, @{thm split_conv}])) 1 + let val simp = + full_simp_tac (put_simpset HOL_basic_ss ctxt + addsimps [@{thm split_paired_all}, @{thm split_conv}]) 1 in - SOME (Goal.prove (Simplifier.the_context ss) [] [] + SOME (Goal.prove ctxt [] [] (Const ("==", T --> T --> propT) $ S $ S') (K (EVERY [rtac eq_reflection 1, rtac @{thm subset_antisym} 1, @@ -69,8 +70,9 @@ val anyt = Free ("t", TFree ("'t", [])); fun strong_ind_simproc tab = - Simplifier.simproc_global_i @{theory HOL} "strong_ind" [anyt] (fn thy => fn ss => fn t => + Simplifier.simproc_global_i @{theory HOL} "strong_ind" [anyt] (fn ctxt => fn t => let + val thy = Proof_Context.theory_of ctxt; fun close p t f = let val vs = Term.add_vars t [] in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs)) @@ -93,14 +95,15 @@ Type (_, [_, Type (_, [T, _])]))) $ p $ S)) = mkop s T (m, p, mk_collect p T (head_of u), S) | decomp _ = NONE; - val simp = full_simp_tac (Simplifier.inherit_context ss - (HOL_basic_ss addsimps [mem_Collect_eq, @{thm split_conv}])) 1; + val simp = + full_simp_tac + (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq, @{thm split_conv}]) 1; fun mk_rew t = (case strip_abs_vars t of [] => NONE | xs => (case decomp (strip_abs_body t) of NONE => NONE | SOME (bop, (m, p, S, S')) => - SOME (close (Goal.prove (Simplifier.the_context ss) [] []) + SOME (close (Goal.prove ctxt [] []) (Logic.mk_equals (t, fold_rev Term.abs xs (m $ p $ (bop $ S $ S')))) (K (EVERY [rtac eq_reflection 1, REPEAT (rtac ext 1), rtac iffI 1, @@ -239,12 +242,13 @@ (list_comb (x', map Bound (length Ts - 1 downto 0)))))) end) fs; -fun mk_to_pred_eq p fs optfs' T thm = +fun mk_to_pred_eq ctxt p fs optfs' T thm = let val thy = theory_of_thm thm; val insts = mk_to_pred_inst thy fs; val thm' = Thm.instantiate ([], insts) thm; - val thm'' = (case optfs' of + val thm'' = + (case optfs' of NONE => thm' RS sym | SOME fs' => let @@ -261,7 +265,7 @@ HOLogic.boolT (Bound 0))))] arg_cong' RS sym) end) in - Simplifier.simplify (HOL_basic_ss addsimps [mem_Collect_eq, @{thm split_conv}] + Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq, @{thm split_conv}] addsimprocs [collect_mem_simproc]) thm'' |> zero_var_indexes |> eta_contract_thm (equal p) end; @@ -278,6 +282,7 @@ @{typ bool} => let val thy = Context.theory_of context; + val ctxt = Context.proof_of context; fun factors_of t fs = case strip_abs_body t of Const (@{const_name Set.member}, _) $ u $ S => if is_Free S orelse is_Var S then @@ -305,7 +310,7 @@ else {to_set_simps = thm :: to_set_simps, to_pred_simps = - mk_to_pred_eq h fs fs' T' thm :: to_pred_simps, + mk_to_pred_eq ctxt h fs fs' T' thm :: to_pred_simps, set_arities = Symtab.insert_list op = (s', (T', (map (AList.lookup op = fs) ts', fs'))) set_arities, pred_arities = Symtab.insert_list op = (s, @@ -332,7 +337,7 @@ let val rules' = map mk_meta_eq rules in Simplifier.simproc_global_i @{theory HOL} "to_pred" [anyt] - (fn thy => K (lookup_rule thy (prop_of #> Logic.dest_equals) rules')) + (fn ctxt => (lookup_rule (Proof_Context.theory_of ctxt) (prop_of #> Logic.dest_equals) rules')) end; fun to_pred_proc thy rules t = case lookup_rule thy I rules t of @@ -341,11 +346,12 @@ SOME (Envir.subst_term (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs); -fun to_pred thms ctxt thm = +fun to_pred thms context thm = let - val thy = Context.theory_of ctxt; + val thy = Context.theory_of context; + val ctxt = Context.proof_of context; val {to_pred_simps, set_arities, pred_arities, ...} = - fold (add ctxt) thms (Data.get ctxt); + fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy set_arities (NONE, prop_of thm) []); (* instantiate each set parameter with {(x, y). p x y} *) @@ -353,7 +359,7 @@ in thm |> Thm.instantiate ([], insts) |> - Simplifier.full_simplify (HOL_basic_ss addsimprocs + Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimprocs [to_pred_simproc (mem_Collect_eq :: @{thm split_conv} :: to_pred_simps)]) |> eta_contract_thm (is_pred pred_arities) |> Rule_Cases.save thm @@ -364,11 +370,12 @@ (**** convert theorem in predicate notation to set notation ****) -fun to_set thms ctxt thm = +fun to_set thms context thm = let - val thy = Context.theory_of ctxt; + val thy = Context.theory_of context; + val ctxt = Context.proof_of context; val {to_set_simps, pred_arities, ...} = - fold (add ctxt) thms (Data.get ctxt); + fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy pred_arities (NONE, prop_of thm) []); (* instantiate each predicate parameter with %x y. (x, y) : s *) @@ -389,7 +396,7 @@ in thm |> Thm.instantiate ([], insts) |> - Simplifier.full_simplify (HOL_basic_ss addsimps to_set_simps + Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps to_set_simps addsimprocs [strong_ind_simproc pred_arities, collect_mem_simproc]) |> Rule_Cases.save thm end; @@ -410,7 +417,7 @@ forall is_none xs) arities) (prop_of thm) then thm |> - Simplifier.full_simplify (HOL_basic_ss addsimprocs + Simplifier.full_simplify (Simplifier.global_context thy HOL_basic_ss addsimprocs [to_pred_simproc (mem_Collect_eq :: @{thm split_conv} :: to_pred_simps)]) |> eta_contract_thm (is_pred pred_arities) else thm @@ -518,7 +525,7 @@ fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs U (map Bound (length fs downto 0)), list_comb (c, params)))))) - (K (REPEAT (rtac ext 1) THEN simp_tac (HOL_basic_ss addsimps + (K (REPEAT (rtac ext 1) THEN simp_tac (put_simpset HOL_basic_ss lthy addsimps [def, mem_Collect_eq, @{thm split_conv}]) 1)) in lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"), @@ -571,7 +578,7 @@ "convert rule to predicate notation" #> Attrib.setup @{binding mono_set} (Attrib.add_del mono_add mono_del) "declaration of monotonicity rule for set operators" #> - Simplifier.map_simpset_global (fn ss => ss addsimprocs [collect_mem_simproc]); + map_theory_simpset (fn ctxt => ctxt addsimprocs [collect_mem_simproc]); (* outer syntax *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/int_arith.ML --- a/src/HOL/Tools/int_arith.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/int_arith.ML Thu Apr 18 17:07:01 2013 +0200 @@ -24,7 +24,7 @@ val lhss0 = [@{cpat "0::?'a::ring"}]; -fun proc0 phi ss ct = +fun proc0 phi ctxt ct = let val T = ctyp_of_term ct in if typ_of T = @{typ int} then NONE else SOME (instantiate' [SOME T] [] zeroth) @@ -38,7 +38,7 @@ val lhss1 = [@{cpat "1::?'a::ring_1"}]; -fun proc1 phi ss ct = +fun proc1 phi ctxt ct = let val T = ctyp_of_term ct in if typ_of T = @{typ int} then NONE else SOME (instantiate' [SOME T] [] oneth) @@ -58,15 +58,17 @@ | check (a $ b) = check a andalso check b | check _ = false; -val conv = - Simplifier.rewrite - (HOL_basic_ss addsimps +val conv_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps ((map (fn th => th RS sym) [@{thm of_int_add}, @{thm of_int_mult}, @{thm of_int_diff}, @{thm of_int_minus}])@ [@{thm of_int_less_iff}, @{thm of_int_le_iff}, @{thm of_int_eq_iff}]) addsimprocs [zero_to_of_int_zero_simproc,one_to_of_int_one_simproc]); -fun sproc phi ss ct = if check (term_of ct) then SOME (conv ct) else NONE +fun sproc phi ctxt ct = + if check (term_of ct) then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct) + else NONE; val lhss' = [@{cpat "(?x::?'a::ring_char_0) = (?y::?'a)"}, diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/legacy_transfer.ML --- a/src/HOL/Tools/legacy_transfer.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/legacy_transfer.ML Thu Apr 18 17:07:01 2013 +0200 @@ -130,7 +130,7 @@ val (hyps, ctxt5) = ctxt4 |> Assumption.add_assumes (map transform c_vars'); val simpset = - Simplifier.context ctxt5 HOL_ss addsimps (inj @ embed @ return) + put_simpset HOL_ss ctxt5 addsimps (inj @ embed @ return) |> fold Simplifier.add_cong cong; val thm' = thm |> Drule.cterm_instantiate (c_vars ~~ c_exprs') diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/lin_arith.ML Thu Apr 18 17:07:01 2013 +0200 @@ -6,10 +6,10 @@ signature LIN_ARITH = sig - val pre_tac: simpset -> int -> tactic + val pre_tac: Proof.context -> int -> tactic val simple_tac: Proof.context -> int -> tactic val tac: Proof.context -> int -> tactic - val simproc: simpset -> term -> thm option + val simproc: Proof.context -> term -> thm option val add_inj_thms: thm list -> Context.generic -> Context.generic val add_lessD: thm -> Context.generic -> Context.generic val add_simps: thm list -> Context.generic -> Context.generic @@ -709,18 +709,17 @@ (* !split_limit splits are possible. *) local - val nnf_simpset = - (empty_ss + fun nnf_simpset ctxt = + (empty_simpset ctxt |> Simplifier.set_mkeqTrue mk_eq_True |> Simplifier.set_mksimps (mksimps mksimps_pairs)) addsimps [@{thm imp_conv_disj}, @{thm iff_conv_conj_imp}, @{thm de_Morgan_disj}, @{thm de_Morgan_conj}, not_all, not_ex, not_not] - fun prem_nnf_tac ss = full_simp_tac (Simplifier.inherit_context ss nnf_simpset) + fun prem_nnf_tac ctxt = full_simp_tac (nnf_simpset ctxt) in -fun split_once_tac ss split_thms = +fun split_once_tac ctxt split_thms = let - val ctxt = Simplifier.the_context ss val thy = Proof_Context.theory_of ctxt val cond_split_tac = SUBGOAL (fn (subgoal, i) => let @@ -743,7 +742,7 @@ REPEAT_DETERM o etac rev_mp, cond_split_tac, rtac ccontr, - prem_nnf_tac ss, + prem_nnf_tac ctxt, TRY o REPEAT_ALL_NEW (DETERM o (eresolve_tac [conjE, exE] ORELSE' etac disjE)) ] end; @@ -755,16 +754,15 @@ (* subgoals and finally attempt to solve them by finding an immediate *) (* contradiction (i.e., a term and its negation) in their premises. *) -fun pre_tac ss i = +fun pre_tac ctxt i = let - val ctxt = Simplifier.the_context ss; val split_thms = filter (is_split_thm ctxt) (#splits (get_arith_data ctxt)) fun is_relevant t = is_some (decomp ctxt t) in DETERM ( TRY (filter_prems_tac is_relevant i) THEN ( - (TRY o REPEAT_ALL_NEW (split_once_tac ss split_thms)) + (TRY o REPEAT_ALL_NEW (split_once_tac ctxt split_thms)) THEN_ALL_NEW (CONVERSION Drule.beta_eta_conversion THEN' @@ -801,7 +799,8 @@ inj_thms = inj_thms, lessD = lessD @ [@{thm "Suc_leI"}], neqE = [@{thm linorder_neqE_nat}, @{thm linorder_neqE_linordered_idom}], - simpset = HOL_basic_ss + simpset = + put_simpset HOL_basic_ss @{context} addsimps @{thms ring_distribs} addsimps [@{thm if_True}, @{thm if_False}] addsimps @@ -819,12 +818,14 @@ addsimprocs [@{simproc nateq_cancel_sums}, @{simproc natless_cancel_sums}, @{simproc natle_cancel_sums}] - |> Simplifier.add_cong @{thm if_weak_cong}, + |> Simplifier.add_cong @{thm if_weak_cong} + |> simpset_of, number_of = number_of}) #> add_discrete_type @{type_name nat}; -fun add_arith_facts ss = - Simplifier.add_prems (Arith_Data.get_arith_facts (Simplifier.the_context ss)) ss; +(* FIXME !?? *) +fun add_arith_facts ctxt = + Simplifier.add_prems (Arith_Data.get_arith_facts ctxt) ctxt; val simproc = add_arith_facts #> Fast_Arith.lin_arith_simproc; @@ -849,17 +850,16 @@ *) local - val nnf_simpset = - (empty_ss + fun nnf_simpset ctxt = + (empty_simpset ctxt |> Simplifier.set_mkeqTrue mk_eq_True |> Simplifier.set_mksimps (mksimps mksimps_pairs)) addsimps [@{thm imp_conv_disj}, @{thm iff_conv_conj_imp}, @{thm de_Morgan_disj}, @{thm de_Morgan_conj}, @{thm not_all}, @{thm not_ex}, @{thm not_not}]; - fun prem_nnf_tac i st = - full_simp_tac (Simplifier.global_context (Thm.theory_of_thm st) nnf_simpset) i st; + fun prem_nnf_tac ctxt = full_simp_tac (nnf_simpset ctxt); in -fun refute_tac test prep_tac ref_tac = +fun refute_tac ctxt test prep_tac ref_tac = let val refute_prems_tac = REPEAT_DETERM (eresolve_tac [@{thm conjE}, @{thm exE}] 1 ORELSE @@ -868,7 +868,7 @@ (DETERM (etac @{thm notE} 1 THEN eq_assume_tac 1) ORELSE ref_tac 1); in EVERY'[TRY o filter_prems_tac test, - REPEAT_DETERM o etac @{thm rev_mp}, prep_tac, rtac @{thm ccontr}, prem_nnf_tac, + REPEAT_DETERM o etac @{thm rev_mp}, prep_tac, rtac @{thm ccontr}, prem_nnf_tac ctxt, SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] end; @@ -887,7 +887,7 @@ (max m n + k <= r) = (m+k <= r & n+k <= r) (l <= min m n + k) = (l <= m+k & l <= n+k) *) - refute_tac (K true) + refute_tac ctxt (K true) (* Splitting is also done inside simple_tac, but not completely -- *) (* split_tac may use split theorems that have not been implemented in *) (* simple_tac (cf. pre_decomp and split_once_items above), and *) @@ -910,11 +910,11 @@ (* context setup *) val setup = - init_arith_data #> - Simplifier.map_ss (fn ss => ss - addSolver (mk_solver "lin_arith" (add_arith_facts #> Fast_Arith.prems_lin_arith_tac))); + init_arith_data; val global_setup = + map_theory_simpset (fn ctxt => ctxt + addSolver (mk_solver "lin_arith" (add_arith_facts #> Fast_Arith.prems_lin_arith_tac))) #> Attrib.setup @{binding arith_split} (Scan.succeed (Thm.declaration_attribute add_split)) "declaration of split rules for arithmetic procedure" #> Method.setup @{binding linarith} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/nat_numeral_simprocs.ML --- a/src/HOL/Tools/nat_numeral_simprocs.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/nat_numeral_simprocs.ML Thu Apr 18 17:07:01 2013 +0200 @@ -5,21 +5,21 @@ signature NAT_NUMERAL_SIMPROCS = sig - val combine_numerals: simpset -> cterm -> thm option - val eq_cancel_numerals: simpset -> cterm -> thm option - val less_cancel_numerals: simpset -> cterm -> thm option - val le_cancel_numerals: simpset -> cterm -> thm option - val diff_cancel_numerals: simpset -> cterm -> thm option - val eq_cancel_numeral_factor: simpset -> cterm -> thm option - val less_cancel_numeral_factor: simpset -> cterm -> thm option - val le_cancel_numeral_factor: simpset -> cterm -> thm option - val div_cancel_numeral_factor: simpset -> cterm -> thm option - val dvd_cancel_numeral_factor: simpset -> cterm -> thm option - val eq_cancel_factor: simpset -> cterm -> thm option - val less_cancel_factor: simpset -> cterm -> thm option - val le_cancel_factor: simpset -> cterm -> thm option - val div_cancel_factor: simpset -> cterm -> thm option - val dvd_cancel_factor: simpset -> cterm -> thm option + val combine_numerals: Proof.context -> cterm -> thm option + val eq_cancel_numerals: Proof.context -> cterm -> thm option + val less_cancel_numerals: Proof.context -> cterm -> thm option + val le_cancel_numerals: Proof.context -> cterm -> thm option + val diff_cancel_numerals: Proof.context -> cterm -> thm option + val eq_cancel_numeral_factor: Proof.context -> cterm -> thm option + val less_cancel_numeral_factor: Proof.context -> cterm -> thm option + val le_cancel_numeral_factor: Proof.context -> cterm -> thm option + val div_cancel_numeral_factor: Proof.context -> cterm -> thm option + val dvd_cancel_numeral_factor: Proof.context -> cterm -> thm option + val eq_cancel_factor: Proof.context -> cterm -> thm option + val less_cancel_factor: Proof.context -> cterm -> thm option + val le_cancel_factor: Proof.context -> cterm -> thm option + val div_cancel_factor: Proof.context -> cterm -> thm option + val dvd_cancel_factor: Proof.context -> cterm -> thm option end; structure Nat_Numeral_Simprocs : NAT_NUMERAL_SIMPROCS = @@ -27,9 +27,8 @@ (*Maps n to #n for n = 1, 2*) val numeral_syms = [@{thm numeral_1_eq_1} RS sym, @{thm numeral_2_eq_2} RS sym]; -val numeral_sym_ss = HOL_basic_ss addsimps numeral_syms; - -val rename_numerals = simplify numeral_sym_ss o Thm.transfer @{theory}; +val numeral_sym_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps numeral_syms); (*Utilities*) @@ -134,6 +133,9 @@ end; +(* FIXME !? *) +val rename_numerals = simplify (put_simpset numeral_sym_ss @{context}) o Thm.transfer @{theory}; + (*Simplify 1*n and n*1 to n*) val add_0s = map rename_numerals [@{thm Nat.add_0}, @{thm Nat.add_0_right}]; val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}]; @@ -162,15 +164,22 @@ val find_first_coeff = find_first_coeff [] val trans_tac = Numeral_Simprocs.trans_tac - val norm_ss1 = Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @ - [@{thm Suc_eq_plus1_left}] @ @{thms add_ac} - val norm_ss2 = Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + val norm_ss1 = + simpset_of (put_simpset Numeral_Simprocs.num_ss @{context} + addsimps numeral_syms @ add_0s @ mult_1s @ + [@{thm Suc_eq_plus1_left}] @ @{thms add_ac}) + val norm_ss2 = + simpset_of (put_simpset Numeral_Simprocs.num_ss @{context} + addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac}) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) - val numeral_simp_ss = HOL_basic_ss addsimps add_0s @ bin_simps; - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)); + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps add_0s @ bin_simps); + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)); val simplify_meta_eq = simplify_meta_eq val prove_conv = Arith_Data.prove_conv end; @@ -207,10 +216,10 @@ val bal_add2 = @{thm nat_diff_add_eq2} RS trans ); -fun eq_cancel_numerals ss ct = EqCancelNumerals.proc ss (term_of ct) -fun less_cancel_numerals ss ct = LessCancelNumerals.proc ss (term_of ct) -fun le_cancel_numerals ss ct = LeCancelNumerals.proc ss (term_of ct) -fun diff_cancel_numerals ss ct = DiffCancelNumerals.proc ss (term_of ct) +fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (term_of ct) +fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (term_of ct) +fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (term_of ct) +fun diff_cancel_numerals ctxt ct = DiffCancelNumerals.proc ctxt (term_of ct) (*** Applying CombineNumeralsFun ***) @@ -228,20 +237,27 @@ val prove_conv = Arith_Data.prove_conv_nohyps val trans_tac = Numeral_Simprocs.trans_tac - val norm_ss1 = Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_plus1}] @ @{thms add_ac} - val norm_ss2 = Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + val norm_ss1 = + simpset_of (put_simpset Numeral_Simprocs.num_ss @{context} + addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_plus1}] @ @{thms add_ac}) + val norm_ss2 = + simpset_of (put_simpset Numeral_Simprocs.num_ss @{context} + addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac}) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) - val numeral_simp_ss = HOL_basic_ss addsimps add_0s @ bin_simps; - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps add_0s @ bin_simps); + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = simplify_meta_eq end; structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); -fun combine_numerals ss ct = CombineNumerals.proc ss (term_of ct) +fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (term_of ct) (*** Applying CancelNumeralFactorFun ***) @@ -252,15 +268,20 @@ val dest_coeff = dest_coeff val trans_tac = Numeral_Simprocs.trans_tac - val norm_ss1 = Numeral_Simprocs.num_ss addsimps - numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_plus1_left}] @ @{thms add_ac} - val norm_ss2 = Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) + val norm_ss1 = + simpset_of (put_simpset Numeral_Simprocs.num_ss @{context} + addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_plus1_left}] @ @{thms add_ac}) + val norm_ss2 = + simpset_of (put_simpset Numeral_Simprocs.num_ss @{context} + addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac}) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) - val numeral_simp_ss = HOL_basic_ss addsimps bin_simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps bin_simps) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = simplify_meta_eq val prove_conv = Arith_Data.prove_conv end; @@ -305,11 +326,11 @@ val neg_exchanges = true ) -fun eq_cancel_numeral_factor ss ct = EqCancelNumeralFactor.proc ss (term_of ct) -fun less_cancel_numeral_factor ss ct = LessCancelNumeralFactor.proc ss (term_of ct) -fun le_cancel_numeral_factor ss ct = LeCancelNumeralFactor.proc ss (term_of ct) -fun div_cancel_numeral_factor ss ct = DivCancelNumeralFactor.proc ss (term_of ct) -fun dvd_cancel_numeral_factor ss ct = DvdCancelNumeralFactor.proc ss (term_of ct) +fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (term_of ct) +fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (term_of ct) +fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (term_of ct) +fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (term_of ct) +fun dvd_cancel_numeral_factor ctxt ct = DvdCancelNumeralFactor.proc ctxt (term_of ct) (*** Applying ExtractCommonTermFun ***) @@ -329,8 +350,8 @@ val simplify_one = Arith_Data.simplify_meta_eq [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}]; -fun cancel_simplify_meta_eq ss cancel_th th = - simplify_one ss (([th, cancel_th]) MRS trans); +fun cancel_simplify_meta_eq ctxt cancel_th th = + simplify_one ctxt (([th, cancel_th]) MRS trans); structure CancelFactorCommon = struct @@ -340,8 +361,11 @@ val dest_coeff = dest_coeff val find_first = find_first_t [] val trans_tac = Numeral_Simprocs.trans_tac - val norm_ss = HOL_basic_ss addsimps mult_1s @ @{thms mult_ac} - fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) + val norm_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps mult_1s @ @{thms mult_ac}) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) val simplify_meta_eq = cancel_simplify_meta_eq fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end; @@ -381,10 +405,10 @@ fun simp_conv _ _ = SOME @{thm nat_mult_dvd_cancel_disj} ); -fun eq_cancel_factor ss ct = EqCancelFactor.proc ss (term_of ct) -fun less_cancel_factor ss ct = LessCancelFactor.proc ss (term_of ct) -fun le_cancel_factor ss ct = LeCancelFactor.proc ss (term_of ct) -fun div_cancel_factor ss ct = DivideCancelFactor.proc ss (term_of ct) -fun dvd_cancel_factor ss ct = DvdCancelFactor.proc ss (term_of ct) +fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (term_of ct) +fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (term_of ct) +fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (term_of ct) +fun div_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (term_of ct) +fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (term_of ct) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/numeral_simprocs.ML Thu Apr 18 17:07:01 2013 +0200 @@ -16,30 +16,30 @@ signature NUMERAL_SIMPROCS = sig - val prep_simproc: theory -> string * string list * (theory -> simpset -> term -> thm option) + val prep_simproc: theory -> string * string list * (Proof.context -> term -> thm option) -> simproc val trans_tac: thm option -> tactic - val assoc_fold: simpset -> cterm -> thm option - val combine_numerals: simpset -> cterm -> thm option - val eq_cancel_numerals: simpset -> cterm -> thm option - val less_cancel_numerals: simpset -> cterm -> thm option - val le_cancel_numerals: simpset -> cterm -> thm option - val eq_cancel_factor: simpset -> cterm -> thm option - val le_cancel_factor: simpset -> cterm -> thm option - val less_cancel_factor: simpset -> cterm -> thm option - val div_cancel_factor: simpset -> cterm -> thm option - val mod_cancel_factor: simpset -> cterm -> thm option - val dvd_cancel_factor: simpset -> cterm -> thm option - val divide_cancel_factor: simpset -> cterm -> thm option - val eq_cancel_numeral_factor: simpset -> cterm -> thm option - val less_cancel_numeral_factor: simpset -> cterm -> thm option - val le_cancel_numeral_factor: simpset -> cterm -> thm option - val div_cancel_numeral_factor: simpset -> cterm -> thm option - val divide_cancel_numeral_factor: simpset -> cterm -> thm option - val field_combine_numerals: simpset -> cterm -> thm option + val assoc_fold: Proof.context -> cterm -> thm option + val combine_numerals: Proof.context -> cterm -> thm option + val eq_cancel_numerals: Proof.context -> cterm -> thm option + val less_cancel_numerals: Proof.context -> cterm -> thm option + val le_cancel_numerals: Proof.context -> cterm -> thm option + val eq_cancel_factor: Proof.context -> cterm -> thm option + val le_cancel_factor: Proof.context -> cterm -> thm option + val less_cancel_factor: Proof.context -> cterm -> thm option + val div_cancel_factor: Proof.context -> cterm -> thm option + val mod_cancel_factor: Proof.context -> cterm -> thm option + val dvd_cancel_factor: Proof.context -> cterm -> thm option + val divide_cancel_factor: Proof.context -> cterm -> thm option + val eq_cancel_numeral_factor: Proof.context -> cterm -> thm option + val less_cancel_numeral_factor: Proof.context -> cterm -> thm option + val le_cancel_numeral_factor: Proof.context -> cterm -> thm option + val div_cancel_numeral_factor: Proof.context -> cterm -> thm option + val divide_cancel_numeral_factor: Proof.context -> cterm -> thm option + val field_combine_numerals: Proof.context -> cterm -> thm option val field_cancel_numeral_factors: simproc list val num_ss: simpset - val field_comp_conv: conv + val field_comp_conv: Proof.context -> conv end; structure Numeral_Simprocs : NUMERAL_SIMPROCS = @@ -172,7 +172,9 @@ fun numtermless tu = (numterm_ord tu = LESS); -val num_ss = HOL_basic_ss |> Simplifier.set_termless numtermless; +val num_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + |> Simplifier.set_termless numtermless); (*Maps 1 to Numeral1 so that arithmetic isn't complicated by the abstract 1.*) val numeral_syms = [@{thm numeral_1_eq_1} RS sym]; @@ -234,10 +236,18 @@ val mult_minus_simps = [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2]; -val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ - diff_simps @ minus_simps @ @{thms add_ac} -val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps -val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac} +val norm_ss1 = + simpset_of (put_simpset num_ss @{context} + addsimps numeral_syms @ add_0s @ mult_1s @ + diff_simps @ minus_simps @ @{thms add_ac}) + +val norm_ss2 = + simpset_of (put_simpset num_ss @{context} + addsimps non_add_simps @ mult_minus_simps) + +val norm_ss3 = + simpset_of (put_simpset num_ss @{context} + addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac}) structure CancelNumeralsCommon = struct @@ -248,13 +258,15 @@ val find_first_coeff = find_first_coeff [] val trans_tac = trans_tac - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) - val numeral_simp_ss = HOL_basic_ss addsimps add_0s @ simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps add_0s @ simps) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq post_simps val prove_conv = Arith_Data.prove_conv end; @@ -283,9 +295,9 @@ val bal_add2 = @{thm le_add_iff2} RS trans ); -fun eq_cancel_numerals ss ct = EqCancelNumerals.proc ss (term_of ct) -fun less_cancel_numerals ss ct = LessCancelNumerals.proc ss (term_of ct) -fun le_cancel_numerals ss ct = LeCancelNumerals.proc ss (term_of ct) +fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (term_of ct) +fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (term_of ct) +fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (term_of ct) structure CombineNumeralsData = struct @@ -300,13 +312,15 @@ val prove_conv = Arith_Data.prove_conv_nohyps val trans_tac = trans_tac - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) - val numeral_simp_ss = HOL_basic_ss addsimps add_0s @ simps - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps add_0s @ simps) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq post_simps end; @@ -326,22 +340,27 @@ val prove_conv = Arith_Data.prove_conv_nohyps val trans_tac = trans_tac - val norm_ss1a = norm_ss1 addsimps inverse_1s @ divide_simps - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1a)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + val norm_ss1a = + simpset_of (put_simpset norm_ss1 @{context} addsimps inverse_1s @ divide_simps) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1a ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) - val numeral_simp_ss = HOL_basic_ss addsimps add_0s @ simps @ [@{thm add_frac_eq}, @{thm not_False_eq_True}] - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps add_0s @ simps @ [@{thm add_frac_eq}, @{thm not_False_eq_True}]) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq field_post_simps end; structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData); -fun combine_numerals ss ct = CombineNumerals.proc ss (term_of ct) +fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (term_of ct) -fun field_combine_numerals ss ct = FieldCombineNumerals.proc ss (term_of ct) +fun field_combine_numerals ctxt ct = FieldCombineNumerals.proc ctxt (term_of ct) + (** Constant folding for multiplication in semirings **) @@ -349,14 +368,14 @@ structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = struct - val assoc_ss = HOL_basic_ss addsimps @{thms mult_ac} + val assoc_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac}) val eq_reflection = eq_reflection val is_numeral = can HOLogic.dest_number end; structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); -fun assoc_fold ss ct = Semiring_Times_Assoc.proc ss (term_of ct) +fun assoc_fold ctxt ct = Semiring_Times_Assoc.proc ctxt (term_of ct) structure CancelNumeralFactorCommon = struct @@ -364,18 +383,23 @@ val dest_coeff = dest_coeff 1 val trans_tac = trans_tac - val norm_ss1 = HOL_basic_ss addsimps minus_from_mult_simps @ mult_1s - val norm_ss2 = HOL_basic_ss addsimps simps @ mult_minus_simps - val norm_ss3 = HOL_basic_ss addsimps @{thms mult_ac} - fun norm_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) + val norm_ss1 = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps minus_from_mult_simps @ mult_1s) + val norm_ss2 = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps simps @ mult_minus_simps) + val norm_ss3 = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac}) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) + THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) (* simp_thms are necessary because some of the cancellation rules below (e.g. mult_less_cancel_left) introduce various logical connectives *) - val numeral_simp_ss = HOL_basic_ss addsimps simps @ @{thms simp_thms} - fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps simps @ @{thms simp_thms}) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq ([@{thm Nat.add_0}, @{thm Nat.add_0_right}] @ post_simps) val prove_conv = Arith_Data.prove_conv @@ -423,18 +447,18 @@ val neg_exchanges = true ) -fun eq_cancel_numeral_factor ss ct = EqCancelNumeralFactor.proc ss (term_of ct) -fun less_cancel_numeral_factor ss ct = LessCancelNumeralFactor.proc ss (term_of ct) -fun le_cancel_numeral_factor ss ct = LeCancelNumeralFactor.proc ss (term_of ct) -fun div_cancel_numeral_factor ss ct = DivCancelNumeralFactor.proc ss (term_of ct) -fun divide_cancel_numeral_factor ss ct = DivideCancelNumeralFactor.proc ss (term_of ct) +fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (term_of ct) +fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (term_of ct) +fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (term_of ct) +fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (term_of ct) +fun divide_cancel_numeral_factor ctxt ct = DivideCancelNumeralFactor.proc ctxt (term_of ct) val field_cancel_numeral_factors = map (prep_simproc @{theory}) [("field_eq_cancel_numeral_factor", ["(l::'a::field) * m = n", "(l::'a::field) = m * n"], - K EqCancelNumeralFactor.proc), + EqCancelNumeralFactor.proc), ("field_cancel_numeral_factor", ["((l::'a::field_inverse_zero) * m) / n", "(l::'a::field_inverse_zero) / (m * n)", @@ -442,7 +466,7 @@ "((numeral v)::'a::field_inverse_zero) / (neg_numeral w)", "((neg_numeral v)::'a::field_inverse_zero) / (numeral w)", "((neg_numeral v)::'a::field_inverse_zero) / (neg_numeral w)"], - K DivideCancelNumeralFactor.proc)] + DivideCancelNumeralFactor.proc)] (** Declarations for ExtractCommonTerm **) @@ -458,22 +482,22 @@ val simplify_one = Arith_Data.simplify_meta_eq [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_1_eq_1}]; -fun cancel_simplify_meta_eq ss cancel_th th = - simplify_one ss (([th, cancel_th]) MRS trans); +fun cancel_simplify_meta_eq ctxt cancel_th th = + simplify_one ctxt (([th, cancel_th]) MRS trans); local val Tp_Eq = Thm.reflexive (Thm.cterm_of @{theory HOL} HOLogic.Trueprop) fun Eq_True_elim Eq = Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI} in -fun sign_conv pos_th neg_th ss t = +fun sign_conv pos_th neg_th ctxt t = let val T = fastype_of t; val zero = Const(@{const_name Groups.zero}, T); val less = Const(@{const_name Orderings.less}, [T,T] ---> HOLogic.boolT); val pos = less $ zero $ t and neg = less $ t $ zero - val thy = Proof_Context.theory_of (Simplifier.the_context ss) + val thy = Proof_Context.theory_of ctxt fun prove p = - SOME (Eq_True_elim (Simplifier.asm_rewrite ss (Thm.cterm_of thy p))) + SOME (Eq_True_elim (Simplifier.asm_rewrite ctxt (Thm.cterm_of thy p))) handle THM _ => NONE in case prove pos of SOME th => SOME(th RS pos_th) @@ -491,8 +515,10 @@ val dest_coeff = dest_coeff val find_first = find_first_t [] val trans_tac = trans_tac - val norm_ss = HOL_basic_ss addsimps mult_1s @ @{thms mult_ac} - fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) + val norm_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps mult_1s @ @{thms mult_ac}) + fun norm_tac ctxt = + ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) val simplify_meta_eq = cancel_simplify_meta_eq fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end; @@ -554,13 +580,13 @@ fun simp_conv _ _ = SOME @{thm mult_divide_mult_cancel_left_if} ); -fun eq_cancel_factor ss ct = EqCancelFactor.proc ss (term_of ct) -fun le_cancel_factor ss ct = LeCancelFactor.proc ss (term_of ct) -fun less_cancel_factor ss ct = LessCancelFactor.proc ss (term_of ct) -fun div_cancel_factor ss ct = DivCancelFactor.proc ss (term_of ct) -fun mod_cancel_factor ss ct = ModCancelFactor.proc ss (term_of ct) -fun dvd_cancel_factor ss ct = DvdCancelFactor.proc ss (term_of ct) -fun divide_cancel_factor ss ct = DivideCancelFactor.proc ss (term_of ct) +fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (term_of ct) +fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (term_of ct) +fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (term_of ct) +fun div_cancel_factor ctxt ct = DivCancelFactor.proc ctxt (term_of ct) +fun mod_cancel_factor ctxt ct = ModCancelFactor.proc ctxt (term_of ct) +fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (term_of ct) +fun divide_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (term_of ct) local val zr = @{cpat "0"} @@ -571,29 +597,29 @@ val add_frac_num = mk_meta_eq @{thm "add_frac_num"} val add_num_frac = mk_meta_eq @{thm "add_num_frac"} - fun prove_nz ss T t = + fun prove_nz ctxt T t = let val z = Thm.instantiate_cterm ([(zT,T)],[]) zr val eq = Thm.instantiate_cterm ([(eqT,T)],[]) geq - val th = Simplifier.rewrite (ss addsimps @{thms simp_thms}) + val th = Simplifier.rewrite (ctxt addsimps @{thms simp_thms}) (Thm.apply @{cterm "Trueprop"} (Thm.apply @{cterm "Not"} (Thm.apply (Thm.apply eq t) z))) in Thm.equal_elim (Thm.symmetric th) TrueI end - fun proc phi ss ct = + fun proc phi ctxt ct = let val ((x,y),(w,z)) = (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct val _ = map (HOLogic.dest_number o term_of) [x,y,z,w] val T = ctyp_of_term x - val [y_nz, z_nz] = map (prove_nz ss T) [y, z] + val [y_nz, z_nz] = map (prove_nz ctxt T) [y, z] val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq in SOME (Thm.implies_elim (Thm.implies_elim th y_nz) z_nz) end handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE - fun proc2 phi ss ct = + fun proc2 phi ctxt ct = let val (l,r) = Thm.dest_binop ct val T = ctyp_of_term l @@ -601,13 +627,13 @@ (Const(@{const_name Fields.divide},_)$_$_, _) => let val (x,y) = Thm.dest_binop l val z = r val _ = map (HOLogic.dest_number o term_of) [x,y,z] - val ynz = prove_nz ss T y + val ynz = prove_nz ctxt T y in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz) end | (_, Const (@{const_name Fields.divide},_)$_$_) => let val (x,y) = Thm.dest_binop r val z = l val _ = map (HOLogic.dest_number o term_of) [x,y,z] - val ynz = prove_nz ss T y + val ynz = prove_nz ctxt T y in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz) end | _ => NONE) @@ -619,7 +645,7 @@ val is_number = is_number o term_of - fun proc3 phi ss ct = + fun proc3 phi ctxt ct = (case term_of ct of Const(@{const_name Orderings.less},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ => let @@ -699,17 +725,21 @@ Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute})))) (@{thm field_divide_inverse} RS sym)] -in - -val field_comp_conv = - Simplifier.rewrite - (HOL_basic_ss addsimps @{thms "semiring_norm"} +val field_comp_ss = + simpset_of + (put_simpset HOL_basic_ss @{context} + addsimps @{thms "semiring_norm"} addsimps ths addsimps @{thms simp_thms} addsimprocs field_cancel_numeral_factors addsimprocs [add_frac_frac_simproc, add_frac_num_simproc, ord_frac_simproc] |> Simplifier.add_cong @{thm "if_weak_cong"}) + +in + +fun field_comp_conv ctxt = + Simplifier.rewrite (put_simpset field_comp_ss ctxt) then_conv - Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}]) + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_1_eq_1}]) end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/recdef.ML --- a/src/HOL/Tools/recdef.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/recdef.ML Thu Apr 18 17:07:01 2013 +0200 @@ -166,16 +166,14 @@ NONE => ctxt0 | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0)); val {simps, congs, wfs} = get_hints ctxt; - val ctxt' = ctxt - |> Simplifier.map_simpset (fn ss => ss addsimps simps |> Simplifier.del_cong @{thm imp_cong}); + val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong}; in (ctxt', rev (map snd congs), wfs) end; fun prepare_hints_i thy () = let val ctxt = Proof_Context.init_global thy; val {simps, congs, wfs} = get_global_hints thy; - val ctxt' = ctxt - |> Simplifier.map_simpset (fn ss => ss addsimps simps |> Simplifier.del_cong @{thm imp_cong}); + val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong}; in (ctxt', rev (map snd congs), wfs) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/record.ML Thu Apr 18 17:07:01 2013 +0200 @@ -42,8 +42,8 @@ val upd_simproc: simproc val split_simproc: (term -> int) -> simproc val ex_sel_eq_simproc: simproc - val split_tac: int -> tactic - val split_simp_tac: thm list -> (term -> int) -> int -> tactic + val split_tac: Proof.context -> int -> tactic + val split_simp_tac: Proof.context -> thm list -> (term -> int) -> int -> tactic val split_wrapper: string * (Proof.context -> wrapper) val codegen: bool Config.T @@ -459,10 +459,9 @@ val is_selector = Symtab.defined o #selectors o get_sel_upd; val get_updates = Symtab.lookup o #updates o get_sel_upd; -fun get_ss_with_context getss thy = Simplifier.global_context thy (getss (get_sel_upd thy)); -val get_simpset = get_ss_with_context #simpset; -val get_sel_upd_defs = get_ss_with_context #defset; +val get_simpset = #simpset o get_sel_upd; +val get_sel_upd_defs = #defset o get_sel_upd; fun get_update_details u thy = let val sel_upd = get_sel_upd thy in @@ -475,6 +474,8 @@ fun put_sel_upd names more depth simps defs thy = let + val ctxt0 = Proof_Context.init_global thy; + val all = names @ [more]; val sels = map (rpair (depth, false)) names @ [(more, (depth, true))]; val upds = map (suffix updateN) all ~~ all; @@ -485,8 +486,8 @@ make_data records {selectors = fold Symtab.update_new sels selectors, updates = fold Symtab.update_new upds updates, - simpset = simpset addsimps simps, - defset = defset addsimps defs} + simpset = simpset_map ctxt0 (fn ctxt => ctxt addsimps simps) simpset, + defset = simpset_map ctxt0 (fn ctxt => ctxt addsimps defs) defset} equalities extinjects extsplit splits extfields fieldext; in Data.put data thy end; @@ -966,10 +967,10 @@ val prop = lhs === rhs; val othm = Goal.prove (Proof_Context.init_global thy) [] [] prop - (fn _ => - simp_tac defset 1 THEN + (fn {context = ctxt, ...} => + simp_tac (put_simpset defset ctxt) 1 THEN REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN - TRY (simp_tac (HOL_ss addsimps @{thms id_apply id_o o_id}) 1)); + TRY (simp_tac (put_simpset HOL_ss ctxt addsimps @{thms id_apply id_o o_id}) 1)); val dest = if is_sel_upd_pair thy acc upd then @{thm o_eq_dest} @@ -990,10 +991,10 @@ val prop = lhs === rhs; val othm = Goal.prove (Proof_Context.init_global thy) [] [] prop - (fn _ => - simp_tac defset 1 THEN + (fn {context = ctxt, ...} => + simp_tac (put_simpset defset ctxt) 1 THEN REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN - TRY (simp_tac (HOL_ss addsimps @{thms id_apply}) 1)); + TRY (simp_tac (put_simpset HOL_ss ctxt addsimps @{thms id_apply}) 1)); val dest = if comp then @{thm o_eq_dest_lhs} else @{thm o_eq_dest}; in Drule.export_without_context (othm RS dest) end; @@ -1031,9 +1032,9 @@ val simps = (if length args = 1 then get_accupd_simps else get_updupd_simps) thy lhs defset; in Goal.prove (Proof_Context.init_global thy) [] [] prop' - (fn _ => - simp_tac (HOL_basic_ss addsimps (simps @ @{thms K_record_comp})) 1 THEN - TRY (simp_tac (HOL_basic_ss addsimps ex_simps addsimprocs ex_simprs) 1)) + (fn {context = ctxt, ...} => + simp_tac (put_simpset HOL_basic_ss ctxt addsimps (simps @ @{thms K_record_comp})) 1 THEN + TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps ex_simps addsimprocs ex_simprs) 1)) end; @@ -1068,63 +1069,65 @@ *) val simproc = Simplifier.simproc_global @{theory HOL} "record_simp" ["x"] - (fn thy => fn _ => fn t => - (case t of - (sel as Const (s, Type (_, [_, rangeS]))) $ - ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) => - if is_selector thy s andalso is_some (get_updates thy u) then - let - val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy; + (fn ctxt => fn t => + let val thy = Proof_Context.theory_of ctxt in + (case t of + (sel as Const (s, Type (_, [_, rangeS]))) $ + ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) => + if is_selector thy s andalso is_some (get_updates thy u) then + let + val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy; - fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) = - (case Symtab.lookup updates u of - NONE => NONE - | SOME u_name => - if u_name = s then - (case mk_eq_terms r of - NONE => - let - val rv = ("r", rT); - val rb = Bound 0; - val (kv, kb) = K_skeleton "k" kT (Bound 1) k; - in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end - | SOME (trm, trm', vars) => - let - val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k; - in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end) - else if has_field extfields u_name rangeS orelse - has_field extfields s (domain_type kT) then NONE - else - (case mk_eq_terms r of - SOME (trm, trm', vars) => - let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k - in SOME (upd $ kb $ trm, trm', kv :: vars) end - | NONE => - let - val rv = ("r", rT); - val rb = Bound 0; - val (kv, kb) = K_skeleton "k" kT (Bound 1) k; - in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end)) - | mk_eq_terms _ = NONE; - in - (case mk_eq_terms (upd $ k $ r) of - SOME (trm, trm', vars) => - SOME - (prove_unfold_defs thy [] [] - (Logic.list_all (vars, Logic.mk_equals (sel $ trm, trm')))) - | NONE => NONE) - end - else NONE - | _ => NONE)); + fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) = + (case Symtab.lookup updates u of + NONE => NONE + | SOME u_name => + if u_name = s then + (case mk_eq_terms r of + NONE => + let + val rv = ("r", rT); + val rb = Bound 0; + val (kv, kb) = K_skeleton "k" kT (Bound 1) k; + in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end + | SOME (trm, trm', vars) => + let + val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k; + in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end) + else if has_field extfields u_name rangeS orelse + has_field extfields s (domain_type kT) then NONE + else + (case mk_eq_terms r of + SOME (trm, trm', vars) => + let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k + in SOME (upd $ kb $ trm, trm', kv :: vars) end + | NONE => + let + val rv = ("r", rT); + val rb = Bound 0; + val (kv, kb) = K_skeleton "k" kT (Bound 1) k; + in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end)) + | mk_eq_terms _ = NONE; + in + (case mk_eq_terms (upd $ k $ r) of + SOME (trm, trm', vars) => + SOME + (prove_unfold_defs thy [] [] + (Logic.list_all (vars, Logic.mk_equals (sel $ trm, trm')))) + | NONE => NONE) + end + else NONE + | _ => NONE) + end); -fun get_upd_acc_cong_thm upd acc thy simpset = +fun get_upd_acc_cong_thm upd acc thy ss = let val insts = [("upd", cterm_of thy upd), ("ac", cterm_of thy acc)]; val prop = Thm.concl_of (named_cterm_instantiate insts updacc_cong_triv); in Goal.prove (Proof_Context.init_global thy) [] [] prop - (fn _ => - simp_tac simpset 1 THEN + (fn {context = ctxt, ...} => + simp_tac (put_simpset ss ctxt) 1 THEN REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN TRY (resolve_tac [updacc_cong_idI] 1)) end; @@ -1142,8 +1145,9 @@ both a more update and an update to a field within it.*) val upd_simproc = Simplifier.simproc_global @{theory HOL} "record_upd_simp" ["x"] - (fn thy => fn _ => fn t => + (fn ctxt => fn t => let + val thy = Proof_Context.theory_of ctxt; (*We can use more-updators with other updators as long as none of the other updators go deeper than any more updator. min here is the depth of the deepest other @@ -1262,12 +1266,12 @@ *) val eq_simproc = Simplifier.simproc_global @{theory HOL} "record_eq_simp" ["r = s"] - (fn thy => fn _ => fn t => + (fn ctxt => fn t => (case t of Const (@{const_name HOL.eq}, Type (_, [T, _])) $ _ $ _ => (case rec_id ~1 T of "" => NONE | name => - (case get_equalities thy name of + (case get_equalities (Proof_Context.theory_of ctxt) name of NONE => NONE | SOME thm => SOME (thm RS @{thm Eq_TrueI}))) | _ => NONE)); @@ -1282,7 +1286,7 @@ P t > 0: split up to given bound of record extensions.*) fun split_simproc P = Simplifier.simproc_global @{theory HOL} "record_split_simp" ["x"] - (fn thy => fn _ => fn t => + (fn ctxt => fn t => (case t of Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ => if quantifier = @{const_name all} orelse @@ -1294,7 +1298,7 @@ | _ => let val split = P t in if split <> 0 then - (case get_splits thy (rec_id split T) of + (case get_splits (Proof_Context.theory_of ctxt) (rec_id split T) of NONE => NONE | SOME (all_thm, All_thm, Ex_thm, _) => SOME @@ -1310,8 +1314,9 @@ val ex_sel_eq_simproc = Simplifier.simproc_global @{theory HOL} "ex_sel_eq_simproc" ["Ex t"] - (fn thy => fn ss => fn t => + (fn ctxt => fn t => let + val thy = Proof_Context.theory_of ctxt; fun mkeq (lr, Teq, (sel, Tsel), x) i = if is_selector thy sel then let @@ -1340,7 +1345,7 @@ (Const (@{const_name Ex}, Tex) $ Abs (s, T, eq), @{term True})); in SOME (Goal.prove_sorry_global thy [] [] prop - (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy) + (fn _ => simp_tac (put_simpset (get_simpset thy) ctxt addsimps @{thms simp_thms} addsimprocs [split_simproc (K ~1)]) 1)) end handle TERM _ => NONE) | _ => NONE) @@ -1356,9 +1361,9 @@ P t = 0: do not split P t = ~1: completely split P t > 0: split up to given bound of record extensions.*) -fun split_simp_tac thms P = CSUBGOAL (fn (cgoal, i) => +fun split_simp_tac ctxt thms P = CSUBGOAL (fn (cgoal, i) => let - val thy = Thm.theory_of_cterm cgoal; + val thy = Proof_Context.theory_of ctxt; val goal = term_of cgoal; val frees = filter (is_recT o #2) (Term.add_frees goal []); @@ -1376,9 +1381,9 @@ val crec = cterm_of thy r; val thm = cterm_instantiate [(crec, cfree)] induct_thm; in - simp_tac (HOL_basic_ss addsimps @{thms induct_atomize}) i THEN + simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) i THEN rtac thm i THEN - simp_tac (HOL_basic_ss addsimps @{thms induct_rulify}) i + simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_rulify}) i end; val split_frees_tacs = @@ -1402,14 +1407,14 @@ val thms' = @{thms o_apply K_record_comp} @ thms; in EVERY split_frees_tacs THEN - full_simp_tac (get_simpset thy addsimps thms' addsimprocs simprocs) i + full_simp_tac (put_simpset (get_simpset thy) ctxt addsimps thms' addsimprocs simprocs) i end); (* split_tac *) (*Split all records in the goal, which are quantified by !! or ALL.*) -val split_tac = CSUBGOAL (fn (cgoal, i) => +fun split_tac ctxt = CSUBGOAL (fn (cgoal, i) => let val goal = term_of cgoal; @@ -1423,7 +1428,7 @@ | is_all _ = 0; in if has_rec goal then - full_simp_tac (HOL_basic_ss addsimprocs [split_simproc is_all]) i + full_simp_tac (put_simpset HOL_basic_ss ctxt addsimprocs [split_simproc is_all]) i else no_tac end); @@ -1431,7 +1436,7 @@ (* wrapper *) val split_name = "record_split_tac"; -val split_wrapper = (split_name, fn _ => fn tac => split_tac ORELSE' tac); +val split_wrapper = (split_name, fn ctxt => fn tac => split_tac ctxt ORELSE' tac); @@ -1581,10 +1586,10 @@ in Logic.mk_equals (Logic.all s (P $ s), fold_rev Logic.all vars_more (P $ ext)) end; val inject = timeit_msg ctxt "record extension inject proof:" (fn () => - simplify HOL_ss + simplify (Simplifier.global_context defs_thy HOL_ss) (Goal.prove_sorry_global defs_thy [] [] inject_prop - (fn _ => - simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN + (fn {context = ctxt, ...} => + simp_tac (put_simpset HOL_basic_ss ctxt addsimps [ext_def]) 1 THEN REPEAT_DETERM (rtac @{thm refl_conj_eq} 1 ORELSE Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE @@ -1603,7 +1608,7 @@ val cterm_ext = cterm_of defs_thy ext; val start = named_cterm_instantiate [("y", cterm_ext)] surject_assist_idE; val tactic1 = - simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN + simp_tac (Simplifier.global_context defs_thy HOL_basic_ss addsimps [ext_def]) 1 THEN REPEAT_ALL_NEW Iso_Tuple_Support.iso_tuple_intros_tac 1; val tactic2 = REPEAT (rtac surject_assistI 1 THEN rtac refl 1); val [halfway] = Seq.list_of (tactic1 start); @@ -1624,10 +1629,10 @@ val induct = timeit_msg ctxt "record extension induct proof:" (fn () => let val (assm, concl) = induct_prop in Goal.prove_sorry_global defs_thy [] [assm] concl - (fn {prems, ...} => + (fn {context = ctxt, prems, ...} => cut_tac (split_meta RS Drule.equal_elim_rule2) 1 THEN resolve_tac prems 2 THEN - asm_simp_tac HOL_ss 1) + asm_simp_tac (put_simpset HOL_ss ctxt) 1) end); val ([(_, [induct']), (_, [inject']), (_, [surjective']), (_, [split_meta'])], thm_thy) = @@ -1934,7 +1939,7 @@ val init_thm = named_cterm_instantiate insts updacc_eq_triv; val terminal = rtac updacc_eq_idI 1 THEN rtac refl 1; val tactic = - simp_tac (HOL_basic_ss addsimps ext_defs) 1 THEN + simp_tac (Simplifier.global_context ext_thy HOL_basic_ss addsimps ext_defs) 1 THEN REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE terminal); val updaccs = Seq.list_of (tactic init_thm); in @@ -2076,31 +2081,34 @@ (* 3rd stage: thms_thy *) val record_ss = get_simpset defs_thy; - val sel_upd_ss = record_ss addsimps (sel_defs @ accessor_thms @ upd_defs @ updator_thms); + val sel_upd_ctxt = + Simplifier.global_context defs_thy record_ss + addsimps (sel_defs @ accessor_thms @ upd_defs @ updator_thms); val (sel_convs, upd_convs) = timeit_msg ctxt "record sel_convs/upd_convs proof:" (fn () => grouped 10 Par_List.map (fn prop => - Goal.prove_sorry_global defs_thy [] [] prop - (K (ALLGOALS (asm_full_simp_tac sel_upd_ss)))) (sel_conv_props @ upd_conv_props)) + Goal.prove_sorry_global defs_thy [] [] prop + (fn _ => ALLGOALS (asm_full_simp_tac sel_upd_ctxt))) + (sel_conv_props @ upd_conv_props)) |> chop (length sel_conv_props); val (fold_congs, unfold_congs) = timeit_msg ctxt "record upd fold/unfold congs:" (fn () => let val symdefs = map Thm.symmetric (sel_defs @ upd_defs); - val fold_ss = HOL_basic_ss addsimps symdefs; - val ua_congs = map (Drule.export_without_context o simplify fold_ss) upd_acc_cong_assists; + val fold_ctxt = Simplifier.global_context defs_thy HOL_basic_ss addsimps symdefs; + val ua_congs = map (Drule.export_without_context o simplify fold_ctxt) upd_acc_cong_assists; in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end); val parent_induct = Option.map #induct_scheme (try List.last parents); val induct_scheme = timeit_msg ctxt "record induct_scheme proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] induct_scheme_prop - (fn _ => + (fn {context = ctxt, ...} => EVERY [case parent_induct of NONE => all_tac | SOME ind => try_param_tac rN ind 1, try_param_tac rN ext_induct 1, - asm_simp_tac HOL_basic_ss 1])); + asm_simp_tac (put_simpset HOL_basic_ss ctxt) 1])); val induct = timeit_msg ctxt "record induct proof:" (fn () => Goal.prove_sorry_global defs_thy [] [#1 induct_prop] (#2 induct_prop) (fn {prems, ...} => @@ -2109,20 +2117,16 @@ THEN resolve_tac prems 1)); val surjective = timeit_msg ctxt "record surjective proof:" (fn () => - let - val leaf_ss = - get_sel_upd_defs defs_thy addsimps (sel_defs @ @{thms o_assoc id_apply id_o o_id}); - val init_ss = HOL_basic_ss addsimps ext_defs; - in - Goal.prove_sorry_global defs_thy [] [] surjective_prop - (fn _ => - EVERY - [rtac surject_assist_idE 1, - simp_tac init_ss 1, - REPEAT - (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE - (rtac surject_assistI 1 THEN simp_tac leaf_ss 1))]) - end); + Goal.prove_sorry_global defs_thy [] [] surjective_prop + (fn {context = ctxt, ...} => + EVERY + [rtac surject_assist_idE 1, + simp_tac (put_simpset HOL_basic_ss ctxt addsimps ext_defs) 1, + REPEAT + (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE + (rtac surject_assistI 1 THEN + simp_tac (put_simpset (get_sel_upd_defs defs_thy) ctxt + addsimps (sel_defs @ @{thms o_assoc id_apply id_o o_id})) 1))])); val cases_scheme = timeit_msg ctxt "record cases_scheme proof:" (fn () => Goal.prove_sorry_global defs_thy [] [#1 cases_scheme_prop] (#2 cases_scheme_prop) @@ -2132,9 +2136,10 @@ val cases = timeit_msg ctxt "record cases proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] cases_prop - (fn _ => + (fn {context = ctxt, ...} => try_param_tac rN cases_scheme 1 THEN - ALLGOALS (asm_full_simp_tac (HOL_basic_ss addsimps @{thms unit_all_eq1})))); + ALLGOALS (asm_full_simp_tac + (put_simpset HOL_basic_ss ctxt addsimps @{thms unit_all_eq1})))); val split_meta = timeit_msg ctxt "record split_meta proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] split_meta_prop @@ -2154,16 +2159,17 @@ val split_ex = timeit_msg ctxt "record split_ex proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] split_ex_prop - (fn _ => + (fn {context = ctxt, ...} => simp_tac - (HOL_basic_ss addsimps + (put_simpset HOL_basic_ss ctxt addsimps (@{lemma "EX x. P x == ~ (ALL x. ~ P x)" by simp} :: @{thms not_not Not_eq_iff})) 1 THEN rtac split_object 1)); val equality = timeit_msg ctxt "record equality proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] equality_prop - (fn _ => asm_full_simp_tac (record_ss addsimps (split_meta :: sel_convs)) 1)); + (fn {context = ctxt, ...} => + asm_full_simp_tac (put_simpset record_ss ctxt addsimps (split_meta :: sel_convs)) 1)); val ([(_, sel_convs'), (_, upd_convs'), (_, sel_defs'), (_, upd_defs'), (_, fold_congs'), (_, unfold_congs'), @@ -2312,7 +2318,7 @@ val setup = Sign.add_trfuns ([], parse_translation, [], []) #> Sign.add_advanced_trfuns ([], advanced_parse_translation, [], []) #> - Simplifier.map_simpset_global (fn ss => ss addsimprocs [simproc, upd_simproc, eq_simproc]); + map_theory_simpset (fn ctxt => ctxt addsimprocs [simproc, upd_simproc, eq_simproc]); (* outer syntax *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/semiring_normalizer.ML --- a/src/HOL/Tools/semiring_normalizer.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/semiring_normalizer.ML Thu Apr 18 17:07:01 2013 +0200 @@ -27,10 +27,20 @@ val semiring_normalizers_conv: cterm list -> cterm list * thm list -> cterm list * thm list -> cterm list * thm list -> (cterm -> bool) * conv * conv * conv -> (cterm -> cterm -> bool) -> - {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv} + {add: Proof.context -> conv, + mul: Proof.context -> conv, + neg: Proof.context -> conv, + main: Proof.context -> conv, + pow: Proof.context -> conv, + sub: Proof.context -> conv} val semiring_normalizers_ord_wrapper: Proof.context -> entry -> (cterm -> cterm -> bool) -> - {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv} + {add: Proof.context -> conv, + mul: Proof.context -> conv, + neg: Proof.context -> conv, + main: Proof.context -> conv, + pow: Proof.context -> conv, + sub: Proof.context -> conv} val setup: theory -> theory end @@ -177,9 +187,9 @@ handle TERM _ => error "ring_dest_const")), mk_const = fn phi => fn cT => fn x => Numeral.mk_cnumber cT (case Rat.quotient_of_rat x of (i, 1) => i | _ => error "int_of_rat: bad int"), - conv = fn phi => fn _ => Simplifier.rewrite (HOL_basic_ss addsimps @{thms semiring_norm}) - then_conv Simplifier.rewrite (HOL_basic_ss addsimps - @{thms numeral_1_eq_1})}; + conv = fn phi => fn ctxt => + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms semiring_norm}) + then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms numeral_1_eq_1})}; fun field_funs key = let @@ -208,7 +218,7 @@ {is_const = K numeral_is_const, dest_const = K dest_const, mk_const = mk_const, - conv = K (K Numeral_Simprocs.field_comp_conv)} + conv = K Numeral_Simprocs.field_comp_conv} end; @@ -236,23 +246,26 @@ val dest_numeral = term_of #> HOLogic.dest_number #> snd; val is_numeral = can dest_numeral; -val numeral01_conv = Simplifier.rewrite - (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}]); -val zero1_numeral_conv = - Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1} RS sym]); -fun zerone_conv cv = zero1_numeral_conv then_conv cv then_conv numeral01_conv; +fun numeral01_conv ctxt = + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_1_eq_1}]); + +fun zero1_numeral_conv ctxt = + Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_1_eq_1} RS sym]); + +fun zerone_conv ctxt cv = + zero1_numeral_conv ctxt then_conv cv then_conv numeral01_conv ctxt; val natarith = [@{thm "numeral_plus_numeral"}, @{thm "diff_nat_numeral"}, @{thm "numeral_times_numeral"}, @{thm "numeral_eq_iff"}, @{thm "numeral_less_iff"}]; -val nat_add_conv = - zerone_conv - (Simplifier.rewrite - (HOL_basic_ss - addsimps @{thms arith_simps} @ natarith @ @{thms rel_simps} - @ [@{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}, - @{thm add_numeral_left}, @{thm Suc_eq_plus1}] - @ map (fn th => th RS sym) @{thms numerals})); +fun nat_add_conv ctxt = + zerone_conv ctxt + (Simplifier.rewrite + (put_simpset HOL_basic_ss ctxt + addsimps @{thms arith_simps} @ natarith @ @{thms rel_simps} + @ [@{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}, + @{thm add_numeral_left}, @{thm Suc_eq_plus1}] + @ map (fn th => th RS sym) @{thms numerals})); val zeron_tm = @{cterm "0::nat"}; val onen_tm = @{cterm "1::nat"}; @@ -316,7 +329,7 @@ (* Also deals with "const * const", but both terms must involve powers of *) (* the same variable, or both be constants, or behaviour may be incorrect. *) - fun powvar_mul_conv tm = + fun powvar_mul_conv ctxt tm = let val (l,r) = dest_mul tm in if is_semiring_constant l andalso is_semiring_constant r @@ -328,16 +341,16 @@ ((let val (rx,rn) = dest_pow r val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29 val (tm1,tm2) = Thm.dest_comb(concl th1) in - Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end) + Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => (let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31 val (tm1,tm2) = Thm.dest_comb(concl th1) in - Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end)) end) + Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end)) end) handle CTERM _ => ((let val (rx,rn) = dest_pow r val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30 val (tm1,tm2) = Thm.dest_comb(concl th1) in - Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv tm2)) end) + Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => inst_thm [(cx,l)] pthm_32 )) @@ -353,7 +366,7 @@ (* Conversion for "(monomial)^n", where n is a numeral. *) - val monomial_pow_conv = + fun monomial_pow_conv ctxt = let fun monomial_pow tm bod ntm = if not(is_comb bod) @@ -374,7 +387,7 @@ then let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34 val (l,r) = Thm.dest_comb(concl th1) - in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv r)) + in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv ctxt r)) end else if opr aconvc mul_tm @@ -405,7 +418,7 @@ end; (* Multiplication of canonical monomials. *) - val monomial_mul_conv = + fun monomial_mul_conv ctxt = let fun powvar tm = if is_semiring_constant tm then one_tm @@ -435,7 +448,7 @@ val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] pthm_15 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 - val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2 + val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 val th3 = Thm.transitive th1 th2 val (tm5,tm6) = Thm.dest_comb(concl th3) val (tm7,tm8) = Thm.dest_comb tm6 @@ -458,7 +471,7 @@ val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_18 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 - val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2 + val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 in Thm.transitive th1 th2 end else @@ -480,7 +493,7 @@ let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 - in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv tm4)) tm2) + in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2) end else if ord > 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22 @@ -493,7 +506,7 @@ handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr - in if ord = 0 then powvar_mul_conv tm + in if ord = 0 then powvar_mul_conv ctxt tm else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09 else Thm.reflexive tm end)) end)) @@ -502,7 +515,7 @@ end; (* Multiplication by monomial of a polynomial. *) - val polynomial_monomial_mul_conv = + fun polynomial_monomial_mul_conv ctxt = let fun pmm_conv tm = let val (l,r) = dest_mul tm @@ -511,10 +524,11 @@ val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 - val th2 = Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv tm4)) (pmm_conv tm2) + val th2 = + Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv ctxt tm4)) (pmm_conv tm2) in Thm.transitive th1 th2 end) - handle CTERM _ => monomial_mul_conv tm) + handle CTERM _ => monomial_mul_conv ctxt tm) end in pmm_conv end; @@ -592,7 +606,7 @@ (* Addition of two polynomials. *) -val polynomial_add_conv = +fun polynomial_add_conv ctxt = let fun dezero_rule th = let @@ -690,25 +704,25 @@ (* Multiplication of two polynomials. *) -val polynomial_mul_conv = +fun polynomial_mul_conv ctxt = let fun pmul tm = let val (l,r) = dest_mul tm in - if not(is_add l) then polynomial_monomial_mul_conv tm + if not(is_add l) then polynomial_monomial_mul_conv ctxt tm else if not(is_add r) then let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09 - in Thm.transitive th1 (polynomial_monomial_mul_conv(concl th1)) + in Thm.transitive th1 (polynomial_monomial_mul_conv ctxt (concl th1)) end else let val (a,b) = dest_add l val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_10 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 - val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv tm4) + val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv ctxt tm4) val th3 = Thm.transitive th1 (Thm.combination th2 (pmul tm2)) - in Thm.transitive th3 (polynomial_add_conv (concl th3)) + in Thm.transitive th3 (polynomial_add_conv ctxt (concl th3)) end end in fn tm => @@ -724,12 +738,12 @@ (* Power of polynomial (optimized for the monomial and trivial cases). *) -fun num_conv n = - nat_add_conv (Thm.apply @{cterm Suc} (Numeral.mk_cnumber @{ctyp nat} (dest_numeral n - 1))) +fun num_conv ctxt n = + nat_add_conv ctxt (Thm.apply @{cterm Suc} (Numeral.mk_cnumber @{ctyp nat} (dest_numeral n - 1))) |> Thm.symmetric; -val polynomial_pow_conv = +fun polynomial_pow_conv ctxt = let fun ppow tm = let val (l,n) = dest_pow tm @@ -737,52 +751,52 @@ if n aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if n aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else - let val th1 = num_conv n + let val th1 = num_conv ctxt n val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38 val (tm1,tm2) = Thm.dest_comb(concl th2) val th3 = Thm.transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2)) val th4 = Thm.transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3 - in Thm.transitive th4 (polynomial_mul_conv (concl th4)) + in Thm.transitive th4 (polynomial_mul_conv ctxt (concl th4)) end end in fn tm => - if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv tm + if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv ctxt tm end; (* Negation. *) -fun polynomial_neg_conv tm = +fun polynomial_neg_conv ctxt tm = let val (l,r) = Thm.dest_comb tm in if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else let val th1 = inst_thm [(cx',r)] neg_mul val th2 = Thm.transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1)) - in Thm.transitive th2 (polynomial_monomial_mul_conv (concl th2)) + in Thm.transitive th2 (polynomial_monomial_mul_conv ctxt (concl th2)) end end; (* Subtraction. *) -fun polynomial_sub_conv tm = +fun polynomial_sub_conv ctxt tm = let val (l,r) = dest_sub tm val th1 = inst_thm [(cx',l),(cy',r)] sub_add val (tm1,tm2) = Thm.dest_comb(concl th1) - val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv tm2) - in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv (concl th2))) + val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv ctxt tm2) + in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv ctxt (concl th2))) end; (* Conversion from HOL term. *) -fun polynomial_conv tm = +fun polynomial_conv ctxt tm = if is_semiring_constant tm then semiring_add_conv tm else if not(is_comb tm) then Thm.reflexive tm else let val (lopr,r) = Thm.dest_comb tm in if lopr aconvc neg_tm then - let val th1 = Drule.arg_cong_rule lopr (polynomial_conv r) - in Thm.transitive th1 (polynomial_neg_conv (concl th1)) + let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) + in Thm.transitive th1 (polynomial_neg_conv ctxt (concl th1)) end else if lopr aconvc inverse_tm then - let val th1 = Drule.arg_cong_rule lopr (polynomial_conv r) + let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (semiring_mul_conv (concl th1)) end else @@ -791,14 +805,14 @@ let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_numeral r then - let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv l)) r - in Thm.transitive th1 (polynomial_pow_conv (concl th1)) + let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) r + in Thm.transitive th1 (polynomial_pow_conv ctxt (concl th1)) end else if opr aconvc divide_tm then - let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv l)) - (polynomial_conv r) - val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv) + let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) + (polynomial_conv ctxt r) + val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv ctxt) (Thm.rhs_of th1) in Thm.transitive th1 th2 end @@ -806,10 +820,11 @@ if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm then let val th1 = - Thm.combination (Drule.arg_cong_rule opr (polynomial_conv l)) (polynomial_conv r) - val f = if opr aconvc add_tm then polynomial_add_conv - else if opr aconvc mul_tm then polynomial_mul_conv - else polynomial_sub_conv + Thm.combination + (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) + val f = if opr aconvc add_tm then polynomial_add_conv ctxt + else if opr aconvc mul_tm then polynomial_mul_conv ctxt + else polynomial_sub_conv ctxt in Thm.transitive th1 (f (concl th1)) end else Thm.reflexive tm @@ -826,8 +841,10 @@ end; val nat_exp_ss = - HOL_basic_ss addsimps (@{thms eval_nat_numeral} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps}) - addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]; + simpset_of + (put_simpset HOL_basic_ss @{context} + addsimps (@{thms eval_nat_numeral} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps}) + addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]); fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS; @@ -838,15 +855,17 @@ {conv, dest_const, mk_const, is_const}) ord = let val pow_conv = - Conv.arg_conv (Simplifier.rewrite nat_exp_ss) + Conv.arg_conv (Simplifier.rewrite (put_simpset nat_exp_ss ctxt)) then_conv Simplifier.rewrite - (HOL_basic_ss addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) + (put_simpset HOL_basic_ss ctxt addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) then_conv conv ctxt val dat = (is_const, conv ctxt, conv ctxt, pow_conv) in semiring_normalizers_conv vars semiring ring field dat ord end; fun semiring_normalize_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) ord = - #main (semiring_normalizers_ord_wrapper ctxt ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal},{conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) ord); + #main (semiring_normalizers_ord_wrapper ctxt + ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal}, + {conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) ord) ctxt; fun semiring_normalize_wrapper ctxt data = semiring_normalize_ord_wrapper ctxt data simple_cterm_ord; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/set_comprehension_pointfree.ML --- a/src/HOL/Tools/set_comprehension_pointfree.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/set_comprehension_pointfree.ML Thu Apr 18 17:07:01 2013 +0200 @@ -7,9 +7,9 @@ signature SET_COMPREHENSION_POINTFREE = sig - val base_simproc : simpset -> cterm -> thm option - val code_simproc : simpset -> cterm -> thm option - val simproc : simpset -> cterm -> thm option + val base_simproc : Proof.context -> cterm -> thm option + val code_simproc : Proof.context -> cterm -> thm option + val simproc : Proof.context -> cterm -> thm option end structure Set_Comprehension_Pointfree : SET_COMPREHENSION_POINTFREE = @@ -314,10 +314,10 @@ val collectI' = @{lemma "\ P a ==> a \ {x. P x}" by auto} val collectE' = @{lemma "a \ {x. P x} ==> (\ P a ==> Q) ==> Q" by auto} -fun elim_Collect_tac ss = dtac @{thm iffD1[OF mem_Collect_eq]} +fun elim_Collect_tac ctxt = dtac @{thm iffD1[OF mem_Collect_eq]} THEN' (REPEAT_DETERM o (eresolve_tac @{thms exE})) THEN' REPEAT_DETERM o etac @{thm conjE} - THEN' TRY o hyp_subst_tac' ss; + THEN' TRY o hyp_subst_tac' ctxt; fun intro_image_tac ctxt = rtac @{thm image_eqI} THEN' (REPEAT_DETERM1 o @@ -328,29 +328,29 @@ (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Conv.rewr_conv (mk_meta_eq prod_case_distrib)))))) ctxt))) -fun elim_image_tac ss = etac @{thm imageE} +fun elim_image_tac ctxt = etac @{thm imageE} THEN' REPEAT_DETERM o CHANGED o - (TRY o Simplifier.full_simp_tac (Simplifier.inherit_context ss (HOL_basic_ss addsimps @{thms split_paired_all prod.cases})) + (TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.cases}) THEN' REPEAT_DETERM o etac @{thm Pair_inject} - THEN' TRY o hyp_subst_tac' ss) + THEN' TRY o hyp_subst_tac' ctxt) -fun tac1_of_formula ss (Int (fm1, fm2)) = +fun tac1_of_formula ctxt (Int (fm1, fm2)) = TRY o etac @{thm conjE} THEN' rtac @{thm IntI} - THEN' (fn i => tac1_of_formula ss fm2 (i + 1)) - THEN' tac1_of_formula ss fm1 - | tac1_of_formula ss (Un (fm1, fm2)) = + THEN' (fn i => tac1_of_formula ctxt fm2 (i + 1)) + THEN' tac1_of_formula ctxt fm1 + | tac1_of_formula ctxt (Un (fm1, fm2)) = etac @{thm disjE} THEN' rtac @{thm UnI1} - THEN' tac1_of_formula ss fm1 + THEN' tac1_of_formula ctxt fm1 THEN' rtac @{thm UnI2} - THEN' tac1_of_formula ss fm2 - | tac1_of_formula ss (Atom _) = + THEN' tac1_of_formula ctxt fm2 + | tac1_of_formula ctxt (Atom _) = REPEAT_DETERM1 o (atac ORELSE' rtac @{thm SigmaI} ORELSE' ((rtac @{thm CollectI} ORELSE' rtac collectI') THEN' - TRY o Simplifier.simp_tac (Simplifier.inherit_context ss (HOL_basic_ss addsimps [@{thm prod.cases}]))) + TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.cases}])) ORELSE' ((rtac @{thm vimageI2} ORELSE' rtac vimageI2') THEN' - TRY o Simplifier.simp_tac (Simplifier.inherit_context ss (HOL_basic_ss addsimps [@{thm prod.cases}]))) + TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.cases}])) ORELSE' (rtac @{thm image_eqI} THEN' (REPEAT_DETERM o (rtac @{thm refl} @@ -359,48 +359,48 @@ ORELSE' rtac @{thm iffD2[OF Compl_iff]} ORELSE' atac) -fun tac2_of_formula ss (Int (fm1, fm2)) = +fun tac2_of_formula ctxt (Int (fm1, fm2)) = TRY o etac @{thm IntE} THEN' TRY o rtac @{thm conjI} - THEN' (fn i => tac2_of_formula ss fm2 (i + 1)) - THEN' tac2_of_formula ss fm1 - | tac2_of_formula ss (Un (fm1, fm2)) = + THEN' (fn i => tac2_of_formula ctxt fm2 (i + 1)) + THEN' tac2_of_formula ctxt fm1 + | tac2_of_formula ctxt (Un (fm1, fm2)) = etac @{thm UnE} THEN' rtac @{thm disjI1} - THEN' tac2_of_formula ss fm1 + THEN' tac2_of_formula ctxt fm1 THEN' rtac @{thm disjI2} - THEN' tac2_of_formula ss fm2 - | tac2_of_formula ss (Atom _) = + THEN' tac2_of_formula ctxt fm2 + | tac2_of_formula ctxt (Atom _) = REPEAT_DETERM o (atac ORELSE' dtac @{thm iffD1[OF mem_Sigma_iff]} ORELSE' etac @{thm conjE} ORELSE' ((etac @{thm CollectE} ORELSE' etac collectE') THEN' - TRY o Simplifier.full_simp_tac (Simplifier.inherit_context ss (HOL_basic_ss addsimps [@{thm prod.cases}])) THEN' - REPEAT_DETERM o etac @{thm Pair_inject} THEN' TRY o hyp_subst_tac' ss THEN' TRY o rtac @{thm refl}) + TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.cases}]) THEN' + REPEAT_DETERM o etac @{thm Pair_inject} THEN' TRY o hyp_subst_tac' ctxt THEN' TRY o rtac @{thm refl}) ORELSE' (etac @{thm imageE} THEN' (REPEAT_DETERM o CHANGED o - (TRY o Simplifier.full_simp_tac (Simplifier.inherit_context ss (HOL_basic_ss addsimps @{thms split_paired_all prod.cases})) + (TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.cases}) THEN' REPEAT_DETERM o etac @{thm Pair_inject} - THEN' TRY o hyp_subst_tac' ss THEN' TRY o rtac @{thm refl}))) + THEN' TRY o hyp_subst_tac' ctxt THEN' TRY o rtac @{thm refl}))) ORELSE' etac @{thm ComplE} ORELSE' ((etac @{thm vimageE} ORELSE' etac vimageE') - THEN' TRY o Simplifier.full_simp_tac (Simplifier.inherit_context ss (HOL_basic_ss addsimps [@{thm prod.cases}])) - THEN' TRY o hyp_subst_tac' ss THEN' TRY o rtac @{thm refl})) + THEN' TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.cases}]) + THEN' TRY o hyp_subst_tac' ctxt THEN' TRY o rtac @{thm refl})) -fun tac ss ctxt fm = +fun tac ctxt fm = let val subset_tac1 = rtac @{thm subsetI} - THEN' elim_Collect_tac ss + THEN' elim_Collect_tac ctxt THEN' intro_image_tac ctxt - THEN' tac1_of_formula ss fm + THEN' tac1_of_formula ctxt fm val subset_tac2 = rtac @{thm subsetI} - THEN' elim_image_tac ss + THEN' elim_image_tac ctxt THEN' rtac @{thm iffD2[OF mem_Collect_eq]} THEN' REPEAT_DETERM o resolve_tac @{thms exI} THEN' (TRY o REPEAT_ALL_NEW (rtac @{thm conjI})) - THEN' (K (TRY (FIRSTGOAL ((TRY o hyp_subst_tac' ss) THEN' rtac @{thm refl})))) + THEN' (K (TRY (FIRSTGOAL ((TRY o hyp_subst_tac' ctxt) THEN' rtac @{thm refl})))) THEN' (fn i => EVERY (rev (map_index (fn (j, f) => - REPEAT_DETERM (etac @{thm IntE} (i + j)) THEN tac2_of_formula ss f (i + j)) (strip_Int fm)))) + REPEAT_DETERM (etac @{thm IntE} (i + j)) THEN tac2_of_formula ctxt f (i + j)) (strip_Int fm)))) in rtac @{thm subset_antisym} THEN' subset_tac1 THEN' subset_tac2 end; @@ -409,48 +409,46 @@ (* preprocessing conversion: rewrites {(x1, ..., xn). P x1 ... xn} to {(x1, ..., xn) | x1 ... xn. P x1 ... xn} *) -fun comprehension_conv ss ct = -let - val ctxt = Simplifier.the_context ss - fun dest_Collect (Const (@{const_name Collect}, T) $ t) = (HOLogic.dest_setT (body_type T), t) - | dest_Collect t = raise TERM ("dest_Collect", [t]) - fun list_ex vs t = fold_rev (fn (x, T) => fn t => HOLogic.exists_const T $ Abs (x, T, t)) vs t - fun mk_term t = - let - val (T, t') = dest_Collect t - val (t'', vs, fp) = case strip_psplits t' of - (_, [_], _) => raise TERM("mk_term", [t']) - | (t'', vs, fp) => (t'', vs, fp) - val Ts = map snd vs - val eq = HOLogic.eq_const T $ Bound (length Ts) $ - (HOLogic.mk_ptuple fp (HOLogic.mk_ptupleT fp Ts) (rev (map_index (fn (i, _) => Bound i) Ts))) - in - HOLogic.Collect_const T $ absdummy T (list_ex vs (HOLogic.mk_conj (eq, t''))) - end; - fun is_eq th = is_some (try (HOLogic.dest_eq o HOLogic.dest_Trueprop) (prop_of th)) - val unfold_thms = @{thms split_paired_all mem_Collect_eq prod.cases} - fun tac ctxt = - rtac @{thm set_eqI} - THEN' Simplifier.simp_tac - (Simplifier.inherit_context ss (HOL_basic_ss addsimps unfold_thms)) - THEN' rtac @{thm iffI} - THEN' REPEAT_DETERM o rtac @{thm exI} - THEN' rtac @{thm conjI} THEN' rtac @{thm refl} THEN' atac - THEN' REPEAT_DETERM o etac @{thm exE} - THEN' etac @{thm conjE} - THEN' REPEAT_DETERM o etac @{thm Pair_inject} - THEN' Subgoal.FOCUS (fn {prems, ...} => - Simplifier.simp_tac - (Simplifier.inherit_context ss (HOL_basic_ss addsimps (filter is_eq prems))) 1) ctxt - THEN' TRY o atac -in - case try mk_term (term_of ct) of - NONE => Thm.reflexive ct - | SOME t' => - Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (term_of ct, t'))) - (fn {context, ...} => tac context 1) - RS @{thm eq_reflection} -end +fun comprehension_conv ctxt ct = + let + fun dest_Collect (Const (@{const_name Collect}, T) $ t) = (HOLogic.dest_setT (body_type T), t) + | dest_Collect t = raise TERM ("dest_Collect", [t]) + fun list_ex vs t = fold_rev (fn (x, T) => fn t => HOLogic.exists_const T $ Abs (x, T, t)) vs t + fun mk_term t = + let + val (T, t') = dest_Collect t + val (t'', vs, fp) = case strip_psplits t' of + (_, [_], _) => raise TERM("mk_term", [t']) + | (t'', vs, fp) => (t'', vs, fp) + val Ts = map snd vs + val eq = HOLogic.eq_const T $ Bound (length Ts) $ + (HOLogic.mk_ptuple fp (HOLogic.mk_ptupleT fp Ts) (rev (map_index (fn (i, _) => Bound i) Ts))) + in + HOLogic.Collect_const T $ absdummy T (list_ex vs (HOLogic.mk_conj (eq, t''))) + end; + fun is_eq th = is_some (try (HOLogic.dest_eq o HOLogic.dest_Trueprop) (prop_of th)) + val unfold_thms = @{thms split_paired_all mem_Collect_eq prod.cases} + fun tac ctxt = + rtac @{thm set_eqI} + THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps unfold_thms) + THEN' rtac @{thm iffI} + THEN' REPEAT_DETERM o rtac @{thm exI} + THEN' rtac @{thm conjI} THEN' rtac @{thm refl} THEN' atac + THEN' REPEAT_DETERM o etac @{thm exE} + THEN' etac @{thm conjE} + THEN' REPEAT_DETERM o etac @{thm Pair_inject} + THEN' Subgoal.FOCUS (fn {prems, ...} => + (* FIXME inner context!? *) + simp_tac (put_simpset HOL_basic_ss ctxt addsimps (filter is_eq prems)) 1) ctxt + THEN' TRY o atac + in + case try mk_term (term_of ct) of + NONE => Thm.reflexive ct + | SOME t' => + Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (term_of ct, t'))) + (fn {context, ...} => tac context 1) + RS @{thm eq_reflection} + end (* main simprocs *) @@ -463,19 +461,17 @@ @{lemma "A \ B \ A \ C = A \ (B \ C)" by auto}, @{lemma "(A \ B \ C \ D) = (A \ C) \ (B \ D)" by auto}] -fun conv ss t = +fun conv ctxt t = let - val ctxt = Simplifier.the_context ss val ([t'], ctxt') = Variable.import_terms true [t] (Variable.declare_term t ctxt) - val ss' = Simplifier.context ctxt' ss val ct = cterm_of (Proof_Context.theory_of ctxt') t' fun unfold_conv thms = Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) - (Raw_Simplifier.inherit_context ss' empty_ss addsimps thms) - val prep_eq = (comprehension_conv ss' then_conv unfold_conv prep_thms) ct + (empty_simpset ctxt' addsimps thms) + val prep_eq = (comprehension_conv ctxt' then_conv unfold_conv prep_thms) ct val t'' = term_of (Thm.rhs_of prep_eq) fun mk_thm (fm, t''') = Goal.prove ctxt' [] [] - (HOLogic.mk_Trueprop (HOLogic.mk_eq (t'', t'''))) (fn {context, ...} => tac ss' context fm 1) + (HOLogic.mk_Trueprop (HOLogic.mk_eq (t'', t'''))) (fn {context, ...} => tac context fm 1) fun unfold th = th RS ((prep_eq RS meta_eq_to_obj_eq) RS @{thm trans}) val post = Conv.fconv_rule @@ -485,11 +481,11 @@ Option.map (export o post o unfold o mk_thm) (rewrite_term t'') end; -fun base_simproc ss redex = +fun base_simproc ctxt redex = let val set_compr = term_of redex in - conv ss set_compr + conv ctxt set_compr |> Option.map (fn thm => thm RS @{thm eq_reflection}) end; @@ -502,24 +498,23 @@ cterm_instantiate [(certify f, certify pred)] arg_cong end; -fun simproc ss redex = +fun simproc ctxt redex = let - val ctxt = Simplifier.the_context ss val pred $ set_compr = term_of redex val arg_cong' = instantiate_arg_cong ctxt pred in - conv ss set_compr + conv ctxt set_compr |> Option.map (fn thm => thm RS arg_cong' RS @{thm eq_reflection}) end; -fun code_simproc ss redex = +fun code_simproc ctxt redex = let fun unfold_conv thms = Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) - (Raw_Simplifier.inherit_context ss empty_ss addsimps thms) + (empty_simpset ctxt addsimps thms) val prep_thm = unfold_conv @{thms eq_equal[symmetric]} redex in - case base_simproc ss (Thm.rhs_of prep_thm) of + case base_simproc ctxt (Thm.rhs_of prep_thm) of SOME rewr_thm => SOME (transitive_thm OF [transitive_thm OF [prep_thm, rewr_thm], unfold_conv @{thms eq_equal} (Thm.rhs_of rewr_thm)]) | NONE => NONE diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/simpdata.ML --- a/src/HOL/Tools/simpdata.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/simpdata.ML Thu Apr 18 17:07:01 2013 +0200 @@ -48,7 +48,7 @@ | _ $ (Const (@{const_name Not}, _) $ _) => th RS @{thm Eq_FalseI} | _ => th RS @{thm Eq_TrueI} -fun mk_eq_True (_: simpset) r = +fun mk_eq_True (_: Proof.context) r = SOME (r RS @{thm meta_eq_to_obj_eq} RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE; (* Produce theorems of the form @@ -79,7 +79,7 @@ end; (*Congruence rules for = (instead of ==)*) -fun mk_meta_cong (_: simpset) rl = zero_var_indexes +fun mk_meta_cong (_: Proof.context) rl = zero_var_indexes (let val rl' = Seq.hd (TRYALL (fn i => fn st => rtac (lift_meta_eq_to_obj_eq i st) i st) rl) in mk_meta_eq rl' handle THM _ => @@ -106,21 +106,21 @@ end; in atoms end; -fun mksimps pairs (_: simpset) = +fun mksimps pairs (_: Proof.context) = map_filter (try mk_eq) o mk_atomize pairs o gen_all; -fun unsafe_solver_tac ss = +fun unsafe_solver_tac ctxt = (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ss), + FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt), atac, etac @{thm FalseE}]; val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; (*No premature instantiation of variables during simplification*) -fun safe_solver_tac ss = +fun safe_solver_tac ctxt = (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ss), + FIRST' [match_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt), eq_assume_tac, ematch_tac @{thms FalseE}]; val safe_solver = mk_solver "HOL safe" safe_solver_tac; @@ -166,19 +166,20 @@ (@{const_name If}, [@{thm if_bool_eq_conj} RS @{thm iffD1}])]; val HOL_basic_ss = - Simplifier.global_context @{theory} empty_ss + empty_simpset @{context} setSSolver safe_solver setSolver unsafe_solver |> Simplifier.set_subgoaler asm_simp_tac |> Simplifier.set_mksimps (mksimps mksimps_pairs) |> Simplifier.set_mkeqTrue mk_eq_True - |> Simplifier.set_mkcong mk_meta_cong; - -fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews); + |> Simplifier.set_mkcong mk_meta_cong + |> simpset_of; -fun unfold_tac ths = - let val ss0 = Simplifier.clear_ss HOL_basic_ss addsimps ths - in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; +fun hol_simplify ctxt rews = + Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps rews); + +fun unfold_tac ths ctxt = + ALLGOALS (full_simp_tac (clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths)); end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Tools/split_rule.ML Thu Apr 18 17:07:01 2013 +0200 @@ -6,9 +6,9 @@ signature SPLIT_RULE = sig - val split_rule_var: term -> thm -> thm - val split_rule: thm -> thm - val complete_split_rule: thm -> thm + val split_rule_var: Proof.context -> term -> thm -> thm + val split_rule: Proof.context -> thm -> thm + val complete_split_rule: Proof.context -> thm -> thm val setup: theory -> theory end; @@ -21,10 +21,11 @@ Const (@{const_name Product_Type.internal_split}, [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc); -val eval_internal_split = - hol_simplify [@{thm internal_split_def}] o hol_simplify [@{thm internal_split_conv}]; +fun eval_internal_split ctxt = + hol_simplify ctxt @{thms internal_split_def} o + hol_simplify ctxt @{thms internal_split_conv}; -val remove_internal_split = eval_internal_split o split_all; +fun remove_internal_split ctxt = eval_internal_split ctxt o split_all ctxt; (*In ap_split S T u, term u expects separate arguments for the factors of S, @@ -84,23 +85,24 @@ | (t, ts) => fold collect_vars ts); -val split_rule_var = (Drule.export_without_context o remove_internal_split) oo split_rule_var'; +fun split_rule_var ctxt = + (Drule.export_without_context o remove_internal_split ctxt) oo split_rule_var'; (*curries ALL function variables occurring in a rule's conclusion*) -fun split_rule rl = +fun split_rule ctxt rl = fold_rev split_rule_var' (Misc_Legacy.term_vars (concl_of rl)) rl - |> remove_internal_split + |> remove_internal_split ctxt |> Drule.export_without_context; (*curries ALL function variables*) -fun complete_split_rule rl = +fun complete_split_rule ctxt rl = let val prop = Thm.prop_of rl; val xs = Term.fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I) prop []; val vars = collect_vars prop []; in fst (fold_rev complete_split_rule_var vars (rl, xs)) - |> remove_internal_split + |> remove_internal_split ctxt |> Drule.export_without_context |> Rule_Cases.save rl end; @@ -110,10 +112,11 @@ val setup = Attrib.setup @{binding split_format} - (Scan.lift - (Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute (K complete_split_rule)))) + (Scan.lift (Args.parens (Args.$$$ "complete") + >> K (Thm.rule_attribute (complete_split_rule o Context.proof_of)))) "split pair-typed subterms in premises, or function arguments" #> - Attrib.setup @{binding split_rule} (Scan.succeed (Thm.rule_attribute (K split_rule))) + Attrib.setup @{binding split_rule} + (Scan.succeed (Thm.rule_attribute (split_rule o Context.proof_of))) "curries ALL function variables occurring in a rule's conclusion"; end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Transitive_Closure.thy Thu Apr 18 17:07:01 2013 +0200 @@ -1262,11 +1262,11 @@ *} setup {* - Simplifier.map_simpset_global (fn ss => ss - addSolver (mk_solver "Trancl" (Trancl_Tac.trancl_tac o Simplifier.the_context)) - addSolver (mk_solver "Rtrancl" (Trancl_Tac.rtrancl_tac o Simplifier.the_context)) - addSolver (mk_solver "Tranclp" (Tranclp_Tac.trancl_tac o Simplifier.the_context)) - addSolver (mk_solver "Rtranclp" (Tranclp_Tac.rtrancl_tac o Simplifier.the_context))) + map_theory_simpset (fn ctxt => ctxt + addSolver (mk_solver "Trancl" Trancl_Tac.trancl_tac) + addSolver (mk_solver "Rtrancl" Trancl_Tac.rtrancl_tac) + addSolver (mk_solver "Tranclp" Tranclp_Tac.trancl_tac) + addSolver (mk_solver "Rtranclp" Tranclp_Tac.rtrancl_tac)) *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/UNITY/Comp/Alloc.thy --- a/src/HOL/UNITY/Comp/Alloc.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/UNITY/Comp/Alloc.thy Thu Apr 18 17:07:01 2013 +0200 @@ -329,11 +329,11 @@ ML {* fun record_auto_tac ctxt = let val ctxt' = - (ctxt addSWrapper Record.split_wrapper) - |> map_simpset (fn ss => ss addsimps + ctxt addSWrapper Record.split_wrapper + addsimps [@{thm sysOfAlloc_def}, @{thm sysOfClient_def}, @{thm client_map_def}, @{thm non_dummy_def}, @{thm funPair_def}, - @{thm o_apply}, @{thm Let_def}]) + @{thm o_apply}, @{thm Let_def}] in auto_tac ctxt' end; *} @@ -425,7 +425,7 @@ apply (rule fst_o_funPair) done -ML {* bind_thms ("fst_o_client_map'", make_o_equivs @{thm fst_o_client_map}) *} +ML {* bind_thms ("fst_o_client_map'", make_o_equivs @{context} @{thm fst_o_client_map}) *} declare fst_o_client_map' [simp] lemma snd_o_client_map: "snd o client_map = clientState_d.dummy" @@ -433,7 +433,7 @@ apply (rule snd_o_funPair) done -ML {* bind_thms ("snd_o_client_map'", make_o_equivs @{thm snd_o_client_map}) *} +ML {* bind_thms ("snd_o_client_map'", make_o_equivs @{context} @{thm snd_o_client_map}) *} declare snd_o_client_map' [simp] @@ -443,28 +443,28 @@ apply record_auto done -ML {* bind_thms ("client_o_sysOfAlloc'", make_o_equivs @{thm client_o_sysOfAlloc}) *} +ML {* bind_thms ("client_o_sysOfAlloc'", make_o_equivs @{context} @{thm client_o_sysOfAlloc}) *} declare client_o_sysOfAlloc' [simp] lemma allocGiv_o_sysOfAlloc_eq: "allocGiv o sysOfAlloc = allocGiv" apply record_auto done -ML {* bind_thms ("allocGiv_o_sysOfAlloc_eq'", make_o_equivs @{thm allocGiv_o_sysOfAlloc_eq}) *} +ML {* bind_thms ("allocGiv_o_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocGiv_o_sysOfAlloc_eq}) *} declare allocGiv_o_sysOfAlloc_eq' [simp] lemma allocAsk_o_sysOfAlloc_eq: "allocAsk o sysOfAlloc = allocAsk" apply record_auto done -ML {* bind_thms ("allocAsk_o_sysOfAlloc_eq'", make_o_equivs @{thm allocAsk_o_sysOfAlloc_eq}) *} +ML {* bind_thms ("allocAsk_o_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocAsk_o_sysOfAlloc_eq}) *} declare allocAsk_o_sysOfAlloc_eq' [simp] lemma allocRel_o_sysOfAlloc_eq: "allocRel o sysOfAlloc = allocRel" apply record_auto done -ML {* bind_thms ("allocRel_o_sysOfAlloc_eq'", make_o_equivs @{thm allocRel_o_sysOfAlloc_eq}) *} +ML {* bind_thms ("allocRel_o_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocRel_o_sysOfAlloc_eq}) *} declare allocRel_o_sysOfAlloc_eq' [simp] @@ -474,49 +474,49 @@ apply record_auto done -ML {* bind_thms ("client_o_sysOfClient'", make_o_equivs @{thm client_o_sysOfClient}) *} +ML {* bind_thms ("client_o_sysOfClient'", make_o_equivs @{context} @{thm client_o_sysOfClient}) *} declare client_o_sysOfClient' [simp] lemma allocGiv_o_sysOfClient_eq: "allocGiv o sysOfClient = allocGiv o snd " apply record_auto done -ML {* bind_thms ("allocGiv_o_sysOfClient_eq'", make_o_equivs @{thm allocGiv_o_sysOfClient_eq}) *} +ML {* bind_thms ("allocGiv_o_sysOfClient_eq'", make_o_equivs @{context} @{thm allocGiv_o_sysOfClient_eq}) *} declare allocGiv_o_sysOfClient_eq' [simp] lemma allocAsk_o_sysOfClient_eq: "allocAsk o sysOfClient = allocAsk o snd " apply record_auto done -ML {* bind_thms ("allocAsk_o_sysOfClient_eq'", make_o_equivs @{thm allocAsk_o_sysOfClient_eq}) *} +ML {* bind_thms ("allocAsk_o_sysOfClient_eq'", make_o_equivs @{context} @{thm allocAsk_o_sysOfClient_eq}) *} declare allocAsk_o_sysOfClient_eq' [simp] lemma allocRel_o_sysOfClient_eq: "allocRel o sysOfClient = allocRel o snd " apply record_auto done -ML {* bind_thms ("allocRel_o_sysOfClient_eq'", make_o_equivs @{thm allocRel_o_sysOfClient_eq}) *} +ML {* bind_thms ("allocRel_o_sysOfClient_eq'", make_o_equivs @{context} @{thm allocRel_o_sysOfClient_eq}) *} declare allocRel_o_sysOfClient_eq' [simp] lemma allocGiv_o_inv_sysOfAlloc_eq: "allocGiv o inv sysOfAlloc = allocGiv" apply (simp add: o_def) done -ML {* bind_thms ("allocGiv_o_inv_sysOfAlloc_eq'", make_o_equivs @{thm allocGiv_o_inv_sysOfAlloc_eq}) *} +ML {* bind_thms ("allocGiv_o_inv_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocGiv_o_inv_sysOfAlloc_eq}) *} declare allocGiv_o_inv_sysOfAlloc_eq' [simp] lemma allocAsk_o_inv_sysOfAlloc_eq: "allocAsk o inv sysOfAlloc = allocAsk" apply (simp add: o_def) done -ML {* bind_thms ("allocAsk_o_inv_sysOfAlloc_eq'", make_o_equivs @{thm allocAsk_o_inv_sysOfAlloc_eq}) *} +ML {* bind_thms ("allocAsk_o_inv_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocAsk_o_inv_sysOfAlloc_eq}) *} declare allocAsk_o_inv_sysOfAlloc_eq' [simp] lemma allocRel_o_inv_sysOfAlloc_eq: "allocRel o inv sysOfAlloc = allocRel" apply (simp add: o_def) done -ML {* bind_thms ("allocRel_o_inv_sysOfAlloc_eq'", make_o_equivs @{thm allocRel_o_inv_sysOfAlloc_eq}) *} +ML {* bind_thms ("allocRel_o_inv_sysOfAlloc_eq'", make_o_equivs @{context} @{thm allocRel_o_inv_sysOfAlloc_eq}) *} declare allocRel_o_inv_sysOfAlloc_eq' [simp] lemma rel_inv_client_map_drop_map: "(rel o inv client_map o drop_map i o inv sysOfClient) = @@ -524,7 +524,7 @@ apply (simp add: o_def drop_map_def) done -ML {* bind_thms ("rel_inv_client_map_drop_map'", make_o_equivs @{thm rel_inv_client_map_drop_map}) *} +ML {* bind_thms ("rel_inv_client_map_drop_map'", make_o_equivs @{context} @{thm rel_inv_client_map_drop_map}) *} declare rel_inv_client_map_drop_map [simp] lemma ask_inv_client_map_drop_map: "(ask o inv client_map o drop_map i o inv sysOfClient) = @@ -532,7 +532,7 @@ apply (simp add: o_def drop_map_def) done -ML {* bind_thms ("ask_inv_client_map_drop_map'", make_o_equivs @{thm ask_inv_client_map_drop_map}) *} +ML {* bind_thms ("ask_inv_client_map_drop_map'", make_o_equivs @{context} @{thm ask_inv_client_map_drop_map}) *} declare ask_inv_client_map_drop_map [simp] @@ -546,7 +546,7 @@ val [Client_Increasing_ask, Client_Increasing_rel, Client_Bounded, Client_Progress, Client_AllowedActs, Client_preserves_giv, Client_preserves_dummy] = - @{thm Client} |> simplify (@{simpset} addsimps @{thms client_spec_simps}) + @{thm Client} |> simplify (@{context} addsimps @{thms client_spec_simps}) |> list_of_Int; bind_thm ("Client_Increasing_ask", Client_Increasing_ask); @@ -576,7 +576,7 @@ val [Network_Ask, Network_Giv, Network_Rel, Network_AllowedActs, Network_preserves_allocGiv, Network_preserves_rel, Network_preserves_ask] = - @{thm Network} |> simplify (@{simpset} addsimps @{thms network_spec_simps}) + @{thm Network} |> simplify (@{context} addsimps @{thms network_spec_simps}) |> list_of_Int; bind_thm ("Network_Ask", Network_Ask); @@ -607,7 +607,7 @@ val [Alloc_Increasing_0, Alloc_Safety, Alloc_Progress, Alloc_AllowedActs, Alloc_preserves_allocRel, Alloc_preserves_allocAsk, Alloc_preserves_dummy] = - @{thm Alloc} |> simplify (@{simpset} addsimps @{thms alloc_spec_simps}) + @{thm Alloc} |> simplify (@{context} addsimps @{thms alloc_spec_simps}) |> list_of_Int; bind_thm ("Alloc_Increasing_0", Alloc_Increasing_0); @@ -709,26 +709,25 @@ *) ML {* -fun rename_client_map_tac ss = +fun rename_client_map_tac ctxt = EVERY [ - simp_tac (ss addsimps [@{thm rename_guarantees_eq_rename_inv}]) 1, + simp_tac (ctxt addsimps [@{thm rename_guarantees_eq_rename_inv}]) 1, rtac @{thm guarantees_PLam_I} 1, assume_tac 2, (*preserves: routine reasoning*) - asm_simp_tac (ss addsimps [@{thm lift_preserves_sub}]) 2, + asm_simp_tac (ctxt addsimps [@{thm lift_preserves_sub}]) 2, (*the guarantee for "lift i (rename client_map Client)" *) asm_simp_tac - (ss addsimps [@{thm lift_guarantees_eq_lift_inv}, + (ctxt addsimps [@{thm lift_guarantees_eq_lift_inv}, @{thm rename_guarantees_eq_rename_inv}, @{thm bij_imp_bij_inv}, @{thm surj_rename}, @{thm inv_inv_eq}]) 1, - asm_simp_tac - (@{simpset} addsimps [@{thm o_def}, @{thm non_dummy_def}, @{thm guarantees_Int_right}]) 1] + asm_simp_tac (* FIXME ctxt !!? *) + (@{context} addsimps [@{thm o_def}, @{thm non_dummy_def}, @{thm guarantees_Int_right}]) 1] *} method_setup rename_client_map = {* - Scan.succeed (fn ctxt => - SIMPLE_METHOD (rename_client_map_tac (simpset_of ctxt))) + Scan.succeed (fn ctxt => SIMPLE_METHOD (rename_client_map_tac ctxt)) *} text{*Lifting @{text Client_Increasing} to @{term systemState}*} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/UNITY/Simple/NSP_Bad.thy --- a/src/HOL/UNITY/Simple/NSP_Bad.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/UNITY/Simple/NSP_Bad.thy Thu Apr 18 17:07:01 2013 +0200 @@ -107,11 +107,11 @@ [REPEAT (etac @{thm Always_ConstrainsI} 1), REPEAT (resolve_tac [@{thm StableI}, @{thm stableI}, @{thm constrains_imp_Constrains}] 1), rtac @{thm ns_constrainsI} 1, - full_simp_tac (simpset_of ctxt) 1, + full_simp_tac ctxt 1, REPEAT (FIRSTGOAL (etac disjE)), ALLGOALS (clarify_tac (ctxt delrules [impI, @{thm impCE}])), REPEAT (FIRSTGOAL analz_mono_contra_tac), - ALLGOALS (asm_simp_tac (simpset_of ctxt))]) i; + ALLGOALS (asm_simp_tac ctxt)]) i; (*Tactic for proving secrecy theorems*) fun ns_induct_tac ctxt = diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/UNITY/UNITY_Main.thy --- a/src/HOL/UNITY/UNITY_Main.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/UNITY/UNITY_Main.thy Thu Apr 18 17:07:01 2013 +0200 @@ -21,9 +21,9 @@ *} "for proving progress properties" setup {* - Simplifier.map_simpset_global (fn ss => ss - addsimps (make_o_equivs @{thm fst_o_funPair} @ make_o_equivs @{thm snd_o_funPair}) - addsimps (make_o_equivs @{thm fst_o_lift_map} @ make_o_equivs @{thm snd_o_lift_map})) + map_theory_simpset (fn ctxt => ctxt + addsimps (make_o_equivs ctxt @{thm fst_o_funPair} @ make_o_equivs ctxt @{thm snd_o_funPair}) + addsimps (make_o_equivs ctxt @{thm fst_o_lift_map} @ make_o_equivs ctxt @{thm snd_o_lift_map})) *} end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/UNITY/UNITY_tactics.ML --- a/src/HOL/UNITY/UNITY_tactics.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/UNITY/UNITY_tactics.ML Thu Apr 18 17:07:01 2013 +0200 @@ -23,12 +23,13 @@ resolve_tac [@{thm StableI}, @{thm stableI}, @{thm constrains_imp_Constrains}] 1), (*for safety, the totalize operator can be ignored*) - simp_tac (HOL_ss addsimps [@{thm mk_total_program_def}, @{thm totalize_constrains_iff}]) 1, + simp_tac (put_simpset HOL_ss ctxt + addsimps [@{thm mk_total_program_def}, @{thm totalize_constrains_iff}]) 1, rtac @{thm constrainsI} 1, - full_simp_tac (simpset_of ctxt) 1, + full_simp_tac ctxt 1, REPEAT (FIRSTGOAL (etac disjE)), ALLGOALS (clarify_tac ctxt), - ALLGOALS (asm_full_simp_tac (simpset_of ctxt))]) i; + ALLGOALS (asm_full_simp_tac ctxt)]) i; (*proves "ensures/leadsTo" properties when the program is specified*) fun ensures_tac ctxt sact = @@ -40,22 +41,22 @@ REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis}, @{thm EnsuresI}, @{thm ensuresI}] 1), (*now there are two subgoals: co & transient*) - simp_tac (simpset_of ctxt addsimps [@{thm mk_total_program_def}]) 2, + simp_tac (ctxt addsimps [@{thm mk_total_program_def}]) 2, res_inst_tac ctxt [(("act", 0), sact)] @{thm totalize_transientI} 2 ORELSE res_inst_tac ctxt [(("act", 0), sact)] @{thm transientI} 2, (*simplify the command's domain*) - simp_tac (simpset_of ctxt addsimps @{thms Domain_unfold}) 3, + simp_tac (ctxt addsimps @{thms Domain_unfold}) 3, constrains_tac ctxt 1, ALLGOALS (clarify_tac ctxt), - ALLGOALS (asm_lr_simp_tac (simpset_of ctxt))]); + ALLGOALS (asm_lr_simp_tac ctxt)]); (*Composition equivalences, from Lift_prog*) -fun make_o_equivs th = - [th, - th RS @{thm o_equiv_assoc} |> simplify (HOL_ss addsimps [@{thm o_assoc}]), - th RS @{thm o_equiv_apply} |> simplify (HOL_ss addsimps [@{thm o_def}, @{thm sub_def}])]; +fun make_o_equivs ctxt th = + [th, + th RS @{thm o_equiv_assoc} |> simplify (put_simpset HOL_ss ctxt addsimps [@{thm o_assoc}]), + th RS @{thm o_equiv_apply} |> simplify (put_simpset HOL_ss ctxt addsimps [@{thm o_def}, @{thm sub_def}])]; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Word/Word.thy Thu Apr 18 17:07:01 2013 +0200 @@ -1447,8 +1447,8 @@ (* uint_arith_tac: reduce to arithmetic on int, try to solve by arith *) ML {* -fun uint_arith_ss_of ss = - ss addsimps @{thms uint_arith_simps} +fun uint_arith_simpset ctxt = + ctxt addsimps @{thms uint_arith_simps} delsimps @{thms word_uint.Rep_inject} |> fold Splitter.add_split @{thms split_if_asm} |> fold Simplifier.add_cong @{thms power_False_cong} @@ -1460,9 +1460,11 @@ handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, - full_simp_tac (uint_arith_ss_of (simpset_of ctxt)) 1, - ALLGOALS (full_simp_tac (HOL_ss |> fold Splitter.add_split @{thms uint_splits} - |> fold Simplifier.add_cong @{thms power_False_cong})), + full_simp_tac (uint_arith_simpset ctxt) 1, + ALLGOALS (full_simp_tac + (put_simpset HOL_ss ctxt + |> fold Splitter.add_split @{thms uint_splits} + |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN REPEAT (etac conjE n) THEN @@ -1946,8 +1948,8 @@ (* unat_arith_tac: tactic to reduce word arithmetic to nat, try to solve via arith *) ML {* -fun unat_arith_ss_of ss = - ss addsimps @{thms unat_arith_simps} +fun unat_arith_simpset ctxt = + ctxt addsimps @{thms unat_arith_simps} delsimps @{thms word_unat.Rep_inject} |> fold Splitter.add_split @{thms split_if_asm} |> fold Simplifier.add_cong @{thms power_False_cong} @@ -1959,9 +1961,11 @@ handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, - full_simp_tac (unat_arith_ss_of (simpset_of ctxt)) 1, - ALLGOALS (full_simp_tac (HOL_ss |> fold Splitter.add_split @{thms unat_splits} - |> fold Simplifier.add_cong @{thms power_False_cong})), + full_simp_tac (unat_arith_simpset ctxt) 1, + ALLGOALS (full_simp_tac + (put_simpset HOL_ss ctxt + |> fold Splitter.add_split @{thms unat_splits} + |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN REPEAT (etac conjE n) THEN diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/Word/WordBitwise.thy --- a/src/HOL/Word/WordBitwise.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/Word/WordBitwise.thy Thu Apr 18 17:07:01 2013 +0200 @@ -507,33 +507,36 @@ (uncurry (Thm.mk_binop @{cterm "Cons :: nat => _"})) @{cterm "[] :: nat list"} ns; -fun upt_conv ct = case term_of ct of - (@{const upt} $ n $ m) => let - val (i, j) = pairself (snd o HOLogic.dest_number) (n, m); - val ns = map (Numeral.mk_cnumber @{ctyp nat}) (i upto (j - 1)) - |> mk_nat_clist; - val prop = Thm.mk_binop @{cterm "op = :: nat list => _"} ct ns - |> Thm.apply @{cterm Trueprop}; - in - try (fn () => - Goal.prove_internal [] prop - (K (REPEAT_DETERM (resolve_tac @{thms upt_eq_list_intros} 1 - ORELSE simp_tac word_ss 1))) |> mk_meta_eq) () - end +fun upt_conv ctxt ct = + case term_of ct of + (@{const upt} $ n $ m) => + let + val (i, j) = pairself (snd o HOLogic.dest_number) (n, m); + val ns = map (Numeral.mk_cnumber @{ctyp nat}) (i upto (j - 1)) + |> mk_nat_clist; + val prop = Thm.mk_binop @{cterm "op = :: nat list => _"} ct ns + |> Thm.apply @{cterm Trueprop}; + in + try (fn () => + Goal.prove_internal [] prop + (K (REPEAT_DETERM (resolve_tac @{thms upt_eq_list_intros} 1 + ORELSE simp_tac (put_simpset word_ss ctxt) 1))) |> mk_meta_eq) () + end | _ => NONE; val expand_upt_simproc = Simplifier.make_simproc {lhss = [@{cpat "upt _ _"}], name = "expand_upt", identifier = [], - proc = K (K upt_conv)}; + proc = K upt_conv}; -fun word_len_simproc_fn ct = case term_of ct of +fun word_len_simproc_fn ctxt ct = + case term_of ct of Const (@{const_name len_of}, _) $ t => (let val T = fastype_of t |> dest_Type |> snd |> the_single val n = Numeral.mk_cnumber @{ctyp nat} (Word_Lib.dest_binT T); val prop = Thm.mk_binop @{cterm "op = :: nat => _"} ct n |> Thm.apply @{cterm Trueprop}; - in Goal.prove_internal [] prop (K (simp_tac word_ss 1)) + in Goal.prove_internal [] prop (K (simp_tac (put_simpset word_ss ctxt) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE | TYPE _ => NONE) | _ => NONE; @@ -541,12 +544,14 @@ val word_len_simproc = Simplifier.make_simproc {lhss = [@{cpat "len_of _"}], name = "word_len", identifier = [], - proc = K (K word_len_simproc_fn)}; + proc = K word_len_simproc_fn}; (* convert 5 or nat 5 to Suc 4 when n_sucs = 1, Suc (Suc 4) when n_sucs = 2, or just 5 (discarding nat) when n_sucs = 0 *) -fun nat_get_Suc_simproc_fn n_sucs thy ct = let +fun nat_get_Suc_simproc_fn n_sucs ctxt ct = + let + val thy = Proof_Context.theory_of ctxt; val (f $ arg) = (term_of ct); val n = (case arg of @{term nat} $ n => n | n => n) |> HOLogic.dest_number |> snd; @@ -558,32 +563,38 @@ fun propfn g = HOLogic.mk_eq (g arg, g arg') |> HOLogic.mk_Trueprop |> cterm_of thy; val eq1 = Goal.prove_internal [] (propfn I) - (K (simp_tac word_ss 1)); + (K (simp_tac (put_simpset word_ss ctxt) 1)); in Goal.prove_internal [] (propfn (curry (op $) f)) - (K (simp_tac (HOL_ss addsimps [eq1]) 1)) + (K (simp_tac (put_simpset HOL_ss ctxt addsimps [eq1]) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE; fun nat_get_Suc_simproc n_sucs thy cts = Simplifier.make_simproc {lhss = map (fn t => Thm.apply t @{cpat "?n :: nat"}) cts, name = "nat_get_Suc", identifier = [], - proc = K (K (nat_get_Suc_simproc_fn n_sucs thy))}; + proc = K (nat_get_Suc_simproc_fn n_sucs)}; -val no_split_ss = Splitter.del_split @{thm split_if} HOL_ss; +val no_split_ss = + simpset_of (put_simpset HOL_ss @{context} + |> Splitter.del_split @{thm split_if}); -val expand_word_eq_sss = (HOL_basic_ss addsimps - @{thms word_eq_rbl_eq word_le_rbl word_less_rbl word_sle_rbl word_sless_rbl}, - [no_split_ss addsimps @{thms rbl_word_plus rbl_word_and rbl_word_or rbl_word_not +val expand_word_eq_sss = + (simpset_of (put_simpset HOL_basic_ss @{context} addsimps + @{thms word_eq_rbl_eq word_le_rbl word_less_rbl word_sle_rbl word_sless_rbl}), + map simpset_of [ + put_simpset no_split_ss @{context} addsimps + @{thms rbl_word_plus rbl_word_and rbl_word_or rbl_word_not rbl_word_neg bl_word_sub rbl_word_xor rbl_word_cat rbl_word_slice rbl_word_scast rbl_word_ucast rbl_shiftl rbl_shiftr rbl_sshiftr rbl_word_if}, - no_split_ss addsimps @{thms to_bl_numeral to_bl_neg_numeral - to_bl_0 rbl_word_1}, - no_split_ss addsimps @{thms rev_rev_ident - rev_replicate rev_map to_bl_upt word_size} + put_simpset no_split_ss @{context} addsimps + @{thms to_bl_numeral to_bl_neg_numeral to_bl_0 rbl_word_1}, + put_simpset no_split_ss @{context} addsimps + @{thms rev_rev_ident rev_replicate rev_map to_bl_upt word_size} addsimprocs [word_len_simproc], - no_split_ss addsimps @{thms list.simps split_conv replicate.simps map.simps + put_simpset no_split_ss @{context} addsimps + @{thms list.simps split_conv replicate.simps map.simps zip_Cons_Cons zip_Nil drop_Suc_Cons drop_0 drop_Nil foldr.simps map2_Cons map2_Nil takefill_Suc_Cons takefill_Suc_Nil takefill.Z rbl_succ2_simps @@ -596,20 +607,24 @@ @{cpat drop}, @{cpat "bin_to_bl"}, @{cpat "takefill_last ?x"}, @{cpat "drop_nonempty ?x"}]], - no_split_ss addsimps @{thms xor3_simps carry_simps if_bool_simps} + put_simpset no_split_ss @{context} addsimps @{thms xor3_simps carry_simps if_bool_simps} ]) -val tac = +fun tac ctxt = let val (ss, sss) = expand_word_eq_sss; - in foldr1 (op THEN_ALL_NEW) ((CHANGED o safe_full_simp_tac ss) :: map safe_full_simp_tac sss) end; + in + foldr1 (op THEN_ALL_NEW) + ((CHANGED o safe_full_simp_tac (put_simpset ss ctxt)) :: + map (fn ss => safe_full_simp_tac (put_simpset ss ctxt)) sss) + end; end *} method_setup word_bitwise = - {* Scan.succeed (K (Method.SIMPLE_METHOD (Word_Bitwise_Tac.tac 1))) *} + {* Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (Word_Bitwise_Tac.tac ctxt 1)) *} "decomposer for word equalities and inequalities into bit propositions" end diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/ex/Binary.thy --- a/src/HOL/ex/Binary.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/ex/Binary.thy Thu Apr 18 17:07:01 2013 +0200 @@ -115,15 +115,17 @@ fun plus m n = @{term "plus :: nat \ nat \ nat"} $ m $ n; fun mult m n = @{term "times :: nat \ nat \ nat"} $ m $ n; - val binary_ss = HOL_basic_ss addsimps @{thms binary_simps}; + val binary_ss = + simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms binary_simps}); fun prove ctxt prop = - Goal.prove ctxt [] [] prop (fn _ => ALLGOALS (full_simp_tac binary_ss)); + Goal.prove ctxt [] [] prop + (fn _ => ALLGOALS (full_simp_tac (put_simpset binary_ss ctxt))); - fun binary_proc proc ss ct = + fun binary_proc proc ctxt ct = (case Thm.term_of ct of _ $ t $ u => (case try (pairself (`dest_binary)) (t, u) of - SOME args => proc (Simplifier.the_context ss) args + SOME args => proc ctxt args | NONE => NONE) | _ => NONE); in @@ -170,16 +172,16 @@ simproc_setup binary_nat_mod ("m mod (n::nat)") = {* K (divmod_proc @{thm binary_divmod(2)}) *} method_setup binary_simp = {* - Scan.succeed (K (SIMPLE_METHOD' + Scan.succeed (fn ctxt => SIMPLE_METHOD' (full_simp_tac - (HOL_basic_ss + (put_simpset HOL_basic_ss ctxt addsimps @{thms binary_simps} addsimprocs [@{simproc binary_nat_less_eq}, @{simproc binary_nat_less}, @{simproc binary_nat_diff}, @{simproc binary_nat_div}, - @{simproc binary_nat_mod}])))) + @{simproc binary_nat_mod}]))) *} diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/ex/Numeral_Representation.thy --- a/src/HOL/ex/Numeral_Representation.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/ex/Numeral_Representation.thy Thu Apr 18 17:07:01 2013 +0200 @@ -817,7 +817,7 @@ ML {* structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = struct - val assoc_ss = HOL_ss addsimps @{thms mult_ac_numeral} + val assoc_ss = simpset_of (put_simpset HOL_ss @{context} addsimps @{thms mult_ac_numeral}) val eq_reflection = eq_reflection fun is_numeral (Const(@{const_name of_num}, _) $ _) = true | is_numeral _ = false; diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/ex/Records.thy --- a/src/HOL/ex/Records.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/ex/Records.thy Thu Apr 18 17:07:01 2013 +0200 @@ -254,37 +254,40 @@ *} lemma "(\r. P (xpos r)) \ (\x. P x)" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.split_simproc (K ~1)]) 1 *}) + apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context} + addsimprocs [Record.split_simproc (K ~1)]) 1 *}) apply simp done lemma "(\r. P (xpos r)) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply simp done lemma "(\r. P (xpos r)) \ (\x. P x)" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.split_simproc (K ~1)]) 1 *}) + apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context} + addsimprocs [Record.split_simproc (K ~1)]) 1 *}) apply simp done lemma "(\r. P (xpos r)) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1 *}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1 *}) apply simp done lemma "\r. P (xpos r) \ (\x. P x)" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.split_simproc (K ~1)]) 1 *}) + apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context} + addsimprocs [Record.split_simproc (K ~1)]) 1 *}) apply auto done lemma "\r. P (xpos r) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply auto done lemma "P (xpos r) \ (\x. P x)" - apply (tactic {* Record.split_simp_tac [] (K ~1) 1*}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1*}) apply auto done @@ -295,7 +298,7 @@ assume pre: "P (xpos r)" then have "\x. P x" apply - - apply (tactic {* Record.split_simp_tac [] (K ~1) 1 *}) + apply (tactic {* Record.split_simp_tac @{context} [] (K ~1) 1 *}) apply auto done } @@ -306,7 +309,8 @@ illustrated by the following lemma. *} lemma "\r. xpos r = x" - apply (tactic {* simp_tac (HOL_basic_ss addsimprocs [Record.ex_sel_eq_simproc]) 1 *}) + apply (tactic {* simp_tac (put_simpset HOL_basic_ss @{context} + addsimprocs [Record.ex_sel_eq_simproc]) 1 *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/HOL/ex/Simproc_Tests.thy --- a/src/HOL/ex/Simproc_Tests.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/HOL/ex/Simproc_Tests.thy Thu Apr 18 17:07:01 2013 +0200 @@ -18,7 +18,8 @@ subsection {* ML bindings *} ML {* - fun test ps = CHANGED (asm_simp_tac (HOL_basic_ss addsimprocs ps) 1) + fun test ctxt ps = + CHANGED (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimprocs ps) 1) *} subsection {* Cancellation simprocs from @{text Nat.thy} *} @@ -27,21 +28,21 @@ fix a b c d :: nat { assume "b = Suc c" have "a + b = Suc (c + a)" - by (tactic {* test [@{simproc nateq_cancel_sums}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_sums}] *}) fact next assume "b < Suc c" have "a + b < Suc (c + a)" - by (tactic {* test [@{simproc natless_cancel_sums}] *}) fact + by (tactic {* test @{context} [@{simproc natless_cancel_sums}] *}) fact next assume "b \ Suc c" have "a + b \ Suc (c + a)" - by (tactic {* test [@{simproc natle_cancel_sums}] *}) fact + by (tactic {* test @{context} [@{simproc natle_cancel_sums}] *}) fact next assume "b - Suc c = d" have "a + b - Suc (c + a) = d" - by (tactic {* test [@{simproc natdiff_cancel_sums}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_sums}] *}) fact } end schematic_lemma "\(y::?'b::size). size (?x::?'a::size) \ size y + size ?x" - by (tactic {* test [@{simproc natle_cancel_sums}] *}) (rule le0) + by (tactic {* test @{context} [@{simproc natle_cancel_sums}] *}) (rule le0) (* TODO: test more simprocs with schematic variables *) subsection {* Abelian group cancellation simprocs *} @@ -50,10 +51,10 @@ fix a b c u :: "'a::ab_group_add" { assume "(a + 0) - (b + 0) = u" have "(a + c) - (b + c) = u" - by (tactic {* test [@{simproc group_cancel_diff}] *}) fact + by (tactic {* test @{context} [@{simproc group_cancel_diff}] *}) fact next assume "a + 0 = b + 0" have "a + c = b + c" - by (tactic {* test [@{simproc group_cancel_eq}] *}) fact + by (tactic {* test @{context} [@{simproc group_cancel_eq}] *}) fact } end (* TODO: more tests for Groups.group_cancel_{add,diff,eq,less,le} *) @@ -69,61 +70,61 @@ fix a b c d oo uu i j k l u v w x y z :: "'a::comm_ring_1" { assume "a + - b = u" have "(a + c) - (b + c) = u" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "10 + (2 * l + oo) = uu" have "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = uu" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "-3 + (i + (j + k)) = y" have "(i + j + 12 + k) - 15 = y" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "7 + (i + (j + k)) = y" have "(i + j + 12 + k) - 5 = y" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "-4 * (u * v) + (2 * x + y) = w" have "(2*x - (u*v) + y) - v*3*u = w" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "2 * x * u * v + y = w" have "(2*x*u*v + (u*v)*4 + y) - v*u*4 = w" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "3 * (u * v) + (2 * x * u * v + y) = w" have "(2*x*u*v + (u*v)*4 + y) - v*u = w" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "-3 * (u * v) + (- (x * u * v) + - y) = w" have "u*v - (x*u*v + (u*v)*4 + y) = w" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "a + - c = d" have "a + -(b+c) + b = d" apply (simp only: minus_add_distrib) - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "-2 * b + (a + - c) = d" have "a + -(b+c) - b = d" apply (simp only: minus_add_distrib) - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "-7 + (i + (j + (k + (- u + - y)))) = z" have "(i + j + -2 + k) - (u + 5 + y) = z" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "-27 + (i + (j + k)) = y" have "(i + j + -12 + k) - 15 = y" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "27 + (i + (j + k)) = y" have "(i + j + 12 + k) - -15 = y" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact next assume "3 + (i + (j + k)) = y" have "(i + j + -12 + k) - -15 = y" - by (tactic {* test [@{simproc int_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}] *}) fact } end @@ -133,20 +134,20 @@ fix i j k u vv w y z w' y' z' :: "'a::comm_ring_1" { assume "u = 0" have "2*u = u" - by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc inteq_cancel_numerals}] *}) fact (* conclusion matches Rings.ring_1_no_zero_divisors_class.mult_cancel_right2 *) next assume "i + (j + k) = 3 + (u + y)" have "(i + j + 12 + k) = u + 15 + y" - by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc inteq_cancel_numerals}] *}) fact next assume "7 + (j + (i + k)) = y" have "(i + j*2 + 12 + k) = j + 5 + y" - by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc inteq_cancel_numerals}] *}) fact next assume "u + (6*z + (4*y + 6*w)) = 6*z' + (4*y' + (6*w' + vv))" have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + vv" - by (tactic {* test [@{simproc int_combine_numerals}, @{simproc inteq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc int_combine_numerals}, @{simproc inteq_cancel_numerals}] *}) fact } end @@ -156,18 +157,18 @@ fix b c i j k u y :: "'a::linordered_idom" { assume "y < 2 * b" have "y - b < b" - by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact next assume "c + y < 4 * b" have "y - (3*b + c) < b - 2*c" - by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact next assume "i + (j + k) < 8 + (u + y)" have "(i + j + -3 + k) < u + 5 + y" - by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact next assume "9 + (i + (j + k)) < u + y" have "(i + j + 3 + k) < u + -6 + y" - by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc intless_cancel_numerals}] *}) fact } end @@ -177,22 +178,22 @@ fix x y :: "'a::{idom,ring_char_0}" { assume "3*x = 4*y" have "9*x = 12 * y" - by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact next assume "-3*x = 4*y" have "-99*x = 132 * y" - by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact next assume "111*x = -44*y" have "999*x = -396 * y" - by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact next assume "11*x = 9*y" have "-99*x = -81 * y" - by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact next assume "2*x = y" have "-2 * x = -1 * y" - by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact next assume "2*x = y" have "-2 * x = -y" - by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_numeral_factor}] *}) fact } end @@ -202,20 +203,20 @@ fix x y z :: "'a::{semiring_div,comm_ring_1,ring_char_0}" { assume "(3*x) div (4*y) = z" have "(9*x) div (12*y) = z" - by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact next assume "(-3*x) div (4*y) = z" have "(-99*x) div (132*y) = z" - by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact next assume "(111*x) div (-44*y) = z" have "(999*x) div (-396*y) = z" - by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact next assume "(11*x) div (9*y) = z" have "(-99*x) div (-81*y) = z" - by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact next assume "(2*x) div y = z" have "(-2 * x) div (-1 * y) = z" - by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_numeral_factors}] *}) fact } end @@ -225,22 +226,22 @@ fix x y :: "'a::linordered_idom" { assume "3*x < 4*y" have "9*x < 12 * y" - by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact next assume "-3*x < 4*y" have "-99*x < 132 * y" - by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact next assume "111*x < -44*y" have "999*x < -396 * y" - by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact next assume "9*y < 11*x" have "-99*x < -81 * y" - by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact next assume "y < 2*x" have "-2 * x < -y" - by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact next assume "23*y < x" have "-x < -23 * y" - by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_less_cancel_numeral_factor}] *}) fact } end @@ -250,28 +251,28 @@ fix x y :: "'a::linordered_idom" { assume "3*x \ 4*y" have "9*x \ 12 * y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "-3*x \ 4*y" have "-99*x \ 132 * y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "111*x \ -44*y" have "999*x \ -396 * y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "9*y \ 11*x" have "-99*x \ -81 * y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "y \ 2*x" have "-2 * x \ -1 * y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "23*y \ x" have "-x \ -23 * y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "y \ 0" have "0 \ y * -2" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact next assume "- x \ y" have "- (2 * x) \ 2*y" - by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_le_cancel_numeral_factor}] *}) fact } end @@ -281,19 +282,19 @@ fix x y z :: "'a::{field_inverse_zero,ring_char_0}" { assume "(3*x) / (4*y) = z" have "(9*x) / (12 * y) = z" - by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact next assume "(-3*x) / (4*y) = z" have "(-99*x) / (132 * y) = z" - by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact next assume "(111*x) / (-44*y) = z" have "(999*x) / (-396 * y) = z" - by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact next assume "(11*x) / (9*y) = z" have "(-99*x) / (-81 * y) = z" - by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact next assume "(2*x) / y = z" have "(-2 * x) / (-1 * y) = z" - by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_numeral_factor}] *}) fact } end @@ -303,22 +304,22 @@ fix a b c d k x y :: "'a::idom" { assume "k = 0 \ x = y" have "x*k = k*y" - by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact next assume "k = 0 \ 1 = y" have "k = k*y" - by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact next assume "b = 0 \ a*c = 1" have "a*(b*c) = b" - by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact next assume "a = 0 \ b = 0 \ c = d*x" have "a*(b*c) = d*b*(x*a)" - by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact next assume "k = 0 \ x = y" have "x*k = k*y" - by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact next assume "k = 0 \ 1 = y" have "k = k*y" - by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc ring_eq_cancel_factor}] *}) fact } end @@ -329,19 +330,19 @@ { assume "(if k = 0 then 0 else x div y) = uu" have "(x*k) div (k*y) = uu" - by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact next assume "(if k = 0 then 0 else 1 div y) = uu" have "(k) div (k*y) = uu" - by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact next assume "(if b = 0 then 0 else a * c) = uu" have "(a*(b*c)) div b = uu" - by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact next assume "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu" have "(a*(b*c)) div (d*b*(x*a)) = uu" - by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc int_div_cancel_factor}] *}) fact } end @@ -355,19 +356,19 @@ { assume "(if k = 0 then 0 else x / y) = uu" have "(x*k) / (k*y) = uu" - by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact next assume "(if k = 0 then 0 else 1 / y) = uu" have "(k) / (k*y) = uu" - by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact next assume "(if b = 0 then 0 else a * c / 1) = uu" have "(a*(b*c)) / b = uu" - by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact next assume "(if a = 0 then 0 else if b = 0 then 0 else c / (d * x)) = uu" have "(a*(b*c)) / (d*b*(x*a)) = uu" - by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc divide_cancel_factor}] *}) fact } end @@ -382,22 +383,22 @@ fix x y z :: "'a::linordered_idom" { assume "0 < z \ x < y" have "0 < z \ x*z < y*z" - by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact next assume "0 < z \ x < y" have "0 < z \ x*z < z*y" - by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact next assume "0 < z \ x < y" have "0 < z \ z*x < y*z" - by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact next assume "0 < z \ x < y" have "0 < z \ z*x < z*y" - by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc linordered_ring_less_cancel_factor}] *}) fact next txt "This simproc now uses the simplifier to prove that terms to be canceled are positive/negative." assume z_pos: "0 < z" assume "x < y" have "z*x < z*y" - by (tactic {* CHANGED (asm_simp_tac (HOL_basic_ss + by (tactic {* CHANGED (asm_simp_tac (put_simpset HOL_basic_ss @{context} addsimprocs [@{simproc linordered_ring_less_cancel_factor}] addsimps [@{thm z_pos}]) 1) *}) fact } @@ -409,10 +410,10 @@ fix x y z :: "'a::linordered_idom" { assume "0 < z \ x \ y" have "0 < z \ x*z \ y*z" - by (tactic {* test [@{simproc linordered_ring_le_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc linordered_ring_le_cancel_factor}] *}) fact next assume "0 < z \ x \ y" have "0 < z \ z*x \ z*y" - by (tactic {* test [@{simproc linordered_ring_le_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc linordered_ring_le_cancel_factor}] *}) fact } end @@ -422,28 +423,28 @@ fix x y z uu :: "'a::{field_inverse_zero,ring_char_0}" { assume "5 / 6 * x = uu" have "x / 2 + x / 3 = uu" - by (tactic {* test [@{simproc field_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact next assume "6 / 9 * x + y = uu" have "x / 3 + y + x / 3 = uu" - by (tactic {* test [@{simproc field_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact next assume "9 / 9 * x = uu" have "2 * x / 3 + x / 3 = uu" - by (tactic {* test [@{simproc field_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact next assume "y + z = uu" have "x / 2 + y - 3 * x / 6 + z = uu" - by (tactic {* test [@{simproc field_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact next assume "1 / 15 * x + y = uu" have "7 * x / 5 + y - 4 * x / 3 = uu" - by (tactic {* test [@{simproc field_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc field_combine_numerals}] *}) fact } end lemma fixes x :: "'a::{linordered_field_inverse_zero}" shows "2/3 * x + x / 3 = uu" -apply (tactic {* test [@{simproc field_combine_numerals}] *})? +apply (tactic {* test @{context} [@{simproc field_combine_numerals}] *})? oops -- "FIXME: test fails" subsection {* @{text nat_combine_numerals} *} @@ -452,25 +453,25 @@ fix i j k m n u :: nat { assume "4*k = u" have "k + 3*k = u" - by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact next (* FIXME "Suc (i + 3) \ i + 4" *) assume "4 * Suc 0 + i = u" have "Suc (i + 3) = u" - by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact next (* FIXME "Suc (i + j + 3 + k) \ i + j + 4 + k" *) assume "4 * Suc 0 + (i + (j + k)) = u" have "Suc (i + j + 3 + k) = u" - by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact next assume "2 * j + 4 * k = u" have "k + j + 3*k + j = u" - by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact next assume "6 * Suc 0 + (5 * (i * j) + (4 * k + i)) = u" have "Suc (j*i + i + k + 5 + 3*k + i*j*4) = u" - by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact next assume "5 * (m * n) = u" have "(2*n*m) + (3*(m*n)) = u" - by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nat_combine_numerals}] *}) fact } end @@ -480,44 +481,44 @@ fix i j k l oo u uu vv w y z w' y' z' :: "nat" { assume "Suc 0 * u = 0" have "2*u = (u::nat)" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "Suc 0 * u = Suc 0" have "2*u = Suc (u)" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "i + (j + k) = 3 * Suc 0 + (u + y)" have "(i + j + 12 + k) = u + 15 + y" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "7 * Suc 0 + (i + (j + k)) = u + y" have "(i + j + 12 + k) = u + 5 + y" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "11 * Suc 0 + (i + (j + k)) = u + y" have "(i + j + 12 + k) = Suc (u + y)" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "i + (j + k) = 2 * Suc 0 + (u + y)" have "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "Suc 0 * u + (2 * y + 3 * z) = Suc 0" have "2*y + 3*z + 2*u = Suc (u)" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "Suc 0 * u + (2 * y + (3 * z + (6 * w + (2 * y + 3 * z)))) = Suc 0" have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "Suc 0 * u + (2 * y + (3 * z + (6 * w + (2 * y + 3 * z)))) = 2 * y' + (3 * z' + (6 * w' + (2 * y' + (3 * z' + vv))))" have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + vv" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact next assume "2 * u + (2 * z + (5 * Suc 0 + 2 * y)) = vv" have "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)" - by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc nateq_cancel_numerals}] *}) fact } end @@ -528,17 +529,17 @@ fix c i j k l m oo u uu vv w y z w' y' z' :: "nat" { assume "0 < j" have "(2*length xs < 2*length xs + j)" - by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact next assume "0 < j" have "(2*length xs < length xs * 2 + j)" - by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact next assume "i + (j + k) < u + y" have "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))" - by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact next assume "0 < Suc 0 * (m * n) + u" have "(2*n*m) < (3*(m*n)) + u" - by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natless_cancel_numerals}] *}) fact } end @@ -550,22 +551,22 @@ { assume "u + y \ 36 * Suc 0 + (i + (j + k))" have "Suc (Suc (Suc (Suc (Suc (u + y))))) \ ((i + j) + 41 + k)" - by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact next assume "5 * Suc 0 + (case length (f c) of 0 \ 0 | Suc k \ k) = 0" have "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) \ Suc 0)" - by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact next assume "6 + length l2 = 0" have "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) \ length l1" - by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact next assume "5 + length l3 = 0" have "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) \ length (compT P E A ST mxr e))" - by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact next assume "5 + length (compT P E (A \ A' e) ST mxr c) = 0" have "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un A' e) ST mxr c))))))) \ length (compT P E A ST mxr e))" - by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natle_cancel_numerals}] *}) fact } end @@ -576,54 +577,54 @@ fix c e i j k l oo u uu vv v w x y z zz w' y' z' :: "nat" { assume "i + (j + k) - 3 * Suc 0 = y" have "(i + j + 12 + k) - 15 = y" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "7 * Suc 0 + (i + (j + k)) - 0 = y" have "(i + j + 12 + k) - 5 = y" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "u - Suc 0 * Suc 0 = y" have "Suc u - 2 = y" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "Suc 0 * Suc 0 + u - 0 = y" have "Suc (Suc (Suc u)) - 2 = y" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "Suc 0 * Suc 0 + (i + (j + k)) - 0 = y" have "(i + j + 2 + k) - 1 = y" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "i + (j + k) - Suc 0 * Suc 0 = y" have "(i + j + 1 + k) - 2 = y" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "2 * x + y - 2 * (u * v) = w" have "(2*x + (u*v) + y) - v*3*u = w" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "2 * x * u * v + (5 + y) - 0 = w" have "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = w" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "3 * (u * v) + (2 * x * u * v + y) - 0 = w" have "(2*x*u*v + (u*v)*4 + y) - v*u = w" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "3 * u + (2 + (2 * x * u * v + y)) - 0 = w" have "Suc (Suc (2*x*u*v + u*4 + y)) - u = w" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "Suc (Suc 0 * (u * v)) - 0 = w" have "Suc ((u*v)*4) - v*3*u = w" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "2 - 0 = w" have "Suc (Suc ((u*v)*3)) - v*3*u = w" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "17 * Suc 0 + (i + (j + k)) - (u + y) = zz" have "(i + j + 32 + k) - (u + 15 + y) = zz" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact next assume "u + y - 0 = v" have "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v" - by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact + by (tactic {* test @{context} [@{simproc natdiff_cancel_numerals}] *}) fact } end @@ -637,68 +638,68 @@ fix a b c d k x y uu :: nat { assume "k = 0 \ x = y" have "x*k = k*y" - by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact next assume "k = 0 \ Suc 0 = y" have "k = k*y" - by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact next assume "b = 0 \ a * c = Suc 0" have "a*(b*c) = b" - by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact next assume "a = 0 \ b = 0 \ c = d * x" have "a*(b*c) = d*b*(x*a)" - by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_eq_cancel_factor}] *}) fact next assume "0 < k \ x < y" have "x*k < k*y" - by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact next assume "0 < k \ Suc 0 < y" have "k < k*y" - by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact next assume "0 < b \ a * c < Suc 0" have "a*(b*c) < b" - by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact next assume "0 < a \ 0 < b \ c < d * x" have "a*(b*c) < d*b*(x*a)" - by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_less_cancel_factor}] *}) fact next assume "0 < k \ x \ y" have "x*k \ k*y" - by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact next assume "0 < k \ Suc 0 \ y" have "k \ k*y" - by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact next assume "0 < b \ a * c \ Suc 0" have "a*(b*c) \ b" - by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact next assume "0 < a \ 0 < b \ c \ d * x" have "a*(b*c) \ d*b*(x*a)" - by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_le_cancel_factor}] *}) fact next assume "(if k = 0 then 0 else x div y) = uu" have "(x*k) div (k*y) = uu" - by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact next assume "(if k = 0 then 0 else Suc 0 div y) = uu" have "k div (k*y) = uu" - by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact next assume "(if b = 0 then 0 else a * c) = uu" have "(a*(b*c)) div (b) = uu" - by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact next assume "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu" have "(a*(b*c)) div (d*b*(x*a)) = uu" - by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_div_cancel_factor}] *}) fact next assume "k = 0 \ x dvd y" have "(x*k) dvd (k*y)" - by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact next assume "k = 0 \ Suc 0 dvd y" have "k dvd (k*y)" - by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact next assume "b = 0 \ a * c dvd Suc 0" have "(a*(b*c)) dvd (b)" - by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact next assume "b = 0 \ Suc 0 dvd a * c" have "b dvd (a*(b*c))" - by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact next assume "a = 0 \ b = 0 \ c dvd d * x" have "(a*(b*c)) dvd (d*b*(x*a))" - by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_dvd_cancel_factor}] *}) fact } end @@ -708,19 +709,19 @@ fix x y z :: nat { assume "3 * x = 4 * y" have "9*x = 12 * y" - by (tactic {* test [@{simproc nat_eq_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_eq_cancel_numeral_factor}] *}) fact next assume "3 * x < 4 * y" have "9*x < 12 * y" - by (tactic {* test [@{simproc nat_less_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_less_cancel_numeral_factor}] *}) fact next assume "3 * x \ 4 * y" have "9*x \ 12 * y" - by (tactic {* test [@{simproc nat_le_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_le_cancel_numeral_factor}] *}) fact next assume "(3 * x) div (4 * y) = z" have "(9*x) div (12 * y) = z" - by (tactic {* test [@{simproc nat_div_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_div_cancel_numeral_factor}] *}) fact next assume "(3 * x) dvd (4 * y)" have "(9*x) dvd (12 * y)" - by (tactic {* test [@{simproc nat_dvd_cancel_numeral_factor}] *}) fact + by (tactic {* test @{context} [@{simproc nat_dvd_cancel_numeral_factor}] *}) fact } end @@ -728,39 +729,39 @@ notepad begin have "(10::int) div 3 = 3" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "(10::int) mod 3 = 1" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "(10::int) div -3 = -4" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "(10::int) mod -3 = -2" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "(-10::int) div 3 = -4" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "(-10::int) mod 3 = 2" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "(-10::int) div -3 = 3" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "(-10::int) mod -3 = -1" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "(8452::int) mod 3 = 1" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "(59485::int) div 434 = 137" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "(1000006::int) mod 10 = 6" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "10000000 div 2 = (5000000::int)" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "10000001 mod 2 = (1::int)" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "10000055 div 32 = (312501::int)" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "10000055 mod 32 = (23::int)" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) have "100094 div 144 = (695::int)" - by (tactic {* test [@{simproc binary_int_div}] *}) + by (tactic {* test @{context} [@{simproc binary_int_div}] *}) have "100094 mod 144 = (14::int)" - by (tactic {* test [@{simproc binary_int_mod}] *}) + by (tactic {* test @{context} [@{simproc binary_int_mod}] *}) end end diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/assoc_fold.ML --- a/src/Provers/Arith/assoc_fold.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/assoc_fold.ML Thu Apr 18 17:07:01 2013 +0200 @@ -16,7 +16,7 @@ signature ASSOC_FOLD = sig - val proc: simpset -> term -> thm option + val proc: Proof.context -> term -> thm option end; functor Assoc_Fold(Data: ASSOC_FOLD_DATA): ASSOC_FOLD = @@ -38,15 +38,14 @@ | _ => (lits, t::others)); (*A simproc to combine all literals in a associative nest*) -fun proc ss lhs = +fun proc ctxt lhs = let val plus = (case lhs of f $ _ $ _ => f | _ => error "Assoc_fold: bad pattern") val (lits, others) = sift_terms plus (lhs, ([],[])) val _ = length lits < 2 andalso raise Assoc_fail (*we can't reduce the number of terms*) val rhs = plus $ mk_sum plus lits $ mk_sum plus others - val th = Goal.prove (Simplifier.the_context ss) [] [] (Logic.mk_equals (lhs, rhs)) - (fn _ => rtac Data.eq_reflection 1 THEN - simp_tac (Simplifier.inherit_context ss Data.assoc_ss) 1) + val th = Goal.prove ctxt [] [] (Logic.mk_equals (lhs, rhs)) + (fn _ => rtac Data.eq_reflection 1 THEN simp_tac (put_simpset Data.assoc_ss ctxt) 1) in SOME th end handle Assoc_fail => NONE; end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/cancel_div_mod.ML --- a/src/Provers/Arith/cancel_div_mod.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/cancel_div_mod.ML Thu Apr 18 17:07:01 2013 +0200 @@ -21,13 +21,13 @@ val div_mod_eqs: thm list (* (n*(m div n) + m mod n) + k == m + k and ((m div n)*n + m mod n) + k == m + k *) - val prove_eq_sums: simpset -> term * term -> thm + val prove_eq_sums: Proof.context -> term * term -> thm (* must prove ac0-equivalence of sums *) end; signature CANCEL_DIV_MOD = sig - val proc: simpset -> cterm -> thm option + val proc: Proof.context -> cterm -> thm option end; @@ -66,11 +66,11 @@ val m = Data.mk_binop Data.mod_name pq in mk_plus(mk_plus(d,m),Data.mk_sum(ts |> remove (op =) d |> remove (op =) m)) end -fun cancel ss t pq = - let val teqt' = Data.prove_eq_sums ss (t, rearrange t pq) +fun cancel ctxt t pq = + let val teqt' = Data.prove_eq_sums ctxt (t, rearrange t pq) in hd (Data.div_mod_eqs RL [teqt' RS transitive_thm]) end; -fun proc ss ct = +fun proc ctxt ct = let val t = term_of ct; val (divs, mods) = coll_div_mod t ([], []); @@ -78,7 +78,7 @@ if null divs orelse null mods then NONE else (case inter (op =) mods divs of - pq :: _ => SOME (cancel ss t pq) + pq :: _ => SOME (cancel ctxt t pq) | [] => NONE) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/cancel_numeral_factor.ML --- a/src/Provers/Arith/cancel_numeral_factor.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/cancel_numeral_factor.ML Thu Apr 18 17:07:01 2013 +0200 @@ -29,27 +29,27 @@ as with < and <= but not = and div*) (*proof tools*) val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option - val trans_tac: thm option -> tactic (*applies the initial lemma*) - val norm_tac: simpset -> tactic (*proves the initial lemma*) - val numeral_simp_tac: simpset -> tactic (*proves the final theorem*) - val simplify_meta_eq: simpset -> thm -> thm (*simplifies the final theorem*) + val trans_tac: thm option -> tactic (*applies the initial lemma*) + val norm_tac: Proof.context -> tactic (*proves the initial lemma*) + val numeral_simp_tac: Proof.context -> tactic (*proves the final theorem*) + val simplify_meta_eq: Proof.context -> thm -> thm (*simplifies the final theorem*) end; functor CancelNumeralFactorFun(Data: CANCEL_NUMERAL_FACTOR_DATA): sig - val proc: simpset -> term -> thm option + val proc: Proof.context -> term -> thm option end = struct (*the simplification procedure*) -fun proc ss t = +fun proc ctxt t = let - val ctxt = Simplifier.the_context ss; - val prems = Simplifier.prems_of ss; + val prems = Simplifier.prems_of ctxt; val ([t'], ctxt') = Variable.import_terms true [t] ctxt val export = singleton (Variable.export ctxt' ctxt) + (* FIXME ctxt vs. ctxt' *) val (t1,t2) = Data.dest_bal t' val (m1, t1') = Data.dest_coeff t1 @@ -67,13 +67,13 @@ else Data.mk_bal (Data.mk_coeff(n1,t1'), Data.mk_coeff(n2,t2')) val reshape = (*Move d to the front and put the rest into standard form i * #m * j == #d * (#n * (j * k)) *) - Data.prove_conv [Data.norm_tac ss] ctxt prems + Data.prove_conv [Data.norm_tac ctxt] ctxt prems (t', Data.mk_bal (newshape(n1,t1'), newshape(n2,t2'))) in - Option.map (export o Data.simplify_meta_eq ss) + Option.map (export o Data.simplify_meta_eq ctxt) (Data.prove_conv [Data.trans_tac reshape, rtac Data.cancel 1, - Data.numeral_simp_tac ss] ctxt prems (t', rhs)) + Data.numeral_simp_tac ctxt] ctxt prems (t', rhs)) end (* FIXME avoid handling of generic exceptions *) handle TERM _ => NONE diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/cancel_numerals.ML --- a/src/Provers/Arith/cancel_numerals.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/cancel_numerals.ML Thu Apr 18 17:07:01 2013 +0200 @@ -37,14 +37,14 @@ (*proof tools*) val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option val trans_tac: thm option -> tactic (*applies the initial lemma*) - val norm_tac: simpset -> tactic (*proves the initial lemma*) - val numeral_simp_tac: simpset -> tactic (*proves the final theorem*) - val simplify_meta_eq: simpset -> thm -> thm (*simplifies the final theorem*) + val norm_tac: Proof.context -> tactic (*proves the initial lemma*) + val numeral_simp_tac: Proof.context -> tactic (*proves the final theorem*) + val simplify_meta_eq: Proof.context -> thm -> thm (*simplifies the final theorem*) end; signature CANCEL_NUMERALS = sig - val proc: simpset -> term -> thm option + val proc: Proof.context -> term -> thm option end; functor CancelNumeralsFun(Data: CANCEL_NUMERALS_DATA): CANCEL_NUMERALS = @@ -65,12 +65,12 @@ in seek terms1 end; (*the simplification procedure*) -fun proc ss t = +fun proc ctxt t = let - val ctxt = Simplifier.the_context ss - val prems = Simplifier.prems_of ss + val prems = Simplifier.prems_of ctxt val ([t'], ctxt') = Variable.import_terms true [t] ctxt val export = singleton (Variable.export ctxt' ctxt) + (* FIXME ctxt cs. ctxt' (!?) *) val (t1,t2) = Data.dest_bal t' val terms1 = Data.dest_sum t1 @@ -86,19 +86,19 @@ i + #m + j + k == #m + i + (j + k) *) if n1=0 orelse n2=0 then (*trivial, so do nothing*) raise TERM("cancel_numerals", []) - else Data.prove_conv [Data.norm_tac ss] ctxt prems + else Data.prove_conv [Data.norm_tac ctxt] ctxt prems (t', Data.mk_bal (newshape(n1,terms1'), newshape(n2,terms2'))) in - Option.map (export o Data.simplify_meta_eq ss) + Option.map (export o Data.simplify_meta_eq ctxt) (if n2 <= n1 then Data.prove_conv [Data.trans_tac reshape, rtac Data.bal_add1 1, - Data.numeral_simp_tac ss] ctxt prems + Data.numeral_simp_tac ctxt] ctxt prems (t', Data.mk_bal (newshape(n1-n2,terms1'), Data.mk_sum T terms2')) else Data.prove_conv [Data.trans_tac reshape, rtac Data.bal_add2 1, - Data.numeral_simp_tac ss] ctxt prems + Data.numeral_simp_tac ctxt] ctxt prems (t', Data.mk_bal (Data.mk_sum T terms1', newshape(n2-n1,terms2')))) end (* FIXME avoid handling of generic exceptions *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/combine_numerals.ML --- a/src/Provers/Arith/combine_numerals.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/combine_numerals.ML Thu Apr 18 17:07:01 2013 +0200 @@ -30,15 +30,15 @@ (*proof tools*) val prove_conv: tactic list -> Proof.context -> term * term -> thm option val trans_tac: thm option -> tactic (*applies the initial lemma*) - val norm_tac: simpset -> tactic (*proves the initial lemma*) - val numeral_simp_tac: simpset -> tactic (*proves the final theorem*) - val simplify_meta_eq: simpset -> thm -> thm (*simplifies the final theorem*) + val norm_tac: Proof.context -> tactic (*proves the initial lemma*) + val numeral_simp_tac: Proof.context -> tactic (*proves the final theorem*) + val simplify_meta_eq: Proof.context -> thm -> thm (*simplifies the final theorem*) end; functor CombineNumeralsFun(Data: COMBINE_NUMERALS_DATA): sig - val proc: simpset -> term -> thm option + val proc: Proof.context -> term -> thm option end = struct @@ -65,11 +65,11 @@ | NONE => find_repeated (tab, t::past, terms); (*the simplification procedure*) -fun proc ss t = +fun proc ctxt t = let - val ctxt = Simplifier.the_context ss; val ([t'], ctxt') = Variable.import_terms true [t] ctxt val export = singleton (Variable.export ctxt' ctxt) + (* FIXME ctxt vs. ctxt' (!?) *) val (u,m,n,terms) = find_repeated (Termtab.empty, [], Data.dest_sum t') val T = Term.fastype_of u @@ -78,13 +78,13 @@ i + #m + j + k == #m + i + (j + k) *) if Data.iszero m orelse Data.iszero n then (*trivial, so do nothing*) raise TERM("combine_numerals", []) - else Data.prove_conv [Data.norm_tac ss] ctxt + else Data.prove_conv [Data.norm_tac ctxt] ctxt (t', Data.mk_sum T ([Data.mk_coeff (m, u), Data.mk_coeff (n, u)] @ terms)) in - Option.map (export o Data.simplify_meta_eq ss) + Option.map (export o Data.simplify_meta_eq ctxt) (Data.prove_conv [Data.trans_tac reshape, rtac Data.left_distrib 1, - Data.numeral_simp_tac ss] ctxt + Data.numeral_simp_tac ctxt] ctxt (t', Data.mk_sum T (Data.mk_coeff(Data.add(m,n), u) :: terms))) end (* FIXME avoid handling of generic exceptions *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/extract_common_term.ML --- a/src/Provers/Arith/extract_common_term.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/extract_common_term.ML Thu Apr 18 17:07:01 2013 +0200 @@ -23,15 +23,15 @@ val find_first: term -> term list -> term list (*proof tools*) val mk_eq: term * term -> term - val norm_tac: simpset -> tactic (*proves the result*) - val simplify_meta_eq: simpset -> thm -> thm -> thm (*simplifies the result*) - val simp_conv: simpset -> term -> thm option (*proves simp thm*) + val norm_tac: Proof.context -> tactic (*proves the result*) + val simplify_meta_eq: Proof.context -> thm -> thm -> thm (*simplifies the result*) + val simp_conv: Proof.context -> term -> thm option (*proves simp thm*) end; functor ExtractCommonTermFun(Data: EXTRACT_COMMON_TERM_DATA): sig - val proc: simpset -> term -> thm option + val proc: Proof.context -> term -> thm option end = struct @@ -46,12 +46,12 @@ in seek terms1 end; (*the simplification procedure*) -fun proc ss t = +fun proc ctxt t = let - val ctxt = Simplifier.the_context ss; - val prems = Simplifier.prems_of ss; + val prems = Simplifier.prems_of ctxt; val ([t'], ctxt') = Variable.import_terms true [t] ctxt val export = singleton (Variable.export ctxt' ctxt) + (* FIXME ctxt vs. ctxt' (!?) *) val (t1,t2) = Data.dest_bal t' val terms1 = Data.dest_sum t1 @@ -59,7 +59,7 @@ val u = find_common (terms1,terms2) val simp_th = - case Data.simp_conv ss u of NONE => raise TERM("no simp", []) + case Data.simp_conv ctxt u of NONE => raise TERM("no simp", []) | SOME th => th val terms1' = Data.find_first u terms1 and terms2' = Data.find_first u terms2 @@ -67,10 +67,10 @@ val t'' = Data.mk_bal (Data.mk_sum T (u::terms1'), Data.mk_sum T (u::terms2')) val reshape = - Goal.prove ctxt [] [] (Data.mk_eq (t', t'')) (K (Data.norm_tac ss)) + Goal.prove ctxt [] [] (Data.mk_eq (t', t'')) (K (Data.norm_tac ctxt)) in - SOME (export (Data.simplify_meta_eq ss simp_th reshape)) + SOME (export (Data.simplify_meta_eq ctxt simp_th reshape)) end (* FIXME avoid handling of generic exceptions *) handle TERM _ => NONE diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/Arith/fast_lin_arith.ML --- a/src/Provers/Arith/fast_lin_arith.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/Arith/fast_lin_arith.ML Thu Apr 18 17:07:01 2013 +0200 @@ -53,7 +53,7 @@ val pre_decomp: Proof.context -> typ list * term list -> (typ list * term list) list (*preprocessing, performed on the goal -- must do the same as 'pre_decomp':*) - val pre_tac: simpset -> int -> tactic + val pre_tac: Proof.context -> int -> tactic (*the limit on the number of ~= allowed; because each ~= is split into two cases, this can lead to an explosion*) @@ -79,22 +79,22 @@ beta-equivalent modifications, as some of the lin. arith. code is not insensitive to them.) -ss must reduce contradictory <= to False. +Simpset must reduce contradictory <= to False. It should also cancel common summands to keep <= reduced; otherwise <= can grow to massive proportions. *) signature FAST_LIN_ARITH = sig - val prems_lin_arith_tac: simpset -> int -> tactic + val prems_lin_arith_tac: Proof.context -> int -> tactic val lin_arith_tac: Proof.context -> bool -> int -> tactic - val lin_arith_simproc: simpset -> term -> thm option + val lin_arith_simproc: Proof.context -> term -> thm option val map_data: ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list, - lessD: thm list, neqE: thm list, simpset: Simplifier.simpset, + lessD: thm list, neqE: thm list, simpset: simpset, number_of: (theory -> typ -> int -> cterm) option} -> {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list, - lessD: thm list, neqE: thm list, simpset: Simplifier.simpset, + lessD: thm list, neqE: thm list, simpset: simpset, number_of: (theory -> typ -> int -> cterm) option}) -> Context.generic -> Context.generic val add_inj_thms: thm list -> Context.generic -> Context.generic @@ -119,12 +119,12 @@ inj_thms: thm list, lessD: thm list, neqE: thm list, - simpset: Simplifier.simpset, + simpset: simpset, number_of: (theory -> typ -> int -> cterm) option}; val empty : T = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [], - lessD = [], neqE = [], simpset = Simplifier.empty_ss, + lessD = [], neqE = [], simpset = empty_ss, number_of = NONE}; val extend = I; fun merge @@ -137,7 +137,7 @@ inj_thms = Thm.merge_thms (inj_thms1, inj_thms2), lessD = Thm.merge_thms (lessD1, lessD2), neqE = Thm.merge_thms (neqE1, neqE2), - simpset = Simplifier.merge_ss (simpset1, simpset2), + simpset = merge_ss (simpset1, simpset2), number_of = merge_options (number_of1, number_of2)}; ); @@ -152,14 +152,16 @@ {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms, lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of}; -fun map_simpset f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} = - {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms, - lessD = lessD, neqE = neqE, simpset = f simpset, number_of = number_of}; +fun map_simpset f context = + map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} => + {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms, + lessD = lessD, neqE = neqE, simpset = simpset_map (Context.proof_of context) f simpset, + number_of = number_of}) context; fun add_inj_thms thms = map_data (map_inj_thms (append thms)); fun add_lessD thm = map_data (map_lessD (fn thms => thms @ [thm])); -fun add_simps thms = map_data (map_simpset (fn simpset => simpset addsimps thms)); -fun add_simprocs procs = map_data (map_simpset (fn simpset => simpset addsimprocs procs)); +fun add_simps thms = map_simpset (fn ctxt => ctxt addsimps thms); +fun add_simprocs procs = map_simpset (fn ctxt => ctxt addsimprocs procs); fun set_number_of f = map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, ...} => @@ -473,13 +475,12 @@ exception FalseE of thm in -fun mkthm ss asms (just: injust) = +fun mkthm ctxt asms (just: injust) = let - val ctxt = Simplifier.the_context ss; val thy = Proof_Context.theory_of ctxt; val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset, ...} = get_data ctxt; val number_of = number_of ctxt; - val simpset' = Simplifier.inherit_context ss simpset; + val simpset_ctxt = put_simpset simpset ctxt; fun only_concl f thm = if Thm.no_prems thm then f (Thm.concl_of thm) else NONE; val atoms = atoms_of (map_filter (only_concl (LA_Data.decomp ctxt)) asms); @@ -507,7 +508,7 @@ let fun mul i th = if i = 1 then th else mul (i - 1) (add_thms thm th) in mul n thm end; - val rewr = Simplifier.rewrite simpset'; + val rewr = Simplifier.rewrite simpset_ctxt; val rewrite_concl = Conv.fconv_rule (Conv.concl_conv ~1 (Conv.arg_conv (Conv.binop_conv rewr))); fun discharge_prem thm = if Thm.nprems_of thm = 0 then thm else @@ -537,7 +538,7 @@ else mult n thm; fun simp thm = - let val thm' = trace_thm ctxt ["Simplified:"] (full_simplify simpset' thm) + let val thm' = trace_thm ctxt ["Simplified:"] (full_simplify simpset_ctxt thm) in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end; fun mk (Asm i) = trace_thm ctxt ["Asm " ^ string_of_int i] (nth asms i) @@ -554,7 +555,7 @@ let val _ = trace_msg ctxt "mkthm"; val thm = trace_thm ctxt ["Final thm:"] (mk just); - val fls = simplify simpset' thm; + val fls = simplify simpset_ctxt thm; val _ = trace_thm ctxt ["After simplification:"] fls; val _ = if LA_Logic.is_False fls then () @@ -799,10 +800,9 @@ (trace_msg ctxt "prove failed (cannot negate conclusion)."; (false, NONE)); -fun refute_tac ss (i, split_neq, justs) = +fun refute_tac ctxt (i, split_neq, justs) = fn state => let - val ctxt = Simplifier.the_context ss; val _ = trace_thm ctxt ["refute_tac (on subgoal " ^ string_of_int i ^ ", with " ^ string_of_int (length justs) ^ " justification(s)):"] state @@ -815,12 +815,12 @@ all_tac) THEN PRIMITIVE (trace_thm ctxt ["State after neqE:"]) THEN (* use theorems generated from the actual justifications *) - Subgoal.FOCUS (fn {prems, ...} => rtac (mkthm ss prems j) 1) ctxt i + Subgoal.FOCUS (fn {prems, ...} => rtac (mkthm ctxt prems j) 1) ctxt i in (* rewrite "[| A1; ...; An |] ==> B" to "[| A1; ...; An; ~B |] ==> False" *) DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i) THEN (* user-defined preprocessing of the subgoal *) - DETERM (LA_Data.pre_tac ss i) THEN + DETERM (LA_Data.pre_tac ctxt i) THEN PRIMITIVE (trace_thm ctxt ["State after pre_tac:"]) THEN (* prove every resulting subgoal, using its justification *) EVERY (map just1 justs) @@ -830,9 +830,8 @@ Fast but very incomplete decider. Only premises and conclusions that are already (negated) (in)equations are taken into account. *) -fun simpset_lin_arith_tac ss show_ex = SUBGOAL (fn (A, i) => +fun simpset_lin_arith_tac ctxt show_ex = SUBGOAL (fn (A, i) => let - val ctxt = Simplifier.the_context ss val params = rev (Logic.strip_params A) val Hs = Logic.strip_assums_hyp A val concl = Logic.strip_assums_concl A @@ -841,15 +840,15 @@ case prove ctxt params show_ex true Hs concl of (_, NONE) => (trace_msg ctxt "Refutation failed."; no_tac) | (split_neq, SOME js) => (trace_msg ctxt "Refutation succeeded."; - refute_tac ss (i, split_neq, js)) + refute_tac ctxt (i, split_neq, js)) end); -fun prems_lin_arith_tac ss = - Method.insert_tac (Simplifier.prems_of ss) THEN' - simpset_lin_arith_tac ss false; +fun prems_lin_arith_tac ctxt = + Method.insert_tac (Simplifier.prems_of ctxt) THEN' + simpset_lin_arith_tac ctxt false; fun lin_arith_tac ctxt = - simpset_lin_arith_tac (Simplifier.context ctxt Simplifier.empty_ss); + simpset_lin_arith_tac (empty_simpset ctxt); @@ -892,25 +891,24 @@ | NONE => elim_neq neqs (asm::asms', asms)) in elim_neq neqE ([], asms) end; -fun fwdproof ss (Tip asms : splittree) (j::js : injust list) = (mkthm ss asms j, js) - | fwdproof ss (Spl (thm, ct1, tree1, ct2, tree2)) js = +fun fwdproof ctxt (Tip asms : splittree) (j::js : injust list) = (mkthm ctxt asms j, js) + | fwdproof ctxt (Spl (thm, ct1, tree1, ct2, tree2)) js = let - val (thm1, js1) = fwdproof ss tree1 js - val (thm2, js2) = fwdproof ss tree2 js1 + val (thm1, js1) = fwdproof ctxt tree1 js + val (thm2, js2) = fwdproof ctxt tree2 js1 val thm1' = Thm.implies_intr ct1 thm1 val thm2' = Thm.implies_intr ct2 thm2 in (thm2' COMP (thm1' COMP thm), js2) end; (* FIXME needs handle THM _ => NONE ? *) -fun prover ss thms Tconcl (js : injust list) split_neq pos : thm option = +fun prover ctxt thms Tconcl (js : injust list) split_neq pos : thm option = let - val ctxt = Simplifier.the_context ss val thy = Proof_Context.theory_of ctxt val nTconcl = LA_Logic.neg_prop Tconcl val cnTconcl = cterm_of thy nTconcl val nTconclthm = Thm.assume cnTconcl val tree = (if split_neq then splitasms ctxt else Tip) (thms @ [nTconclthm]) - val (Falsethm, _) = fwdproof ss tree js + val (Falsethm, _) = fwdproof ctxt tree js val contr = if pos then LA_Logic.ccontr else LA_Logic.notI val concl = Thm.implies_intr cnTconcl Falsethm COMP contr in SOME (trace_thm ctxt ["Proved by lin. arith. prover:"] (LA_Logic.mk_Eq concl)) end @@ -923,19 +921,18 @@ 2. lin_arith_simproc is applied by the Simplifier which dives into terms and will thus try the non-negated concl anyway. *) -fun lin_arith_simproc ss concl = +fun lin_arith_simproc ctxt concl = let - val ctxt = Simplifier.the_context ss - val thms = maps LA_Logic.atomize (Simplifier.prems_of ss) + val thms = maps LA_Logic.atomize (Simplifier.prems_of ctxt) val Hs = map Thm.prop_of thms val Tconcl = LA_Logic.mk_Trueprop concl in case prove ctxt [] false false Hs Tconcl of (* concl provable? *) - (split_neq, SOME js) => prover ss thms Tconcl js split_neq true + (split_neq, SOME js) => prover ctxt thms Tconcl js split_neq true | (_, NONE) => let val nTconcl = LA_Logic.neg_prop Tconcl in case prove ctxt [] false false Hs nTconcl of (* ~concl provable? *) - (split_neq, SOME js) => prover ss thms nTconcl js split_neq false + (split_neq, SOME js) => prover ctxt thms nTconcl js split_neq false | (_, NONE) => NONE end end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/clasimp.ML --- a/src/Provers/clasimp.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/clasimp.ML Thu Apr 18 17:07:01 2013 +0200 @@ -44,7 +44,8 @@ (* simp as classical wrapper *) -fun clasimp f name tac ctxt = f (ctxt, (name, CHANGED o tac (simpset_of ctxt))); +(* FIXME !? *) +fun clasimp f name tac ctxt = f (ctxt, (name, fn _ => CHANGED o tac ctxt)); (*Add a simpset to the claset!*) (*Caution: only one simpset added can be added by each of addSss and addss*) @@ -111,7 +112,7 @@ (* tactics *) fun clarsimp_tac ctxt = - Simplifier.safe_asm_full_simp_tac (simpset_of ctxt) THEN_ALL_NEW + Simplifier.safe_asm_full_simp_tac ctxt THEN_ALL_NEW Classical.clarify_tac (addSss ctxt); @@ -145,7 +146,7 @@ ORELSE' (CHANGED o nodup_depth_tac (addss ctxt) n); (* slower but more general *) in - PARALLEL_GOALS (ALLGOALS (Simplifier.asm_full_simp_tac (simpset_of ctxt))) THEN + PARALLEL_GOALS (ALLGOALS (Simplifier.asm_full_simp_tac ctxt)) THEN TRY (Classical.safe_tac ctxt) THEN REPEAT_DETERM (FIRSTGOAL main_tac) THEN TRY (Classical.safe_tac (addSss ctxt)) THEN @@ -162,7 +163,7 @@ let val ctxt' = addss ctxt in SELECT_GOAL (Classical.clarify_tac ctxt' 1 THEN - IF_UNSOLVED (Simplifier.asm_full_simp_tac (simpset_of ctxt) 1) THEN + IF_UNSOLVED (Simplifier.asm_full_simp_tac ctxt 1) THEN ALLGOALS (Classical.first_best_tac ctxt')) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/classical.ML --- a/src/Provers/classical.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/classical.ML Thu Apr 18 17:07:01 2013 +0200 @@ -45,10 +45,10 @@ val delSWrapper: Proof.context * string -> Proof.context val addWrapper: Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context val delWrapper: Proof.context * string -> Proof.context - val addSbefore: Proof.context * (string * (int -> tactic)) -> Proof.context - val addSafter: Proof.context * (string * (int -> tactic)) -> Proof.context - val addbefore: Proof.context * (string * (int -> tactic)) -> Proof.context - val addafter: Proof.context * (string * (int -> tactic)) -> Proof.context + val addSbefore: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context + val addSafter: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context + val addbefore: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context + val addafter: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context val addD2: Proof.context * (string * thm) -> Proof.context val addE2: Proof.context * (string * thm) -> Proof.context val addSD2: Proof.context * (string * thm) -> Proof.context @@ -670,17 +670,21 @@ (map_uwrappers (delete_warn ("No such unsafe wrapper in claset: " ^ name) name)); (* compose a safe tactic alternatively before/after safe_step_tac *) -fun ctxt addSbefore (name, tac1) = ctxt addSWrapper (name, fn _ => fn tac2 => tac1 ORELSE' tac2); -fun ctxt addSafter (name, tac2) = ctxt addSWrapper (name, fn _ => fn tac1 => tac1 ORELSE' tac2); +fun ctxt addSbefore (name, tac1) = + ctxt addSWrapper (name, fn ctxt => fn tac2 => tac1 ctxt ORELSE' tac2); +fun ctxt addSafter (name, tac2) = + ctxt addSWrapper (name, fn ctxt => fn tac1 => tac1 ORELSE' tac2 ctxt); (*compose a tactic alternatively before/after the step tactic *) -fun ctxt addbefore (name, tac1) = ctxt addWrapper (name, fn _ => fn tac2 => tac1 APPEND' tac2); -fun ctxt addafter (name, tac2) = ctxt addWrapper (name, fn _ => fn tac1 => tac1 APPEND' tac2); +fun ctxt addbefore (name, tac1) = + ctxt addWrapper (name, fn ctxt => fn tac2 => tac1 ctxt APPEND' tac2); +fun ctxt addafter (name, tac2) = + ctxt addWrapper (name, fn ctxt => fn tac1 => tac1 APPEND' tac2 ctxt); -fun ctxt addD2 (name, thm) = ctxt addafter (name, dtac thm THEN' assume_tac); -fun ctxt addE2 (name, thm) = ctxt addafter (name, etac thm THEN' assume_tac); -fun ctxt addSD2 (name, thm) = ctxt addSafter (name, dmatch_tac [thm] THEN' eq_assume_tac); -fun ctxt addSE2 (name, thm) = ctxt addSafter (name, ematch_tac [thm] THEN' eq_assume_tac); +fun ctxt addD2 (name, thm) = ctxt addafter (name, fn _ => dtac thm THEN' assume_tac); +fun ctxt addE2 (name, thm) = ctxt addafter (name, fn _ => etac thm THEN' assume_tac); +fun ctxt addSD2 (name, thm) = ctxt addSafter (name, fn _ => dmatch_tac [thm] THEN' eq_assume_tac); +fun ctxt addSE2 (name, thm) = ctxt addSafter (name, fn _ => ematch_tac [thm] THEN' eq_assume_tac); diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/hypsubst.ML --- a/src/Provers/hypsubst.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/hypsubst.ML Thu Apr 18 17:07:01 2013 +0200 @@ -48,7 +48,7 @@ sig val bound_hyp_subst_tac : int -> tactic val hyp_subst_tac : int -> tactic - val hyp_subst_tac' : simpset -> int -> tactic + val hyp_subst_tac' : Proof.context -> int -> tactic val blast_hyp_subst_tac : bool -> int -> tactic val stac : thm -> int -> tactic val hypsubst_setup : theory -> theory @@ -127,15 +127,16 @@ (*Select a suitable equality assumption; substitute throughout the subgoal If bnd is true, then it replaces Bound variables only. *) - fun gen_hyp_subst_tac opt_ss bnd = + fun gen_hyp_subst_tac opt_ctxt bnd = let fun tac i st = SUBGOAL (fn (Bi, _) => let val (k, _) = eq_var bnd true Bi - val map_simpset = case opt_ss of - NONE => Simplifier.global_context (Thm.theory_of_thm st) - | SOME ss => Simplifier.inherit_context ss - val hyp_subst_ss = map_simpset empty_ss |> Simplifier.set_mksimps (K (mk_eqs bnd)) - in EVERY [rotate_tac k i, asm_lr_simp_tac hyp_subst_ss i, + val ctxt = + (case opt_ctxt of + NONE => Proof_Context.init_global (Thm.theory_of_thm st) + | SOME ctxt => ctxt) + val hyp_subst_ctxt = empty_simpset ctxt |> Simplifier.set_mksimps (K (mk_eqs bnd)) + in EVERY [rotate_tac k i, asm_lr_simp_tac hyp_subst_ctxt i, etac thin_rl i, rotate_tac (~k) i] end handle THM _ => no_tac | EQ_VAR => no_tac) i st in REPEAT_DETERM1 o tac end; @@ -200,8 +201,8 @@ val hyp_subst_tac = FIRST' [ematch_tac [Data.thin_refl], gen_hyp_subst_tac NONE false, vars_gen_hyp_subst_tac false]; -fun hyp_subst_tac' ss = FIRST' [ematch_tac [Data.thin_refl], - gen_hyp_subst_tac (SOME ss) false, vars_gen_hyp_subst_tac false]; +fun hyp_subst_tac' ctxt = FIRST' [ematch_tac [Data.thin_refl], + gen_hyp_subst_tac (SOME ctxt) false, vars_gen_hyp_subst_tac false]; (*Substitutes for Bound variables only -- this is always safe*) val bound_hyp_subst_tac = diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/quantifier1.ML --- a/src/Provers/quantifier1.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/quantifier1.ML Thu Apr 18 17:07:01 2013 +0200 @@ -66,11 +66,11 @@ sig val prove_one_point_all_tac: tactic val prove_one_point_ex_tac: tactic - val rearrange_all: simpset -> cterm -> thm option - val rearrange_ex: simpset -> cterm -> thm option - val rearrange_ball: tactic -> simpset -> cterm -> thm option - val rearrange_bex: tactic -> simpset -> cterm -> thm option - val rearrange_Collect: tactic -> simpset -> cterm -> thm option + val rearrange_all: Proof.context -> cterm -> thm option + val rearrange_ex: Proof.context -> cterm -> thm option + val rearrange_ball: tactic -> Proof.context -> cterm -> thm option + val rearrange_bex: tactic -> Proof.context -> cterm -> thm option + val rearrange_Collect: tactic -> Proof.context -> cterm -> thm option end; functor Quantifier1(Data: QUANTIFIER1_DATA): QUANTIFIER1 = @@ -120,9 +120,8 @@ | exqu xs P = extract (null xs) xs P in exqu [] end; -fun prove_conv tac ss tu = +fun prove_conv tac ctxt tu = let - val ctxt = Simplifier.the_context ss; val (goal, ctxt') = yield_singleton (Variable.import_terms true) (Logic.mk_equals tu) ctxt; val thm = @@ -172,17 +171,17 @@ val Q = if n = 0 then P else renumber 0 n P; in quant xs (qC $ Abs (x, T, Q)) end; -fun rearrange_all ss ct = +fun rearrange_all ctxt ct = (case term_of ct of F as (all as Const (q, _)) $ Abs (x, T, P) => (case extract_quant extract_imp q P of NONE => NONE | SOME (xs, eq, Q) => let val R = quantify all x T xs (Data.imp $ eq $ Q) - in SOME (prove_conv prove_one_point_all_tac ss (F, R)) end) + in SOME (prove_conv prove_one_point_all_tac ctxt (F, R)) end) | _ => NONE); -fun rearrange_ball tac ss ct = +fun rearrange_ball tac ctxt ct = (case term_of ct of F as Ball $ A $ Abs (x, T, P) => (case extract_imp true [] P of @@ -191,37 +190,37 @@ if not (null xs) then NONE else let val R = Data.imp $ eq $ Q - in SOME (prove_conv tac ss (F, Ball $ A $ Abs (x, T, R))) end) + in SOME (prove_conv tac ctxt (F, Ball $ A $ Abs (x, T, R))) end) | _ => NONE); -fun rearrange_ex ss ct = +fun rearrange_ex ctxt ct = (case term_of ct of F as (ex as Const (q, _)) $ Abs (x, T, P) => (case extract_quant extract_conj q P of NONE => NONE | SOME (xs, eq, Q) => let val R = quantify ex x T xs (Data.conj $ eq $ Q) - in SOME (prove_conv prove_one_point_ex_tac ss (F, R)) end) + in SOME (prove_conv prove_one_point_ex_tac ctxt (F, R)) end) | _ => NONE); -fun rearrange_bex tac ss ct = +fun rearrange_bex tac ctxt ct = (case term_of ct of F as Bex $ A $ Abs (x, T, P) => (case extract_conj true [] P of NONE => NONE | SOME (xs, eq, Q) => if not (null xs) then NONE - else SOME (prove_conv tac ss (F, Bex $ A $ Abs (x, T, Data.conj $ eq $ Q)))) + else SOME (prove_conv tac ctxt (F, Bex $ A $ Abs (x, T, Data.conj $ eq $ Q)))) | _ => NONE); -fun rearrange_Collect tac ss ct = +fun rearrange_Collect tac ctxt ct = (case term_of ct of F as Collect $ Abs (x, T, P) => (case extract_conj true [] P of NONE => NONE | SOME (_, eq, Q) => let val R = Collect $ Abs (x, T, Data.conj $ eq $ Q) - in SOME (prove_conv tac ss (F, R)) end) + in SOME (prove_conv tac ctxt (F, R)) end) | _ => NONE); end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Provers/splitter.ML --- a/src/Provers/splitter.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Provers/splitter.ML Thu Apr 18 17:07:01 2013 +0200 @@ -32,8 +32,8 @@ val split_tac : thm list -> int -> tactic val split_inside_tac: thm list -> int -> tactic val split_asm_tac : thm list -> int -> tactic - val add_split: thm -> simpset -> simpset - val del_split: thm -> simpset -> simpset + val add_split: thm -> Proof.context -> Proof.context + val del_split: thm -> Proof.context -> Proof.context val split_add: attribute val split_del: attribute val split_modifiers : Method.modifier parser list @@ -426,15 +426,15 @@ fun split_name (name, T) asm = "split " ^ (if asm then "asm " else "") ^ name ^ " :: " ^ string_of_typ T; -fun add_split split ss = +fun add_split split ctxt = let val (name, asm) = split_thm_info split val tac = (if asm then split_asm_tac else split_tac) [split] - in Simplifier.addloop (ss, (split_name name asm, tac)) end; + in Simplifier.addloop (ctxt, (split_name name asm, tac)) end; -fun del_split split ss = +fun del_split split ctxt = let val (name, asm) = split_thm_info split - in Simplifier.delloop (ss, split_name name asm) end; + in Simplifier.delloop (ctxt, split_name name asm) end; (* attributes *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/Isar/code.ML Thu Apr 18 17:07:01 2013 +0200 @@ -874,7 +874,7 @@ fun rewrite_eqn thy conv ss = let val ctxt = Proof_Context.init_global thy; - val rewrite = Conv.fconv_rule (conv (Simplifier.rewrite ss)); + val rewrite = Conv.fconv_rule (conv (Simplifier.rewrite (put_simpset ss ctxt))); in singleton (Variable.trade (K (map rewrite)) ctxt) end; fun cert_of_eqns_preprocess thy functrans ss c = diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/Isar/isar_cmd.ML Thu Apr 18 17:07:01 2013 +0200 @@ -193,7 +193,7 @@ fun simproc_setup name lhss (txt, pos) identifier = ML_Lex.read pos txt |> ML_Context.expression pos - "val proc: Morphism.morphism -> Simplifier.simpset -> cterm -> thm option" + "val proc: Morphism.morphism -> Proof.context -> cterm -> thm option" ("Context.map_proof (Simplifier.def_simproc_cmd {name = " ^ ML_Syntax.make_binding name ^ ", \ \lhss = " ^ ML_Syntax.print_strings lhss ^ ", proc = proc, \ \identifier = Library.maps ML_Context.thms " ^ ML_Syntax.print_strings identifier ^ "})") diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/Isar/isar_syn.ML Thu Apr 18 17:07:01 2013 +0200 @@ -839,9 +839,7 @@ Outer_Syntax.improper_command @{command_spec "print_simpset"} "print context of Simplifier" (Scan.succeed (Toplevel.unknown_context o - Toplevel.keep (fn state => - let val ctxt = Toplevel.context_of state - in Pretty.writeln (Simplifier.pretty_ss ctxt (simpset_of ctxt)) end))); + Toplevel.keep (Pretty.writeln o Simplifier.pretty_simpset o Toplevel.context_of))); val _ = Outer_Syntax.improper_command @{command_spec "print_rules"} "print intro/elim rules" diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/Isar/local_defs.ML --- a/src/Pure/Isar/local_defs.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/Isar/local_defs.ML Thu Apr 18 17:07:01 2013 +0200 @@ -191,7 +191,7 @@ fun meta_rewrite_conv ctxt = Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) - (Raw_Simplifier.context ctxt empty_ss + (empty_simpset ctxt addsimps (Rules.get (Context.Proof ctxt)) |> Raw_Simplifier.add_eqcong Drule.equals_cong); (*protect meta-level equality*) diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/Tools/find_theorems.ML Thu Apr 18 17:07:01 2013 +0200 @@ -334,7 +334,7 @@ fun filter_simp ctxt t (Internal (_, thm)) = let - val mksimps = Simplifier.mksimps (simpset_of ctxt); + val mksimps = Simplifier.mksimps ctxt; val extract_simp = (map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl); val ss = is_matching_thm extract_simp ctxt false t thm; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/raw_simplifier.ML --- a/src/Pure/raw_simplifier.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/raw_simplifier.ML Thu Apr 18 17:07:01 2013 +0200 @@ -18,10 +18,10 @@ type cong_name = bool * string type rrule val eq_rrule: rrule * rrule -> bool - type simpset type proc type solver - val mk_solver: string -> (simpset -> int -> tactic) -> solver + val mk_solver: string -> (Proof.context -> int -> tactic) -> solver + type simpset val empty_ss: simpset val merge_ss: simpset * simpset -> simpset val dest_ss: simpset -> @@ -36,21 +36,28 @@ val eq_simproc: simproc * simproc -> bool val transform_simproc: morphism -> simproc -> simproc val make_simproc: {name: string, lhss: cterm list, - proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> simproc - val mk_simproc: string -> cterm list -> (theory -> simpset -> term -> thm option) -> simproc - val addsimps: simpset * thm list -> simpset - val delsimps: simpset * thm list -> simpset - val addsimprocs: simpset * simproc list -> simpset - val delsimprocs: simpset * simproc list -> simpset - val setloop': simpset * (simpset -> int -> tactic) -> simpset - val setloop: simpset * (int -> tactic) -> simpset - val addloop': simpset * (string * (simpset -> int -> tactic)) -> simpset - val addloop: simpset * (string * (int -> tactic)) -> simpset - val delloop: simpset * string -> simpset - val setSSolver: simpset * solver -> simpset - val addSSolver: simpset * solver -> simpset - val setSolver: simpset * solver -> simpset - val addSolver: simpset * solver -> simpset + proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> simproc + val mk_simproc: string -> cterm list -> (Proof.context -> term -> thm option) -> simproc + val simpset_of: Proof.context -> simpset + val put_simpset: simpset -> Proof.context -> Proof.context + val global_context: theory -> simpset -> Proof.context + val simpset_map: Proof.context -> (Proof.context -> Proof.context) -> simpset -> simpset + val map_theory_simpset: (Proof.context -> Proof.context) -> theory -> theory + val empty_simpset: Proof.context -> Proof.context + val clear_simpset: Proof.context -> Proof.context + val addsimps: Proof.context * thm list -> Proof.context + val delsimps: Proof.context * thm list -> Proof.context + val addsimprocs: Proof.context * simproc list -> Proof.context + val delsimprocs: Proof.context * simproc list -> Proof.context + val setloop': Proof.context * (Proof.context -> int -> tactic) -> Proof.context + val setloop: Proof.context * (int -> tactic) -> Proof.context + val addloop': Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context + val addloop: Proof.context * (string * (int -> tactic)) -> Proof.context + val delloop: Proof.context * string -> Proof.context + val setSSolver: Proof.context * solver -> Proof.context + val addSSolver: Proof.context * solver -> Proof.context + val setSolver: Proof.context * solver -> Proof.context + val addSolver: Proof.context * solver -> Proof.context val rewrite_rule: thm list -> thm -> thm val rewrite_goals_rule: thm list -> thm -> thm @@ -71,62 +78,58 @@ {rules: rrule Net.net, prems: thm list, bounds: int * ((string * typ) * string) list, - depth: int * bool Unsynchronized.ref, - context: Proof.context option} * + depth: int * bool Unsynchronized.ref} * {congs: (cong_name * thm) list * cong_name list, procs: proc Net.net, mk_rews: - {mk: simpset -> thm -> thm list, - mk_cong: simpset -> thm -> thm, - mk_sym: simpset -> thm -> thm option, - mk_eq_True: simpset -> thm -> thm option, - reorient: theory -> term list -> term -> term -> bool}, + {mk: Proof.context -> thm -> thm list, + mk_cong: Proof.context -> thm -> thm, + mk_sym: Proof.context -> thm -> thm option, + mk_eq_True: Proof.context -> thm -> thm option, + reorient: Proof.context -> term list -> term -> term -> bool}, termless: term * term -> bool, - subgoal_tac: simpset -> int -> tactic, - loop_tacs: (string * (simpset -> int -> tactic)) list, + subgoal_tac: Proof.context -> int -> tactic, + loop_tacs: (string * (Proof.context -> int -> tactic)) list, solvers: solver list * solver list} - val prems_of: simpset -> thm list - val add_simp: thm -> simpset -> simpset - val del_simp: thm -> simpset -> simpset - val add_eqcong: thm -> simpset -> simpset - val del_eqcong: thm -> simpset -> simpset - val add_cong: thm -> simpset -> simpset - val del_cong: thm -> simpset -> simpset - val mksimps: simpset -> thm -> thm list - val set_mksimps: (simpset -> thm -> thm list) -> simpset -> simpset - val set_mkcong: (simpset -> thm -> thm) -> simpset -> simpset - val set_mksym: (simpset -> thm -> thm option) -> simpset -> simpset - val set_mkeqTrue: (simpset -> thm -> thm option) -> simpset -> simpset - val set_termless: (term * term -> bool) -> simpset -> simpset - val set_subgoaler: (simpset -> int -> tactic) -> simpset -> simpset - val solver: simpset -> solver -> int -> tactic + val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic + val prems_of: Proof.context -> thm list + val add_simp: thm -> Proof.context -> Proof.context + val del_simp: thm -> Proof.context -> Proof.context + val add_eqcong: thm -> Proof.context -> Proof.context + val del_eqcong: thm -> Proof.context -> Proof.context + val add_cong: thm -> Proof.context -> Proof.context + val del_cong: thm -> Proof.context -> Proof.context + val mksimps: Proof.context -> thm -> thm list + val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context + val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context + val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context + val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context + val set_termless: (term * term -> bool) -> Proof.context -> Proof.context + val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context + val solver: Proof.context -> solver -> int -> tactic val simp_depth_limit_raw: Config.raw - val clear_ss: simpset -> simpset - val default_mk_sym: simpset -> thm -> thm option - val simproc_global_i: theory -> string -> term list - -> (theory -> simpset -> term -> thm option) -> simproc - val simproc_global: theory -> string -> string list - -> (theory -> simpset -> term -> thm option) -> simproc + val default_mk_sym: Proof.context -> thm -> thm option + val simproc_global_i: theory -> string -> term list -> + (Proof.context -> term -> thm option) -> simproc + val simproc_global: theory -> string -> string list -> + (Proof.context -> term -> thm option) -> simproc val simp_trace_depth_limit_raw: Config.raw val simp_trace_depth_limit_default: int Unsynchronized.ref val simp_trace_default: bool Unsynchronized.ref val simp_trace_raw: Config.raw val simp_debug_raw: Config.raw - val add_prems: thm list -> simpset -> simpset - val inherit_context: simpset -> simpset -> simpset - val the_context: simpset -> Proof.context - val context: Proof.context -> simpset -> simpset - val global_context: theory -> simpset -> simpset - val with_context: Proof.context -> (simpset -> simpset) -> simpset -> simpset + val add_prems: thm list -> Proof.context -> Proof.context val debug_bounds: bool Unsynchronized.ref - val set_reorient: (theory -> term list -> term -> term -> bool) -> simpset -> simpset - val set_solvers: solver list -> simpset -> simpset - val rewrite_cterm: bool * bool * bool -> (simpset -> thm -> thm option) -> simpset -> conv + val set_reorient: (Proof.context -> term list -> term -> term -> bool) -> + Proof.context -> Proof.context + val set_solvers: solver list -> Proof.context -> Proof.context + val rewrite_cterm: bool * bool * bool -> + (Proof.context -> thm -> thm option) -> Proof.context -> conv val rewrite_term: theory -> thm list -> (term -> term option) list -> term -> term val rewrite_thm: bool * bool * bool -> - (simpset -> thm -> thm option) -> simpset -> thm -> thm + (Proof.context -> thm -> thm option) -> Proof.context -> thm -> thm val generic_rewrite_goal_tac: bool * bool * bool -> - (simpset -> tactic) -> simpset -> int -> tactic + (Proof.context -> tactic) -> Proof.context -> int -> tactic val rewrite: bool -> thm list -> conv val simplify: bool -> thm list -> thm -> thm end; @@ -170,223 +173,6 @@ fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) = Thm.eq_thm_prop (thm1, thm2); - -(* simplification sets, procedures, and solvers *) - -(*A simpset contains data required during conversion: - rules: discrimination net of rewrite rules; - prems: current premises; - bounds: maximal index of bound variables already used - (for generating new names when rewriting under lambda abstractions); - depth: simp_depth and exceeded flag; - congs: association list of congruence rules and - a list of `weak' congruence constants. - A congruence is `weak' if it avoids normalization of some argument. - procs: discrimination net of simplification procedures - (functions that prove rewrite rules on the fly); - mk_rews: - mk: turn simplification thms into rewrite rules; - mk_cong: prepare congruence rules; - mk_sym: turn == around; - mk_eq_True: turn P into P == True; - termless: relation for ordered rewriting;*) - -datatype simpset = - Simpset of - {rules: rrule Net.net, - prems: thm list, - bounds: int * ((string * typ) * string) list, - depth: int * bool Unsynchronized.ref, - context: Proof.context option} * - {congs: (cong_name * thm) list * cong_name list, - procs: proc Net.net, - mk_rews: - {mk: simpset -> thm -> thm list, - mk_cong: simpset -> thm -> thm, - mk_sym: simpset -> thm -> thm option, - mk_eq_True: simpset -> thm -> thm option, - reorient: theory -> term list -> term -> term -> bool}, - termless: term * term -> bool, - subgoal_tac: simpset -> int -> tactic, - loop_tacs: (string * (simpset -> int -> tactic)) list, - solvers: solver list * solver list} -and proc = - Proc of - {name: string, - lhs: cterm, - proc: simpset -> cterm -> thm option, - id: stamp * thm list} -and solver = - Solver of - {name: string, - solver: simpset -> int -> tactic, - id: stamp}; - - -fun internal_ss (Simpset args) = args; - -fun make_ss1 (rules, prems, bounds, depth, context) = - {rules = rules, prems = prems, bounds = bounds, depth = depth, context = context}; - -fun map_ss1 f {rules, prems, bounds, depth, context} = - make_ss1 (f (rules, prems, bounds, depth, context)); - -fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) = - {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless, - subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers}; - -fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} = - make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); - -fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2); - -fun map_simpset1 f (Simpset (r1, r2)) = Simpset (map_ss1 f r1, r2); -fun map_simpset2 f (Simpset (r1, r2)) = Simpset (r1, map_ss2 f r2); - -fun prems_of (Simpset ({prems, ...}, _)) = prems; - -fun eq_procid ((s1: stamp, ths1: thm list), (s2, ths2)) = - s1 = s2 andalso eq_list Thm.eq_thm (ths1, ths2); -fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = eq_procid (id1, id2); - -fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()}; - -fun solver_name (Solver {name, ...}) = name; -fun solver ss (Solver {solver = tac, ...}) = tac ss; -fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2); - - -(* simp depth *) - -val simp_depth_limit_raw = Config.declare "simp_depth_limit" (K (Config.Int 100)); -val simp_depth_limit = Config.int simp_depth_limit_raw; - -val simp_trace_depth_limit_default = Unsynchronized.ref 1; -val simp_trace_depth_limit_raw = Config.declare "simp_trace_depth_limit" - (fn _ => Config.Int (! simp_trace_depth_limit_default)); -val simp_trace_depth_limit = Config.int simp_trace_depth_limit_raw; - -fun simp_trace_depth_limit_of NONE = ! simp_trace_depth_limit_default - | simp_trace_depth_limit_of (SOME ctxt) = Config.get ctxt simp_trace_depth_limit; - -fun trace_depth (Simpset ({depth = (depth, exceeded), context, ...}, _)) msg = - if depth > simp_trace_depth_limit_of context then - if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true) - else - (tracing (enclose "[" "]" (string_of_int depth) ^ msg); exceeded := false); - -val inc_simp_depth = map_simpset1 (fn (rules, prems, bounds, (depth, exceeded), context) => - (rules, prems, bounds, - (depth + 1, - if depth = simp_trace_depth_limit_of context - then Unsynchronized.ref false else exceeded), context)); - -fun simp_depth (Simpset ({depth = (depth, _), ...}, _)) = depth; - - -(* diagnostics *) - -exception SIMPLIFIER of string * thm; - -val simp_debug_raw = Config.declare "simp_debug" (K (Config.Bool false)); -val simp_debug = Config.bool simp_debug_raw; - -val simp_trace_default = Unsynchronized.ref false; -val simp_trace_raw = Config.declare "simp_trace" (fn _ => Config.Bool (! simp_trace_default)); -val simp_trace = Config.bool simp_trace_raw; - -fun if_enabled (Simpset ({context, ...}, _)) flag f = - (case context of - SOME ctxt => if Config.get ctxt flag then f ctxt else () - | NONE => ()) - -fun if_visible (Simpset ({context, ...}, _)) f x = - (case context of - SOME ctxt => Context_Position.if_visible ctxt f x - | NONE => ()); - -local - -fun prnt ss warn a = if warn then warning a else trace_depth ss a; - -fun show_bounds (Simpset ({bounds = (_, bs), ...}, _)) t = - let - val names = Term.declare_term_names t Name.context; - val xs = rev (#1 (fold_map Name.variant (rev (map #2 bs)) names)); - fun subst (((b, T), _), x') = (Free (b, T), Syntax_Trans.mark_bound_abs (x', T)); - in Term.subst_atomic (ListPair.map subst (bs, xs)) t end; - -fun print_term ss warn a t ctxt = prnt ss warn (a () ^ "\n" ^ - Syntax.string_of_term ctxt - (if Config.get ctxt simp_debug then t else show_bounds ss t)); - -in - -fun print_term_global ss warn a thy t = - print_term ss warn (K a) t (Proof_Context.init_global thy); - -fun debug warn a ss = if_enabled ss simp_debug (fn _ => prnt ss warn (a ())); -fun trace warn a ss = if_enabled ss simp_trace (fn _ => prnt ss warn (a ())); - -fun debug_term warn a ss t = if_enabled ss simp_debug (print_term ss warn a t); -fun trace_term warn a ss t = if_enabled ss simp_trace (print_term ss warn a t); - -fun trace_cterm warn a ss ct = - if_enabled ss simp_trace (print_term ss warn a (Thm.term_of ct)); - -fun trace_thm a ss th = - if_enabled ss simp_trace (print_term ss false a (Thm.full_prop_of th)); - -fun trace_named_thm a ss (th, name) = - if_enabled ss simp_trace (print_term ss false - (fn () => if name = "" then a () else a () ^ " " ^ quote name ^ ":") - (Thm.full_prop_of th)); - -fun warn_thm a ss th = - print_term_global ss true a (Thm.theory_of_thm th) (Thm.full_prop_of th); - -fun cond_warn_thm a ss th = if_visible ss (fn () => warn_thm a ss th) (); - -end; - - - -(** simpset operations **) - -(* context *) - -fun eq_bound (x: string, (y, _)) = x = y; - -fun add_bound bound = map_simpset1 (fn (rules, prems, (count, bounds), depth, context) => - (rules, prems, (count + 1, bound :: bounds), depth, context)); - -fun add_prems ths = map_simpset1 (fn (rules, prems, bounds, depth, context) => - (rules, ths @ prems, bounds, depth, context)); - -fun inherit_context (Simpset ({bounds, depth, context, ...}, _)) = - map_simpset1 (fn (rules, prems, _, _, _) => (rules, prems, bounds, depth, context)); - -fun the_context (Simpset ({context = SOME ctxt, ...}, _)) = ctxt - | the_context _ = raise Fail "Simplifier: no proof context in simpset"; - -fun context ctxt = - map_simpset1 (fn (rules, prems, bounds, depth, _) => (rules, prems, bounds, depth, SOME ctxt)); - -val global_context = context o Proof_Context.init_global; - -fun activate_context thy ss = - let - val ctxt = the_context ss; - val ctxt' = ctxt - |> Context.raw_transfer (Theory.merge (thy, Proof_Context.theory_of ctxt)) - |> Context_Position.set_visible false; - in context ctxt' ss end; - -fun with_context ctxt f ss = inherit_context ss (f (context ctxt ss)); - - -(* maintain simp rules *) - (* FIXME: it seems that the conditions on extra variables are too liberal if prems are nonempty: does solving the prems really guarantee instantiation of all its Vars? Better: a dynamic check each time a rule is applied. @@ -413,30 +199,8 @@ val extra = rrule_extra_vars elhs thm; in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end; -fun del_rrule (rrule as {thm, elhs, ...}) ss = - ss |> map_simpset1 (fn (rules, prems, bounds, depth, context) => - (Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, bounds, depth, context)) - handle Net.DELETE => (cond_warn_thm "Rewrite rule not in simpset:" ss thm; ss); - -fun insert_rrule (rrule as {thm, name, ...}) ss = - (trace_named_thm (fn () => "Adding rewrite rule") ss (thm, name); - ss |> map_simpset1 (fn (rules, prems, bounds, depth, context) => - let - val rrule2 as {elhs, ...} = mk_rrule2 rrule; - val rules' = Net.insert_term eq_rrule (term_of elhs, rrule2) rules; - in (rules', prems, bounds, depth, context) end) - handle Net.INSERT => (cond_warn_thm "Ignoring duplicate rewrite rule:" ss thm; ss)); - -fun vperm (Var _, Var _) = true - | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) - | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) - | vperm (t, u) = (t = u); - -fun var_perm (t, u) = - vperm (t, u) andalso eq_set (op =) (Term.add_vars t [], Term.add_vars u []); - (*simple test for looping rewrite rules and stupid orientations*) -fun default_reorient thy prems lhs rhs = +fun default_reorient ctxt prems lhs rhs = rewrite_rule_extra_vars prems lhs rhs orelse is_Var (head_of lhs) @@ -449,13 +213,338 @@ *) exists (fn t => Logic.occs (lhs, t)) (rhs :: prems) orelse - null prems andalso Pattern.matches thy (lhs, rhs) + null prems andalso Pattern.matches (Proof_Context.theory_of ctxt) (lhs, rhs) (*the condition "null prems" is necessary because conditional rewrites with extra variables in the conditions may terminate although the rhs is an instance of the lhs; example: ?m < ?n ==> f(?n) == f(?m)*) orelse is_Const lhs andalso not (is_Const rhs); + +(* simplification procedures *) + +datatype proc = + Proc of + {name: string, + lhs: cterm, + proc: Proof.context -> cterm -> thm option, + id: stamp * thm list}; + +fun eq_procid ((s1: stamp, ths1: thm list), (s2, ths2)) = + s1 = s2 andalso eq_list Thm.eq_thm (ths1, ths2); + +fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = eq_procid (id1, id2); + + +(* solvers *) + +datatype solver = + Solver of + {name: string, + solver: Proof.context -> int -> tactic, + id: stamp}; + +fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()}; + +fun solver_name (Solver {name, ...}) = name; +fun solver ctxt (Solver {solver = tac, ...}) = tac ctxt; +fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2); + + +(* simplification sets *) + +(*A simpset contains data required during conversion: + rules: discrimination net of rewrite rules; + prems: current premises; + bounds: maximal index of bound variables already used + (for generating new names when rewriting under lambda abstractions); + depth: simp_depth and exceeded flag; + congs: association list of congruence rules and + a list of `weak' congruence constants. + A congruence is `weak' if it avoids normalization of some argument. + procs: discrimination net of simplification procedures + (functions that prove rewrite rules on the fly); + mk_rews: + mk: turn simplification thms into rewrite rules; + mk_cong: prepare congruence rules; + mk_sym: turn == around; + mk_eq_True: turn P into P == True; + termless: relation for ordered rewriting;*) + +datatype simpset = + Simpset of + {rules: rrule Net.net, + prems: thm list, + bounds: int * ((string * typ) * string) list, + depth: int * bool Unsynchronized.ref} * + {congs: (cong_name * thm) list * cong_name list, + procs: proc Net.net, + mk_rews: + {mk: Proof.context -> thm -> thm list, + mk_cong: Proof.context -> thm -> thm, + mk_sym: Proof.context -> thm -> thm option, + mk_eq_True: Proof.context -> thm -> thm option, + reorient: Proof.context -> term list -> term -> term -> bool}, + termless: term * term -> bool, + subgoal_tac: Proof.context -> int -> tactic, + loop_tacs: (string * (Proof.context -> int -> tactic)) list, + solvers: solver list * solver list}; + +fun internal_ss (Simpset args) = args; + +fun make_ss1 (rules, prems, bounds, depth) = + {rules = rules, prems = prems, bounds = bounds, depth = depth}; + +fun map_ss1 f {rules, prems, bounds, depth} = + make_ss1 (f (rules, prems, bounds, depth)); + +fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) = + {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless, + subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers}; + +fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} = + make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); + +fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2); + +fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) = + {simps = Net.entries rules + |> map (fn {name, thm, ...} => (name, thm)), + procs = Net.entries procs + |> map (fn Proc {name, lhs, id, ...} => ((name, lhs), id)) + |> partition_eq (eq_snd eq_procid) + |> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)), + congs = #1 congs, + weak_congs = #2 congs, + loopers = map fst loop_tacs, + unsafe_solvers = map solver_name (#1 solvers), + safe_solvers = map solver_name (#2 solvers)}; + + +(* empty *) + +fun init_ss mk_rews termless subgoal_tac solvers = + make_simpset ((Net.empty, [], (0, []), (0, Unsynchronized.ref false)), + (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers)); + +fun default_mk_sym _ th = SOME (th RS Drule.symmetric_thm); + +val empty_ss = + init_ss + {mk = fn _ => fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [], + mk_cong = K I, + mk_sym = default_mk_sym, + mk_eq_True = K (K NONE), + reorient = default_reorient} + Term_Ord.termless (K (K no_tac)) ([], []); + + +(* merge *) (*NOTE: ignores some fields of 2nd simpset*) + +fun merge_ss (ss1, ss2) = + if pointer_eq (ss1, ss2) then ss1 + else + let + val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth = depth1}, + {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac, + loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; + val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = depth2}, + {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _, + loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; + + val rules' = Net.merge eq_rrule (rules1, rules2); + val prems' = Thm.merge_thms (prems1, prems2); + val bounds' = if #1 bounds1 < #1 bounds2 then bounds2 else bounds1; + val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1; + val congs' = merge (Thm.eq_thm_prop o pairself #2) (congs1, congs2); + val weak' = merge (op =) (weak1, weak2); + val procs' = Net.merge eq_proc (procs1, procs2); + val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2); + val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2); + val solvers' = merge eq_solver (solvers1, solvers2); + in + make_simpset ((rules', prems', bounds', depth'), ((congs', weak'), procs', + mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) + end; + + + +(** context data **) + +structure Simpset = Generic_Data +( + type T = simpset; + val empty = empty_ss; + val extend = I; + val merge = merge_ss; +); + +val simpset_of = Simpset.get o Context.Proof; + +fun map_simpset f = Context.proof_map (Simpset.map f); +fun map_simpset1 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (map_ss1 f ss1, ss2)); +fun map_simpset2 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (ss1, map_ss2 f ss2)); + +fun simpset_map ctxt f ss = ctxt |> map_simpset (K ss) |> f |> Context.Proof |> Simpset.get; + +fun put_simpset (Simpset ({rules, prems, ...}, ss2)) = (* FIXME prems from context (!?) *) + map_simpset (fn Simpset ({bounds, depth, ...}, _) => + Simpset (make_ss1 (rules, prems, bounds, depth), ss2)); + +fun global_context thy ss = Proof_Context.init_global thy |> put_simpset ss; + +val empty_simpset = put_simpset empty_ss; + +fun map_theory_simpset f thy = + let + val ctxt' = f (Proof_Context.init_global thy); + val thy' = Proof_Context.theory_of ctxt'; + in Context.theory_map (Simpset.map (K (simpset_of ctxt'))) thy' end; + +fun map_ss f = Context.mapping (map_theory_simpset f) f; + +val clear_simpset = + map_simpset (fn Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...}) => + init_ss mk_rews termless subgoal_tac solvers); + + +(* simp depth *) + +val simp_depth_limit_raw = Config.declare "simp_depth_limit" (K (Config.Int 100)); +val simp_depth_limit = Config.int simp_depth_limit_raw; + +val simp_trace_depth_limit_default = Unsynchronized.ref 1; +val simp_trace_depth_limit_raw = Config.declare "simp_trace_depth_limit" + (fn _ => Config.Int (! simp_trace_depth_limit_default)); +val simp_trace_depth_limit = Config.int simp_trace_depth_limit_raw; + +fun trace_depth ctxt msg = + let + val Simpset ({depth = (depth, exceeded), ...}, _) = simpset_of ctxt; + val depth_limit = Config.get ctxt simp_trace_depth_limit; + in + if depth > depth_limit then + if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true) + else (tracing (enclose "[" "]" (string_of_int depth) ^ msg); exceeded := false) + end; + +fun inc_simp_depth ctxt = + ctxt |> map_simpset1 (fn (rules, prems, bounds, (depth, exceeded)) => + (rules, prems, bounds, + (depth + 1, + if depth = Config.get ctxt simp_trace_depth_limit + then Unsynchronized.ref false else exceeded))); + +fun simp_depth ctxt = + let val Simpset ({depth = (depth, _), ...}, _) = simpset_of ctxt + in depth end; + + +(* diagnostics *) + +exception SIMPLIFIER of string * thm; + +val simp_debug_raw = Config.declare "simp_debug" (K (Config.Bool false)); +val simp_debug = Config.bool simp_debug_raw; + +val simp_trace_default = Unsynchronized.ref false; +val simp_trace_raw = Config.declare "simp_trace" (fn _ => Config.Bool (! simp_trace_default)); +val simp_trace = Config.bool simp_trace_raw; + +fun if_enabled ctxt flag f = if Config.get ctxt flag then f ctxt else (); + +local + +fun prnt ctxt warn a = if warn then warning a else trace_depth ctxt a; + +fun show_bounds ctxt t = + let + val Simpset ({bounds = (_, bs), ...}, _) = simpset_of ctxt; + val names = Term.declare_term_names t Name.context; + val xs = rev (#1 (fold_map Name.variant (rev (map #2 bs)) names)); + fun subst (((b, T), _), x') = (Free (b, T), Syntax_Trans.mark_bound_abs (x', T)); + in Term.subst_atomic (ListPair.map subst (bs, xs)) t end; + +in + +fun print_term ctxt warn a t = + prnt ctxt warn (a () ^ "\n" ^ + Syntax.string_of_term ctxt (if Config.get ctxt simp_debug then t else show_bounds ctxt t)); + +fun debug ctxt warn a = if_enabled ctxt simp_debug (fn _ => prnt ctxt warn (a ())); +fun trace ctxt warn a = if_enabled ctxt simp_trace (fn _ => prnt ctxt warn (a ())); + +fun debug_term ctxt warn a t = if_enabled ctxt simp_debug (fn _ => print_term ctxt warn a t); +fun trace_term ctxt warn a t = if_enabled ctxt simp_trace (fn _ => print_term ctxt warn a t); + +fun trace_cterm ctxt warn a ct = + if_enabled ctxt simp_trace (fn _ => print_term ctxt warn a (Thm.term_of ct)); + +fun trace_thm ctxt a th = + if_enabled ctxt simp_trace (fn _ => print_term ctxt false a (Thm.full_prop_of th)); + +fun trace_named_thm ctxt a (th, name) = + if_enabled ctxt simp_trace (fn _ => + print_term ctxt false + (fn () => if name = "" then a () else a () ^ " " ^ quote name ^ ":") + (Thm.full_prop_of th)); + +fun warn_thm ctxt a th = print_term ctxt true a (Thm.full_prop_of th); +fun cond_warn_thm ctxt a th = Context_Position.if_visible ctxt (fn () => warn_thm ctxt a th) (); + +end; + + + +(** simpset operations **) + +(* context *) + +fun eq_bound (x: string, (y, _)) = x = y; + +fun add_bound bound = + map_simpset1 (fn (rules, prems, (count, bounds), depth) => + (rules, prems, (count + 1, bound :: bounds), depth)); + +fun prems_of ctxt = + let val Simpset ({prems, ...}, _) = simpset_of ctxt in prems end; + +fun add_prems ths = + map_simpset1 (fn (rules, prems, bounds, depth) => (rules, ths @ prems, bounds, depth)); + +fun activate_context thy ctxt = ctxt (* FIXME ?? *) + |> Context.raw_transfer (Theory.merge (thy, Proof_Context.theory_of ctxt)) + |> Context_Position.set_visible false; + + +(* maintain simp rules *) + +fun del_rrule (rrule as {thm, elhs, ...}) ctxt = + ctxt |> map_simpset1 (fn (rules, prems, bounds, depth) => + (Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, bounds, depth)) + handle Net.DELETE => (cond_warn_thm ctxt (fn () => "Rewrite rule not in simpset:") thm; ctxt); + +fun insert_rrule (rrule as {thm, name, ...}) ctxt = + (trace_named_thm ctxt (fn () => "Adding rewrite rule") (thm, name); + ctxt |> map_simpset1 (fn (rules, prems, bounds, depth) => + let + val rrule2 as {elhs, ...} = mk_rrule2 rrule; + val rules' = Net.insert_term eq_rrule (term_of elhs, rrule2) rules; + in (rules', prems, bounds, depth) end) + handle Net.INSERT => (cond_warn_thm ctxt (fn () => "Ignoring duplicate rewrite rule:") thm; ctxt)); + +local + +fun vperm (Var _, Var _) = true + | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) + | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) + | vperm (t, u) = (t = u); + +fun var_perm (t, u) = + vperm (t, u) andalso eq_set (op =) (Term.add_vars t [], Term.add_vars u []); + +in + fun decomp_simp thm = let val thy = Thm.theory_of_thm thm; @@ -472,79 +561,83 @@ not (is_Var (term_of elhs)); in (thy, prems, term_of lhs, elhs, term_of rhs, perm) end; +end; + fun decomp_simp' thm = let val (_, _, lhs, _, rhs, _) = decomp_simp thm in if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", thm) else (lhs, rhs) end; -fun mk_eq_True (ss as Simpset (_, {mk_rews = {mk_eq_True, ...}, ...})) (thm, name) = - (case mk_eq_True ss thm of - NONE => [] - | SOME eq_True => - let - val (_, _, lhs, elhs, _, _) = decomp_simp eq_True; - in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end); +fun mk_eq_True ctxt (thm, name) = + let val Simpset (_, {mk_rews = {mk_eq_True, ...}, ...}) = simpset_of ctxt in + (case mk_eq_True ctxt thm of + NONE => [] + | SOME eq_True => + let val (_, _, lhs, elhs, _, _) = decomp_simp eq_True; + in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end) + end; (*create the rewrite rule and possibly also the eq_True variant, in case there are extra vars on the rhs*) -fun rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm2) = +fun rrule_eq_True (thm, name, lhs, elhs, rhs, ctxt, thm2) = let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in if rewrite_rule_extra_vars [] lhs rhs then - mk_eq_True ss (thm2, name) @ [rrule] + mk_eq_True ctxt (thm2, name) @ [rrule] else [rrule] end; -fun mk_rrule ss (thm, name) = +fun mk_rrule ctxt (thm, name) = let val (_, prems, lhs, elhs, rhs, perm) = decomp_simp thm in if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] else (*weak test for loops*) if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs) - then mk_eq_True ss (thm, name) - else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm) + then mk_eq_True ctxt (thm, name) + else rrule_eq_True (thm, name, lhs, elhs, rhs, ctxt, thm) end; -fun orient_rrule ss (thm, name) = +fun orient_rrule ctxt (thm, name) = let val (thy, prems, lhs, elhs, rhs, perm) = decomp_simp thm; - val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = ss; + val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = simpset_of ctxt; in if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] - else if reorient thy prems lhs rhs then - if reorient thy prems rhs lhs - then mk_eq_True ss (thm, name) + else if reorient ctxt prems lhs rhs then + if reorient ctxt prems rhs lhs + then mk_eq_True ctxt (thm, name) else - (case mk_sym ss thm of + (case mk_sym ctxt thm of NONE => [] | SOME thm' => let val (_, _, lhs', elhs', rhs', _) = decomp_simp thm' - in rrule_eq_True (thm', name, lhs', elhs', rhs', ss, thm) end) - else rrule_eq_True (thm, name, lhs, elhs, rhs, ss, thm) + in rrule_eq_True (thm', name, lhs', elhs', rhs', ctxt, thm) end) + else rrule_eq_True (thm, name, lhs, elhs, rhs, ctxt, thm) end; -fun extract_rews (ss as Simpset (_, {mk_rews = {mk, ...}, ...}), thms) = - maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ss thm)) thms; +fun extract_rews (ctxt, thms) = + let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt + in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms end; -fun extract_safe_rrules (ss, thm) = - maps (orient_rrule ss) (extract_rews (ss, [thm])); +fun extract_safe_rrules (ctxt, thm) = + maps (orient_rrule ctxt) (extract_rews (ctxt, [thm])); (* add/del rules explicitly *) -fun comb_simps comb mk_rrule (ss, thms) = +fun comb_simps comb mk_rrule (ctxt, thms) = let - val rews = extract_rews (ss, thms); - in fold (fold comb o mk_rrule) rews ss end; + val rews = extract_rews (ctxt, thms); + in fold (fold comb o mk_rrule) rews ctxt end; -fun ss addsimps thms = - comb_simps insert_rrule (mk_rrule ss) (ss, thms); +fun ctxt addsimps thms = + comb_simps insert_rrule (mk_rrule ctxt) (ctxt, thms); -fun ss delsimps thms = - comb_simps del_rrule (map mk_rrule2 o mk_rrule ss) (ss, thms); +fun ctxt delsimps thms = + comb_simps del_rrule (map mk_rrule2 o mk_rrule ctxt) (ctxt, thms); -fun add_simp thm ss = ss addsimps [thm]; -fun del_simp thm ss = ss delsimps [thm]; +fun add_simp thm ctxt = ctxt addsimps [thm]; +fun del_simp thm ctxt = ctxt delsimps [thm]; (* congs *) @@ -574,11 +667,13 @@ is_full_cong_prems prems (xs ~~ ys) end; -fun mk_cong (ss as Simpset (_, {mk_rews = {mk_cong = f, ...}, ...})) = f ss; +fun mk_cong ctxt = + let val Simpset (_, {mk_rews = {mk_cong = f, ...}, ...}) = simpset_of ctxt + in f ctxt end; in -fun add_eqcong thm ss = ss |> +fun add_eqcong thm ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => let val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) @@ -588,14 +683,15 @@ raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm); val (xs, weak) = congs; val _ = - if AList.defined (op =) xs a - then if_visible ss warning ("Overwriting congruence rule for " ^ quote (#2 a)) + if AList.defined (op =) xs a then + Context_Position.if_visible ctxt + warning ("Overwriting congruence rule for " ^ quote (#2 a)) else (); val xs' = AList.update (op =) (a, thm) xs; val weak' = if is_full_cong thm then weak else a :: weak; in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); -fun del_eqcong thm ss = ss |> +fun del_eqcong thm ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => let val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) @@ -609,8 +705,8 @@ if is_full_cong thm then NONE else SOME a); in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end); -fun add_cong thm ss = add_eqcong (mk_cong ss thm) ss; -fun del_cong thm ss = del_eqcong (mk_cong ss thm) ss; +fun add_cong thm ctxt = add_eqcong (mk_cong ctxt thm) ctxt; +fun del_cong thm ctxt = del_eqcong (mk_cong ctxt thm) ctxt; end; @@ -621,7 +717,7 @@ Simproc of {name: string, lhss: cterm list, - proc: morphism -> simpset -> cterm -> thm option, + proc: morphism -> Proof.context -> cterm -> thm option, id: stamp * thm list}; fun eq_simproc (Simproc {id = id1, ...}, Simproc {id = id2, ...}) = eq_procid (id1, id2); @@ -637,8 +733,8 @@ Simproc {name = name, lhss = lhss, proc = proc, id = (stamp (), identifier)}; fun mk_simproc name lhss proc = - make_simproc {name = name, lhss = lhss, proc = fn _ => fn ss => fn ct => - proc (Proof_Context.theory_of (the_context ss)) ss (Thm.term_of ct), identifier = []}; + make_simproc {name = name, lhss = lhss, proc = fn _ => fn ctxt => fn ct => + proc ctxt (Thm.term_of ct), identifier = []}; (* FIXME avoid global thy and Logic.varify_global *) fun simproc_global_i thy name = mk_simproc name o map (Thm.cterm_of thy o Logic.varify_global); @@ -647,28 +743,30 @@ local -fun add_proc (proc as Proc {name, lhs, ...}) ss = - (trace_cterm false (fn () => "Adding simplification procedure " ^ quote name ^ " for") ss lhs; - map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => +fun add_proc (proc as Proc {name, lhs, ...}) ctxt = + (trace_cterm ctxt false (fn () => "Adding simplification procedure " ^ quote name ^ " for") lhs; + ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => (congs, Net.insert_term eq_proc (term_of lhs, proc) procs, - mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss + mk_rews, termless, subgoal_tac, loop_tacs, solvers)) handle Net.INSERT => - (if_visible ss warning ("Ignoring duplicate simplification procedure " ^ quote name); ss)); + (Context_Position.if_visible ctxt + warning ("Ignoring duplicate simplification procedure " ^ quote name); ctxt)); -fun del_proc (proc as Proc {name, lhs, ...}) ss = - map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => +fun del_proc (proc as Proc {name, lhs, ...}) ctxt = + ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => (congs, Net.delete_term eq_proc (term_of lhs, proc) procs, - mk_rews, termless, subgoal_tac, loop_tacs, solvers)) ss + mk_rews, termless, subgoal_tac, loop_tacs, solvers)) handle Net.DELETE => - (if_visible ss warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ss); + (Context_Position.if_visible ctxt + warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ctxt); fun prep_procs (Simproc {name, lhss, proc, id}) = lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, id = id}); in -fun ss addsimprocs ps = fold (fold add_proc o prep_procs) ps ss; -fun ss delsimprocs ps = fold (fold del_proc o prep_procs) ps ss; +fun ctxt addsimprocs ps = fold (fold add_proc o prep_procs) ps ctxt; +fun ctxt delsimprocs ps = fold (fold del_proc o prep_procs) ps ctxt; end; @@ -688,7 +786,9 @@ in -fun mksimps (ss as Simpset (_, {mk_rews = {mk, ...}, ...})) = mk ss; +fun mksimps ctxt = + let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt + in mk ctxt end; fun set_mksimps mk = map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); @@ -721,39 +821,41 @@ map_simpset2 (fn (congs, procs, mk_rews, termless, _, loop_tacs, solvers) => (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers)); -fun ss setloop' tac = ss |> +fun ctxt setloop' tac = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, _, solvers) => (congs, procs, mk_rews, termless, subgoal_tac, [("", tac)], solvers)); -fun ss setloop tac = ss setloop' (K tac); +fun ctxt setloop tac = ctxt setloop' (K tac); -fun ss addloop' (name, tac) = ss |> +fun ctxt addloop' (name, tac) = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, termless, subgoal_tac, AList.update (op =) (name, tac) loop_tacs, solvers)); -fun ss addloop (name, tac) = ss addloop' (name, K tac); +fun ctxt addloop (name, tac) = ctxt addloop' (name, K tac); -fun ss delloop name = ss |> +fun ctxt delloop name = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, termless, subgoal_tac, (if AList.defined (op =) loop_tacs name then () - else if_visible ss warning ("No such looper in simpset: " ^ quote name); - AList.delete (op =) name loop_tacs), solvers)); + else + Context_Position.if_visible ctxt + warning ("No such looper in simpset: " ^ quote name); + AList.delete (op =) name loop_tacs), solvers)); -fun ss setSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, +fun ctxt setSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, _)) => (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, [solver]))); -fun ss addSSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, +fun ctxt addSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, insert eq_solver solver solvers))); -fun ss setSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, +fun ctxt setSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, ([solver], solvers))); -fun ss addSolver solver = ss |> map_simpset2 (fn (congs, procs, mk_rews, termless, +fun ctxt addSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (insert eq_solver solver unsafe_solvers, solvers))); @@ -762,73 +864,6 @@ subgoal_tac, loop_tacs, (solvers, solvers))); -(* empty *) - -fun init_ss mk_rews termless subgoal_tac solvers = - make_simpset ((Net.empty, [], (0, []), (0, Unsynchronized.ref false), NONE), - (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers)); - -fun clear_ss (ss as Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...})) = - init_ss mk_rews termless subgoal_tac solvers - |> inherit_context ss; - -fun default_mk_sym _ th = SOME (th RS Drule.symmetric_thm); - -val empty_ss = - init_ss - {mk = fn _ => fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [], - mk_cong = K I, - mk_sym = default_mk_sym, - mk_eq_True = K (K NONE), - reorient = default_reorient} - Term_Ord.termless (K (K no_tac)) ([], []); - - -(* merge *) (*NOTE: ignores some fields of 2nd simpset*) - -fun merge_ss (ss1, ss2) = - if pointer_eq (ss1, ss2) then ss1 - else - let - val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth = depth1, context = _}, - {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac, - loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; - val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = depth2, context = _}, - {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _, - loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; - - val rules' = Net.merge eq_rrule (rules1, rules2); - val prems' = Thm.merge_thms (prems1, prems2); - val bounds' = if #1 bounds1 < #1 bounds2 then bounds2 else bounds1; - val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1; - val congs' = merge (Thm.eq_thm_prop o pairself #2) (congs1, congs2); - val weak' = merge (op =) (weak1, weak2); - val procs' = Net.merge eq_proc (procs1, procs2); - val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2); - val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2); - val solvers' = merge eq_solver (solvers1, solvers2); - in - make_simpset ((rules', prems', bounds', depth', NONE), ((congs', weak'), procs', - mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) - end; - - -(* dest_ss *) - -fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) = - {simps = Net.entries rules - |> map (fn {name, thm, ...} => (name, thm)), - procs = Net.entries procs - |> map (fn Proc {name, lhs, id, ...} => ((name, lhs), id)) - |> partition_eq (eq_snd eq_procid) - |> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)), - congs = #1 congs, - weak_congs = #2 congs, - loopers = map fst loop_tacs, - unsafe_solvers = map solver_name (#1 solvers), - safe_solvers = map solver_name (#2 solvers)}; - - (** rewriting **) @@ -838,28 +873,28 @@ Science of Computer Programming 3 (1983), pages 119-149. *) -fun check_conv msg ss thm thm' = +fun check_conv ctxt msg thm thm' = let val thm'' = Thm.transitive thm thm' handle THM _ => Thm.transitive thm (Thm.transitive (Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm') - in if msg then trace_thm (fn () => "SUCCEEDED") ss thm' else (); SOME thm'' end + in if msg then trace_thm ctxt (fn () => "SUCCEEDED") thm' else (); SOME thm'' end handle THM _ => let val _ $ _ $ prop0 = Thm.prop_of thm; in - trace_thm (fn () => "Proved wrong thm (Check subgoaler?)") ss thm'; - trace_term false (fn () => "Should have proved:") ss prop0; + trace_thm ctxt (fn () => "Proved wrong thm (Check subgoaler?)") thm'; + trace_term ctxt false (fn () => "Should have proved:") prop0; NONE end; (* mk_procrule *) -fun mk_procrule ss thm = +fun mk_procrule ctxt thm = let val (_, prems, lhs, elhs, rhs, _) = decomp_simp thm in if rewrite_rule_extra_vars prems lhs rhs - then (cond_warn_thm "Extra vars on rhs:" ss thm; []) + then (cond_warn_thm ctxt (fn () => "Extra vars on rhs:") thm; []) else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}] end; @@ -905,10 +940,9 @@ IMPORTANT: rewrite rules must not introduce new Vars or TVars! *) -fun rewritec (prover, thyt, maxt) ss t = +fun rewritec (prover, thyt, maxt) ctxt t = let - val ctxt = the_context ss; - val Simpset ({rules, ...}, {congs, procs, termless, ...}) = ss; + val Simpset ({rules, ...}, {congs, procs, termless, ...}) = simpset_of ctxt; val eta_thm = Thm.eta_conversion t; val eta_t' = Thm.rhs_of eta_thm; val eta_t = term_of eta_t'; @@ -927,36 +961,37 @@ val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop') in if perm andalso not (termless (rhs', lhs')) - then (trace_named_thm (fn () => "Cannot apply permutative rewrite rule") ss (thm, name); - trace_thm (fn () => "Term does not become smaller:") ss thm'; NONE) - else (trace_named_thm (fn () => "Applying instance of rewrite rule") ss (thm, name); + then (trace_named_thm ctxt (fn () => "Cannot apply permutative rewrite rule") (thm, name); + trace_thm ctxt (fn () => "Term does not become smaller:") thm'; NONE) + else (trace_named_thm ctxt (fn () => "Applying instance of rewrite rule") (thm, name); if unconditional then - (trace_thm (fn () => "Rewriting:") ss thm'; + (trace_thm ctxt (fn () => "Rewriting:") thm'; let val lr = Logic.dest_equals prop; - val SOME thm'' = check_conv false ss eta_thm thm'; + val SOME thm'' = check_conv ctxt false eta_thm thm'; in SOME (thm'', uncond_skel (congs, lr)) end) else - (trace_thm (fn () => "Trying to rewrite:") ss thm'; - if simp_depth ss > Config.get ctxt simp_depth_limit + (trace_thm ctxt (fn () => "Trying to rewrite:") thm'; + if simp_depth ctxt > Config.get ctxt simp_depth_limit then let val s = "simp_depth_limit exceeded - giving up"; - val _ = trace false (fn () => s) ss; - val _ = if_visible ss warning s; + val _ = trace ctxt false (fn () => s); + val _ = Context_Position.if_visible ctxt warning s; in NONE end else - case prover ss thm' of - NONE => (trace_thm (fn () => "FAILED") ss thm'; NONE) + case prover ctxt thm' of + NONE => (trace_thm ctxt (fn () => "FAILED") thm'; NONE) | SOME thm2 => - (case check_conv true ss eta_thm thm2 of - NONE => NONE | - SOME thm2' => - let val concl = Logic.strip_imp_concl prop - val lr = Logic.dest_equals concl - in SOME (thm2', cond_skel (congs, lr)) end))) - end + (case check_conv ctxt true eta_thm thm2 of + NONE => NONE + | SOME thm2' => + let + val concl = Logic.strip_imp_concl prop; + val lr = Logic.dest_equals concl; + in SOME (thm2', cond_skel (congs, lr)) end))) + end; fun rews [] = NONE | rews (rrule :: rrules) = @@ -979,15 +1014,15 @@ fun proc_rews [] = NONE | proc_rews (Proc {name, proc, lhs, ...} :: ps) = if Pattern.matches thyt (Thm.term_of lhs, Thm.term_of t) then - (debug_term false (fn () => "Trying procedure " ^ quote name ^ " on:") ss eta_t; - case proc ss eta_t' of - NONE => (debug false (fn () => "FAILED") ss; proc_rews ps) + (debug_term ctxt false (fn () => "Trying procedure " ^ quote name ^ " on:") eta_t; + case proc ctxt eta_t' of + NONE => (debug ctxt false (fn () => "FAILED"); proc_rews ps) | SOME raw_thm => - (trace_thm (fn () => "Procedure " ^ quote name ^ " produced rewrite rule:") - ss raw_thm; - (case rews (mk_procrule ss raw_thm) of - NONE => (trace_cterm true (fn () => "IGNORED result of simproc " ^ quote name ^ - " -- does not match") ss t; proc_rews ps) + (trace_thm ctxt (fn () => "Procedure " ^ quote name ^ " produced rewrite rule:") + raw_thm; + (case rews (mk_procrule ctxt raw_thm) of + NONE => (trace_cterm ctxt true (fn () => "IGNORED result of simproc " ^ quote name ^ + " -- does not match") t; proc_rews ps) | some => some))) else proc_rews ps; in @@ -1002,20 +1037,21 @@ (* conversion to apply a congruence rule to a term *) -fun congc prover ss maxt cong t = - let val rthm = Thm.incr_indexes (maxt + 1) cong; - val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (cprop_of rthm))); - val insts = Thm.match (rlhs, t) - (* Thm.match can raise Pattern.MATCH; - is handled when congc is called *) - val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm); - val _ = trace_thm (fn () => "Applying congruence rule:") ss thm'; - fun err (msg, thm) = (trace_thm (fn () => msg) ss thm; NONE) +fun congc prover ctxt maxt cong t = + let + val rthm = Thm.incr_indexes (maxt + 1) cong; + val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (cprop_of rthm))); + val insts = Thm.match (rlhs, t) + (* Thm.match can raise Pattern.MATCH; + is handled when congc is called *) + val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm); + val _ = trace_thm ctxt (fn () => "Applying congruence rule:") thm'; + fun err (msg, thm) = (trace_thm ctxt (fn () => msg) thm; NONE); in (case prover thm' of NONE => err ("Congruence proof failed. Could not prove", thm') | SOME thm2 => - (case check_conv true ss (Drule.beta_eta_conversion t) thm2 of + (case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of NONE => err ("Congruence proof failed. Should not have proved", thm2) | SOME thm2' => if op aconv (pairself term_of (Thm.dest_equals (cprop_of thm2'))) @@ -1035,108 +1071,111 @@ fun bottomc ((simprem, useprem, mutsimp), prover, thy, maxidx) = let - fun botc skel ss t = + fun botc skel ctxt t = if is_Var skel then NONE else - (case subc skel ss t of + (case subc skel ctxt t of some as SOME thm1 => - (case rewritec (prover, thy, maxidx) ss (Thm.rhs_of thm1) of + (case rewritec (prover, thy, maxidx) ctxt (Thm.rhs_of thm1) of SOME (thm2, skel2) => transitive2 (Thm.transitive thm1 thm2) - (botc skel2 ss (Thm.rhs_of thm2)) + (botc skel2 ctxt (Thm.rhs_of thm2)) | NONE => some) | NONE => - (case rewritec (prover, thy, maxidx) ss t of + (case rewritec (prover, thy, maxidx) ctxt t of SOME (thm2, skel2) => transitive2 thm2 - (botc skel2 ss (Thm.rhs_of thm2)) + (botc skel2 ctxt (Thm.rhs_of thm2)) | NONE => NONE)) - and try_botc ss t = - (case botc skel0 ss t of - SOME trec1 => trec1 | NONE => (Thm.reflexive t)) + and try_botc ctxt t = + (case botc skel0 ctxt t of + SOME trec1 => trec1 + | NONE => (Thm.reflexive t)) - and subc skel (ss as Simpset ({bounds, ...}, {congs, ...})) t0 = - (case term_of t0 of - Abs (a, T, _) => - let - val b = Name.bound (#1 bounds); - val (v, t') = Thm.dest_abs (SOME b) t0; - val b' = #1 (Term.dest_Free (Thm.term_of v)); - val _ = - if b <> b' then - warning ("Simplifier: renamed bound variable " ^ - quote b ^ " to " ^ quote b' ^ Position.here (Position.thread_data ())) - else (); - val ss' = add_bound ((b', T), a) ss; - val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0; - in case botc skel' ss' t' of - SOME thm => SOME (Thm.abstract_rule a v thm) - | NONE => NONE - end - | t $ _ => (case t of - Const ("==>", _) $ _ => impc t0 ss - | Abs _ => - let val thm = Thm.beta_conversion false t0 - in case subc skel0 ss (Thm.rhs_of thm) of - NONE => SOME thm - | SOME thm' => SOME (Thm.transitive thm thm') - end - | _ => - let fun appc () = - let - val (tskel, uskel) = case skel of - tskel $ uskel => (tskel, uskel) - | _ => (skel0, skel0); - val (ct, cu) = Thm.dest_comb t0 - in - (case botc tskel ss ct of - SOME thm1 => - (case botc uskel ss cu of - SOME thm2 => SOME (Thm.combination thm1 thm2) - | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu))) - | NONE => - (case botc uskel ss cu of - SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1) - | NONE => NONE)) - end - val (h, ts) = strip_comb t - in case cong_name h of - SOME a => - (case AList.lookup (op =) (fst congs) a of - NONE => appc () - | SOME cong => - (*post processing: some partial applications h t1 ... tj, j <= length ts, - may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*) - (let - val thm = congc (prover ss) ss maxidx cong t0; - val t = the_default t0 (Option.map Thm.rhs_of thm); - val (cl, cr) = Thm.dest_comb t - val dVar = Var(("", 0), dummyT) - val skel = - list_comb (h, replicate (length ts) dVar) - in case botc skel ss cl of - NONE => thm - | SOME thm' => transitive3 thm - (Thm.combination thm' (Thm.reflexive cr)) - end handle Pattern.MATCH => appc ())) - | _ => appc () - end) - | _ => NONE) + and subc skel ctxt t0 = + let val Simpset ({bounds, ...}, {congs, ...}) = simpset_of ctxt in + (case term_of t0 of + Abs (a, T, _) => + let + val b = Name.bound (#1 bounds); + val (v, t') = Thm.dest_abs (SOME b) t0; + val b' = #1 (Term.dest_Free (Thm.term_of v)); + val _ = + if b <> b' then + warning ("Simplifier: renamed bound variable " ^ + quote b ^ " to " ^ quote b' ^ Position.here (Position.thread_data ())) + else (); + val ctxt' = add_bound ((b', T), a) ctxt; + val skel' = case skel of Abs (_, _, sk) => sk | _ => skel0; + in + (case botc skel' ctxt' t' of + SOME thm => SOME (Thm.abstract_rule a v thm) + | NONE => NONE) + end + | t $ _ => (case t of + Const ("==>", _) $ _ => impc t0 ctxt + | Abs _ => + let val thm = Thm.beta_conversion false t0 + in case subc skel0 ctxt (Thm.rhs_of thm) of + NONE => SOME thm + | SOME thm' => SOME (Thm.transitive thm thm') + end + | _ => + let fun appc () = + let + val (tskel, uskel) = case skel of + tskel $ uskel => (tskel, uskel) + | _ => (skel0, skel0); + val (ct, cu) = Thm.dest_comb t0 + in + (case botc tskel ctxt ct of + SOME thm1 => + (case botc uskel ctxt cu of + SOME thm2 => SOME (Thm.combination thm1 thm2) + | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu))) + | NONE => + (case botc uskel ctxt cu of + SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1) + | NONE => NONE)) + end + val (h, ts) = strip_comb t + in case cong_name h of + SOME a => + (case AList.lookup (op =) (fst congs) a of + NONE => appc () + | SOME cong => + (*post processing: some partial applications h t1 ... tj, j <= length ts, + may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*) + (let + val thm = congc (prover ctxt) ctxt maxidx cong t0; + val t = the_default t0 (Option.map Thm.rhs_of thm); + val (cl, cr) = Thm.dest_comb t + val dVar = Var(("", 0), dummyT) + val skel = + list_comb (h, replicate (length ts) dVar) + in case botc skel ctxt cl of + NONE => thm + | SOME thm' => transitive3 thm + (Thm.combination thm' (Thm.reflexive cr)) + end handle Pattern.MATCH => appc ())) + | _ => appc () + end) + | _ => NONE) + end + and impc ct ctxt = + if mutsimp then mut_impc0 [] ct [] [] ctxt else nonmut_impc ct ctxt - and impc ct ss = - if mutsimp then mut_impc0 [] ct [] [] ss else nonmut_impc ct ss - - and rules_of_prem ss prem = + and rules_of_prem ctxt prem = if maxidx_of_term (term_of prem) <> ~1 - then (trace_cterm true + then (trace_cterm ctxt true (fn () => "Cannot add premise as rewrite rule because it contains (type) unknowns:") - ss prem; ([], NONE)) + prem; ([], NONE)) else let val asm = Thm.assume prem - in (extract_safe_rrules (ss, asm), SOME asm) end + in (extract_safe_rrules (ctxt, asm), SOME asm) end - and add_rrules (rrss, asms) ss = - (fold o fold) insert_rrule rrss ss |> add_prems (map_filter I asms) + and add_rrules (rrss, asms) ctxt = + (fold o fold) insert_rrule rrss ctxt |> add_prems (map_filter I asms) and disch r prem eq = let @@ -1157,44 +1196,44 @@ end and rebuild [] _ _ _ _ eq = eq - | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ss eq = + | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ctxt eq = let - val ss' = add_rrules (rev rrss, rev asms) ss; + val ctxt' = add_rrules (rev rrss, rev asms) ctxt; val concl' = Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq)); val dprem = Option.map (disch false prem) in - (case rewritec (prover, thy, maxidx) ss' concl' of - NONE => rebuild prems concl' rrss asms ss (dprem eq) + (case rewritec (prover, thy, maxidx) ctxt' concl' of + NONE => rebuild prems concl' rrss asms ctxt (dprem eq) | SOME (eq', _) => transitive2 (fold (disch false) prems (the (transitive3 (dprem eq) eq'))) - (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ss)) + (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ctxt)) end - and mut_impc0 prems concl rrss asms ss = + and mut_impc0 prems concl rrss asms ctxt = let val prems' = strip_imp_prems concl; - val (rrss', asms') = split_list (map (rules_of_prem ss) prems') + val (rrss', asms') = split_list (map (rules_of_prem ctxt) prems') in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss') - (asms @ asms') [] [] [] [] ss ~1 ~1 + (asms @ asms') [] [] [] [] ctxt ~1 ~1 end - and mut_impc [] concl [] [] prems' rrss' asms' eqns ss changed k = + and mut_impc [] concl [] [] prems' rrss' asms' eqns ctxt changed k = transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1 (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE) (if changed > 0 then mut_impc (rev prems') concl (rev rrss') (rev asms') - [] [] [] [] ss ~1 changed - else rebuild prems' concl rrss' asms' ss - (botc skel0 (add_rrules (rev rrss', rev asms') ss) concl)) + [] [] [] [] ctxt ~1 changed + else rebuild prems' concl rrss' asms' ctxt + (botc skel0 (add_rrules (rev rrss', rev asms') ctxt) concl)) | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms) - prems' rrss' asms' eqns ss changed k = + prems' rrss' asms' eqns ctxt changed k = case (if k = 0 then NONE else botc skel0 (add_rrules - (rev rrss' @ rrss, rev asms' @ asms) ss) prem) of + (rev rrss' @ rrss, rev asms' @ asms) ctxt) prem) of NONE => mut_impc prems concl rrss asms (prem :: prems') - (rrs :: rrss') (asm :: asms') (NONE :: eqns) ss changed + (rrs :: rrss') (asm :: asms') (NONE :: eqns) ctxt changed (if k = 0 then 0 else k - 1) | SOME eqn => let @@ -1202,26 +1241,26 @@ val tprems = map term_of prems; val i = 1 + fold Integer.max (map (fn p => find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1; - val (rrs', asm') = rules_of_prem ss prem' + val (rrs', asm') = rules_of_prem ctxt prem'; in mut_impc prems concl rrss asms (prem' :: prems') (rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true) (take i prems) (Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies (drop i prems, concl))))) :: eqns) - ss (length prems') ~1 + ctxt (length prems') ~1 end (*legacy code - only for backwards compatibility*) - and nonmut_impc ct ss = + and nonmut_impc ct ctxt = let val (prem, conc) = Thm.dest_implies ct; - val thm1 = if simprem then botc skel0 ss prem else NONE; + val thm1 = if simprem then botc skel0 ctxt prem else NONE; val prem1 = the_default prem (Option.map Thm.rhs_of thm1); - val ss1 = - if not useprem then ss - else add_rrules (apsnd single (apfst single (rules_of_prem ss prem1))) ss + val ctxt1 = + if not useprem then ctxt + else add_rrules (apsnd single (apfst single (rules_of_prem ctxt prem1))) ctxt in - (case botc skel0 ss1 conc of + (case botc skel0 ctxt1 conc of NONE => (case thm1 of NONE => NONE @@ -1251,42 +1290,47 @@ val debug_bounds = Unsynchronized.ref false; -fun check_bounds ss ct = +fun check_bounds ctxt ct = if ! debug_bounds then let - val Simpset ({bounds = (_, bounds), ...}, _) = ss; - val bs = fold_aterms (fn Free (x, _) => - if Name.is_bound x andalso not (AList.defined eq_bound bounds x) - then insert (op =) x else I - | _ => I) (term_of ct) []; + val Simpset ({bounds = (_, bounds), ...}, _) = simpset_of ctxt; + val bs = + fold_aterms + (fn Free (x, _) => + if Name.is_bound x andalso not (AList.defined eq_bound bounds x) + then insert (op =) x else I + | _ => I) (term_of ct) []; in if null bs then () - else print_term_global ss true ("Simplifier: term contains loose bounds: " ^ commas_quote bs) - (Thm.theory_of_cterm ct) (Thm.term_of ct) + else + print_term ctxt true + (fn () => "Simplifier: term contains loose bounds: " ^ commas_quote bs) + (Thm.term_of ct) end else (); -fun rewrite_cterm mode prover raw_ss raw_ct = +fun rewrite_cterm mode prover raw_ctxt raw_ct = let val thy = Thm.theory_of_cterm raw_ct; val ct = Thm.adjust_maxidx_cterm ~1 raw_ct; val {maxidx, ...} = Thm.rep_cterm ct; - val ss = inc_simp_depth (activate_context thy raw_ss); - val depth = simp_depth ss; + val ctxt = inc_simp_depth (activate_context thy raw_ctxt); + val depth = simp_depth ctxt; val _ = if depth mod 20 = 0 then - if_visible ss warning ("Simplification depth " ^ string_of_int depth) + Context_Position.if_visible ctxt warning ("Simplification depth " ^ string_of_int depth) else (); - val _ = trace_cterm false (fn () => "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:") ss ct; - val _ = check_bounds ss ct; - in bottomc (mode, Option.map Drule.flexflex_unique oo prover, thy, maxidx) ss ct end; + val _ = trace_cterm ctxt false (fn () => "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:") ct; + val _ = check_bounds ctxt ct; + in bottomc (mode, Option.map Drule.flexflex_unique oo prover, thy, maxidx) ctxt ct end; val simple_prover = - SINGLE o (fn ss => ALLGOALS (resolve_tac (prems_of ss))); + SINGLE o (fn ctxt => ALLGOALS (resolve_tac (prems_of ctxt))); fun rewrite _ [] ct = Thm.reflexive ct - | rewrite full thms ct = rewrite_cterm (full, false, false) simple_prover - (global_context (Thm.theory_of_cterm ct) empty_ss addsimps thms) ct; + | rewrite full thms ct = + rewrite_cterm (full, false, false) simple_prover + (global_context (Thm.theory_of_cterm ct) empty_ss addsimps thms) ct; fun simplify full thms = Conv.fconv_rule (rewrite full thms); val rewrite_rule = simplify true; @@ -1295,7 +1339,7 @@ fun rewrite_term thy rules procs = Pattern.rewrite_term thy (map decomp_simp' rules) procs; -fun rewrite_thm mode prover ss = Conv.fconv_rule (rewrite_cterm mode prover ss); +fun rewrite_thm mode prover ctxt = Conv.fconv_rule (rewrite_cterm mode prover ctxt); (*Rewrite the subgoals of a proof state (represented by a theorem)*) fun rewrite_goals_rule thms th = @@ -1309,16 +1353,14 @@ fun rewrite_goals_tac defs = PRIMITIVE (rewrite_goals_rule defs); (*Rewrite one subgoal*) -fun generic_rewrite_goal_tac mode prover_tac ss i thm = +fun generic_rewrite_goal_tac mode prover_tac ctxt i thm = if 0 < i andalso i <= Thm.nprems_of thm then - Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ss) i thm) + Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ctxt) i thm) else Seq.empty; -fun rewrite_goal_tac rews = - let val ss = empty_ss addsimps rews in - fn i => fn st => generic_rewrite_goal_tac (true, false, false) (K no_tac) - (global_context (Thm.theory_of_thm st) ss) i st - end; +fun rewrite_goal_tac rews i st = + generic_rewrite_goal_tac (true, false, false) (K no_tac) + (global_context (Thm.theory_of_thm st) empty_ss addsimps rews) i st; (*Prunes all redundant parameters from the proof state by rewriting.*) val prune_params_tac = rewrite_goals_tac [Drule.triv_forall_equality]; @@ -1352,17 +1394,25 @@ fun gen_norm_hhf ss th = (if Drule.is_norm_hhf (Thm.prop_of th) then th - else Conv.fconv_rule - (rewrite_cterm (true, false, false) (K (K NONE)) (global_context (Thm.theory_of_thm th) ss)) th) + else + Conv.fconv_rule + (rewrite_cterm (true, false, false) (K (K NONE)) + (global_context (Thm.theory_of_thm th) ss)) th) |> Thm.adjust_maxidx_thm ~1 |> Drule.gen_all; -val hhf_ss = empty_ss addsimps Drule.norm_hhf_eqs; +val hhf_ss = + simpset_of (empty_simpset (Context.proof_of (Context.the_thread_data ())) + addsimps Drule.norm_hhf_eqs); + +val hhf_protect_ss = + simpset_of (empty_simpset (Context.proof_of (Context.the_thread_data ())) + addsimps Drule.norm_hhf_eqs |> add_eqcong Drule.protect_cong); in val norm_hhf = gen_norm_hhf hhf_ss; -val norm_hhf_protect = gen_norm_hhf (hhf_ss |> add_eqcong Drule.protect_cong); +val norm_hhf_protect = gen_norm_hhf hhf_protect_ss; end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Pure/simplifier.ML --- a/src/Pure/simplifier.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Pure/simplifier.ML Thu Apr 18 17:07:01 2013 +0200 @@ -8,67 +8,57 @@ signature BASIC_SIMPLIFIER = sig include BASIC_RAW_SIMPLIFIER - val map_simpset: (simpset -> simpset) -> Proof.context -> Proof.context - val simpset_of: Proof.context -> simpset val Addsimprocs: simproc list -> unit val Delsimprocs: simproc list -> unit - val simp_tac: simpset -> int -> tactic - val asm_simp_tac: simpset -> int -> tactic - val full_simp_tac: simpset -> int -> tactic - val asm_lr_simp_tac: simpset -> int -> tactic - val asm_full_simp_tac: simpset -> int -> tactic - val safe_simp_tac: simpset -> int -> tactic - val safe_asm_simp_tac: simpset -> int -> tactic - val safe_full_simp_tac: simpset -> int -> tactic - val safe_asm_lr_simp_tac: simpset -> int -> tactic - val safe_asm_full_simp_tac: simpset -> int -> tactic - val simplify: simpset -> thm -> thm - val asm_simplify: simpset -> thm -> thm - val full_simplify: simpset -> thm -> thm - val asm_lr_simplify: simpset -> thm -> thm - val asm_full_simplify: simpset -> thm -> thm + val simp_tac: Proof.context -> int -> tactic + val asm_simp_tac: Proof.context -> int -> tactic + val full_simp_tac: Proof.context -> int -> tactic + val asm_lr_simp_tac: Proof.context -> int -> tactic + val asm_full_simp_tac: Proof.context -> int -> tactic + val safe_simp_tac: Proof.context -> int -> tactic + val safe_asm_simp_tac: Proof.context -> int -> tactic + val safe_full_simp_tac: Proof.context -> int -> tactic + val safe_asm_lr_simp_tac: Proof.context -> int -> tactic + val safe_asm_full_simp_tac: Proof.context -> int -> tactic + val simplify: Proof.context -> thm -> thm + val asm_simplify: Proof.context -> thm -> thm + val full_simplify: Proof.context -> thm -> thm + val asm_lr_simplify: Proof.context -> thm -> thm + val asm_full_simplify: Proof.context -> thm -> thm end; signature SIMPLIFIER = sig include BASIC_SIMPLIFIER - val map_simpset_global: (simpset -> simpset) -> theory -> theory - val pretty_ss: Proof.context -> simpset -> Pretty.T - val clear_ss: simpset -> simpset - val default_mk_sym: simpset -> thm -> thm option + val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic + val pretty_simpset: Proof.context -> Pretty.T + val default_mk_sym: Proof.context -> thm -> thm option val debug_bounds: bool Unsynchronized.ref - val prems_of: simpset -> thm list - val add_simp: thm -> simpset -> simpset - val del_simp: thm -> simpset -> simpset - val add_eqcong: thm -> simpset -> simpset - val del_eqcong: thm -> simpset -> simpset - val add_cong: thm -> simpset -> simpset - val del_cong: thm -> simpset -> simpset - val add_prems: thm list -> simpset -> simpset - val mksimps: simpset -> thm -> thm list - val set_mksimps: (simpset -> thm -> thm list) -> simpset -> simpset - val set_mkcong: (simpset -> thm -> thm) -> simpset -> simpset - val set_mksym: (simpset -> thm -> thm option) -> simpset -> simpset - val set_mkeqTrue: (simpset -> thm -> thm option) -> simpset -> simpset - val set_termless: (term * term -> bool) -> simpset -> simpset - val set_subgoaler: (simpset -> int -> tactic) -> simpset -> simpset - val inherit_context: simpset -> simpset -> simpset - val the_context: simpset -> Proof.context - val context: Proof.context -> simpset -> simpset - val global_context: theory -> simpset -> simpset - val with_context: Proof.context -> (simpset -> simpset) -> simpset -> simpset + val prems_of: Proof.context -> thm list + val add_simp: thm -> Proof.context -> Proof.context + val del_simp: thm -> Proof.context -> Proof.context + val add_eqcong: thm -> Proof.context -> Proof.context + val del_eqcong: thm -> Proof.context -> Proof.context + val add_cong: thm -> Proof.context -> Proof.context + val del_cong: thm -> Proof.context -> Proof.context + val add_prems: thm list -> Proof.context -> Proof.context + val mksimps: Proof.context -> thm -> thm list + val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context + val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context + val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context + val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context + val set_termless: (term * term -> bool) -> Proof.context -> Proof.context + val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context val simproc_global_i: theory -> string -> term list -> - (theory -> simpset -> term -> thm option) -> simproc + (Proof.context -> term -> thm option) -> simproc val simproc_global: theory -> string -> string list -> - (theory -> simpset -> term -> thm option) -> simproc - val rewrite: simpset -> conv - val asm_rewrite: simpset -> conv - val full_rewrite: simpset -> conv - val asm_lr_rewrite: simpset -> conv - val asm_full_rewrite: simpset -> conv - val get_ss: Context.generic -> simpset - val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic - val attrib: (thm -> simpset -> simpset) -> attribute + (Proof.context -> term -> thm option) -> simproc + val rewrite: Proof.context -> conv + val asm_rewrite: Proof.context -> conv + val full_rewrite: Proof.context -> conv + val asm_lr_rewrite: Proof.context -> conv + val asm_full_rewrite: Proof.context -> conv + val attrib: (thm -> Proof.context -> Proof.context) -> attribute val simp_add: attribute val simp_del: attribute val cong_add: attribute @@ -76,10 +66,10 @@ val check_simproc: Proof.context -> xstring * Position.T -> string val the_simproc: Proof.context -> string -> simproc val def_simproc: {name: binding, lhss: term list, - proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> + proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> local_theory -> local_theory val def_simproc_cmd: {name: binding, lhss: string list, - proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> + proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> local_theory -> local_theory val cong_modifiers: Method.modifier parser list val simp_modifiers': Method.modifier parser list @@ -94,29 +84,9 @@ open Raw_Simplifier; -(** data **) - -structure Data = Generic_Data -( - type T = simpset * simproc Name_Space.table; - val empty : T = (empty_ss, Name_Space.empty_table "simproc"); - fun extend (ss, tab) = (Raw_Simplifier.inherit_context empty_ss ss, tab); - fun merge ((ss1, tab1), (ss2, tab2)) = - (merge_ss (ss1, ss2), Name_Space.merge_tables (tab1, tab2)); -); - -val get_ss = fst o Data.get; - -fun map_ss f context = - Data.map (apfst ((Raw_Simplifier.with_context (Context.proof_of context) f))) context; - -val get_simprocs = snd o Data.get o Context.Proof; - - - (** pretty printing **) -fun pretty_ss ctxt ss = +fun pretty_simpset ctxt = let val pretty_term = Syntax.pretty_term ctxt; val pretty_thm = Display.pretty_thm ctxt; @@ -130,7 +100,8 @@ fun pretty_cong (name, thm) = Pretty.block [pretty_cong_name name, Pretty.str ":", Pretty.brk 1, pretty_thm thm]; - val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss ss; + val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = + dest_ss (simpset_of ctxt); in [Pretty.big_list "simplification rules:" (map (pretty_thm_item o #2) simps), Pretty.big_list "simplification procedures:" (map pretty_proc (sort_wrt #1 procs)), @@ -143,7 +114,7 @@ -(** simpset data **) +(** declarations **) (* attributes *) @@ -155,29 +126,31 @@ val cong_del = attrib del_cong; -(* local simpset *) - -fun map_simpset f = Context.proof_map (map_ss f); -fun simpset_of ctxt = Raw_Simplifier.context ctxt (get_ss (Context.Proof ctxt)); +(* global simprocs *) -val _ = Context.>> (Context.map_theory - (ML_Antiquote.value (Binding.name "simpset") - (Scan.succeed "Simplifier.simpset_of ML_context"))); - +fun Addsimprocs args = + Context.>> (Context.map_theory (map_theory_simpset (fn ctxt => ctxt addsimprocs args))); -(* global simpset *) - -fun map_simpset_global f = Context.theory_map (map_ss f); - -fun Addsimprocs args = Context.>> (map_ss (fn ss => ss addsimprocs args)); -fun Delsimprocs args = Context.>> (map_ss (fn ss => ss delsimprocs args)); +fun Delsimprocs args = + Context.>> (Context.map_theory (map_theory_simpset (fn ctxt => ctxt delsimprocs args))); (** named simprocs **) +structure Simprocs = Generic_Data +( + type T = simproc Name_Space.table; + val empty : T = Name_Space.empty_table "simproc"; + val extend = I; + fun merge data : T = Name_Space.merge_tables data; +); + + (* get simprocs *) +val get_simprocs = Simprocs.get o Context.Proof; + fun check_simproc ctxt = Name_Space.check (Context.Proof ctxt) (get_simprocs ctxt) #> #1; val the_simproc = Name_Space.get o get_simprocs; @@ -213,8 +186,8 @@ val simproc' = transform_simproc phi simproc; in context - |> Data.map (fn (ss, tab) => - (ss addsimprocs [simproc'], #2 (Name_Space.define context true (b', simproc') tab))) + |> Simprocs.map (#2 o Name_Space.define context true (b', simproc')) + |> map_ss (fn ctxt => ctxt addsimprocs [simproc']) end) end; @@ -229,33 +202,34 @@ (** simplification tactics and rules **) -fun solve_all_tac solvers ss = +fun solve_all_tac solvers ctxt = let - val (_, {subgoal_tac, ...}) = Raw_Simplifier.internal_ss ss; - val solve_tac = subgoal_tac (Raw_Simplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac); + val (_, {subgoal_tac, ...}) = Raw_Simplifier.internal_ss (simpset_of ctxt); + val solve_tac = subgoal_tac (Raw_Simplifier.set_solvers solvers ctxt) THEN_ALL_NEW (K no_tac); in DEPTH_SOLVE (solve_tac 1) end; (*NOTE: may instantiate unknowns that appear also in other subgoals*) -fun generic_simp_tac safe mode ss = +fun generic_simp_tac safe mode ctxt = let - val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = Raw_Simplifier.internal_ss ss; - val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs)); - val solve_tac = FIRST' (map (Raw_Simplifier.solver ss) + val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = + Raw_Simplifier.internal_ss (simpset_of ctxt); + val loop_tac = FIRST' (map (fn (_, tac) => tac ctxt) (rev loop_tacs)); + val solve_tac = FIRST' (map (Raw_Simplifier.solver ctxt) (rev (if safe then solvers else unsafe_solvers))); fun simp_loop_tac i = - Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN + Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ctxt i THEN (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); in SELECT_GOAL (simp_loop_tac 1) end; local -fun simp rew mode ss thm = +fun simp rew mode ctxt thm = let - val (_, {solvers = (unsafe_solvers, _), ...}) = Raw_Simplifier.internal_ss ss; + val (_, {solvers = (unsafe_solvers, _), ...}) = Raw_Simplifier.internal_ss (simpset_of ctxt); val tacf = solve_all_tac (rev unsafe_solvers); fun prover s th = Option.map #1 (Seq.pull (tacf s th)); - in rew mode prover ss thm end; + in rew mode prover ctxt thm end; in @@ -318,7 +292,7 @@ (Args.del -- Args.colon >> K (op delsimprocs) || Scan.option (Args.add -- Args.colon) >> K (op addsimprocs)) >> (fn f => fn simproc => fn phi => Thm.declaration_attribute - (K (map_ss (fn ss => f (ss, [transform_simproc phi simproc]))))); + (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc]))))); in @@ -345,8 +319,8 @@ val simplified = conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute (fn context => - f ((if null ths then I else Raw_Simplifier.clear_ss) - (simpset_of (Context.proof_of context)) addsimps ths))); + f ((if null ths then I else Raw_Simplifier.clear_simpset) + (Context.proof_of context) addsimps ths))); end; @@ -375,15 +349,13 @@ [Args.$$$ simpN -- Args.colon >> K (I, simp_add), Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add), Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del), - Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon - >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)] + Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >> K (Raw_Simplifier.clear_simpset, simp_add)] @ cong_modifiers; val simp_modifiers' = [Args.add -- Args.colon >> K (I, simp_add), Args.del -- Args.colon >> K (I, simp_del), - Args.$$$ onlyN -- Args.colon - >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)] + Args.$$$ onlyN -- Args.colon >> K (Raw_Simplifier.clear_simpset, simp_add)] @ cong_modifiers; val simp_options = @@ -406,25 +378,25 @@ Method.setup (Binding.name simpN) (simp_method more_mods (fn ctxt => fn tac => fn facts => HEADGOAL (Method.insert_tac facts THEN' - (CHANGED_PROP oo tac) (simpset_of ctxt)))) + (CHANGED_PROP oo tac) ctxt))) "simplification" #> Method.setup (Binding.name "simp_all") (simp_method more_mods (fn ctxt => fn tac => fn facts => ALLGOALS (Method.insert_tac facts) THEN - (CHANGED_PROP o PARALLEL_GOALS o ALLGOALS o tac) (simpset_of ctxt))) + (CHANGED_PROP o PARALLEL_GOALS o ALLGOALS o tac) ctxt)) "simplification (all goals)"; -fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ => +fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn ctxt0 => let val trivialities = Drule.reflexive_thm :: trivs; - fun unsafe_solver_tac ss = - FIRST' [resolve_tac (trivialities @ Raw_Simplifier.prems_of ss), assume_tac]; + fun unsafe_solver_tac ctxt = + FIRST' [resolve_tac (trivialities @ Raw_Simplifier.prems_of ctxt), assume_tac]; val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac; (*no premature instantiation of variables during simplification*) - fun safe_solver_tac ss = - FIRST' [match_tac (trivialities @ Raw_Simplifier.prems_of ss), eq_assume_tac]; + fun safe_solver_tac ctxt = + FIRST' [match_tac (trivialities @ Raw_Simplifier.prems_of ctxt), eq_assume_tac]; val safe_solver = mk_solver "easy safe" safe_solver_tac; fun mk_eq thm = @@ -433,7 +405,7 @@ fun mksimps thm = mk_eq (Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm); in - empty_ss + empty_simpset ctxt0 setSSolver safe_solver setSolver unsafe_solver |> set_subgoaler asm_simp_tac diff -r 19b47bfac6ef -r 9e7d1c139569 src/Sequents/LK.thy --- a/src/Sequents/LK.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Sequents/LK.thy Thu Apr 18 17:07:01 2013 +0200 @@ -215,20 +215,20 @@ done ML_file "simpdata.ML" -setup {* Simplifier.map_simpset_global (K LK_ss) *} +setup {* map_theory_simpset (put_simpset LK_ss) *} text {* To create substition rules *} lemma eq_imp_subst: "|- a=b ==> $H, A(a), $G |- $E, A(b), $F" - apply (tactic {* asm_simp_tac LK_basic_ss 1 *}) + apply (tactic {* asm_simp_tac (put_simpset LK_basic_ss @{context}) 1 *}) done lemma split_if: "|- P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))" apply (rule_tac P = Q in cut) - apply (tactic {* simp_tac (@{simpset} addsimps @{thms if_P}) 2 *}) + apply (tactic {* simp_tac (@{context} addsimps @{thms if_P}) 2 *}) apply (rule_tac P = "~Q" in cut) - apply (tactic {* simp_tac (@{simpset} addsimps @{thms if_not_P}) 2 *}) + apply (tactic {* simp_tac (@{context} addsimps @{thms if_not_P}) 2 *}) apply (tactic {* fast_tac LK_pack 1 *}) done diff -r 19b47bfac6ef -r 9e7d1c139569 src/Sequents/LK/Nat.thy --- a/src/Sequents/LK/Nat.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Sequents/LK/Nat.thy Thu Apr 18 17:07:01 2013 +0200 @@ -36,7 +36,7 @@ lemma Suc_n_not_n: "|- Suc(k) ~= k" apply (rule_tac n = k in induct) - apply (tactic {* simp_tac (LK_ss addsimps @{thms Suc_neq_0}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms Suc_neq_0}) 1 *}) apply (tactic {* fast_tac (LK_pack add_safes @{thms Suc_inject_rule}) 1 *}) done @@ -54,26 +54,26 @@ lemma add_assoc: "|- (k+m)+n = k+(m+n)" apply (rule_tac n = "k" in induct) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_0}) 1 *}) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_Suc}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *}) done lemma add_0_right: "|- m+0 = m" apply (rule_tac n = "m" in induct) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_0}) 1 *}) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_Suc}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *}) done lemma add_Suc_right: "|- m+Suc(n) = Suc(m+n)" apply (rule_tac n = "m" in induct) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_0}) 1 *}) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_Suc}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *}) done lemma "(!!n. |- f(Suc(n)) = Suc(f(n))) ==> |- f(i+j) = i+f(j)" apply (rule_tac n = "i" in induct) - apply (tactic {* simp_tac (LK_ss addsimps @{thms add_0}) 1 *}) - apply (tactic {* asm_simp_tac (LK_ss addsimps @{thms add_Suc}) 1 *}) + apply (tactic {* simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_0}) 1 *}) + apply (tactic {* asm_simp_tac (put_simpset LK_ss @{context} addsimps @{thms add_Suc}) 1 *}) done end diff -r 19b47bfac6ef -r 9e7d1c139569 src/Sequents/simpdata.ML --- a/src/Sequents/simpdata.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Sequents/simpdata.ML Thu Apr 18 17:07:01 2013 +0200 @@ -47,8 +47,8 @@ (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}])); (*Congruence rules for = or <-> (instead of ==)*) -fun mk_meta_cong ss rl = - Drule.zero_var_indexes (mk_meta_eq (mk_meta_prems (Simplifier.the_context ss) rl)) +fun mk_meta_cong ctxt rl = + Drule.zero_var_indexes (mk_meta_eq (mk_meta_prems ctxt rl)) handle THM _ => error("Premises and conclusion of congruence rules must use =-equality or <->"); @@ -58,21 +58,22 @@ val triv_rls = [@{thm FalseL}, @{thm TrueR}, @{thm basic}, @{thm refl}, @{thm iff_refl}, reflexive_thm]; -fun unsafe_solver ss = - FIRST' [resolve_tac (triv_rls @ Simplifier.prems_of ss), assume_tac]; +fun unsafe_solver ctxt = + FIRST' [resolve_tac (triv_rls @ Simplifier.prems_of ctxt), assume_tac]; (*No premature instantiation of variables during simplification*) -fun safe_solver ss = - FIRST' [fn i => DETERM (match_tac (triv_rls @ Simplifier.prems_of ss) i), eq_assume_tac]; +fun safe_solver ctxt = + FIRST' [fn i => DETERM (match_tac (triv_rls @ Simplifier.prems_of ctxt) i), eq_assume_tac]; (*No simprules, but basic infrastructure for simplification*) val LK_basic_ss = - Simplifier.global_context @{theory} empty_ss + empty_simpset @{context} setSSolver (mk_solver "safe" safe_solver) setSolver (mk_solver "unsafe" unsafe_solver) |> Simplifier.set_subgoaler asm_simp_tac |> Simplifier.set_mksimps (K (map mk_meta_eq o atomize o gen_all)) - |> Simplifier.set_mkcong mk_meta_cong; + |> Simplifier.set_mkcong mk_meta_cong + |> simpset_of; val LK_simps = [@{thm triv_forall_equality}, (* prunes params *) @@ -84,7 +85,8 @@ @{thms LK_extra_simps}; val LK_ss = - LK_basic_ss addsimps LK_simps + put_simpset LK_basic_ss @{context} addsimps LK_simps |> Simplifier.add_eqcong @{thm left_cong} - |> Simplifier.add_cong @{thm imp_cong}; + |> Simplifier.add_cong @{thm imp_cong} + |> simpset_of; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Tools/Code/code_preproc.ML --- a/src/Tools/Code/code_preproc.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Tools/Code/code_preproc.ML Thu Apr 18 17:07:01 2013 +0200 @@ -7,8 +7,8 @@ signature CODE_PREPROC = sig - val map_pre: (simpset -> simpset) -> theory -> theory - val map_post: (simpset -> simpset) -> theory -> theory + val map_pre: (Proof.context -> Proof.context) -> theory -> theory + val map_post: (Proof.context -> Proof.context) -> theory -> theory val add_unfold: thm -> theory -> theory val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory val del_functrans: string -> theory -> theory @@ -79,8 +79,11 @@ val map_data = Code_Preproc_Data.map o map_thmproc; val map_pre_post = map_data o apfst; -val map_pre = map_pre_post o apfst; -val map_post = map_pre_post o apsnd; + +fun map_simpset which f thy = + map_pre_post (which (simpset_map (Proof_Context.init_global thy) f)) thy; +val map_pre = map_simpset apfst; +val map_post = map_simpset apsnd; val add_unfold = map_pre o Simplifier.add_simp; val del_unfold = map_pre o Simplifier.del_simp; @@ -89,11 +92,13 @@ fun add_code_abbrev raw_thm thy = let - val thm = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy) raw_thm; + val ctxt = Proof_Context.init_global thy; + val thm = Local_Defs.meta_rewrite_rule ctxt raw_thm; val thm_sym = Thm.symmetric thm; in thy |> map_pre_post (fn (pre, post) => - (pre |> Simplifier.add_simp thm_sym, post |> Simplifier.add_simp thm)) + (pre |> simpset_map ctxt (Simplifier.add_simp thm_sym), + post |> simpset_map ctxt (Simplifier.add_simp thm))) end; fun add_functrans (name, f) = (map_data o apsnd) @@ -169,12 +174,12 @@ Pretty.block [ Pretty.str "preprocessing simpset:", Pretty.fbrk, - Simplifier.pretty_ss ctxt pre + Simplifier.pretty_simpset (put_simpset pre ctxt) ], Pretty.block [ Pretty.str "postprocessing simpset:", Pretty.fbrk, - Simplifier.pretty_ss ctxt post + Simplifier.pretty_simpset (put_simpset post ctxt) ], Pretty.block ( Pretty.str "function transformers:" @@ -261,7 +266,7 @@ | NONE => let val functrans = (map (fn (_, (_, f)) => f thy) o #functrans o the_thmproc) thy; - val pre = (Simplifier.global_context thy o #pre o the_thmproc) thy; + val {pre, ...} = the_thmproc thy; val cert = Code.get_cert thy { functrans = functrans, ss = pre } c; val (lhs, rhss) = Code.typargs_deps_of_cert thy cert; in ((lhs, rhss), cert) end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/Tools/Code/code_simp.ML --- a/src/Tools/Code/code_simp.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Tools/Code/code_simp.ML Thu Apr 18 17:07:01 2013 +0200 @@ -6,7 +6,7 @@ signature CODE_SIMP = sig - val map_ss: (simpset -> simpset) -> theory -> theory + val map_ss: (Proof.context -> Proof.context) -> theory -> theory val dynamic_conv: theory -> conv val dynamic_tac: theory -> int -> tactic val dynamic_value: theory -> term -> term @@ -24,11 +24,11 @@ ( type T = simpset; val empty = empty_ss; - fun extend ss = Simplifier.inherit_context empty_ss ss; + val extend = I; val merge = merge_ss; ); -val map_ss = Simpset.map; +fun map_ss f thy = Simpset.map (simpset_map (Proof_Context.init_global thy) f) thy; fun simpset_default thy = Simplifier.global_context thy o the_default (Simpset.get thy); diff -r 19b47bfac6ef -r 9e7d1c139569 src/Tools/eqsubst.ML --- a/src/Tools/eqsubst.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Tools/eqsubst.ML Thu Apr 18 17:07:01 2013 +0200 @@ -68,7 +68,7 @@ (* changes object "=" to meta "==" which prepares a given rewrite rule *) fun prep_meta_eq ctxt = - Simplifier.mksimps (simpset_of ctxt) #> map Drule.zero_var_indexes; + Simplifier.mksimps ctxt #> map Drule.zero_var_indexes; (* a type abriviation for match information *) diff -r 19b47bfac6ef -r 9e7d1c139569 src/Tools/induct.ML --- a/src/Tools/induct.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/Tools/induct.ML Thu Apr 18 17:07:01 2013 +0200 @@ -45,7 +45,7 @@ val coinduct_type: string -> attribute val coinduct_pred: string -> attribute val coinduct_del: attribute - val map_simpset: (simpset -> simpset) -> Context.generic -> Context.generic + val map_simpset: (Proof.context -> Proof.context) -> Context.generic -> Context.generic val induct_simp_add: attribute val induct_simp_del: attribute val no_simpN: string @@ -164,8 +164,7 @@ val rearrange_eqs_simproc = Simplifier.simproc_global (Thm.theory_of_thm Drule.swap_prems_eq) "rearrange_eqs" ["all t"] - (fn thy => fn ss => fn t => - mk_swap_rrule (Simplifier.the_context ss) (cterm_of thy t)); + (fn ctxt => fn t => mk_swap_rrule ctxt (cterm_of (Proof_Context.theory_of ctxt) t)); (* rotate k premises to the left by j, skipping over first j premises *) @@ -225,7 +224,8 @@ ((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)), (init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)), (init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)), - empty_ss addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq]); + simpset_of (empty_simpset @{context} + addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq])); val extend = I; fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1), ((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) = @@ -331,12 +331,11 @@ val coinduct_pred = mk_att add_coinductP consumes1; val coinduct_del = del_att map3; -fun map_simpset f = Data.map (map4 f); +fun map_simpset f context = + context |> (Data.map o map4 o Simplifier.simpset_map (Context.proof_of context)) f; fun induct_simp f = - Thm.declaration_attribute (fn thm => fn context => - map_simpset - (Simplifier.with_context (Context.proof_of context) (fn ss => f (ss, [thm]))) context); + Thm.declaration_attribute (fn thm => map_simpset (fn ctxt => f (ctxt, [thm]))); val induct_simp_add = induct_simp (op addsimps); val induct_simp_del = induct_simp (op delsimps); @@ -444,7 +443,7 @@ (* simplify *) fun simplify_conv' ctxt = - Simplifier.full_rewrite (Simplifier.context ctxt (#4 (get_local ctxt))); + Simplifier.full_rewrite (put_simpset (#4 (get_local ctxt)) ctxt); fun simplify_conv ctxt ct = if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) then diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/Datatype_ZF.thy --- a/src/ZF/Datatype_ZF.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/Datatype_ZF.thy Thu Apr 18 17:07:01 2013 +0200 @@ -68,20 +68,21 @@ Balanced_Tree.make FOLogic.mk_conj (map FOLogic.mk_eq (ListPair.zip (largs,rargs))); - val datatype_ss = @{simpset}; + val datatype_ss = simpset_of @{context}; - fun proc sg ss old = - let val _ = - if !trace then writeln ("data_free: OLD = " ^ Syntax.string_of_term_global sg old) + fun proc ctxt old = + let val thy = Proof_Context.theory_of ctxt + val _ = + if !trace then writeln ("data_free: OLD = " ^ Syntax.string_of_term ctxt old) else () val (lhs,rhs) = FOLogic.dest_eq old val (lhead, largs) = strip_comb lhs and (rhead, rargs) = strip_comb rhs val lname = #1 (dest_Const lhead) handle TERM _ => raise Match; val rname = #1 (dest_Const rhead) handle TERM _ => raise Match; - val lcon_info = the (Symtab.lookup (ConstructorsData.get sg) lname) + val lcon_info = the (Symtab.lookup (ConstructorsData.get thy) lname) handle Option => raise Match; - val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname) + val rcon_info = the (Symtab.lookup (ConstructorsData.get thy) rname) handle Option => raise Match; val new = if #big_rec_name lcon_info = #big_rec_name rcon_info @@ -90,14 +91,14 @@ else Const(@{const_name False},FOLogic.oT) else raise Match val _ = - if !trace then writeln ("NEW = " ^ Syntax.string_of_term_global sg new) + if !trace then writeln ("NEW = " ^ Syntax.string_of_term ctxt new) else (); val goal = Logic.mk_equals (old, new) - val thm = Goal.prove (Simplifier.the_context ss) [] [] goal + val thm = Goal.prove ctxt [] [] goal (fn _ => rtac @{thm iff_reflection} 1 THEN - simp_tac (Simplifier.inherit_context ss datatype_ss addsimps #free_iffs lcon_info) 1) + simp_tac (put_simpset datatype_ss ctxt addsimps #free_iffs lcon_info) 1) handle ERROR msg => - (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Syntax.string_of_term_global sg goal); + (warning (msg ^ "\ndata_free simproc:\nfailed to prove " ^ Syntax.string_of_term ctxt goal); raise Match) in SOME thm end handle Match => NONE; diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/OrdQuant.thy Thu Apr 18 17:07:01 2013 +0200 @@ -370,15 +370,15 @@ simproc_setup defined_rex ("\x[M]. P(x) & Q(x)") = {* let val unfold_rex_tac = unfold_tac @{thms rex_def}; - fun prove_rex_tac ss = unfold_rex_tac ss THEN Quantifier1.prove_one_point_ex_tac; - in fn _ => fn ss => Quantifier1.rearrange_bex (prove_rex_tac ss) ss end + fun prove_rex_tac ctxt = unfold_rex_tac ctxt THEN Quantifier1.prove_one_point_ex_tac; + in fn _ => fn ctxt => Quantifier1.rearrange_bex (prove_rex_tac ctxt) ctxt end *} simproc_setup defined_rall ("\x[M]. P(x) \ Q(x)") = {* let val unfold_rall_tac = unfold_tac @{thms rall_def}; - fun prove_rall_tac ss = unfold_rall_tac ss THEN Quantifier1.prove_one_point_all_tac; - in fn _ => fn ss => Quantifier1.rearrange_ball (prove_rall_tac ss) ss end + fun prove_rall_tac ctxt = unfold_rall_tac ctxt THEN Quantifier1.prove_one_point_all_tac; + in fn _ => fn ctxt => Quantifier1.rearrange_ball (prove_rall_tac ctxt) ctxt end *} end diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/Tools/datatype_package.ML --- a/src/ZF/Tools/datatype_package.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/Tools/datatype_package.ML Thu Apr 18 17:07:01 2013 +0200 @@ -330,10 +330,10 @@ Goal.prove_global thy1 [] [] (Ind_Syntax.traceIt "next recursor equation = " thy1 (mk_recursor_eqn arg)) (*Proves a single recursor equation.*) - (fn _ => EVERY + (fn {context = ctxt, ...} => EVERY [rtac recursor_trans 1, - simp_tac (rank_ss addsimps case_eqns) 1, - IF_UNSOLVED (simp_tac (rank_ss addsimps tl con_defs) 1)]); + simp_tac (put_simpset rank_ss ctxt addsimps case_eqns) 1, + IF_UNSOLVED (simp_tac (put_simpset rank_ss ctxt addsimps tl con_defs) 1)]); in map prove_recursor_eqn (flat con_ty_lists ~~ recursor_cases) end diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/Tools/inductive_package.ML Thu Apr 18 17:07:01 2013 +0200 @@ -269,7 +269,7 @@ Proposition A should have the form t:Si where Si is an inductive set*) fun make_cases ctxt A = rule_by_tactic ctxt - (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac (simpset_of ctxt)) THEN basic_elim_tac) + (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac ctxt) THEN basic_elim_tac) (Thm.assume A RS elim) |> Drule.export_without_context_open; @@ -330,7 +330,7 @@ (Simplifier.global_context thy empty_ss |> Simplifier.set_mksimps (K (map mk_eq o ZF_atomize o gen_all))) setSolver (mk_solver "minimal" - (fn ss => resolve_tac (triv_rls @ Simplifier.prems_of ss) + (fn ctxt => resolve_tac (triv_rls @ Simplifier.prems_of ctxt) ORELSE' assume_tac ORELSE' etac @{thm FalseE})); diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/Tools/typechk.ML --- a/src/ZF/Tools/typechk.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/Tools/typechk.ML Thu Apr 18 17:07:01 2013 +0200 @@ -106,8 +106,8 @@ APPEND typecheck_step_tac (tcset_of ctxt)))); val type_solver = - Simplifier.mk_solver "ZF typecheck" (fn ss => - type_solver_tac (Simplifier.the_context ss) (Simplifier.prems_of ss)); + Simplifier.mk_solver "ZF typecheck" (fn ctxt => + type_solver_tac ctxt (Simplifier.prems_of ctxt)); (* concrete syntax *) @@ -131,6 +131,6 @@ val setup = Attrib.setup @{binding TC} (Attrib.add_del TC_add TC_del) "declaration of type-checking rule" #> typecheck_setup #> - Simplifier.map_simpset_global (fn ss => ss setSolver type_solver); + map_theory_simpset (fn ctxt => ctxt setSolver type_solver); end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/UNITY/Constrains.thy Thu Apr 18 17:07:01 2013 +0200 @@ -471,7 +471,6 @@ (*proves "co" properties when the program is specified*) fun constrains_tac ctxt = - let val ss = simpset_of ctxt in SELECT_GOAL (EVERY [REPEAT (Always_Int_tac 1), REPEAT (etac @{thm Always_ConstrainsI} 1 @@ -482,15 +481,14 @@ (* Three subgoals *) rewrite_goal_tac [@{thm st_set_def}] 3, REPEAT (force_tac ctxt 2), - full_simp_tac (ss addsimps (Program_Defs.get ctxt)) 1, + full_simp_tac (ctxt addsimps (Program_Defs.get ctxt)) 1, ALLGOALS (clarify_tac ctxt), REPEAT (FIRSTGOAL (etac @{thm disjE})), ALLGOALS (clarify_tac ctxt), REPEAT (FIRSTGOAL (etac @{thm disjE})), ALLGOALS (clarify_tac ctxt), - ALLGOALS (asm_full_simp_tac ss), - ALLGOALS (clarify_tac ctxt)]) - end; + ALLGOALS (asm_full_simp_tac ctxt), + ALLGOALS (clarify_tac ctxt)]); (*For proving invariants*) fun always_tac ctxt i = diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/UNITY/SubstAx.thy Thu Apr 18 17:07:01 2013 +0200 @@ -351,29 +351,27 @@ ML {* (*proves "ensures/leadsTo" properties when the program is specified*) fun ensures_tac ctxt sact = - let val ss = simpset_of ctxt in - SELECT_GOAL - (EVERY [REPEAT (Always_Int_tac 1), - etac @{thm Always_LeadsTo_Basis} 1 - ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*) - REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis}, - @{thm EnsuresI}, @{thm ensuresI}] 1), - (*now there are two subgoals: co & transient*) - simp_tac (ss addsimps (Program_Defs.get ctxt)) 2, - res_inst_tac ctxt [(("act", 0), sact)] @{thm transientI} 2, - (*simplify the command's domain*) - simp_tac (ss addsimps [@{thm domain_def}]) 3, - (* proving the domain part *) - clarify_tac ctxt 3, dtac @{thm swap} 3, force_tac ctxt 4, - rtac @{thm ReplaceI} 3, force_tac ctxt 3, force_tac ctxt 4, - asm_full_simp_tac ss 3, rtac @{thm conjI} 3, simp_tac ss 4, - REPEAT (rtac @{thm state_update_type} 3), - constrains_tac ctxt 1, - ALLGOALS (clarify_tac ctxt), - ALLGOALS (asm_full_simp_tac (ss addsimps [@{thm st_set_def}])), - ALLGOALS (clarify_tac ctxt), - ALLGOALS (asm_lr_simp_tac ss)]) - end; + SELECT_GOAL + (EVERY [REPEAT (Always_Int_tac 1), + etac @{thm Always_LeadsTo_Basis} 1 + ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*) + REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis}, + @{thm EnsuresI}, @{thm ensuresI}] 1), + (*now there are two subgoals: co & transient*) + simp_tac (ctxt addsimps (Program_Defs.get ctxt)) 2, + res_inst_tac ctxt [(("act", 0), sact)] @{thm transientI} 2, + (*simplify the command's domain*) + simp_tac (ctxt addsimps [@{thm domain_def}]) 3, + (* proving the domain part *) + clarify_tac ctxt 3, dtac @{thm swap} 3, force_tac ctxt 4, + rtac @{thm ReplaceI} 3, force_tac ctxt 3, force_tac ctxt 4, + asm_full_simp_tac ctxt 3, rtac @{thm conjI} 3, simp_tac ctxt 4, + REPEAT (rtac @{thm state_update_type} 3), + constrains_tac ctxt 1, + ALLGOALS (clarify_tac ctxt), + ALLGOALS (asm_full_simp_tac (ctxt addsimps [@{thm st_set_def}])), + ALLGOALS (clarify_tac ctxt), + ALLGOALS (asm_lr_simp_tac ctxt)]); *} method_setup ensures = {* diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/Univ.thy --- a/src/ZF/Univ.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/Univ.thy Thu Apr 18 17:07:01 2013 +0200 @@ -789,8 +789,9 @@ ML {* -val rank_ss = @{simpset} addsimps [@{thm VsetI}] - addsimps @{thms rank_rls} @ (@{thms rank_rls} RLN (2, [@{thm lt_trans}])); +val rank_ss = + simpset_of (@{context} addsimps [@{thm VsetI}] + addsimps @{thms rank_rls} @ (@{thms rank_rls} RLN (2, [@{thm lt_trans}]))); *} end diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/arith_data.ML --- a/src/ZF/arith_data.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/arith_data.ML Thu Apr 18 17:07:01 2013 +0200 @@ -11,7 +11,7 @@ (*tools for use in similar applications*) val gen_trans_tac: thm -> thm option -> tactic val prove_conv: string -> tactic list -> Proof.context -> thm list -> term * term -> thm option - val simplify_meta_eq: thm list -> simpset -> thm -> thm + val simplify_meta_eq: thm list -> Proof.context -> thm -> thm (*debugging*) structure EqCancelNumeralsData : CANCEL_NUMERALS_DATA structure LessCancelNumeralsData : CANCEL_NUMERALS_DATA @@ -125,13 +125,13 @@ @{thm diff_natify1}, @{thm diff_natify2}]; (*Final simplification: cancel + and **) -fun simplify_meta_eq rules = - let val ss0 = - FOL_ss +fun simplify_meta_eq rules ctxt = + let val ctxt' = + put_simpset FOL_ss ctxt delsimps @{thms iff_simps} (*these could erase the whole rule!*) addsimps rules |> fold Simplifier.add_eqcong [@{thm eq_cong2}, @{thm iff_cong2}] - in fn ss => mk_meta_eq o simplify (Simplifier.inherit_context ss ss0) end; + in mk_meta_eq o simplify ctxt' end; val final_rules = add_0s @ mult_1s @ [@{thm mult_0}, @{thm mult_0_right}]; @@ -143,15 +143,18 @@ val dest_coeff = dest_coeff val find_first_coeff = find_first_coeff [] - val norm_ss1 = ZF_ss addsimps add_0s @ add_succs @ mult_1s @ @{thms add_ac} - val norm_ss2 = ZF_ss addsimps add_0s @ mult_1s @ @{thms add_ac} @ - @{thms mult_ac} @ tc_rules @ natifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) - val numeral_simp_ss = ZF_ss addsimps add_0s @ tc_rules @ natifys - fun numeral_simp_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val norm_ss1 = + simpset_of (put_simpset ZF_ss @{context} addsimps add_0s @ add_succs @ mult_1s @ @{thms add_ac}) + val norm_ss2 = + simpset_of (put_simpset ZF_ss @{context} addsimps add_0s @ mult_1s @ @{thms add_ac} @ + @{thms mult_ac} @ tc_rules @ natifys) + fun norm_tac ctxt = + ALLGOALS (asm_simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss2 ctxt)) + val numeral_simp_ss = + simpset_of (put_simpset ZF_ss @{context} addsimps add_0s @ tc_rules @ natifys) + fun numeral_simp_tac ctxt = + ALLGOALS (asm_simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = simplify_meta_eq final_rules end; @@ -204,17 +207,17 @@ ["l #+ m = n", "l = m #+ n", "l #* m = n", "l = m #* n", "succ(m) = n", "m = succ(n)"], - (K EqCancelNumerals.proc)), + EqCancelNumerals.proc), ("natless_cancel_numerals", ["l #+ m < n", "l < m #+ n", "l #* m < n", "l < m #* n", "succ(m) < n", "m < succ(n)"], - (K LessCancelNumerals.proc)), + LessCancelNumerals.proc), ("natdiff_cancel_numerals", ["(l #+ m) #- n", "l #- (m #+ n)", "(l #* m) #- n", "l #- (m #* n)", "succ(m) #- n", "m #- succ(n)"], - (K DiffCancelNumerals.proc))]; + DiffCancelNumerals.proc)]; end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/int_arith.ML --- a/src/ZF/int_arith.ML Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/int_arith.ML Thu Apr 18 17:07:01 2013 +0200 @@ -158,18 +158,26 @@ val find_first_coeff = find_first_coeff [] val trans_tac = ArithData.gen_trans_tac @{thm iff_trans} - val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ @{thms zadd_ac} - val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys - val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ @{thms zadd_ac} @ @{thms zmult_ac} @ tc_rules @ intifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss3)) + val norm_ss1 = + simpset_of (put_simpset ZF_ss @{context} + addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ @{thms zadd_ac}) + val norm_ss2 = + simpset_of (put_simpset ZF_ss @{context} + addsimps bin_simps @ int_mult_minus_simps @ intifys) + val norm_ss3 = + simpset_of (put_simpset ZF_ss @{context} + addsimps int_minus_from_mult_simps @ @{thms zadd_ac} @ @{thms zmult_ac} @ tc_rules @ intifys) + fun norm_tac ctxt = + ALLGOALS (asm_simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss2 ctxt)) + THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss3 ctxt)) - val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys - fun numeral_simp_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - THEN ALLGOALS (asm_simp_tac (simpset_of (Simplifier.the_context ss))) + val numeral_simp_ss = + simpset_of (put_simpset ZF_ss @{context} + addsimps add_0s @ bin_simps @ tc_rules @ intifys) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) + THEN ALLGOALS (asm_simp_tac ctxt) val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) end; @@ -207,17 +215,17 @@ ["l $+ m = n", "l = m $+ n", "l $- m = n", "l = m $- n", "l $* m = n", "l = m $* n"], - K EqCancelNumerals.proc), + EqCancelNumerals.proc), ("intless_cancel_numerals", ["l $+ m $< n", "l $< m $+ n", "l $- m $< n", "l $< m $- n", "l $* m $< n", "l $< m $* n"], - K LessCancelNumerals.proc), + LessCancelNumerals.proc), ("intle_cancel_numerals", ["l $+ m $<= n", "l $<= m $+ n", "l $- m $<= n", "l $<= m $- n", "l $* m $<= n", "l $<= m $* n"], - K LeCancelNumerals.proc)]; + LeCancelNumerals.proc)]; (*version without the hyps argument*) @@ -236,17 +244,24 @@ val prove_conv = prove_conv_nohyps "int_combine_numerals" val trans_tac = ArithData.gen_trans_tac @{thm trans} - val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ @{thms zadd_ac} @ intifys - val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys - val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ @{thms zadd_ac} @ @{thms zmult_ac} @ tc_rules @ intifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss3)) + val norm_ss1 = + simpset_of (put_simpset ZF_ss @{context} + addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ @{thms zadd_ac} @ intifys) + val norm_ss2 = + simpset_of (put_simpset ZF_ss @{context} + addsimps bin_simps @ int_mult_minus_simps @ intifys) + val norm_ss3 = + simpset_of (put_simpset ZF_ss @{context} + addsimps int_minus_from_mult_simps @ @{thms zadd_ac} @ @{thms zmult_ac} @ tc_rules @ intifys) + fun norm_tac ctxt = + ALLGOALS (asm_simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss2 ctxt)) + THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss3 ctxt)) - val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys - fun numeral_simp_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val numeral_simp_ss = + simpset_of (put_simpset ZF_ss @{context} addsimps add_0s @ bin_simps @ tc_rules @ intifys) + fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) end; @@ -254,7 +269,7 @@ val combine_numerals = prep_simproc @{theory} - ("int_combine_numerals", ["i $+ j", "i $- j"], K CombineNumerals.proc); + ("int_combine_numerals", ["i $+ j", "i $- j"], CombineNumerals.proc); @@ -265,7 +280,7 @@ structure CombineNumeralsProdData = - struct +struct type coeff = int val iszero = (fn x => x = 0) val add = op * @@ -280,25 +295,28 @@ -val norm_ss1 = ZF_ss addsimps mult_1s @ diff_simps @ zminus_simps - val norm_ss2 = ZF_ss addsimps [@{thm zmult_zminus_right} RS @{thm sym}] @ - bin_simps @ @{thms zmult_ac} @ tc_rules @ intifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) +val norm_ss1 = + simpset_of (put_simpset ZF_ss @{context} addsimps mult_1s @ diff_simps @ zminus_simps) +val norm_ss2 = + simpset_of (put_simpset ZF_ss @{context} addsimps [@{thm zmult_zminus_right} RS @{thm sym}] @ + bin_simps @ @{thms zmult_ac} @ tc_rules @ intifys) +fun norm_tac ctxt = + ALLGOALS (asm_simp_tac (put_simpset norm_ss1 ctxt)) + THEN ALLGOALS (asm_simp_tac (put_simpset norm_ss2 ctxt)) - val numeral_simp_ss = ZF_ss addsimps bin_simps @ tc_rules @ intifys - fun numeral_simp_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = ArithData.simplify_meta_eq (mult_1s); - end; +val numeral_simp_ss = + simpset_of (put_simpset ZF_ss @{context} addsimps bin_simps @ tc_rules @ intifys) +fun numeral_simp_tac ctxt = + ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) +val simplify_meta_eq = ArithData.simplify_meta_eq (mult_1s); +end; structure CombineNumeralsProd = CombineNumeralsFun(CombineNumeralsProdData); val combine_numerals_prod = prep_simproc @{theory} - ("int_combine_numerals_prod", ["i $* j"], K CombineNumeralsProd.proc); + ("int_combine_numerals_prod", ["i $* j"], CombineNumeralsProd.proc); end; diff -r 19b47bfac6ef -r 9e7d1c139569 src/ZF/pair.thy --- a/src/ZF/pair.thy Tue Apr 16 17:54:14 2013 +0200 +++ b/src/ZF/pair.thy Thu Apr 18 17:07:01 2013 +0200 @@ -11,25 +11,25 @@ ML_file "simpdata.ML" setup {* - Simplifier.map_simpset_global + map_theory_simpset (Simplifier.set_mksimps (K (map mk_eq o ZF_atomize o gen_all)) #> Simplifier.add_cong @{thm if_weak_cong}) *} -ML {* val ZF_ss = @{simpset} *} +ML {* val ZF_ss = simpset_of @{context} *} simproc_setup defined_Bex ("\x\A. P(x) & Q(x)") = {* let val unfold_bex_tac = unfold_tac @{thms Bex_def}; - fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac; - in fn _ => fn ss => Quantifier1.rearrange_bex (prove_bex_tac ss) ss end + fun prove_bex_tac ctxt = unfold_bex_tac ctxt THEN Quantifier1.prove_one_point_ex_tac; + in fn _ => fn ctxt => Quantifier1.rearrange_bex (prove_bex_tac ctxt) ctxt end *} simproc_setup defined_Ball ("\x\A. P(x) \ Q(x)") = {* let val unfold_ball_tac = unfold_tac @{thms Ball_def}; - fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac; - in fn _ => fn ss => Quantifier1.rearrange_ball (prove_ball_tac ss) ss end + fun prove_ball_tac ctxt = unfold_ball_tac ctxt THEN Quantifier1.prove_one_point_all_tac; + in fn _ => fn ctxt => Quantifier1.rearrange_ball (prove_ball_tac ctxt) ctxt end *}