proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
occasionally clarified use of context;
--- a/NEWS Tue Feb 10 14:29:36 2015 +0100
+++ b/NEWS Tue Feb 10 14:48:26 2015 +0100
@@ -248,8 +248,8 @@
Input.source (with formal position information).
* Proper context for various elementary tactics: assume_tac,
-match_tac, compose_tac, Splitter.split_tac etc. Minor
-INCOMPATIBILITY.
+resolve_tac, eresolve_tac, dresolve_tac, forward_tac, match_tac,
+compose_tac, Splitter.split_tac etc. INCOMPATIBILITY.
* Tactical PARALLEL_ALLGOALS is the most common way to refer to
PARALLEL_GOALS.
--- a/src/CCL/CCL.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CCL/CCL.thy Tue Feb 10 14:48:26 2015 +0100
@@ -205,7 +205,7 @@
val inj_lemmas = maps mk_inj_lemmas rews
in
CHANGED (REPEAT (ares_tac [@{thm iffI}, @{thm allI}, @{thm conjI}] i ORELSE
- eresolve_tac inj_lemmas i ORELSE
+ eresolve_tac ctxt inj_lemmas i ORELSE
asm_simp_tac (ctxt addsimps rews) i))
end;
*}
@@ -267,7 +267,7 @@
let
fun mk_raw_dstnct_thm rls s =
Goal.prove_global @{theory} [] [] (Syntax.read_prop_global @{theory} s)
- (fn _ => rtac @{thm notI} 1 THEN eresolve_tac rls 1)
+ (fn {context = ctxt, ...} => rtac @{thm notI} 1 THEN eresolve_tac ctxt rls 1)
in map (mk_raw_dstnct_thm caseB_lemmas)
(mk_dstnct_rls @{theory} ["bot","true","false","pair","lambda"]) end
--- a/src/CCL/Term.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CCL/Term.thy Tue Feb 10 14:48:26 2015 +0100
@@ -203,7 +203,7 @@
method_setup beta_rl = {*
Scan.succeed (fn ctxt =>
SIMPLE_METHOD' (CHANGED o
- simp_tac (ctxt addsimps @{thms rawBs} setloop (fn _ => stac @{thm letrecB}))))
+ simp_tac (ctxt addsimps @{thms rawBs} setloop (fn _ => stac ctxt @{thm letrecB}))))
*}
lemma ifBtrue: "if true then t else u = t"
--- a/src/CCL/Type.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CCL/Type.thy Tue Feb 10 14:48:26 2015 +0100
@@ -126,8 +126,8 @@
ML {*
fun mk_ncanT_tac top_crls crls =
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
+ resolve_tac ctxt ([major] RL top_crls) 1 THEN
+ REPEAT_SOME (eresolve_tac ctxt (crls @ [@{thm exE}, @{thm bexE}, @{thm conjE}, @{thm disjE}])) THEN
ALLGOALS (asm_simp_tac ctxt) THEN
ALLGOALS (ares_tac (prems RL [@{thm lem}]) ORELSE' etac @{thm bspec}) THEN
safe_tac (ctxt addSIs prems))
@@ -436,7 +436,7 @@
fun POgen_tac ctxt (rla, rlb) i =
SELECT_GOAL (safe_tac ctxt) i THEN
rtac (rlb RS (rla RS @{thm ssubst_pair})) i THEN
- (REPEAT (resolve_tac
+ (REPEAT (resolve_tac ctxt
(@{thms POgenIs} @ [@{thm PO_refl} RS (@{thm POgen_mono} RS @{thm ci3_AI})] @
(@{thms POgenIs} RL [@{thm POgen_mono} RS @{thm ci3_AgenI}]) @
[@{thm POgen_mono} RS @{thm ci3_RI}]) i))
@@ -467,8 +467,8 @@
unfolding data_defs by (genIs EQgenXH EQgen_mono)+
ML {*
-fun EQgen_raw_tac i =
- (REPEAT (resolve_tac (@{thms EQgenIs} @
+fun EQgen_raw_tac ctxt i =
+ (REPEAT (resolve_tac ctxt (@{thms EQgenIs} @
[@{thm EQ_refl} RS (@{thm EQgen_mono} RS @{thm ci3_AI})] @
(@{thms EQgenIs} RL [@{thm EQgen_mono} RS @{thm ci3_AgenI}]) @
[@{thm EQgen_mono} RS @{thm ci3_RI}]) i))
@@ -480,9 +480,9 @@
fun EQgen_tac ctxt rews i =
SELECT_GOAL
(TRY (safe_tac ctxt) THEN
- resolve_tac ((rews @ [@{thm refl}]) RL ((rews @ [@{thm refl}]) RL [@{thm ssubst_pair}])) i THEN
+ resolve_tac ctxt ((rews @ [@{thm refl}]) RL ((rews @ [@{thm refl}]) RL [@{thm ssubst_pair}])) i THEN
ALLGOALS (simp_tac ctxt) THEN
- ALLGOALS EQgen_raw_tac) i
+ ALLGOALS (EQgen_raw_tac ctxt)) i
*}
method_setup EQgen = {*
--- a/src/CCL/Wfd.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CCL/Wfd.thy Tue Feb 10 14:48:26 2015 +0100
@@ -448,7 +448,7 @@
fun rcall_tac ctxt i =
let fun tac ps rl i = res_inst_tac ctxt ps rl i THEN atac i
in IHinst tac @{thms rcallTs} i end
- THEN eresolve_tac @{thms rcall_lemmas} i
+ THEN eresolve_tac ctxt @{thms rcall_lemmas} i
fun raw_step_tac ctxt prems i =
ares_tac (prems@type_rls) i ORELSE
@@ -467,7 +467,7 @@
fun clean_ccs_tac ctxt =
let fun tac ps rl i = eres_inst_tac ctxt ps rl i THEN atac i in
TRY (REPEAT_FIRST (IHinst tac @{thms hyprcallTs} ORELSE'
- eresolve_tac ([asm_rl, @{thm SubtypeE}] @ @{thms rmIHs}) ORELSE'
+ eresolve_tac ctxt ([asm_rl, @{thm SubtypeE}] @ @{thms rmIHs}) ORELSE'
hyp_subst_tac ctxt))
end
@@ -498,7 +498,7 @@
fun eval_tac ths =
Subgoal.FOCUS_PREMS (fn {context = ctxt, prems, ...} =>
let val eval_rules = Named_Theorems.get ctxt @{named_theorems eval}
- in DEPTH_SOLVE_1 (resolve_tac (ths @ prems @ rev eval_rules) 1) end)
+ in DEPTH_SOLVE_1 (resolve_tac ctxt (ths @ prems @ rev eval_rules) 1) end)
*}
method_setup eval = {*
@@ -521,7 +521,7 @@
apply (unfold let_def)
apply (rule 1 [THEN canonical])
apply (tactic {*
- REPEAT (DEPTH_SOLVE_1 (resolve_tac (@{thms assms} @ @{thms eval_rls}) 1 ORELSE
+ REPEAT (DEPTH_SOLVE_1 (resolve_tac @{context} (@{thms assms} @ @{thms eval_rls}) 1 ORELSE
etac @{thm substitute} 1)) *})
done
--- a/src/CTT/Arith.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CTT/Arith.thy Tue Feb 10 14:48:26 2015 +0100
@@ -177,7 +177,7 @@
(Arith_simp.norm_tac ctxt (congr_rls, prems))
fun hyp_arith_rew_tac ctxt prems = make_rew_tac ctxt
- (Arith_simp.cond_norm_tac ctxt (prove_cond_tac, congr_rls, prems))
+ (Arith_simp.cond_norm_tac ctxt (prove_cond_tac ctxt, congr_rls, prems))
end
*}
--- a/src/CTT/CTT.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CTT/CTT.thy Tue Feb 10 14:48:26 2015 +0100
@@ -453,13 +453,13 @@
fun safestep_tac ctxt thms i =
form_tac ctxt ORELSE
- resolve_tac thms i ORELSE
- biresolve_tac safe0_brls i ORELSE mp_tac ctxt i ORELSE
- DETERM (biresolve_tac safep_brls i)
+ resolve_tac ctxt thms i ORELSE
+ biresolve_tac ctxt safe0_brls i ORELSE mp_tac ctxt i ORELSE
+ DETERM (biresolve_tac ctxt safep_brls i)
fun safe_tac ctxt thms i = DEPTH_SOLVE_1 (safestep_tac ctxt thms i)
-fun step_tac ctxt thms = safestep_tac ctxt thms ORELSE' biresolve_tac unsafe_brls
+fun step_tac ctxt thms = safestep_tac ctxt thms ORELSE' biresolve_tac ctxt unsafe_brls
(*Fails unless it solves the goal!*)
fun pc_tac ctxt thms = DEPTH_SOLVE_1 o (step_tac ctxt thms)
--- a/src/CTT/ex/Elimination.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CTT/ex/Elimination.thy Tue Feb 10 14:48:26 2015 +0100
@@ -182,7 +182,7 @@
and "\<And>z. z:A*B \<Longrightarrow> C(z) type"
shows "?a : (SUM z:A*B. C(z)) --> (SUM u:A. SUM v:B. C(<u,v>))"
apply (rule intr_rls)
-apply (tactic {* biresolve_tac safe_brls 2 *})
+apply (tactic {* biresolve_tac @{context} safe_brls 2 *})
(*Now must convert assumption C(z) into antecedent C(<kd,ke>) *)
apply (rule_tac [2] a = "y" in ProdE)
apply (typechk assms)
--- a/src/CTT/ex/Synthesis.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CTT/ex/Synthesis.thy Tue Feb 10 14:48:26 2015 +0100
@@ -92,12 +92,12 @@
* (PROD y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))"
apply intr
apply eqintr
-apply (tactic "resolve_tac [TSimp.split_eqn] 3")
+apply (tactic "resolve_tac @{context} [TSimp.split_eqn] 3")
apply (tactic "SELECT_GOAL (rew_tac @{context} []) 4")
-apply (tactic "resolve_tac [TSimp.split_eqn] 3")
+apply (tactic "resolve_tac @{context} [TSimp.split_eqn] 3")
apply (tactic "SELECT_GOAL (rew_tac @{context} []) 4")
apply (rule_tac [3] p = "y" in NC_succ)
- (** by (resolve_tac comp_rls 3); caused excessive branching **)
+ (** by (resolve_tac @{context} comp_rls 3); caused excessive branching **)
apply rew
done
--- a/src/CTT/ex/Typechecking.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CTT/ex/Typechecking.thy Tue Feb 10 14:48:26 2015 +0100
@@ -66,7 +66,8 @@
(*Proofs involving arbitrary types.
For concreteness, every type variable left over is forced to be N*)
-method_setup N = {* Scan.succeed (fn _ => SIMPLE_METHOD (TRYALL (resolve_tac @{thms NF}))) *}
+method_setup N =
+ {* Scan.succeed (fn ctxt => SIMPLE_METHOD (TRYALL (resolve_tac ctxt @{thms NF}))) *}
schematic_lemma "lam w. <w,w> : ?A"
apply typechk
--- a/src/CTT/rew.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/CTT/rew.ML Tue Feb 10 14:48:26 2015 +0100
@@ -11,7 +11,7 @@
| peEs n = @{thm EqE} :: map (curry (op RS) @{thm ProdE}) (peEs (n-1));
(*Tactic used for proving conditions for the cond_rls*)
-val prove_cond_tac = eresolve_tac (peEs 5);
+fun prove_cond_tac ctxt = eresolve_tac ctxt (peEs 5);
structure TSimp_data: TSIMP_DATA =
@@ -32,11 +32,11 @@
(*Make a rewriting tactic from a normalization tactic*)
fun make_rew_tac ctxt ntac =
- TRY (eqintr_tac ctxt) THEN TRYALL (resolve_tac [TSimp.split_eqn]) THEN
+ TRY (eqintr_tac ctxt) THEN TRYALL (resolve_tac ctxt [TSimp.split_eqn]) THEN
ntac;
fun rew_tac ctxt thms = make_rew_tac ctxt
(TSimp.norm_tac ctxt (standard_congr_rls, thms));
fun hyp_rew_tac ctxt thms = make_rew_tac ctxt
- (TSimp.cond_norm_tac ctxt (prove_cond_tac, standard_congr_rls, thms));
+ (TSimp.cond_norm_tac ctxt (prove_cond_tac ctxt, standard_congr_rls, thms));
--- a/src/Cube/Example.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Cube/Example.thy Tue Feb 10 14:48:26 2015 +0100
@@ -18,9 +18,9 @@
(DEPTH_SOLVE_1 (HEADGOAL (ares_tac (facts @ thms)))))))\<close>
method_setup strip_asms =
- \<open>Attrib.thms >> (fn thms => K (METHOD (fn facts =>
- REPEAT (resolve_tac [@{thm strip_b}, @{thm strip_s}] 1 THEN
- DEPTH_SOLVE_1 (ares_tac (facts @ thms) 1)))))\<close>
+ \<open>Attrib.thms >> (fn thms => fn ctxt => METHOD (fn facts =>
+ REPEAT (resolve_tac ctxt [@{thm strip_b}, @{thm strip_s}] 1 THEN
+ DEPTH_SOLVE_1 (ares_tac (facts @ thms) 1))))\<close>
subsection \<open>Simple types\<close>
--- a/src/Doc/Implementation/Isar.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Implementation/Isar.thy Tue Feb 10 14:48:26 2015 +0100
@@ -366,8 +366,8 @@
have "A \<and> B"
apply (tactic \<open>rtac @{thm conjI} 1\<close>)
- using a apply (tactic \<open>resolve_tac facts 1\<close>)
- using b apply (tactic \<open>resolve_tac facts 1\<close>)
+ using a apply (tactic \<open>resolve_tac @{context} facts 1\<close>)
+ using b apply (tactic \<open>resolve_tac @{context} facts 1\<close>)
done
have "A \<and> B"
--- a/src/Doc/Implementation/Tactic.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Implementation/Tactic.thy Tue Feb 10 14:48:26 2015 +0100
@@ -277,11 +277,11 @@
text %mlref \<open>
\begin{mldecls}
- @{index_ML resolve_tac: "thm list -> int -> tactic"} \\
- @{index_ML eresolve_tac: "thm list -> int -> tactic"} \\
- @{index_ML dresolve_tac: "thm list -> int -> tactic"} \\
- @{index_ML forward_tac: "thm list -> int -> tactic"} \\
- @{index_ML biresolve_tac: "(bool * thm) list -> int -> tactic"} \\[1ex]
+ @{index_ML resolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{index_ML eresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{index_ML dresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{index_ML forward_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{index_ML biresolve_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\[1ex]
@{index_ML assume_tac: "Proof.context -> int -> tactic"} \\
@{index_ML eq_assume_tac: "int -> tactic"} \\[1ex]
@{index_ML match_tac: "Proof.context -> thm list -> int -> tactic"} \\
@@ -292,20 +292,20 @@
\begin{description}
- \item @{ML resolve_tac}~@{text "thms i"} refines the goal state
+ \item @{ML resolve_tac}~@{text "ctxt thms i"} refines the goal state
using the given theorems, which should normally be introduction
rules. The tactic resolves a rule's conclusion with subgoal @{text
i}, replacing it by the corresponding versions of the rule's
premises.
- \item @{ML eresolve_tac}~@{text "thms i"} performs elim-resolution
+ \item @{ML eresolve_tac}~@{text "ctxt thms i"} performs elim-resolution
with the given theorems, which are normally be elimination rules.
- Note that @{ML "eresolve_tac [asm_rl]"} is equivalent to @{ML
- assume_tac}, which facilitates mixing of assumption steps with
+ Note that @{ML_text "eresolve_tac ctxt [asm_rl]"} is equivalent to @{ML_text
+ "assume_tac ctxt"}, which facilitates mixing of assumption steps with
genuine eliminations.
- \item @{ML dresolve_tac}~@{text "thms i"} performs
+ \item @{ML dresolve_tac}~@{text "ctxt thms i"} performs
destruct-resolution with the given theorems, which should normally
be destruction rules. This replaces an assumption by the result of
applying one of the rules.
@@ -314,7 +314,7 @@
selected assumption is not deleted. It applies a rule to an
assumption, adding the result as a new assumption.
- \item @{ML biresolve_tac}~@{text "brls i"} refines the proof state
+ \item @{ML biresolve_tac}~@{text "ctxt brls i"} refines the proof state
by resolution or elim-resolution on each rule, as indicated by its
flag. It affects subgoal @{text "i"} of the proof state.
--- a/src/Doc/Isar_Ref/Generic.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Isar_Ref/Generic.thy Tue Feb 10 14:48:26 2015 +0100
@@ -1054,7 +1054,7 @@
ML \<open>
fun subgoaler_tac ctxt =
assume_tac ctxt ORELSE'
- resolve_tac (Simplifier.prems_of ctxt) ORELSE'
+ resolve_tac ctxt (Simplifier.prems_of ctxt) ORELSE'
asm_simp_tac ctxt
\<close>
--- a/src/Doc/Isar_Ref/ML_Tactic.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Isar_Ref/ML_Tactic.thy Tue Feb 10 14:48:26 2015 +0100
@@ -62,11 +62,11 @@
\medskip
\begin{tabular}{lll}
@{ML rtac}~@{text "a 1"} & & @{text "rule a"} \\
- @{ML resolve_tac}~@{text "[a\<^sub>1, \<dots>] 1"} & & @{text "rule a\<^sub>1 \<dots>"} \\
+ @{ML resolve_tac}~@{text "ctxt [a\<^sub>1, \<dots>] 1"} & & @{text "rule a\<^sub>1 \<dots>"} \\
@{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \<dots>] a 1"} & &
@{text "rule_tac x\<^sub>1 = t\<^sub>1 \<AND> \<dots> \<IN> a"} \\[0.5ex]
@{ML rtac}~@{text "a i"} & & @{text "rule_tac [i] a"} \\
- @{ML resolve_tac}~@{text "[a\<^sub>1, \<dots>] i"} & & @{text "rule_tac [i] a\<^sub>1 \<dots>"} \\
+ @{ML resolve_tac}~@{text "ctxt [a\<^sub>1, \<dots>] i"} & & @{text "rule_tac [i] a\<^sub>1 \<dots>"} \\
@{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \<dots>] a i"} & &
@{text "rule_tac [i] x\<^sub>1 = t\<^sub>1 \<AND> \<dots> \<IN> a"} \\
\end{tabular}
--- a/src/Doc/Tutorial/Protocol/Event.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Tutorial/Protocol/Event.thy Tue Feb 10 14:48:26 2015 +0100
@@ -262,7 +262,7 @@
{*
fun analz_mono_contra_tac ctxt =
rtac @{thm analz_impI} THEN'
- REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
+ REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
THEN' mp_tac ctxt
*}
@@ -333,7 +333,7 @@
fun synth_analz_mono_contra_tac ctxt =
rtac @{thm syan_impI} THEN'
REPEAT1 o
- (dresolve_tac
+ (dresolve_tac ctxt
[@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
@{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
@{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
--- a/src/Doc/Tutorial/Protocol/Message.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Tutorial/Protocol/Message.thy Tue Feb 10 14:48:26 2015 +0100
@@ -837,13 +837,13 @@
*)
fun impOfSubs th = th RSN (2, @{thm rev_subsetD})
-val Fake_insert_tac =
- dresolve_tac [impOfSubs Fake_analz_insert,
+fun Fake_insert_tac ctxt =
+ dresolve_tac ctxt [impOfSubs Fake_analz_insert,
impOfSubs Fake_parts_insert] THEN'
- eresolve_tac [asm_rl, @{thm synth.Inj}];
+ eresolve_tac ctxt [asm_rl, @{thm synth.Inj}];
fun Fake_insert_simp_tac ctxt i =
- REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ctxt i;
+ REPEAT (Fake_insert_tac ctxt i) THEN asm_full_simp_tac ctxt i;
fun atomic_spy_analz_tac ctxt =
SELECT_GOAL
@@ -859,7 +859,7 @@
(res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1),
(*...allowing further simplifications*)
simp_tac ctxt 1,
- REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
+ REPEAT (FIRSTGOAL (resolve_tac ctxt [allI,impI,notI,conjI,iffI])),
DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i);
*}
--- a/src/Doc/Tutorial/Protocol/Public.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Doc/Tutorial/Protocol/Public.thy Tue Feb 10 14:48:26 2015 +0100
@@ -162,7 +162,7 @@
(ALLGOALS (simp_tac (ctxt delsimps [used_Says]))
THEN
REPEAT_FIRST (eq_assume_tac ORELSE'
- resolve_tac [refl, conjI, @{thm Nonce_supply}]));
+ resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}]));
*}
method_setup possibility = {* Scan.succeed (SIMPLE_METHOD o possibility_tac) *}
--- a/src/FOL/IFOL.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOL/IFOL.thy Tue Feb 10 14:48:26 2015 +0100
@@ -208,8 +208,8 @@
(*Finds P-->Q and P in the assumptions, replaces implication by Q *)
ML {*
- fun mp_tac i = eresolve_tac [@{thm notE}, @{thm impE}] i THEN atac i
- fun eq_mp_tac i = eresolve_tac [@{thm notE}, @{thm impE}] i THEN eq_assume_tac i
+ fun mp_tac i = eresolve0_tac [@{thm notE}, @{thm impE}] i THEN atac i
+ fun eq_mp_tac i = eresolve0_tac [@{thm notE}, @{thm impE}] i THEN eq_assume_tac i
*}
@@ -305,8 +305,8 @@
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*)
ML {*
fun iff_tac prems i =
- resolve_tac (prems RL @{thms iffE}) i THEN
- REPEAT1 (eresolve_tac [@{thm asm_rl}, @{thm mp}] i)
+ resolve0_tac (prems RL @{thms iffE}) i THEN
+ REPEAT1 (eresolve0_tac [@{thm asm_rl}, @{thm mp}] i)
*}
lemma conj_cong:
--- a/src/FOL/ex/Miniscope.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOL/ex/Miniscope.thy Tue Feb 10 14:48:26 2015 +0100
@@ -65,7 +65,8 @@
ML {*
val mini_ss = simpset_of (@{context} addsimps @{thms mini_simps});
-fun mini_tac ctxt = resolve_tac @{thms ccontr} THEN' asm_full_simp_tac (put_simpset mini_ss ctxt);
+fun mini_tac ctxt =
+ resolve_tac ctxt @{thms ccontr} THEN' asm_full_simp_tac (put_simpset mini_ss ctxt);
*}
end
--- a/src/FOL/ex/Prolog.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOL/ex/Prolog.thy Tue Feb 10 14:48:26 2015 +0100
@@ -63,26 +63,29 @@
(*backtracking version*)
ML {*
-val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) (resolve_tac (@{thms rules}) 1)
+fun prolog_tac ctxt =
+ DEPTH_FIRST (has_fewer_prems 1) (resolve_tac ctxt @{thms rules} 1)
*}
schematic_lemma "rev(?x, a:b:c:Nil)"
-apply (tactic prolog_tac)
+apply (tactic \<open>prolog_tac @{context}\<close>)
done
schematic_lemma "rev(a:?x:c:?y:Nil, d:?z:b:?u)"
-apply (tactic prolog_tac)
+apply (tactic \<open>prolog_tac @{context}\<close>)
done
(*rev([a..p], ?w) requires 153 inferences *)
schematic_lemma "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"
-apply (tactic {* DEPTH_SOLVE (resolve_tac ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
+apply (tactic {*
+ DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
done
(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences
total inferences = 2 + 1 + 17 + 561 = 581*)
schematic_lemma "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"
-apply (tactic {* DEPTH_SOLVE (resolve_tac ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
+apply (tactic {*
+ DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
done
end
--- a/src/FOL/intprover.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOL/intprover.ML Tue Feb 10 14:48:26 2015 +0100
@@ -76,12 +76,14 @@
(*These steps could instantiate variables and are therefore unsafe.*)
fun inst_step_tac ctxt =
assume_tac ctxt APPEND' mp_tac APPEND'
- biresolve_tac (safe0_brls @ safep_brls);
+ biresolve_tac ctxt (safe0_brls @ safep_brls);
(*One safe or unsafe step. *)
-fun step_tac ctxt i = FIRST [safe_tac ctxt, inst_step_tac ctxt i, biresolve_tac haz_brls i];
+fun step_tac ctxt i =
+ FIRST [safe_tac ctxt, inst_step_tac ctxt i, biresolve_tac ctxt haz_brls i];
-fun step_dup_tac ctxt i = FIRST [safe_tac ctxt, inst_step_tac ctxt i, biresolve_tac haz_dup_brls i];
+fun step_dup_tac ctxt i =
+ FIRST [safe_tac ctxt, inst_step_tac ctxt i, biresolve_tac ctxt haz_dup_brls i];
(*Dumb but fast*)
fun fast_tac ctxt = SELECT_GOAL (DEPTH_SOLVE (step_tac ctxt 1));
--- a/src/FOL/simpdata.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOL/simpdata.ML Tue Feb 10 14:48:26 2015 +0100
@@ -23,7 +23,7 @@
(*Replace premises x=y, X<->Y by X==Y*)
fun mk_meta_prems ctxt =
rule_by_tactic ctxt
- (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
+ (REPEAT_FIRST (resolve_tac ctxt [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
(*Congruence rules for = or <-> (instead of ==)*)
fun mk_meta_cong ctxt rl =
@@ -107,9 +107,9 @@
val triv_rls = [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl}, @{thm notFalseI}];
fun unsafe_solver ctxt =
- FIRST' [resolve_tac (triv_rls @ Simplifier.prems_of ctxt),
+ FIRST' [resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt),
assume_tac ctxt,
- eresolve_tac @{thms FalseE}];
+ eresolve_tac ctxt @{thms FalseE}];
(*No premature instantiation of variables during simplification*)
fun safe_solver ctxt =
--- a/src/FOLP/FOLP.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOLP/FOLP.thy Tue Feb 10 14:48:26 2015 +0100
@@ -57,7 +57,7 @@
shows "?p : R"
apply (rule excluded_middle [THEN disjE])
apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE
- resolve_tac [@{thm r1}, @{thm r2}, @{thm major} RS @{thm mp}] 1) *})
+ resolve_tac @{context} [@{thm r1}, @{thm r2}, @{thm major} RS @{thm mp}] 1) *})
done
(*Double negation law*)
@@ -81,8 +81,8 @@
apply (unfold iff_def)
apply (rule conjE)
apply (tactic {* DEPTH_SOLVE_1 (etac @{thm impCE} 1 ORELSE
- eresolve_tac [@{thm notE}, @{thm impE}] 1 THEN atac 1 ORELSE atac 1 ORELSE
- resolve_tac [@{thm r1}, @{thm r2}] 1) *})+
+ eresolve_tac @{context} [@{thm notE}, @{thm impE}] 1 THEN atac 1 ORELSE atac 1 ORELSE
+ resolve_tac @{context} [@{thm r1}, @{thm r2}] 1) *})+
done
--- a/src/FOLP/IFOLP.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOLP/IFOLP.thy Tue Feb 10 14:48:26 2015 +0100
@@ -273,13 +273,13 @@
(*Finds P-->Q and P in the assumptions, replaces implication by Q *)
ML {*
fun mp_tac ctxt i =
- eresolve_tac [@{thm notE}, make_elim @{thm mp}] i THEN assume_tac ctxt i
+ eresolve_tac ctxt [@{thm notE}, make_elim @{thm mp}] i THEN assume_tac ctxt i
*}
(*Like mp_tac but instantiates no variables*)
ML {*
fun int_uniq_mp_tac ctxt i =
- eresolve_tac [@{thm notE}, @{thm impE}] i THEN uniq_assume_tac ctxt i
+ eresolve_tac ctxt [@{thm notE}, @{thm impE}] i THEN uniq_assume_tac ctxt i
*}
@@ -361,8 +361,8 @@
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*)
ML {*
fun iff_tac prems i =
- resolve_tac (prems RL [@{thm iffE}]) i THEN
- REPEAT1 (eresolve_tac [asm_rl, @{thm mp}] i)
+ resolve0_tac (prems RL [@{thm iffE}]) i THEN
+ REPEAT1 (eresolve0_tac [asm_rl, @{thm mp}] i)
*}
schematic_lemma conj_cong:
@@ -503,17 +503,20 @@
schematic_lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')"
apply (rule iffI)
- apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
+ apply (tactic {*
+ DEPTH_SOLVE (atac 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
done
schematic_lemma pred2_cong: "[| p:a=a'; q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')"
apply (rule iffI)
- apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
+ apply (tactic {*
+ DEPTH_SOLVE (atac 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
done
schematic_lemma pred3_cong: "[| p:a=a'; q:b=b'; r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')"
apply (rule iffI)
- apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
+ apply (tactic {*
+ DEPTH_SOLVE (atac 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
done
lemmas pred_congs = pred1_cong pred2_cong pred3_cong
@@ -541,7 +544,7 @@
and minor: "!!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R"
shows "?p:R"
apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE
- resolve_tac [@{thm disjI1}, @{thm disjI2}, @{thm impI},
+ resolve_tac @{context} [@{thm disjI1}, @{thm disjI2}, @{thm impI},
@{thm major} RS @{thm mp}, @{thm minor}] 1) *})
done
--- a/src/FOLP/classical.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOLP/classical.ML Tue Feb 10 14:48:26 2015 +0100
@@ -73,7 +73,7 @@
fun contr_tac ctxt = etac not_elim THEN' assume_tac ctxt;
(*Finds P-->Q and P in the assumptions, replaces implication by Q *)
-fun mp_tac ctxt i = eresolve_tac ([not_elim,imp_elim]) i THEN assume_tac ctxt i;
+fun mp_tac ctxt i = eresolve_tac ctxt ([not_elim,imp_elim]) i THEN assume_tac ctxt i;
(*Like mp_tac but instantiates no variables*)
fun uniq_mp_tac ctxt i = ematch_tac ctxt ([not_elim,imp_elim]) i THEN uniq_assume_tac ctxt i;
@@ -151,9 +151,9 @@
fun safe_step_tac ctxt (CS{safe0_brls,safep_brls,...}) =
FIRST' [uniq_assume_tac ctxt,
uniq_mp_tac ctxt,
- biresolve_tac safe0_brls,
+ biresolve_tac ctxt safe0_brls,
FIRST' hyp_subst_tacs,
- biresolve_tac safep_brls] ;
+ biresolve_tac ctxt safep_brls] ;
(*Repeatedly attack subgoals using safe inferences*)
fun safe_tac ctxt cs = DETERM (REPEAT_FIRST (safe_step_tac ctxt cs));
@@ -165,7 +165,7 @@
fun step_tac ctxt (cs as (CS{haz_brls,...})) i =
FIRST [safe_tac ctxt cs,
inst_step_tac ctxt i,
- biresolve_tac haz_brls i];
+ biresolve_tac ctxt haz_brls i];
(*** The following tactics all fail unless they solve one goal ***)
@@ -179,7 +179,7 @@
(*Using a "safe" rule to instantiate variables is unsafe. This tactic
allows backtracking from "safe" rules to "unsafe" rules here.*)
fun slow_step_tac ctxt (cs as (CS{haz_brls,...})) i =
- safe_tac ctxt cs ORELSE (assume_tac ctxt i APPEND biresolve_tac haz_brls i);
+ safe_tac ctxt cs ORELSE (assume_tac ctxt i APPEND biresolve_tac ctxt haz_brls i);
end;
end;
--- a/src/FOLP/intprover.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOLP/intprover.ML Tue Feb 10 14:48:26 2015 +0100
@@ -54,9 +54,9 @@
(*Attack subgoals using safe inferences*)
fun safe_step_tac ctxt = FIRST' [uniq_assume_tac ctxt,
int_uniq_mp_tac ctxt,
- biresolve_tac safe0_brls,
+ biresolve_tac ctxt safe0_brls,
hyp_subst_tac,
- biresolve_tac safep_brls] ;
+ biresolve_tac ctxt safep_brls] ;
(*Repeatedly attack subgoals using safe inferences*)
fun safe_tac ctxt = DETERM (REPEAT_FIRST (safe_step_tac ctxt));
@@ -65,7 +65,7 @@
fun inst_step_tac ctxt = assume_tac ctxt APPEND' mp_tac ctxt;
(*One safe or unsafe step. *)
-fun step_tac ctxt i = FIRST [safe_tac ctxt, inst_step_tac ctxt i, biresolve_tac haz_brls i];
+fun step_tac ctxt i = FIRST [safe_tac ctxt, inst_step_tac ctxt i, biresolve_tac ctxt haz_brls i];
(*Dumb but fast*)
fun fast_tac ctxt = SELECT_GOAL (DEPTH_SOLVE (step_tac ctxt 1));
--- a/src/FOLP/simp.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/FOLP/simp.ML Tue Feb 10 14:48:26 2015 +0100
@@ -75,12 +75,12 @@
rewrite rules are not ordered.*)
fun net_tac net =
SUBGOAL(fn (prem,i) =>
- resolve_tac (Net.unify_term net (Logic.strip_assums_concl prem)) i);
+ resolve0_tac (Net.unify_term net (Logic.strip_assums_concl prem)) i);
(*match subgoal i against possible theorems indexed by lhs in the net*)
fun lhs_net_tac net =
SUBGOAL(fn (prem,i) =>
- biresolve_tac (Net.unify_term net
+ biresolve0_tac (Net.unify_term net
(lhs_of (Logic.strip_assums_concl prem))) i);
fun nth_subgoal i thm = nth (prems_of thm) (i - 1);
@@ -119,7 +119,7 @@
fun lhs_is_NORM(thm,i) = case lhs_of_eq i thm of
Const(s,_)$_ => member (op =) norms s | _ => false;
-val refl_tac = resolve_tac refl_thms;
+val refl_tac = resolve0_tac refl_thms;
fun find_res thms thm =
let fun find [] = error "Check Simp_Data"
@@ -138,7 +138,7 @@
SOME(thm',_) => thm'
| NONE => raise THM("Simplifier: could not continue", 0, [thm]);
-fun res1(thm,thms,i) = one_result(resolve_tac thms i,thm);
+fun res1(thm,thms,i) = one_result(resolve0_tac thms i,thm);
(**** Adding "NORM" tags ****)
@@ -204,10 +204,10 @@
fun norm_step_tac st = st |>
(case head_of(rhs_of_eq 1 st) of
Var(ixn,_) => if member (op =) hvs ixn then refl1_tac
- else resolve_tac normI_thms 1 ORELSE refl1_tac
- | Const _ => resolve_tac normI_thms 1 ORELSE
- resolve_tac congs 1 ORELSE refl1_tac
- | Free _ => resolve_tac congs 1 ORELSE refl1_tac
+ else resolve0_tac normI_thms 1 ORELSE refl1_tac
+ | Const _ => resolve0_tac normI_thms 1 ORELSE
+ resolve0_tac congs 1 ORELSE refl1_tac
+ | Free _ => resolve0_tac congs 1 ORELSE refl1_tac
| _ => refl1_tac)
val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) norm_step_tac
val SOME(thm'',_) = Seq.pull(add_norm_tac thm')
--- a/src/HOL/Auth/Event.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Event.thy Tue Feb 10 14:48:26 2015 +0100
@@ -272,7 +272,7 @@
{*
fun analz_mono_contra_tac ctxt =
rtac @{thm analz_impI} THEN'
- REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
+ REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
THEN' (mp_tac ctxt)
*}
@@ -289,7 +289,7 @@
fun synth_analz_mono_contra_tac ctxt =
rtac @{thm syan_impI} THEN'
REPEAT1 o
- (dresolve_tac
+ (dresolve_tac ctxt
[@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
@{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
@{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
--- a/src/HOL/Auth/Message.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Message.thy Tue Feb 10 14:48:26 2015 +0100
@@ -845,13 +845,13 @@
(*Apply rules to break down assumptions of the form
Y \<in> parts(insert X H) and Y \<in> analz(insert X H)
*)
-val Fake_insert_tac =
- dresolve_tac [impOfSubs @{thm Fake_analz_insert},
+fun Fake_insert_tac ctxt =
+ dresolve_tac ctxt [impOfSubs @{thm Fake_analz_insert},
impOfSubs @{thm Fake_parts_insert}] THEN'
- eresolve_tac [asm_rl, @{thm synth.Inj}];
+ eresolve_tac ctxt [asm_rl, @{thm synth.Inj}];
fun Fake_insert_simp_tac ctxt i =
- REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ctxt i;
+ REPEAT (Fake_insert_tac ctxt i) THEN asm_full_simp_tac ctxt i;
fun atomic_spy_analz_tac ctxt =
SELECT_GOAL
@@ -869,7 +869,7 @@
(res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1),
(*...allowing further simplifications*)
simp_tac ctxt 1,
- REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
+ REPEAT (FIRSTGOAL (resolve_tac ctxt [allI,impI,notI,conjI,iffI])),
DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i);
*}
--- a/src/HOL/Auth/OtwayReesBella.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/OtwayReesBella.thy Tue Feb 10 14:48:26 2015 +0100
@@ -126,14 +126,14 @@
ML
{*
-fun parts_explicit_tac i =
- forward_tac [@{thm Oops_parts_knows_Spy}] (i+7) THEN
- forward_tac [@{thm OR4_parts_knows_Spy}] (i+6) THEN
- forward_tac [@{thm OR2_parts_knows_Spy}] (i+4)
+fun parts_explicit_tac ctxt i =
+ forward_tac ctxt [@{thm Oops_parts_knows_Spy}] (i+7) THEN
+ forward_tac ctxt [@{thm OR4_parts_knows_Spy}] (i+6) THEN
+ forward_tac ctxt [@{thm OR2_parts_knows_Spy}] (i+4)
*}
method_setup parts_explicit = {*
- Scan.succeed (K (SIMPLE_METHOD' parts_explicit_tac)) *}
+ Scan.succeed (SIMPLE_METHOD' o parts_explicit_tac) *}
"to explicitly state that some message components belong to parts knows Spy"
@@ -249,7 +249,7 @@
method_setup analz_freshCryptK = {*
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
- (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+ (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshCryptK_lemma}),
ALLGOALS (asm_simp_tac
(put_simpset OtwayReesBella.analz_image_freshK_ss ctxt))]))) *}
@@ -259,7 +259,7 @@
method_setup disentangle = {*
Scan.succeed
(fn ctxt => SIMPLE_METHOD
- (REPEAT_FIRST (eresolve_tac [asm_rl, conjE, disjE]
+ (REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE, disjE]
ORELSE' hyp_subst_tac ctxt))) *}
"for eliminating conjunctions, disjunctions and the like"
--- a/src/HOL/Auth/Public.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Public.thy Tue Feb 10 14:48:26 2015 +0100
@@ -417,7 +417,7 @@
(ALLGOALS (simp_tac (ctxt setSolver safe_solver delsimps [@{thm used_Says}]))
THEN
REPEAT_FIRST (eq_assume_tac ORELSE'
- resolve_tac [refl, conjI, @{thm Nonce_supply}]))
+ resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}]))
(*For harder protocols (such as Recur) where we have to set up some
nonces and keys initially*)
@@ -425,7 +425,7 @@
REPEAT
(ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver))
THEN
- REPEAT_FIRST (resolve_tac [refl, conjI]))
+ REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
end
*}
@@ -433,7 +433,7 @@
method_setup analz_freshK = {*
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
- (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+ (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
ALLGOALS (asm_simp_tac (put_simpset Public.analz_image_freshK_ss ctxt))]))) *}
"for proving the Session Key Compromise theorem"
--- a/src/HOL/Auth/Shared.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Shared.thy Tue Feb 10 14:48:26 2015 +0100
@@ -205,7 +205,7 @@
setSolver safe_solver))
THEN
REPEAT_FIRST (eq_assume_tac ORELSE'
- resolve_tac [refl, conjI, @{thm Nonce_supply}])))
+ resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}])))
(*For harder protocols (such as Recur) where we have to set up some
nonces and keys initially*)
@@ -213,7 +213,7 @@
REPEAT
(ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver))
THEN
- REPEAT_FIRST (resolve_tac [refl, conjI]))
+ REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
val analz_image_freshK_ss =
@@ -237,7 +237,7 @@
method_setup analz_freshK = {*
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
- (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+ (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
ALLGOALS (asm_simp_tac (put_simpset Shared.analz_image_freshK_ss ctxt))]))) *}
"for proving the Session Key Compromise theorem"
--- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Tue Feb 10 14:48:26 2015 +0100
@@ -736,32 +736,32 @@
structure ShoupRubin =
struct
-val prepare_tac =
- (*SR8*) forward_tac [@{thm Outpts_B_Card_form_7}] 14 THEN
- eresolve_tac [exE] 15 THEN eresolve_tac [exE] 15 THEN
- (*SR9*) forward_tac [@{thm Outpts_A_Card_form_4}] 16 THEN
- (*SR11*) forward_tac [@{thm Outpts_A_Card_form_10}] 21 THEN
- eresolve_tac [exE] 22 THEN eresolve_tac [exE] 22
+fun prepare_tac ctxt =
+ (*SR8*) forward_tac ctxt [@{thm Outpts_B_Card_form_7}] 14 THEN
+ eresolve_tac ctxt [exE] 15 THEN eresolve_tac ctxt [exE] 15 THEN
+ (*SR9*) forward_tac ctxt [@{thm Outpts_A_Card_form_4}] 16 THEN
+ (*SR11*) forward_tac ctxt [@{thm Outpts_A_Card_form_10}] 21 THEN
+ eresolve_tac ctxt [exE] 22 THEN eresolve_tac ctxt [exE] 22
fun parts_prepare_tac ctxt =
- prepare_tac THEN
- (*SR9*) dresolve_tac [@{thm Gets_imp_knows_Spy_parts_Snd}] 18 THEN
- (*SR9*) dresolve_tac [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN
- (*Oops1*) dresolve_tac [@{thm Outpts_B_Card_form_7}] 25 THEN
- (*Oops2*) dresolve_tac [@{thm Outpts_A_Card_form_10}] 27 THEN
+ prepare_tac ctxt THEN
+ (*SR9*) dresolve_tac ctxt [@{thm Gets_imp_knows_Spy_parts_Snd}] 18 THEN
+ (*SR9*) dresolve_tac ctxt [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN
+ (*Oops1*) dresolve_tac ctxt [@{thm Outpts_B_Card_form_7}] 25 THEN
+ (*Oops2*) dresolve_tac ctxt [@{thm Outpts_A_Card_form_10}] 27 THEN
(*Base*) (force_tac ctxt) 1
fun analz_prepare_tac ctxt =
- prepare_tac THEN
+ prepare_tac ctxt THEN
dtac @{thm Gets_imp_knows_Spy_analz_Snd} 18 THEN
(*SR9*) dtac @{thm Gets_imp_knows_Spy_analz_Snd} 19 THEN
- REPEAT_FIRST (eresolve_tac [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
+ REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
end
*}
method_setup prepare = {*
- Scan.succeed (K (SIMPLE_METHOD ShoupRubin.prepare_tac)) *}
+ Scan.succeed (SIMPLE_METHOD o ShoupRubin.prepare_tac) *}
"to launch a few simple facts that will help the simplifier"
method_setup parts_prepare = {*
@@ -817,7 +817,7 @@
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
(EVERY [REPEAT_FIRST
- (resolve_tac [allI, ballI, impI]),
+ (resolve_tac ctxt [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt
addsimps [@{thm knows_Spy_Inputs_secureM_sr_Spy},
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Tue Feb 10 14:48:26 2015 +0100
@@ -747,24 +747,24 @@
struct
fun prepare_tac ctxt =
- (*SR_U8*) forward_tac [@{thm Outpts_B_Card_form_7}] 14 THEN
+ (*SR_U8*) forward_tac ctxt [@{thm Outpts_B_Card_form_7}] 14 THEN
(*SR_U8*) clarify_tac ctxt 15 THEN
- (*SR_U9*) forward_tac [@{thm Outpts_A_Card_form_4}] 16 THEN
- (*SR_U11*) forward_tac [@{thm Outpts_A_Card_form_10}] 21
+ (*SR_U9*) forward_tac ctxt [@{thm Outpts_A_Card_form_4}] 16 THEN
+ (*SR_U11*) forward_tac ctxt [@{thm Outpts_A_Card_form_10}] 21
fun parts_prepare_tac ctxt =
prepare_tac ctxt THEN
- (*SR_U9*) dresolve_tac [@{thm Gets_imp_knows_Spy_parts_Snd}] 18 THEN
- (*SR_U9*) dresolve_tac [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN
- (*Oops1*) dresolve_tac [@{thm Outpts_B_Card_form_7}] 25 THEN
- (*Oops2*) dresolve_tac [@{thm Outpts_A_Card_form_10}] 27 THEN
+ (*SR_U9*) dresolve_tac ctxt [@{thm Gets_imp_knows_Spy_parts_Snd}] 18 THEN
+ (*SR_U9*) dresolve_tac ctxt [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN
+ (*Oops1*) dresolve_tac ctxt [@{thm Outpts_B_Card_form_7}] 25 THEN
+ (*Oops2*) dresolve_tac ctxt [@{thm Outpts_A_Card_form_10}] 27 THEN
(*Base*) (force_tac ctxt) 1
fun analz_prepare_tac ctxt =
prepare_tac ctxt THEN
dtac (@{thm Gets_imp_knows_Spy_analz_Snd}) 18 THEN
(*SR_U9*) dtac (@{thm Gets_imp_knows_Spy_analz_Snd}) 19 THEN
- REPEAT_FIRST (eresolve_tac [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
+ REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
end
*}
@@ -826,7 +826,7 @@
method_setup sc_analz_freshK = {*
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
- (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+ (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt
addsimps [@{thm knows_Spy_Inputs_secureM_srb_Spy},
--- a/src/HOL/Auth/Smartcard/Smartcard.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy Tue Feb 10 14:48:26 2015 +0100
@@ -374,7 +374,7 @@
setSolver safe_solver))
THEN
REPEAT_FIRST (eq_assume_tac ORELSE'
- resolve_tac [refl, conjI, @{thm Nonce_supply}])))
+ resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}])))
(*For harder protocols (such as Recur) where we have to set up some
nonces and keys initially*)
@@ -382,7 +382,7 @@
REPEAT
(ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver))
THEN
- REPEAT_FIRST (resolve_tac [refl, conjI]))
+ REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
val analz_image_freshK_ss =
simpset_of
@@ -403,7 +403,7 @@
method_setup analz_freshK = {*
Scan.succeed (fn ctxt =>
(SIMPLE_METHOD
- (EVERY [REPEAT_FIRST (resolve_tac [allI, ballI, impI]),
+ (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt))]))) *}
"for proving the Session Key Compromise theorem"
--- a/src/HOL/Bali/AxExample.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Bali/AxExample.thy Tue Feb 10 14:48:26 2015 +0100
@@ -46,9 +46,9 @@
SOME i => Rule_Insts.instantiate_tac ctxt [((s, i), t)] xs st
| NONE => Seq.empty);
-val ax_tac =
+fun ax_tac ctxt =
REPEAT o rtac allI THEN'
- resolve_tac (@{thm ax_Skip} :: @{thm ax_StatRef} :: @{thm ax_MethdN} :: @{thm ax_Alloc} ::
+ resolve_tac ctxt (@{thm ax_Skip} :: @{thm ax_StatRef} :: @{thm ax_MethdN} :: @{thm ax_Alloc} ::
@{thm ax_Alloc_Arr} :: @{thm ax_SXAlloc_Normal} :: @{thms ax_derivs.intros(8-)});
*}
@@ -58,11 +58,11 @@
.test [Class Base].
{\<lambda>Y s Z. abrupt s = Some (Xcpt (Std IndOutBound))}"
apply (unfold test_def arr_viewed_from_def)
-apply (tactic "ax_tac 1" (*;;*))
+apply (tactic "ax_tac @{context} 1" (*;;*))
defer (* We begin with the last assertion, to synthesise the intermediate
assertions, like in the fashion of the weakest
precondition. *)
-apply (tactic "ax_tac 1" (* Try *))
+apply (tactic "ax_tac @{context} 1" (* Try *))
defer
apply (tactic {* inst1_tac @{context} "Q"
"\<lambda>Y s Z. arr_inv (snd s) \<and> tprg,s\<turnstile>catch SXcpt NullPointer" [] *})
@@ -74,26 +74,26 @@
apply (rule_tac Q' = "(\<lambda>Y s Z. ?Q Y s Z)\<leftarrow>=False\<down>=\<diamondsuit>" in conseq2)
prefer 2
apply simp
-apply (tactic "ax_tac 1" (* While *))
+apply (tactic "ax_tac @{context} 1" (* While *))
prefer 2
apply (rule ax_impossible [THEN conseq1], clarsimp)
apply (rule_tac P' = "Normal ?P" in conseq1)
prefer 2
apply clarsimp
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1" (* AVar *))
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1" (* AVar *))
prefer 2
apply (rule ax_subst_Val_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>a. Normal (PP a\<leftarrow>x)" ["PP", "x"] *})
apply (simp del: avar_def2 peek_and_def2)
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
(* just for clarification: *)
apply (rule_tac Q' = "Normal (\<lambda>Var:(v, f) u ua. fst (snd (avar tprg (Intg 2) v u)) = Some (Xcpt (Std IndOutBound)))" in conseq2)
prefer 2
apply (clarsimp simp add: split_beta)
-apply (tactic "ax_tac 1" (* FVar *))
-apply (tactic "ax_tac 2" (* StatRef *))
+apply (tactic "ax_tac @{context} 1" (* FVar *))
+apply (tactic "ax_tac @{context} 2" (* StatRef *))
apply (rule ax_derivs.Done [THEN conseq1])
apply (clarsimp simp add: arr_inv_def inited_def in_bounds_def)
defer
@@ -101,20 +101,20 @@
apply (rule_tac Q' = "(\<lambda>Y (x, s) Z. x = Some (Xcpt (Std NullPointer)) \<and> arr_inv s) \<and>. heap_free two" in conseq2)
prefer 2
apply (simp add: arr_inv_new_obj)
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply (rule_tac C = "Ext" in ax_Call_known_DynT)
apply (unfold DynT_prop_def)
apply (simp (no_asm))
apply (intro strip)
apply (rule_tac P' = "Normal ?P" in conseq1)
-apply (tactic "ax_tac 1" (* Methd *))
+apply (tactic "ax_tac @{context} 1" (* Methd *))
apply (rule ax_thin [OF _ empty_subsetI])
apply (simp (no_asm) add: body_def2)
-apply (tactic "ax_tac 1" (* Body *))
+apply (tactic "ax_tac @{context} 1" (* Body *))
(* apply (rule_tac [2] ax_derivs.Abrupt) *)
defer
apply (simp (no_asm))
-apply (tactic "ax_tac 1") (* Comp *)
+apply (tactic "ax_tac @{context} 1") (* Comp *)
(* The first statement in the composition
((Ext)z).vee = 1; Return Null
will throw an exception (since z is null). So we can handle
@@ -122,7 +122,7 @@
apply (rule_tac [2] ax_derivs.Abrupt)
apply (rule ax_derivs.Expr) (* Expr *)
-apply (tactic "ax_tac 1") (* Ass *)
+apply (tactic "ax_tac @{context} 1") (* Ass *)
prefer 2
apply (rule ax_subst_Var_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>a vs l vf. PP a vs l vf\<leftarrow>x \<and>. p" ["PP", "x", "p"] *})
@@ -130,9 +130,9 @@
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 *))
-apply (tactic "ax_tac 2", tactic "ax_tac 2", tactic "ax_tac 2")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1" (* FVar *))
+apply (tactic "ax_tac @{context} 2", tactic "ax_tac @{context} 2", tactic "ax_tac @{context} 2")
+apply (tactic "ax_tac @{context} 1")
apply (tactic {* inst1_tac @{context} "R" "\<lambda>a'. Normal ((\<lambda>Vals:vs (x, s) Z. arr_inv s \<and> inited Ext (globs s) \<and> a' \<noteq> Null \<and> vs = [Null]) \<and>. heap_free two)" [] *})
apply fastforce
prefer 4
@@ -140,15 +140,15 @@
apply (rule ax_subst_Val_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>a. Normal (PP a\<leftarrow>x)" ["PP", "x"] *})
apply (simp (no_asm) del: peek_and_def2 heap_free_def2 normal_def2 o_apply)
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
prefer 2
apply (rule ax_subst_Val_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>aa v. Normal (QQ aa v\<leftarrow>y)" ["QQ", "y"] *})
apply (simp del: peek_and_def2 heap_free_def2 normal_def2)
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
(* end method call *)
apply (simp (no_asm))
(* just for clarification: *)
@@ -158,14 +158,14 @@
in conseq2)
prefer 2
apply clarsimp
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
defer
apply (rule ax_subst_Var_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>vf. Normal (PP vf \<and>. p)" ["PP", "p"] *})
apply (simp (no_asm) del: split_paired_All peek_and_def2 initd_def2 heap_free_def2 normal_def2)
-apply (tactic "ax_tac 1" (* NewC *))
-apply (tactic "ax_tac 1" (* ax_Alloc *))
+apply (tactic "ax_tac @{context} 1" (* NewC *))
+apply (tactic "ax_tac @{context} 1" (* ax_Alloc *))
(* just for clarification: *)
apply (rule_tac Q' = "Normal ((\<lambda>Y s Z. arr_inv (store s) \<and> vf=lvar (VName e) (store s)) \<and>. heap_free three \<and>. initd Ext)" in conseq2)
prefer 2
@@ -187,19 +187,19 @@
apply (rule allI)
apply (rule_tac P' = "Normal ?P" in conseq1)
apply (tactic {* simp_tac (@{context} delloop "split_all_tac") 1 *})
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
apply (rule_tac [2] ax_subst_Var_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>vf l vfa. Normal (P vf l vfa)" ["P"] *})
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 "ax_tac @{context} 2" (* NewA *))
+apply (tactic "ax_tac @{context} 3" (* ax_Alloc_Arr *))
+apply (tactic "ax_tac @{context} 3")
apply (tactic {* inst1_tac @{context} "P" "\<lambda>vf l vfa. Normal (P vf l vfa\<leftarrow>\<diamondsuit>)" ["P"] *})
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 *))
+apply (tactic "ax_tac @{context} 2")
+apply (tactic "ax_tac @{context} 1" (* FVar *))
+apply (tactic "ax_tac @{context} 2" (* StatRef *))
apply (rule ax_derivs.Done [THEN conseq1])
apply (tactic {* inst1_tac @{context} "Q" "\<lambda>vf. Normal ((\<lambda>Y s Z. vf=lvar (VName e) (snd s)) \<and>. heap_free four \<and>. initd Base \<and>. initd Ext)" [] *})
apply (clarsimp split del: split_if)
@@ -217,7 +217,7 @@
apply clarsimp
(* end init *)
apply (rule conseq1)
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply clarsimp
done
@@ -234,36 +234,36 @@
(Expr (Ass (LVar i) (Acc (LVar j))))). {Q}"
apply (rule_tac P' = "Q" and Q' = "Q\<leftarrow>=False\<down>=\<diamondsuit>" in conseq12)
apply safe
-apply (tactic "ax_tac 1" (* Loop *))
+apply (tactic "ax_tac @{context} 1" (* Loop *))
apply (rule ax_Normal_cases)
prefer 2
apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
apply (rule conseq1)
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply clarsimp
prefer 2
apply clarsimp
-apply (tactic "ax_tac 1" (* If *))
+apply (tactic "ax_tac @{context} 1" (* If *))
apply (tactic
{* inst1_tac @{context} "P'" "Normal (\<lambda>s.. (\<lambda>Y s Z. True)\<down>=Val (the (locals s i)))" [] *})
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply (rule conseq1)
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply clarsimp
apply (rule allI)
apply (rule ax_escape)
apply auto
apply (rule conseq1)
-apply (tactic "ax_tac 1" (* Throw *))
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1" (* Throw *))
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
apply clarsimp
apply (rule_tac Q' = "Normal (\<lambda>Y s Z. True)" in conseq2)
prefer 2
apply clarsimp
apply (rule conseq1)
-apply (tactic "ax_tac 1")
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
+apply (tactic "ax_tac @{context} 1")
prefer 2
apply (rule ax_subst_Var_allI)
apply (tactic {* inst1_tac @{context} "P'" "\<lambda>b Y ba Z vf. \<lambda>Y (x,s) Z. x=None \<and> snd vf = snd (lvar i s)" [] *})
@@ -271,11 +271,11 @@
apply (rule_tac P' = "Normal ?P" in conseq1)
prefer 2
apply clarsimp
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply (rule conseq1)
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply clarsimp
-apply (tactic "ax_tac 1")
+apply (tactic "ax_tac @{context} 1")
apply clarsimp
done
--- a/src/HOL/Bali/AxSem.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Bali/AxSem.thy Tue Feb 10 14:48:26 2015 +0100
@@ -672,7 +672,7 @@
(* 37 subgoals *)
prefer 18 (* Methd *)
apply (rule ax_derivs.Methd, drule spec, erule mp, fast)
-apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) @{thms ax_derivs.intros})) *})
+apply (tactic {* TRYALL (resolve_tac @{context} ((funpow 5 tl) @{thms ax_derivs.intros})) *})
apply auto
done
@@ -691,7 +691,7 @@
"G,(A::'a triple set)|\<turnstile>(ts'::'a triple set) \<Longrightarrow> !ts. ts \<subseteq> ts' \<longrightarrow> G,A|\<turnstile>ts"
apply (erule ax_derivs.induct)
(*42 subgoals*)
-apply (tactic "ALLGOALS strip_tac")
+apply (tactic "ALLGOALS (strip_tac @{context})")
apply (tactic {* ALLGOALS(REPEAT o (EVERY'[dtac @{thm subset_singletonD},
etac disjE, fast_tac (@{context} addSIs @{thms ax_derivs.empty})]))*})
apply (tactic "TRYALL (hyp_subst_tac @{context})")
@@ -703,7 +703,7 @@
apply (fast intro: ax_derivs.weaken)
apply (rule ax_derivs.conseq, clarify, tactic "smp_tac @{context} 3 1", blast(* unused *))
(*37 subgoals*)
-apply (tactic {* TRYALL (resolve_tac ((funpow 5 tl) @{thms ax_derivs.intros})
+apply (tactic {* TRYALL (resolve_tac @{context} ((funpow 5 tl) @{thms ax_derivs.intros})
THEN_ALL_NEW fast_tac @{context}) *})
(*1 subgoal*)
apply (clarsimp simp add: subset_mtriples_iff)
--- a/src/HOL/Bali/Basis.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Bali/Basis.thy Tue Feb 10 14:48:26 2015 +0100
@@ -9,7 +9,7 @@
subsubsection "misc"
-ML {* fun strip_tac i = REPEAT (resolve_tac [impI, allI] i) *}
+ML {* fun strip_tac ctxt i = REPEAT (resolve_tac ctxt [impI, allI] i) *}
declare split_if_asm [split] option.split [split] option.split_asm [split]
setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
--- a/src/HOL/Bali/Eval.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Bali/Eval.thy Tue Feb 10 14:48:26 2015 +0100
@@ -1163,7 +1163,7 @@
"G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w, s') \<Longrightarrow> (\<forall>w' s''. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w', s'') \<longrightarrow> w' = w \<and> s'' = s')"
apply (erule eval_induct)
apply (tactic {* ALLGOALS (EVERY'
- [strip_tac, rotate_tac ~1, eresolve_tac @{thms eval_elim_cases}]) *})
+ [strip_tac @{context}, rotate_tac ~1, eresolve_tac @{context} @{thms eval_elim_cases}]) *})
(* 31 subgoals *)
prefer 28 (* Try *)
apply (simp (no_asm_use) only: split add: split_if_asm)
--- a/src/HOL/Bali/Evaln.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Bali/Evaln.thy Tue Feb 10 14:48:26 2015 +0100
@@ -448,9 +448,9 @@
lemma evaln_nonstrict [rule_format (no_asm), elim]:
"G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (w, s') \<Longrightarrow> \<forall>m. n\<le>m \<longrightarrow> G\<turnstile>s \<midarrow>t\<succ>\<midarrow>m\<rightarrow> (w, s')"
apply (erule evaln.induct)
-apply (tactic {* ALLGOALS (EVERY'[strip_tac, TRY o etac @{thm Suc_le_D_lemma},
+apply (tactic {* ALLGOALS (EVERY' [strip_tac @{context}, TRY o etac @{thm Suc_le_D_lemma},
REPEAT o smp_tac @{context} 1,
- resolve_tac @{thms evaln.intros} THEN_ALL_NEW TRY o atac]) *})
+ resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW TRY o atac]) *})
(* 3 subgoals *)
apply (auto split del: split_if)
done
--- a/src/HOL/Bali/WellType.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Bali/WellType.thy Tue Feb 10 14:48:26 2015 +0100
@@ -665,7 +665,7 @@
thin_tac @{context} "?E,dt\<Turnstile>e2\<Colon>-?T2"] i
else thin_tac @{context} "All ?P" i) *})
(*apply (safe del: disjE elim!: wt_elim_cases)*)
-apply (tactic {*ALLGOALS (eresolve_tac @{thms wt_elim_cases})*})
+apply (tactic {*ALLGOALS (eresolve_tac @{context} @{thms wt_elim_cases})*})
apply (simp_all (no_asm_use) split del: split_if_asm)
apply (erule_tac [12] V = "All ?P" in thin_rl) (* Call *)
apply (blast del: equalityCE dest: sym [THEN trans])+
--- a/src/HOL/Datatype_Examples/Stream_Processor.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Datatype_Examples/Stream_Processor.thy Tue Feb 10 14:48:26 2015 +0100
@@ -136,7 +136,7 @@
BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac,
rtac @{thm conversepI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks,
- REPEAT_DETERM o eresolve_tac @{thms relcomppE conversepE GrpE},
+ REPEAT_DETERM o eresolve_tac @{context} @{thms relcomppE conversepE GrpE},
hyp_subst_tac @{context}, atac])
end
*})
--- a/src/HOL/Decision_Procs/Approximation.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy Tue Feb 10 14:48:26 2015 +0100
@@ -3699,7 +3699,7 @@
etac @{thm meta_eqE},
rtac @{thm impI}] i)
THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) ctxt i
- THEN DETERM (TRY (filter_prems_tac (K false) i))
+ THEN DETERM (TRY (filter_prems_tac ctxt (K false) i))
THEN DETERM (Reification.tac ctxt form_equations NONE i)
THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
THEN gen_eval_tac (approximation_conv ctxt) ctxt i))
--- a/src/HOL/Decision_Procs/approximation.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Decision_Procs/approximation.ML Tue Feb 10 14:48:26 2015 +0100
@@ -100,7 +100,7 @@
fun prepare_form ctxt term = apply_tactic ctxt term (
REPEAT (FIRST' [etac @{thm intervalE}, etac @{thm meta_eqE}, rtac @{thm impI}] 1)
THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems 1) ctxt 1
- THEN DETERM (TRY (filter_prems_tac (K false) 1)))
+ THEN DETERM (TRY (filter_prems_tac ctxt (K false) 1)))
fun reify_form ctxt term = apply_tactic ctxt term
(Reification.tac ctxt form_equations NONE 1)
--- a/src/HOL/Fun.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Fun.thy Tue Feb 10 14:48:26 2015 +0100
@@ -839,8 +839,8 @@
| (T, SOME rhs) =>
SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs))
(fn _ =>
- resolve_tac [eq_reflection] 1 THEN
- resolve_tac @{thms ext} 1 THEN
+ resolve_tac ctxt [eq_reflection] 1 THEN
+ resolve_tac ctxt @{thms ext} 1 THEN
simp_tac (put_simpset ss ctxt) 1))
end
in proc end
--- a/src/HOL/HOL.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOL.thy Tue Feb 10 14:48:26 2015 +0100
@@ -905,7 +905,7 @@
apply (rule ex1E [OF major])
apply (rule prem)
apply (tactic {* ares_tac @{thms allI} 1 *})+
-apply (tactic {* eresolve_tac [Classical.dup_elim NONE @{thm allE}] 1 *})
+apply (tactic {* eresolve_tac @{context} [Classical.dup_elim NONE @{thm allE}] 1 *})
apply iprover
done
@@ -1822,7 +1822,7 @@
proof
assume "PROP ?ofclass"
show "PROP ?equal"
- by (tactic {* ALLGOALS (resolve_tac [Thm.unconstrainT @{thm eq_equal}]) *})
+ by (tactic {* ALLGOALS (resolve_tac @{context} [Thm.unconstrainT @{thm eq_equal}]) *})
(fact `PROP ?ofclass`)
next
assume "PROP ?equal"
@@ -1923,7 +1923,7 @@
let val conv = Code_Runtime.dynamic_holds_conv ctxt
in
CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 conv)) ctxt) THEN'
- resolve_tac [TrueI]
+ resolve_tac ctxt [TrueI]
end
in
Scan.succeed (SIMPLE_METHOD' o eval_tac)
@@ -1935,7 +1935,7 @@
SIMPLE_METHOD'
(CHANGED_PROP o
(CONVERSION (Nbe.dynamic_conv ctxt)
- THEN_ALL_NEW (TRY o resolve_tac [TrueI]))))
+ THEN_ALL_NEW (TRY o resolve_tac ctxt [TrueI]))))
*} "solve goal by normalization"
@@ -1975,7 +1975,7 @@
val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of);
in
fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]);
- fun smp_tac ctxt j = EVERY'[dresolve_tac (smp j), assume_tac ctxt];
+ fun smp_tac ctxt j = EVERY' [dresolve_tac ctxt (smp j), assume_tac ctxt];
end;
local
--- a/src/HOL/HOLCF/IOA/NTP/Impl.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Impl.thy Tue Feb 10 14:48:26 2015 +0100
@@ -208,19 +208,19 @@
txt {* 10 - 7 *}
apply (tactic "EVERY1 [tac2,tac2,tac2,tac2]")
txt {* 6 *}
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv1_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv1_def}]
(@{thm raw_inv1} RS @{thm invariantE}) RS conjunct1] 1 *})
txt {* 6 - 5 *}
apply (tactic "EVERY1 [tac2,tac2]")
txt {* 4 *}
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv1_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv1_def}]
(@{thm raw_inv1} RS @{thm invariantE}) RS conjunct1] 1 *})
apply (tactic "tac2 1")
txt {* 3 *}
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv1_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv1_def}]
(@{thm raw_inv1} RS @{thm invariantE})] 1 *})
apply (tactic "tac2 1")
@@ -230,7 +230,7 @@
txt {* 2 *}
apply (tactic "tac2 1")
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv1_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv1_def}]
(@{thm raw_inv1} RS @{thm invariantE}) RS conjunct1] 1 *})
apply (intro strip)
apply (erule conjE)+
@@ -238,7 +238,7 @@
txt {* 1 *}
apply (tactic "tac2 1")
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv1_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv1_def}]
(@{thm raw_inv1} RS @{thm invariantE}) RS conjunct2] 1 *})
apply (intro strip)
apply (erule conjE)+
@@ -287,13 +287,13 @@
apply (intro strip, (erule conjE)+)
apply (rule imp_disjL [THEN iffD1])
apply (rule impI)
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv2_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv2_def}]
(@{thm raw_inv2} RS @{thm invariantE})] 1 *})
apply simp
apply (erule conjE)+
apply (rule_tac j = "count (ssent (sen s)) (~sbit (sen s))" and
k = "count (rsent (rec s)) (sbit (sen s))" in le_trans)
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm inv1_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm inv1_def}]
(@{thm raw_inv1} RS @{thm invariantE}) RS conjunct2] 1 *})
apply (simp add: hdr_sum_def Multiset.count_def)
apply (rule add_le_mono)
@@ -308,7 +308,7 @@
apply (intro strip, (erule conjE)+)
apply (rule imp_disjL [THEN iffD1])
apply (rule impI)
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv2_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv2_def}]
(@{thm raw_inv2} RS @{thm invariantE})] 1 *})
apply simp
done
@@ -334,7 +334,7 @@
txt {* 2 b *}
apply (intro strip, (erule conjE)+)
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv2_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv2_def}]
(@{thm raw_inv2} RS @{thm invariantE})] 1 *})
apply simp
@@ -342,9 +342,9 @@
apply (tactic "tac4 1")
apply (intro strip, (erule conjE)+)
apply (rule ccontr)
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv2_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv2_def}]
(@{thm raw_inv2} RS @{thm invariantE})] 1 *})
- apply (tactic {* forward_tac [rewrite_rule @{context} [@{thm Impl.inv3_def}]
+ apply (tactic {* forward_tac @{context} [rewrite_rule @{context} [@{thm Impl.inv3_def}]
(@{thm raw_inv3} RS @{thm invariantE})] 1 *})
apply simp
apply (rename_tac m, erule_tac x = "m" in allE)
--- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Tue Feb 10 14:48:26 2015 +0100
@@ -244,9 +244,9 @@
let
val rules = @{thms compact_sinl compact_sinr compact_spair
compact_up compact_ONE}
- val tacs =
+ fun tacs ctxt =
[rtac (iso_locale RS @{thm iso.compact_abs}) 1,
- REPEAT (resolve_tac rules 1 ORELSE atac 1)]
+ REPEAT (resolve_tac ctxt rules 1 ORELSE atac 1)]
fun con_compact (con, args) =
let
val vs = vars_of args
@@ -255,7 +255,7 @@
val assms = map (mk_trp o mk_compact) vs
val goal = Logic.list_implies (assms, concl)
in
- prove thy con_betas goal (K tacs)
+ prove thy con_betas goal (tacs o #context)
end
in
map con_compact spec'
@@ -500,8 +500,8 @@
val defs = case_beta :: map mk_meta_eq [rep_strict, @{thm cfcomp2}]
val goal = mk_trp (mk_strict case_app)
val rules = @{thms sscase1 ssplit1 strictify1 one_case1}
- val tacs = [resolve_tac rules 1]
- in prove thy defs goal (K tacs) end
+ fun tacs ctxt = [resolve_tac ctxt rules 1]
+ in prove thy defs goal (tacs o #context) end
(* prove rewrites for case combinator *)
local
--- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Tue Feb 10 14:48:26 2015 +0100
@@ -180,7 +180,7 @@
TRY (safe_tac (put_claset HOL_cs ctxt))]
fun con_tac _ =
asm_simp_tac (put_simpset take_ss ctxt) 1 THEN
- (resolve_tac prems' THEN_ALL_NEW etac spec) 1
+ (resolve_tac ctxt prems' THEN_ALL_NEW etac spec) 1
fun cases_tacs (cons, exhaust) =
res_inst_tac ctxt [(("y", 0), "x")] exhaust 1 ::
asm_simp_tac (put_simpset take_ss ctxt addsimps prems) 1 ::
@@ -201,8 +201,8 @@
let
fun finite_tac (take_induct, fin_ind) =
rtac take_induct 1 THEN
- (if is_finite then all_tac else resolve_tac prems 1) THEN
- (rtac fin_ind THEN_ALL_NEW solve_tac prems) 1
+ (if is_finite then all_tac else resolve_tac ctxt prems 1) THEN
+ (rtac fin_ind THEN_ALL_NEW solve_tac ctxt prems) 1
val fin_inds = Project_Rule.projections ctxt finite_ind
in
TRY (safe_tac (put_claset HOL_cs ctxt)) THEN
--- a/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML Tue Feb 10 14:48:26 2015 +0100
@@ -311,11 +311,11 @@
EVERY
[rewrite_goals_tac ctxt map_apply_thms,
rtac (map_cont_thm RS @{thm cont_fix_ind}) 1,
- REPEAT (resolve_tac adm_rules 1),
+ REPEAT (resolve_tac ctxt adm_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)])
+ REPEAT (resolve_tac ctxt (deflation_rules @ prems) 1 ORELSE atac 1)])
end
fun conjuncts [] _ = []
| conjuncts (n::[]) thm = [(n, thm)]
@@ -516,7 +516,7 @@
Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems domain_defl_simps}
fun tac ctxt =
rewrite_goals_tac ctxt (map mk_meta_eq (rev DEFL_simps))
- THEN TRY (resolve_tac defl_unfold_thms 1)
+ THEN TRY (resolve_tac ctxt defl_unfold_thms 1)
in
Goal.prove_global thy [] [] goal (tac o #context)
end
@@ -625,12 +625,12 @@
[rewrite_goals_tac ctxt (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),
+ REPEAT (resolve_tac ctxt adm_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,
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)])
+ REPEAT (resolve_tac ctxt (isodefl_rules @ prems) 1 ORELSE atac 1)])
end
val isodefl_binds = map (Binding.prefix_name "isodefl_") dbinds
fun conjuncts [] _ = []
@@ -654,13 +654,13 @@
fun is_cpo T = Sign.of_sort thy (T, @{sort cpo})
val lhs = list_ccomb (map_const, map mk_ID (filter is_cpo Ts))
val goal = mk_eqs (lhs, mk_ID lhsT)
- val tac = EVERY
+ fun tac ctxt = EVERY
[rtac @{thm isodefl_DEFL_imp_ID} 1,
- stac DEFL_thm 1,
+ stac ctxt DEFL_thm 1,
rtac isodefl_thm 1,
- REPEAT (resolve_tac @{thms isodefl_ID_DEFL isodefl_LIFTDEFL} 1)]
+ REPEAT (resolve_tac ctxt @{thms isodefl_ID_DEFL isodefl_LIFTDEFL} 1)]
in
- Goal.prove_global thy [] [] goal (K tac)
+ Goal.prove_global thy [] [] goal (tac o #context)
end
val map_ID_binds = map (Binding.suffix_name "_map_ID") dbinds
val map_ID_thms =
--- a/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML Tue Feb 10 14:48:26 2015 +0100
@@ -324,7 +324,7 @@
simp_tac (put_simpset HOL_basic_ss ctxt addsimps bottom_rules) 1,
asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps take_Suc_thms) 1,
REPEAT (etac @{thm conjE} 1
- ORELSE resolve_tac deflation_rules 1
+ ORELSE resolve_tac ctxt deflation_rules 1
ORELSE atac 1)])
end
fun conjuncts [] _ = []
@@ -462,9 +462,9 @@
EVERY [
rewrite_goals_tac ctxt finite_defs,
rtac @{thm lub_ID_finite} 1,
- resolve_tac chain_take_thms 1,
- resolve_tac lub_take_thms 1,
- resolve_tac decisive_thms 1]
+ resolve_tac ctxt chain_take_thms 1,
+ resolve_tac ctxt lub_take_thms 1,
+ resolve_tac ctxt decisive_thms 1]
in
Goal.prove_global thy [] [] goal (tac o #context)
end
--- a/src/HOL/HOLCF/Tools/cont_proc.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/HOLCF/Tools/cont_proc.ML Tue Feb 10 14:48:26 2015 +0100
@@ -7,7 +7,7 @@
val is_lcf_term: term -> bool
val cont_thms: term -> thm list
val all_cont_thms: term -> thm list
- val cont_tac: int -> tactic
+ val cont_tac: Proof.context -> int -> tactic
val cont_proc: theory -> simproc
val setup: theory -> theory
end
@@ -95,7 +95,7 @@
conditional rewrite rule with the unsolved subgoals as premises.
*)
-val cont_tac =
+fun cont_tac ctxt =
let
val rules = [cont_K, cont_I, cont_R, cont_A, cont_L]
@@ -110,7 +110,7 @@
in
if is_lcf_term f'
then new_cont_tac f'
- else REPEAT_ALL_NEW (resolve_tac rules)
+ else REPEAT_ALL_NEW (resolve_tac ctxt rules)
end
| cont_tac_of_term _ = K no_tac
in
@@ -123,7 +123,7 @@
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 Option.map fst (Seq.pull (cont_tac ctxt 1 tr)) end
in
fun cont_proc thy =
Simplifier.simproc_global thy "cont_proc" ["cont f"] solve_cont
--- a/src/HOL/IMPP/Hoare.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/IMPP/Hoare.thy Tue Feb 10 14:48:26 2015 +0100
@@ -218,7 +218,7 @@
apply (rule hoare_derivs.conseq, intro strip, tactic "smp_tac @{context} 2 1", clarify, tactic "smp_tac @{context} 1 1",rule exI, rule exI, erule (1) conjI)
prefer 7
apply (rule_tac hoare_derivs.Body, drule_tac spec, erule_tac mp, fast)
-apply (tactic {* ALLGOALS (resolve_tac ((funpow 5 tl) @{thms hoare_derivs.intros}) THEN_ALL_NEW (fast_tac @{context})) *})
+apply (tactic {* ALLGOALS (resolve_tac @{context} ((funpow 5 tl) @{thms hoare_derivs.intros}) THEN_ALL_NEW (fast_tac @{context})) *})
done
lemma weak_Body: "G|-{P}. the (body pn) .{Q} ==> G|-{P}. BODY pn .{Q}"
@@ -278,7 +278,7 @@
lemma hoare_sound: "G||-ts ==> G||=ts"
apply (erule hoare_derivs.induct)
-apply (tactic {* TRYALL (eresolve_tac [@{thm Loop_sound_lemma}, @{thm Body_sound_lemma}] THEN_ALL_NEW atac) *})
+apply (tactic {* TRYALL (eresolve_tac @{context} [@{thm Loop_sound_lemma}, @{thm Body_sound_lemma}] THEN_ALL_NEW atac) *})
apply (unfold hoare_valids_def)
apply blast
apply blast
--- a/src/HOL/IMPP/Natural.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/IMPP/Natural.thy Tue Feb 10 14:48:26 2015 +0100
@@ -112,7 +112,7 @@
lemma evaln_evalc: "<c,s> -n-> t ==> <c,s> -c-> t"
apply (erule evaln.induct)
-apply (tactic {* ALLGOALS (resolve_tac @{thms evalc.intros} THEN_ALL_NEW atac) *})
+apply (tactic {* ALLGOALS (resolve_tac @{context} @{thms evalc.intros} THEN_ALL_NEW atac) *})
done
lemma Suc_le_D_lemma: "[| Suc n <= m'; (!!m. n <= m ==> P (Suc m)) |] ==> P m'"
@@ -140,8 +140,9 @@
apply (erule evalc.induct)
apply (tactic {* ALLGOALS (REPEAT o etac exE) *})
apply (tactic {* TRYALL (EVERY' [dtac @{thm evaln_max2}, assume_tac @{context},
- REPEAT o eresolve_tac [exE, conjE]]) *})
-apply (tactic {* ALLGOALS (rtac exI THEN' resolve_tac @{thms evaln.intros} THEN_ALL_NEW atac) *})
+ REPEAT o eresolve_tac @{context} [exE, conjE]]) *})
+apply (tactic
+ {* ALLGOALS (rtac exI THEN' resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW atac) *})
done
lemma eval_eq: "<c,s> -c-> t = (? n. <c,s> -n-> t)"
--- a/src/HOL/IOA/Solve.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/IOA/Solve.thy Tue Feb 10 14:48:26 2015 +0100
@@ -145,7 +145,7 @@
apply force
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
+ REPEAT((resolve_tac @{context} [conjI, impI] 1 ORELSE etac conjE 1) THEN
asm_full_simp_tac(@{context} addsimps [@{thm comp1_reachable}, @{thm comp2_reachable}]) 1) *})
done
--- a/src/HOL/Library/Countable.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Library/Countable.thy Tue Feb 10 14:48:26 2015 +0100
@@ -190,7 +190,7 @@
[rtac @{thm countable_datatype} i,
rtac typedef_thm i,
etac induct_thm' i,
- REPEAT (resolve_tac rules i ORELSE atac i)]) 1
+ REPEAT (resolve_tac ctxt rules i ORELSE atac i)]) 1
end)
*}
--- a/src/HOL/Library/Old_SMT/old_z3_proof_methods.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Library/Old_SMT/old_z3_proof_methods.ML Tue Feb 10 14:48:26 2015 +0100
@@ -72,7 +72,7 @@
in
Goal.init (mk_inj_prop ctxt' (Thm.dest_arg ct))
|> apply (rtac @{thm injI})
- |> apply (Tactic.solve_tac [rule, rule RS @{thm sym}])
+ |> apply (Tactic.solve_tac ctxt' [rule, rule RS @{thm sym}])
|> Goal.norm_result ctxt' o Goal.finish ctxt'
|> singleton (Variable.export ctxt' ctxt)
end
--- a/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML Tue Feb 10 14:48:26 2015 +0100
@@ -636,10 +636,10 @@
fun skolemize vars =
apfst Thm oo close vars (yield_singleton Assumption.add_assumes)
-fun discharge_sk_tac i st = (
+fun discharge_sk_tac ctxt i st = (
rtac @{thm trans} i
- THEN resolve_tac sk_rules i
- THEN (rtac @{thm refl} ORELSE' discharge_sk_tac) (i+1)
+ THEN resolve_tac ctxt sk_rules i
+ THEN (rtac @{thm refl} ORELSE' discharge_sk_tac ctxt) (i+1)
THEN rtac @{thm refl} i) st
end
@@ -847,13 +847,13 @@
fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl},
@{thm reflexive}, Old_Z3_Proof_Literals.true_thm]
- fun discharge_assms_tac rules =
- REPEAT (HEADGOAL (resolve_tac rules ORELSE' SOLVED' discharge_sk_tac))
+ fun discharge_assms_tac ctxt rules =
+ REPEAT (HEADGOAL (resolve_tac ctxt rules ORELSE' SOLVED' (discharge_sk_tac ctxt)))
fun discharge_assms ctxt rules thm =
if Thm.nprems_of thm = 0 then Goal.norm_result ctxt thm
else
- (case Seq.pull (discharge_assms_tac rules thm) of
+ (case Seq.pull (discharge_assms_tac ctxt rules thm) of
SOME (thm', _) => Goal.norm_result ctxt thm'
| NONE => raise THM ("failed to discharge premise", 1, [thm]))
--- a/src/HOL/Library/bnf_lfp_countable.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Library/bnf_lfp_countable.ML Tue Feb 10 14:48:26 2015 +0100
@@ -25,7 +25,7 @@
fun nchotomy_tac nchotomy =
HEADGOAL (rtac (nchotomy RS @{thm all_reg[rotated]}) THEN'
- REPEAT_ALL_NEW (resolve_tac [allI, impI] ORELSE' eresolve_tac [exE, disjE]));
+ REPEAT_ALL_NEW (resolve0_tac [allI, impI] ORELSE' eresolve0_tac [exE, disjE]));
fun meta_spec_mp_tac 0 = K all_tac
| meta_spec_mp_tac depth =
@@ -43,7 +43,7 @@
HEADGOAL (asm_full_simp_tac
(ss_only (injects @ recs @ map_congs' @ same_ctr_simps) ctxt) THEN_MAYBE'
TRY o REPEAT_ALL_NEW (rtac conjI) THEN_ALL_NEW
- REPEAT_ALL_NEW (eresolve_tac (conjE :: inj_map_strongs')) THEN_ALL_NEW
+ REPEAT_ALL_NEW (eresolve_tac ctxt (conjE :: inj_map_strongs')) THEN_ALL_NEW
(atac ORELSE' use_induction_hypothesis_tac));
fun distinct_ctrs_tac ctxt recs =
@@ -64,7 +64,7 @@
fun endgame_tac ctxt encode_injectives =
unfold_thms_tac ctxt @{thms inj_on_def ball_UNIV} THEN
- ALLGOALS (rtac exI THEN' rtac allI THEN' resolve_tac encode_injectives);
+ ALLGOALS (rtac exI THEN' rtac allI THEN' resolve_tac ctxt encode_injectives);
fun encode_sumN n k t =
Balanced_Tree.access {init = t,
@@ -188,6 +188,6 @@
end;
fun countable_datatype_tac ctxt =
- TRY (Class.intro_classes_tac []) THEN core_countable_datatype_tac ctxt;
+ TRY (Class.intro_classes_tac ctxt []) THEN core_countable_datatype_tac ctxt;
end;
--- a/src/HOL/MicroJava/J/Example.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/MicroJava/J/Example.thy Tue Feb 10 14:48:26 2015 +0100
@@ -372,34 +372,34 @@
apply (auto simp add: appl_methds_foo_Base)
done
-ML {* val t = resolve_tac @{thms ty_expr_ty_exprs_wt_stmt.intros} 1 *}
+lemmas t = ty_expr_ty_exprs_wt_stmt.intros
schematic_lemma wt_test: "(tprg, empty(e\<mapsto>Class Base))\<turnstile>
Expr(e::=NewC Ext);; Expr({Base}LAcc e..foo({?pTs'}[Lit Null]))\<surd>"
-apply (tactic t) -- ";;"
-apply (tactic t) -- "Expr"
-apply (tactic t) -- "LAss"
+apply (rule ty_expr_ty_exprs_wt_stmt.intros) -- ";;"
+apply (rule t) -- "Expr"
+apply (rule t) -- "LAss"
apply simp -- {* @{text "e \<noteq> This"} *}
-apply (tactic t) -- "LAcc"
+apply (rule t) -- "LAcc"
apply (simp (no_asm))
apply (simp (no_asm))
-apply (tactic t) -- "NewC"
+apply (rule t) -- "NewC"
apply (simp (no_asm))
apply (simp (no_asm))
-apply (tactic t) -- "Expr"
-apply (tactic t) -- "Call"
-apply (tactic t) -- "LAcc"
+apply (rule t) -- "Expr"
+apply (rule t) -- "Call"
+apply (rule t) -- "LAcc"
apply (simp (no_asm))
apply (simp (no_asm))
-apply (tactic t) -- "Cons"
-apply (tactic t) -- "Lit"
+apply (rule t) -- "Cons"
+apply (rule t) -- "Lit"
apply (simp (no_asm))
-apply (tactic t) -- "Nil"
+apply (rule t) -- "Nil"
apply (simp (no_asm))
apply (rule max_spec_foo_Base)
done
-ML {* val e = resolve_tac (@{thm NewCI} :: @{thms eval_evals_exec.intros}) 1 *}
+lemmas e = NewCI eval_evals_exec.intros
declare split_if [split del]
declare init_vars_def [simp] c_hupd_def [simp] cast_ok_def [simp]
schematic_lemma exec_test:
@@ -407,37 +407,37 @@
tprg\<turnstile>s0 -test-> ?s"
apply (unfold test_def)
-- "?s = s3 "
-apply (tactic e) -- ";;"
-apply (tactic e) -- "Expr"
-apply (tactic e) -- "LAss"
-apply (tactic e) -- "NewC"
+apply (rule e) -- ";;"
+apply (rule e) -- "Expr"
+apply (rule e) -- "LAss"
+apply (rule e) -- "NewC"
apply force
apply force
apply (simp (no_asm))
apply (erule thin_rl)
-apply (tactic e) -- "Expr"
-apply (tactic e) -- "Call"
-apply (tactic e) -- "LAcc"
+apply (rule e) -- "Expr"
+apply (rule e) -- "Call"
+apply (rule e) -- "LAcc"
apply force
-apply (tactic e) -- "Cons"
-apply (tactic e) -- "Lit"
-apply (tactic e) -- "Nil"
+apply (rule e) -- "Cons"
+apply (rule e) -- "Lit"
+apply (rule e) -- "Nil"
apply (simp (no_asm))
apply (force simp add: foo_Ext_def)
apply (simp (no_asm))
-apply (tactic e) -- "Expr"
-apply (tactic e) -- "FAss"
-apply (tactic e) -- "Cast"
-apply (tactic e) -- "LAcc"
+apply (rule e) -- "Expr"
+apply (rule e) -- "FAss"
+apply (rule e) -- "Cast"
+apply (rule e) -- "LAcc"
apply (simp (no_asm))
apply (simp (no_asm))
apply (simp (no_asm))
-apply (tactic e) -- "XcptE"
+apply (rule e) -- "XcptE"
apply (simp (no_asm))
apply (rule surjective_pairing [symmetric, THEN[2]trans], subst Pair_eq, force)
apply (simp (no_asm))
apply (simp (no_asm))
-apply (tactic e) -- "XcptE"
+apply (rule e) -- "XcptE"
done
end
--- a/src/HOL/MicroJava/J/JTypeSafe.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/MicroJava/J/JTypeSafe.thy Tue Feb 10 14:48:26 2015 +0100
@@ -199,8 +199,8 @@
-- "several simplifications, XcptE, XcptEs, XcptS, Skip, Nil??"
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}]
+apply( tactic "ALLGOALS (REPEAT o resolve_tac @{context} [impI, allI])")
+apply( tactic {* ALLGOALS (eresolve_tac @{context} [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
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 @{context}])")
@@ -240,7 +240,7 @@
apply( fast elim: conforms_localD [THEN lconfD])
-- "for FAss"
-apply( tactic {* EVERY'[eresolve_tac [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
+apply( tactic {* EVERY'[eresolve_tac @{context} [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
THEN_ALL_NEW (full_simp_tac @{context}), REPEAT o (etac conjE), hyp_subst_tac @{context}] 3*})
-- "for if"
@@ -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}]
+apply( tactic {* (eresolve_tac @{context} [@{thm ty_expr.cases}, @{thm ty_exprs.cases}, @{thm wt_stmt.cases}]
THEN_ALL_NEW (full_simp_tac @{context})) 1 *})
apply (intro strip)
apply (case_tac E)
--- a/src/HOL/NSA/transfer.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/NSA/transfer.ML Tue Feb 10 14:48:26 2015 +0100
@@ -63,8 +63,8 @@
rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN
ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN
match_tac ctxt [transitive_thm] 1 THEN
- resolve_tac [@{thm transfer_start}] 1 THEN
- REPEAT_ALL_NEW (resolve_tac intros) 1 THEN
+ resolve_tac ctxt [@{thm transfer_start}] 1 THEN
+ REPEAT_ALL_NEW (resolve_tac ctxt intros) 1 THEN
match_tac ctxt [reflexive_thm] 1
in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end
--- a/src/HOL/NanoJava/Equivalence.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/NanoJava/Equivalence.thy Tue Feb 10 14:48:26 2015 +0100
@@ -104,7 +104,7 @@
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}]) *})
+apply (tactic {* ALLGOALS (REPEAT o dresolve_tac @{context} [@{thm all_conjunct2}, @{thm all3_conjunct2}]) *})
apply (tactic {* ALLGOALS (REPEAT o thin_tac @{context} "hoare ?x ?y") *})
apply (tactic {* ALLGOALS (REPEAT o thin_tac @{context} "ehoare ?x ?y") *})
apply (simp_all only: cnvalid1_eq cenvalid_def2)
--- a/src/HOL/Nominal/nominal_atoms.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Nominal/nominal_atoms.ML Tue Feb 10 14:48:26 2015 +0100
@@ -503,20 +503,20 @@
val cls_name = Sign.full_bname thy' ("pt_"^ak_name);
val at_inst = Global_Theory.get_thm thy' ("at_" ^ ak_name' ^ "_inst");
- val proof1 = EVERY [Class.intro_classes_tac [],
+ fun proof1 ctxt = EVERY [Class.intro_classes_tac ctxt [],
rtac ((at_inst RS at_pt_inst) RS pt1) 1,
rtac ((at_inst RS at_pt_inst) RS pt2) 1,
rtac ((at_inst RS at_pt_inst) RS pt3) 1,
atac 1];
fun proof2 ctxt =
- Class.intro_classes_tac [] THEN
+ Class.intro_classes_tac ctxt [] 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 ctxt => if ak_name = ak_name' then proof1 else proof2 ctxt)
+ (fn ctxt => if ak_name = ak_name' then proof1 ctxt else proof2 ctxt)
end) ak_names thy) ak_names thy12d;
(* show that *)
@@ -536,7 +536,7 @@
val pt_inst = Global_Theory.get_thm thy ("pt_"^ak_name^"_inst");
fun pt_proof thm ctxt =
- EVERY [Class.intro_classes_tac [],
+ EVERY [Class.intro_classes_tac ctxt [],
rtac (thm RS pt1) 1, rtac (thm RS pt2) 1, rtac (thm RS pt3) 1, atac 1];
val pt_thm_fun = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst));
@@ -582,13 +582,13 @@
(if ak_name = ak_name'
then
let val at_thm = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst");
- in EVERY [Class.intro_classes_tac [],
+ in EVERY [Class.intro_classes_tac ctxt [],
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 =
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 EVERY [Class.intro_classes_tac ctxt [], asm_simp_tac simp_s 1] end)
in
Axclass.prove_arity (qu_name,[],[qu_class]) proof thy'
end) ak_names thy) ak_names thy18;
@@ -605,7 +605,7 @@
let
val cls_name = Sign.full_bname thy ("fs_"^ak_name);
val fs_inst = Global_Theory.get_thm thy ("fs_"^ak_name^"_inst");
- fun fs_proof thm ctxt = EVERY [Class.intro_classes_tac [], rtac (thm RS fs1) 1];
+ fun fs_proof thm ctxt = EVERY [Class.intro_classes_tac ctxt [], rtac (thm RS fs1) 1];
val fs_thm_unit = fs_unit_inst;
val fs_thm_prod = fs_inst RS (fs_inst RS fs_prod_inst);
@@ -651,7 +651,7 @@
val pt_inst = Global_Theory.get_thm thy'' ("pt_"^ak_name''^"_inst");
val at_inst = Global_Theory.get_thm thy'' ("at_"^ak_name''^"_inst");
in
- EVERY [Class.intro_classes_tac [],
+ EVERY [Class.intro_classes_tac ctxt [],
rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1]
end)
else
@@ -663,7 +663,7 @@
[ak_name' ^"_prm_"^ak_name^"_def",
ak_name''^"_prm_"^ak_name^"_def"]));
in
- EVERY [Class.intro_classes_tac [], simp_tac simp_s 1]
+ EVERY [Class.intro_classes_tac ctxt [], simp_tac simp_s 1]
end))
in
Axclass.prove_arity (name,[],[cls_name]) proof thy''
@@ -686,7 +686,7 @@
val pt_inst = Global_Theory.get_thm thy' ("pt_"^ak_name^"_inst");
val at_inst = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst");
- fun cp_proof thm ctxt = EVERY [Class.intro_classes_tac [],rtac (thm RS cp1) 1];
+ fun cp_proof thm ctxt = EVERY [Class.intro_classes_tac ctxt [], rtac (thm RS cp1) 1];
val cp_thm_unit = cp_unit_inst;
val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst);
@@ -717,7 +717,7 @@
let
val qu_class = Sign.full_bname thy ("pt_"^ak_name);
fun proof ctxt =
- Class.intro_classes_tac [] THEN
+ Class.intro_classes_tac ctxt [] 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]) proof thy
@@ -729,7 +729,7 @@
val qu_class = Sign.full_bname thy ("fs_"^ak_name);
val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"};
fun proof ctxt =
- Class.intro_classes_tac [] THEN
+ Class.intro_classes_tac ctxt [] THEN
asm_simp_tac (put_simpset HOL_ss ctxt
addsimps [supp_def, Collect_const, finite_emptyI, Simpdata.mk_eq defn]) 1;
in
@@ -742,7 +742,7 @@
val qu_class = Sign.full_bname thy ("cp_"^ak_name^"_"^ak_name');
val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"};
fun proof ctxt =
- Class.intro_classes_tac [] THEN
+ Class.intro_classes_tac ctxt [] THEN
asm_simp_tac (put_simpset HOL_ss ctxt addsimps [Simpdata.mk_eq defn]) 1;
in
Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy
--- a/src/HOL/Nominal/nominal_datatype.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML Tue Feb 10 14:48:26 2015 +0100
@@ -427,7 +427,8 @@
in
fold (fn (s, tvs) => fn thy => Axclass.prove_arity
(s, map (inter_sort thy sort o snd) tvs, [cp_class])
- (fn _ => Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
+ (fn ctxt' =>
+ Class.intro_classes_tac ctxt' [] THEN ALLGOALS (resolve_tac ctxt' thms)) thy)
(full_new_type_names' ~~ tyvars) thy
end;
@@ -439,10 +440,10 @@
fold (fn (s, tvs) => fn thy => Axclass.prove_arity
(s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
(fn ctxt => EVERY
- [Class.intro_classes_tac [],
- resolve_tac perm_empty_thms 1,
- resolve_tac perm_append_thms 1,
- resolve_tac perm_eq_thms 1, assume_tac ctxt 1]) thy)
+ [Class.intro_classes_tac ctxt [],
+ resolve_tac ctxt perm_empty_thms 1,
+ resolve_tac ctxt perm_append_thms 1,
+ resolve_tac ctxt perm_eq_thms 1, assume_tac ctxt 1]) thy)
(full_new_type_names' ~~ tyvars) thy
end) atoms |>
Global_Theory.add_thmss
@@ -562,7 +563,7 @@
[Old_Datatype_Aux.ind_tac rep_induct [] 1,
ALLGOALS (simp_tac (ctxt addsimps
(Thm.symmetric perm_fun_def :: abs_perm))),
- ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac ctxt)])),
+ ALLGOALS (resolve_tac ctxt rep_intrs THEN_ALL_NEW assume_tac ctxt)])),
length new_type_names));
val perm_closed_thmss = map mk_perm_closed atoms;
@@ -578,9 +579,9 @@
(name, map (fn (v, _) => (v, dummyS)) tvs, mx) (* FIXME keep constraints!? *)
(Const (@{const_name Collect}, (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
Const (cname, U --> HOLogic.boolT)) NONE
- (fn _ => rtac exI 1 THEN rtac CollectI 1 THEN
+ (fn ctxt => rtac exI 1 THEN rtac CollectI 1 THEN
QUIET_BREADTH_FIRST (has_fewer_prems 1)
- (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
+ (resolve_tac ctxt rep_intrs 1)) thy |> (fn ((_, r), thy) =>
let
val permT = mk_permT
(TFree (singleton (Name.variant_list (map fst tvs)) "'a", @{sort type}));
@@ -619,7 +620,7 @@
in Axclass.prove_arity
(Sign.intern_type thy name,
map (inter_sort thy sort o snd) tvs, [pt_class])
- (fn ctxt => EVERY [Class.intro_classes_tac [],
+ (fn ctxt => EVERY [Class.intro_classes_tac ctxt [],
rewrite_goals_tac ctxt [perm_def],
asm_full_simp_tac (ctxt addsimps [Rep_inverse]) 1,
asm_full_simp_tac (ctxt addsimps
@@ -648,7 +649,7 @@
Axclass.prove_arity
(Sign.intern_type thy name,
map (inter_sort thy sort o snd) tvs, [cp_class])
- (fn ctxt => EVERY [Class.intro_classes_tac [],
+ (fn ctxt => EVERY [Class.intro_classes_tac ctxt [],
rewrite_goals_tac ctxt [perm_def],
asm_full_simp_tac (ctxt addsimps
((Rep RS perm_closed1 RS Abs_inverse) ::
@@ -801,11 +802,11 @@
val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
in Goal.prove_global_future thy8 [] [] eqn (fn {context = ctxt, ...} => EVERY
- [resolve_tac inj_thms 1,
+ [resolve_tac ctxt inj_thms 1,
rewrite_goals_tac ctxt rewrites,
rtac refl 3,
- resolve_tac rep_intrs 2,
- REPEAT (resolve_tac Rep_thms 1)])
+ resolve_tac ctxt rep_intrs 2,
+ REPEAT (resolve_tac ctxt Rep_thms 1)])
end;
val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
@@ -1041,7 +1042,7 @@
REPEAT (EVERY
[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])]);
+ etac mp 1, resolve_tac ctxt Rep_thms 1])]);
val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
@@ -1058,7 +1059,7 @@
[rtac indrule_lemma' 1,
(Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms' 1),
+ [REPEAT (eresolve_tac ctxt Abs_inverse_thms' 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))]);
@@ -1116,7 +1117,7 @@
val sort = Sign.minimize_sort thy (Sign.certify_sort thy (class :: pt_cp_sort));
in fold (fn Type (s, Ts) => Axclass.prove_arity
(s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
- (fn _ => Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
+ (fn ctxt => Class.intro_classes_tac ctxt [] THEN resolve_tac ctxt ths 1)) newTs thy
end) (atoms ~~ finite_supp_thms);
(**** strong induction theorem ****)
@@ -1243,7 +1244,7 @@
fresh_const T (fastype_of p) $
Bound 0 $ p)))
(fn _ => EVERY
- [resolve_tac exists_fresh' 1,
+ [resolve_tac ctxt exists_fresh' 1,
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
@@ -1354,22 +1355,22 @@
val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
[simp_tac (ind_ss6 addsimps rename_eq) 1,
cut_facts_tac iprems 1,
- (resolve_tac prems THEN_ALL_NEW
+ (resolve_tac context2 prems THEN_ALL_NEW
SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
_ $ (Const (@{const_name Nominal.fresh}, _) $ _ $ _) =>
simp_tac ind_ss1' i
| _ $ (Const (@{const_name Not}, _) $ _) =>
- resolve_tac freshs2' i
+ resolve_tac context2 freshs2' i
| _ => 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
- resolve_tac final 1
+ resolve_tac context2 final 1
end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
in
EVERY
[cut_facts_tac [th] 1,
- REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1),
+ REPEAT (eresolve_tac context [conjE, @{thm allE_Nil}] 1),
REPEAT (etac allE 1),
REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
end);
@@ -1390,8 +1391,8 @@
(augment_sort thy9 fs_cp_sort ind_concl)
(fn {prems, context = ctxt} => EVERY
[rtac induct_aux' 1,
- REPEAT (resolve_tac fs_atoms 1),
- REPEAT ((resolve_tac prems THEN_ALL_NEW
+ REPEAT (resolve_tac ctxt fs_atoms 1),
+ REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW
(etac @{thm meta_spec} ORELSE'
full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [fresh_def]))) 1)])
@@ -1538,7 +1539,7 @@
(simp_tac (put_simpset HOL_basic_ss ctxt
addsimps flat perm_simps'
addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
- (resolve_tac rec_intrs THEN_ALL_NEW
+ (resolve_tac ctxt rec_intrs THEN_ALL_NEW
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 [] []
@@ -1630,7 +1631,7 @@
supports_fresh) 1,
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 (resolve_tac context [allI, impI] 1),
REPEAT_DETERM (etac conjE 1),
rtac unique 1,
SUBPROOF (fn {prems = prems', params = [(_, a), (_, b)], ...} => EVERY
@@ -1683,8 +1684,8 @@
fresh_const T (fastype_of p) $ Bound 0 $ p)))
(fn _ => EVERY
[cut_facts_tac ths 1,
- REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
- resolve_tac exists_fresh' 1,
+ REPEAT_DETERM (dresolve_tac ctxt (the (AList.lookup op = rec_fin_supp T)) 1),
+ resolve_tac ctxt exists_fresh' 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
@@ -1723,7 +1724,7 @@
rec_sets ~~ rec_preds)))))
(fn {context = ctxt, ...} =>
rtac rec_induct 1 THEN
- REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac ctxt) 1))));
+ REPEAT ((resolve_tac ctxt P_ind_ths THEN_ALL_NEW assume_tac ctxt) 1))));
val rec_fin_supp_thms' = map
(fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
(rec_fin_supp_thms ~~ finite_thss);
@@ -1735,9 +1736,9 @@
(finite_thss ~~ finite_ctxt_ths) @
maps (fn ((_, idxss), elim) => maps (fn idxs =>
[full_simp_tac (put_simpset HOL_ss context addsimps [Thm.symmetric fresh_def, supp_prod, Un_iff]) 1,
- REPEAT_DETERM (eresolve_tac @{thms conjE ex1E} 1),
+ REPEAT_DETERM (eresolve_tac context @{thms conjE ex1E} 1),
rtac @{thm ex1I} 1,
- (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
+ (resolve_tac context rec_intrs THEN_ALL_NEW atac) 1,
rotate_tac ~1 1,
((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
(put_simpset HOL_ss context addsimps flat distinct_thms)) 1] @
@@ -1907,7 +1908,7 @@
map (fn th => rtac th 1)
(snd (nth finite_thss l)) @
[rtac rec_prem 1, rtac ih 1,
- REPEAT_DETERM (resolve_tac fresh_prems 1)]))
+ REPEAT_DETERM (resolve_tac context fresh_prems 1)]))
end) atoms
end) (rec_prems1 ~~ ihs);
@@ -1917,13 +1918,13 @@
Goal.prove context'' [] []
(HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
(fn _ => EVERY
- [resolve_tac fcbs 1,
- REPEAT_DETERM (resolve_tac
+ [resolve_tac context'' fcbs 1,
+ REPEAT_DETERM (resolve_tac context''
(fresh_prems @ rec_freshs) 1),
- REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
- THEN resolve_tac rec_prems 1),
- resolve_tac P_ind_ths 1,
- REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
+ REPEAT_DETERM (resolve_tac context'' (maps snd rec_fin_supp_thms') 1
+ THEN resolve_tac context'' rec_prems 1),
+ resolve_tac context'' P_ind_ths 1,
+ REPEAT_DETERM (resolve_tac context'' (P_ths @ rec_prems) 1)]);
val fresh_results'' = map prove_fresh_result boundsl;
@@ -1956,7 +1957,7 @@
(HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
(fn _ => EVERY
[cut_facts_tac recs 1,
- REPEAT_DETERM (dresolve_tac
+ REPEAT_DETERM (dresolve_tac context''
(the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
NominalPermeq.fresh_guess_tac
(put_simpset HOL_ss context'' addsimps (freshs2 @
@@ -1986,7 +1987,7 @@
val final' = Proof_Context.export context'' context' [final];
val _ = warning "finished!"
in
- resolve_tac final' 1
+ resolve_tac context' final' 1
end) context 1])) idxss) (ndescr ~~ rec_elims))
end));
@@ -2025,8 +2026,8 @@
val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
val prems' = flat finite_premss @ finite_ctxt_prems @
rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
- fun solve rules prems = resolve_tac rules THEN_ALL_NEW
- (resolve_tac prems THEN_ALL_NEW atac)
+ fun solve ctxt rules prems = resolve_tac ctxt rules THEN_ALL_NEW
+ (resolve_tac ctxt prems THEN_ALL_NEW atac)
in
Goal.prove_global_future thy12 []
(map (augment_sort thy12 fs_cp_sort) prems')
@@ -2034,9 +2035,9 @@
(fn {context = ctxt, prems} => EVERY
[rewrite_goals_tac ctxt reccomb_defs,
rtac @{thm the1_equality} 1,
- solve rec_unique_thms prems 1,
- resolve_tac rec_intrs 1,
- REPEAT (solve (prems @ rec_total_thms) prems 1)])
+ solve ctxt rec_unique_thms prems 1,
+ resolve_tac ctxt rec_intrs 1,
+ REPEAT (solve ctxt (prems @ rec_total_thms) prems 1)])
end) (rec_eq_prems ~~
Old_Datatype_Prop.make_primrecs reccomb_names descr' thy12);
--- a/src/HOL/Nominal/nominal_inductive.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Nominal/nominal_inductive.ML Tue Feb 10 14:48:26 2015 +0100
@@ -130,7 +130,7 @@
fun prove t =
Goal.prove ctxt [] [] t (fn _ =>
EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
- REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
+ REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)),
REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
in Option.map prove (map_term f prop (the_default prop opt)) end;
@@ -296,8 +296,8 @@
NominalDatatype.fresh_const T (fastype_of p) $
Bound 0 $ p)))
(fn _ => EVERY
- [resolve_tac exists_fresh' 1,
- resolve_tac fs_atoms 1]);
+ [resolve_tac ctxt exists_fresh' 1,
+ resolve_tac ctxt fs_atoms 1]);
val (([(_, cx)], ths), ctxt') = Obtain.result
(fn ctxt' => EVERY
[etac exE 1,
@@ -388,17 +388,17 @@
(simp_tac (put_simpset HOL_basic_ss ctxt''
addsimps [inductive_forall_def']
addsimprocs [NominalDatatype.perm_simproc]) 1 THEN
- resolve_tac gprems2 1)]));
+ resolve_tac ctxt'' gprems2 1)]));
val final = Goal.prove ctxt'' [] [] (term_of concl)
(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];
- in resolve_tac final' 1 end) context 1])
+ in resolve_tac ctxt' final' 1 end) context 1])
(prems ~~ thss ~~ ihyps ~~ prems'')))
in
cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN
- REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN
+ REPEAT (REPEAT (resolve_tac ctxt [conjI, impI] 1) THEN
etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN
asm_full_simp_tac ctxt 1)
end) |> singleton (Proof_Context.export ctxt' ctxt);
@@ -532,10 +532,10 @@
in
simp_tac case_simpset 1 THEN
REPEAT_DETERM (TRY (rtac conjI 1) THEN
- resolve_tac case_hyps' 1)
+ resolve_tac ctxt4 case_hyps' 1)
end) ctxt4 1)
val final = Proof_Context.export ctxt3 ctxt2 [th]
- in resolve_tac final 1 end) ctxt1 1)
+ in resolve_tac ctxt2 final 1 end) ctxt1 1)
(thss ~~ hyps ~~ prems))) |>
singleton (Proof_Context.export ctxt' ctxt))
@@ -634,7 +634,7 @@
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)
intr
- in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1
+ in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac ctxt'' prems'')) 1
end) ctxt' 1 st
in
case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of
--- a/src/HOL/Nominal/nominal_inductive2.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML Tue Feb 10 14:48:26 2015 +0100
@@ -135,7 +135,7 @@
fun prove t =
Goal.prove ctxt [] [] t (fn _ =>
EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
- REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
+ REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)),
REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
in Option.map prove (map_term f prop (the_default prop opt)) end;
@@ -424,17 +424,17 @@
(simp_tac (put_simpset HOL_basic_ss ctxt''
addsimps [inductive_forall_def']
addsimprocs [NominalDatatype.perm_simproc]) 1 THEN
- resolve_tac gprems2 1)]));
+ resolve_tac ctxt'' gprems2 1)]));
val final = Goal.prove ctxt'' [] [] (term_of concl)
(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];
- in resolve_tac final' 1 end) context 1])
+ in resolve_tac ctxt' final' 1 end) context 1])
(prems ~~ thss ~~ vc_compat' ~~ ihyps ~~ prems'')))
in
cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN
- REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN
+ REPEAT (REPEAT (resolve_tac ctxt' [conjI, impI] 1) THEN
etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN
asm_full_simp_tac ctxt 1)
end) |>
--- a/src/HOL/Old_Number_Theory/Chinese.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Old_Number_Theory/Chinese.thy Tue Feb 10 14:48:26 2015 +0100
@@ -166,11 +166,11 @@
"0 < n ==> i \<le> n ==> m_cond n mf ==> km_cond n kf mf
==> \<exists>!x. 0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)"
apply (rule zcong_lineq_unique)
- apply (tactic {* stac @{thm zgcd_zmult_cancel} 2 *})
+ apply (tactic {* stac @{context} @{thm zgcd_zmult_cancel} 2 *})
apply (unfold m_cond_def km_cond_def mhf_def)
apply (simp_all (no_asm_simp))
apply safe
- apply (tactic {* stac @{thm zgcd_zmult_cancel} 3 *})
+ apply (tactic {* stac @{context} @{thm zgcd_zmult_cancel} 3 *})
apply (rule_tac [!] funprod_zgcd)
apply safe
apply simp_all
@@ -228,12 +228,12 @@
apply (rule_tac x = "x_sol n kf bf mf mod funprod mf 0 n" in exI)
apply (unfold lincong_sol_def)
apply safe
- apply (tactic {* stac @{thm zcong_zmod} 3 *})
- apply (tactic {* stac @{thm mod_mult_eq} 3 *})
- apply (tactic {* stac @{thm mod_mod_cancel} 3 *})
- apply (tactic {* stac @{thm x_sol_lin} 4 *})
- apply (tactic {* stac (@{thm mod_mult_eq} RS sym) 6 *})
- apply (tactic {* stac (@{thm zcong_zmod} RS sym) 6 *})
+ apply (tactic {* stac @{context} @{thm zcong_zmod} 3 *})
+ apply (tactic {* stac @{context} @{thm mod_mult_eq} 3 *})
+ apply (tactic {* stac @{context} @{thm mod_mod_cancel} 3 *})
+ apply (tactic {* stac @{context} @{thm x_sol_lin} 4 *})
+ apply (tactic {* stac @{context} (@{thm mod_mult_eq} RS sym) 6 *})
+ apply (tactic {* stac @{context} (@{thm zcong_zmod} RS sym) 6 *})
apply (subgoal_tac [6]
"0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i
\<and> [kf i * mhf mf n i * xilin_sol i n kf bf mf = bf i] (mod mf i)")
--- a/src/HOL/Old_Number_Theory/IntPrimes.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Old_Number_Theory/IntPrimes.thy Tue Feb 10 14:48:26 2015 +0100
@@ -399,7 +399,7 @@
zgcd a n = 1 ==> \<exists>!x. 0 \<le> x \<and> x < n \<and> [a * x = b] (mod n)"
apply auto
apply (rule_tac [2] zcong_zless_imp_eq)
- apply (tactic {* stac (@{thm zcong_cancel2} RS sym) 6 *})
+ apply (tactic {* stac @{context} (@{thm zcong_cancel2} RS sym) 6 *})
apply (rule_tac [8] zcong_trans)
apply (simp_all (no_asm_simp))
prefer 2
--- a/src/HOL/Old_Number_Theory/WilsonBij.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Old_Number_Theory/WilsonBij.thy Tue Feb 10 14:48:26 2015 +0100
@@ -139,9 +139,9 @@
apply (unfold inj_on_def)
apply auto
apply (rule zcong_zless_imp_eq)
- apply (tactic {* stac (@{thm zcong_cancel} RS sym) 5 *})
+ apply (tactic {* stac @{context} (@{thm zcong_cancel} RS sym) 5 *})
apply (rule_tac [7] zcong_trans)
- apply (tactic {* stac @{thm zcong_sym} 8 *})
+ apply (tactic {* stac @{context} @{thm zcong_sym} 8 *})
apply (erule_tac [7] inv_is_inv)
apply (tactic "asm_simp_tac @{context} 9")
apply (erule_tac [9] inv_is_inv)
@@ -192,15 +192,15 @@
apply (unfold reciR_def uniqP_def)
apply auto
apply (rule zcong_zless_imp_eq)
- apply (tactic {* stac (@{thm zcong_cancel2} RS sym) 5 *})
+ apply (tactic {* stac @{context} (@{thm zcong_cancel2} RS sym) 5 *})
apply (rule_tac [7] zcong_trans)
- apply (tactic {* stac @{thm zcong_sym} 8 *})
+ apply (tactic {* stac @{context} @{thm zcong_sym} 8 *})
apply (rule_tac [6] zless_zprime_imp_zrelprime)
apply auto
apply (rule zcong_zless_imp_eq)
- apply (tactic {* stac (@{thm zcong_cancel} RS sym) 5 *})
+ apply (tactic {* stac @{context} (@{thm zcong_cancel} RS sym) 5 *})
apply (rule_tac [7] zcong_trans)
- apply (tactic {* stac @{thm zcong_sym} 8 *})
+ apply (tactic {* stac @{context} @{thm zcong_sym} 8 *})
apply (rule_tac [6] zless_zprime_imp_zrelprime)
apply auto
done
--- a/src/HOL/Old_Number_Theory/WilsonRuss.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy Tue Feb 10 14:48:26 2015 +0100
@@ -252,7 +252,7 @@
apply (subst wset.simps)
apply (auto, unfold Let_def, auto)
apply (subst setprod.insert)
- apply (tactic {* stac @{thm setprod.insert} 3 *})
+ apply (tactic {* stac @{context} @{thm setprod.insert} 3 *})
apply (subgoal_tac [5]
"zcong (a * inv p a * (\<Prod>x\<in>wset (a - 1) p. x)) (1 * 1) p")
prefer 5
--- a/src/HOL/Probability/measurable.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Probability/measurable.ML Tue Feb 10 14:48:26 2015 +0100
@@ -208,9 +208,9 @@
fun r_tac msg =
if Config.get ctxt debug
then FIRST' o
- map (fn thm => resolve_tac [thm]
+ map (fn thm => resolve_tac ctxt [thm]
THEN' K (debug_tac ctxt (debug_fact (msg ^ " resolved using") thm) all_tac))
- else resolve_tac
+ else resolve_tac ctxt
val elem_congI = @{lemma "A = B \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A" by simp}
@@ -249,7 +249,8 @@
val f = dest_measurable_fun (HOLogic.dest_Trueprop t)
fun cert f = map (Option.map (f (Proof_Context.theory_of ctxt)))
fun inst (ts, Ts) =
- Drule.instantiate' (cert ctyp_of Ts) (cert cterm_of ts) @{thm measurable_compose_countable}
+ Drule.instantiate' (cert ctyp_of Ts) (cert cterm_of ts)
+ @{thm measurable_compose_countable}
in r_tac "split countable" (cnt_prefixes ctxt f |> map inst) i end
handle TERM _ => no_tac) 1)
--- a/src/HOL/Product_Type.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Product_Type.thy Tue Feb 10 14:48:26 2015 +0100
@@ -1327,10 +1327,10 @@
SOME (Goal.prove ctxt [] []
(Const (@{const_name Pure.eq}, T --> T --> propT) $ S $ S')
(K (EVERY
- [resolve_tac [eq_reflection] 1,
- resolve_tac @{thms subset_antisym} 1,
- resolve_tac [subsetI] 1, dresolve_tac [CollectD] 1, simp,
- resolve_tac [subsetI] 1, resolve_tac [CollectI] 1, simp])))
+ [resolve_tac ctxt [eq_reflection] 1,
+ resolve_tac ctxt @{thms subset_antisym} 1,
+ resolve_tac ctxt [subsetI] 1, dresolve_tac ctxt [CollectD] 1, simp,
+ resolve_tac ctxt [subsetI] 1, resolve_tac ctxt [CollectI] 1, simp])))
end
else NONE)
| _ => NONE)
--- a/src/HOL/Prolog/prolog.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Prolog/prolog.ML Tue Feb 10 14:48:26 2015 +0100
@@ -119,8 +119,8 @@
asm_full_simp_tac (put_simpset atomize_ss ctxt) THEN' (* atomize the asms *)
(REPEAT_DETERM o (etac conjE)) (* split the asms *)
])
- ORELSE' resolve_tac [disjI1,disjI2] (* "P ==> P | Q","Q ==> P | Q"*)
- ORELSE' ((resolve_tac proga APPEND' hyp_resolve_tac)
+ ORELSE' resolve_tac ctxt [disjI1,disjI2] (* "P ==> P | Q","Q ==> P | Q"*)
+ ORELSE' ((resolve_tac ctxt proga APPEND' hyp_resolve_tac)
THEN' (fn _ => check_HOHH_tac2))
end;
--- a/src/HOL/SET_Protocol/Event_SET.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/SET_Protocol/Event_SET.thy Tue Feb 10 14:48:26 2015 +0100
@@ -184,7 +184,7 @@
{*
fun analz_mono_contra_tac ctxt =
rtac @{thm analz_impI} THEN'
- REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
+ REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
THEN' mp_tac ctxt
*}
--- a/src/HOL/SET_Protocol/Message_SET.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/SET_Protocol/Message_SET.thy Tue Feb 10 14:48:26 2015 +0100
@@ -850,13 +850,13 @@
(*Apply rules to break down assumptions of the form
Y \<in> parts(insert X H) and Y \<in> analz(insert X H)
*)
-val Fake_insert_tac =
- dresolve_tac [impOfSubs @{thm Fake_analz_insert},
+fun Fake_insert_tac ctxt =
+ dresolve_tac ctxt [impOfSubs @{thm Fake_analz_insert},
impOfSubs @{thm Fake_parts_insert}] THEN'
- eresolve_tac [asm_rl, @{thm synth.Inj}];
+ eresolve_tac ctxt [asm_rl, @{thm synth.Inj}];
fun Fake_insert_simp_tac ctxt i =
- REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ctxt i;
+ REPEAT (Fake_insert_tac ctxt i) THEN asm_full_simp_tac ctxt i;
fun atomic_spy_analz_tac ctxt =
SELECT_GOAL
@@ -874,7 +874,7 @@
(res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1),
(*...allowing further simplifications*)
simp_tac ctxt 1,
- REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
+ REPEAT (FIRSTGOAL (resolve_tac ctxt [allI,impI,notI,conjI,iffI])),
DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i);
*}
(*>*)
--- a/src/HOL/SET_Protocol/Public_SET.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/SET_Protocol/Public_SET.thy Tue Feb 10 14:48:26 2015 +0100
@@ -347,7 +347,7 @@
(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}]))
+ resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}]))
(*For harder protocols (such as SET_CR!), where we have to set up some
nonces and keys initially*)
@@ -355,7 +355,7 @@
REPEAT
(ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver))
THEN
- REPEAT_FIRST (resolve_tac [refl, conjI]))
+ REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
*}
method_setup possibility = {*
--- a/src/HOL/SPARK/Tools/spark_vcs.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML Tue Feb 10 14:48:26 2015 +0100
@@ -242,12 +242,12 @@
delsimps @{thms atLeastLessThan_iff}) 1);
val lthy' =
- Class.prove_instantiation_instance (fn _ =>
- Class.intro_classes_tac [] THEN
+ Class.prove_instantiation_instance (fn ctxt =>
+ Class.intro_classes_tac ctxt [] THEN
rtac finite_UNIV 1 THEN
rtac range_pos 1 THEN
- simp_tac (put_simpset HOL_basic_ss lthy addsimps [def3]) 1 THEN
- simp_tac (put_simpset HOL_basic_ss lthy addsimps [def2]) 1) lthy;
+ simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def3]) 1 THEN
+ simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def2]) 1) lthy;
val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) =>
let
--- a/src/HOL/Set.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Set.thy Tue Feb 10 14:48:26 2015 +0100
@@ -70,11 +70,11 @@
simproc_setup defined_Collect ("{x. P x & Q x}") = {*
fn _ => Quantifier1.rearrange_Collect
- (fn _ =>
- resolve_tac @{thms Collect_cong} 1 THEN
- resolve_tac @{thms iffI} 1 THEN
+ (fn ctxt =>
+ resolve_tac ctxt @{thms Collect_cong} 1 THEN
+ resolve_tac ctxt @{thms iffI} 1 THEN
ALLGOALS
- (EVERY' [REPEAT_DETERM o eresolve_tac @{thms conjE},
+ (EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE},
DEPTH_SOLVE_1 o ares_tac @{thms conjI}]))
*}
@@ -359,14 +359,14 @@
fn _ => Quantifier1.rearrange_bex
(fn ctxt =>
unfold_tac ctxt @{thms Bex_def} THEN
- Quantifier1.prove_one_point_ex_tac)
+ Quantifier1.prove_one_point_ex_tac ctxt)
*}
simproc_setup defined_All ("ALL x:A. P x --> Q x") = {*
fn _ => Quantifier1.rearrange_ball
(fn ctxt =>
unfold_tac ctxt @{thms Ball_def} THEN
- Quantifier1.prove_one_point_all_tac)
+ Quantifier1.prove_one_point_all_tac ctxt)
*}
lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x"
@@ -383,7 +383,7 @@
setup {*
map_theory_claset (fn ctxt =>
- ctxt addbefore ("bspec", fn ctxt' => dresolve_tac @{thms bspec} THEN' assume_tac ctxt'))
+ ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt'))
*}
ML {*
--- a/src/HOL/Statespace/state_space.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Statespace/state_space.ML Tue Feb 10 14:48:26 2015 +0100
@@ -354,7 +354,7 @@
val expr = ([(suffix valuetypesN name,
(prefix, Expression.Positional (map SOME pars)))],[]);
in
- prove_interpretation_in (ALLGOALS o solve_tac o Assumption.all_prems_of)
+ prove_interpretation_in (fn ctxt => ALLGOALS (solve_tac ctxt (Assumption.all_prems_of ctxt)))
(suffix valuetypesN name, expr) thy
end;
--- a/src/HOL/TLA/Action.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TLA/Action.thy Tue Feb 10 14:48:26 2015 +0100
@@ -263,9 +263,9 @@
*)
fun action_simp_tac ctxt intros elims =
asm_full_simp_tac
- (ctxt setloop (fn _ => (resolve_tac ((map (action_use ctxt) intros)
+ (ctxt setloop (fn _ => (resolve_tac ctxt ((map (action_use ctxt) intros)
@ [refl,impI,conjI,@{thm actionI},@{thm intI},allI]))
- ORELSE' (eresolve_tac ((map (action_use ctxt) elims)
+ ORELSE' (eresolve_tac ctxt ((map (action_use ctxt) elims)
@ [conjE,disjE,exE]))));
*}
--- a/src/HOL/TLA/Memory/MemoryImplementation.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Tue Feb 10 14:48:26 2015 +0100
@@ -800,7 +800,7 @@
(TRY (rtac @{thm actionI} 1) THEN
Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" 1 THEN
rewrite_goals_tac ctxt @{thms action_rews} THEN
- forward_tac [temp_use ctxt @{thm Step1_4_7}] 1 THEN
+ forward_tac ctxt [temp_use ctxt @{thm Step1_4_7}] 1 THEN
asm_full_simp_tac ctxt 1);
*}
--- a/src/HOL/TLA/TLA.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TLA/TLA.thy Tue Feb 10 14:48:26 2015 +0100
@@ -697,7 +697,8 @@
(* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *)
ML {*
-val box_fair_tac = SELECT_GOAL (REPEAT (dresolve_tac [@{thm BoxWFI}, @{thm BoxSFI}] 1))
+fun box_fair_tac ctxt =
+ SELECT_GOAL (REPEAT (dresolve_tac ctxt [@{thm BoxWFI}, @{thm BoxSFI}] 1))
*}
--- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Tue Feb 10 14:48:26 2015 +0100
@@ -623,7 +623,7 @@
(*Prove subgoals of form A ==> B1 | ... | A | ... | Bn*)
val clause_breaker =
- (REPEAT o (resolve_tac [@{thm "disjI1"}, @{thm "disjI2"}, @{thm "conjI"}]))
+ (REPEAT o (resolve0_tac [@{thm "disjI1"}, @{thm "disjI2"}, @{thm "conjI"}]))
THEN' atac
(*
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Tue Feb 10 14:48:26 2015 +0100
@@ -474,7 +474,7 @@
THEN' (REPEAT_DETERM o etac @{thm conjE})
THEN' (TRY o (expander_animal ctxt))
in
- default_tac ORELSE' resolve_tac @{thms flip}
+ default_tac ORELSE' resolve_tac ctxt @{thms flip}
end
*}
@@ -1154,7 +1154,7 @@
ASAP
(rtac @{thm disjI1} APPEND' rtac @{thm disjI2})
(FIRST' (map closure
- [dresolve_tac @{thms dec_commut_eq},
+ [dresolve_tac ctxt @{thms dec_commut_eq},
dtac @{thm dec_commut_disj},
elim_tac]))
in
@@ -1527,7 +1527,7 @@
val ex_free =
if loop_can_feature [Existential_Free] feats andalso consts_diff = [] then
- eresolve_tac @{thms polar_exE}
+ eresolve_tac ctxt @{thms polar_exE}
else K no_tac
in
ex_var APPEND' ex_free
@@ -1613,8 +1613,8 @@
| Not_neg => trace_tac' "mark: not_neg" (dtac @{thm leo2_rules(10)})
| Or_pos => trace_tac' "mark: or_pos" (dtac @{thm leo2_rules(5)}) (*could add (6) for negated conjunction*)
| Or_neg => trace_tac' "mark: or_neg" (dtac @{thm leo2_rules(7)})
- | Equal_pos => trace_tac' "mark: equal_pos" (dresolve_tac (@{thms eq_pos_bool} @ [@{thm leo2_rules(3)}, @{thm eq_pos_func}]))
- | Equal_neg => trace_tac' "mark: equal_neg" (dresolve_tac [@{thm eq_neg_bool}, @{thm leo2_rules(4)}])
+ | Equal_pos => trace_tac' "mark: equal_pos" (dresolve_tac ctxt (@{thms eq_pos_bool} @ [@{thm leo2_rules(3)}, @{thm eq_pos_func}]))
+ | Equal_neg => trace_tac' "mark: equal_neg" (dresolve_tac ctxt [@{thm eq_neg_bool}, @{thm leo2_rules(4)}])
| Donkey_Cong => trace_tac' "mark: donkey_cong" (simper_animal ctxt THEN' ex_expander_tac ctxt)
| Extuni_Bool2 => trace_tac' "mark: extuni_bool2" (dtac @{thm extuni_bool2})
@@ -1624,7 +1624,7 @@
| Extuni_Dec => trace_tac' "mark: extuni_dec_tac" (extuni_dec_tac ctxt)
| Extuni_FlexRigid => trace_tac' "mark: extuni_flex_rigid" (atac ORELSE' asm_full_simp_tac ctxt)
| Extuni_Func => trace_tac' "mark: extuni_func" (dtac @{thm extuni_func})
- | Polarity_switch => trace_tac' "mark: polarity_switch" (eresolve_tac @{thms polarity_switch})
+ | Polarity_switch => trace_tac' "mark: polarity_switch" (eresolve_tac ctxt @{thms polarity_switch})
| Forall_special_pos => trace_tac' "mark: dorall_special_pos" extcnf_forall_special_pos_tac
val core_tac =
@@ -1804,7 +1804,7 @@
member (op =) (Term.add_frees hyp_lit_body []) (hd hyp_lit_prefix) then
no_tac st
else
- dresolve_tac @{thms drop_redundant_literal_qtfr} i st
+ dresolve_tac ctxt @{thms drop_redundant_literal_qtfr} i st
end
end
end
--- a/src/HOL/TPTP/atp_problem_import.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TPTP/atp_problem_import.ML Tue Feb 10 14:48:26 2015 +0100
@@ -178,7 +178,7 @@
val (outcome, _) =
Nitpick.pick_nits_in_term state params Nitpick.Normal i n step subst [] [] conj
in
- if outcome = "none" then ALLGOALS Skip_Proof.cheat_tac th else Seq.empty
+ if outcome = "none" then ALLGOALS (Skip_Proof.cheat_tac ctxt) th else Seq.empty
end
end
--- a/src/HOL/TPTP/sledgehammer_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/TPTP/sledgehammer_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -70,7 +70,7 @@
("minimize", "false")]
val xs = run_prover override_params fact_override chained i i ctxt th
in
- if is_some xs then ALLGOALS Skip_Proof.cheat_tac th else Seq.empty
+ if is_some xs then ALLGOALS (Skip_Proof.cheat_tac ctxt) th else Seq.empty
end
end;
--- a/src/HOL/Tools/BNF/bnf_comp_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_comp_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -171,11 +171,11 @@
unfold_thms_tac ctxt [collect_set_map] THEN
unfold_thms_tac ctxt comp_wit_thms THEN
REPEAT_DETERM ((atac ORELSE'
- REPEAT_DETERM o eresolve_tac @{thms UnionE UnE} THEN'
- etac imageE THEN' TRY o dresolve_tac Gwit_thms THEN'
+ REPEAT_DETERM o eresolve_tac ctxt @{thms UnionE UnE} THEN'
+ etac imageE THEN' TRY o dresolve_tac ctxt Gwit_thms THEN'
(etac FalseE ORELSE'
hyp_subst_tac ctxt THEN'
- dresolve_tac Fwit_thms THEN'
+ dresolve_tac ctxt Fwit_thms THEN'
(etac FalseE ORELSE' atac))) 1);
@@ -231,7 +231,7 @@
fun mk_simple_rel_OO_Grp_tac rel_OO_Grp in_alt_thm =
rtac (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]) 1;
-fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve_tac (@{thm emptyE} :: wit_thms));
+fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve0_tac (@{thm emptyE} :: wit_thms));
val csum_thms =
@{thms csum_cong1 csum_cong2 csum_cong csum_dup[OF natLeq_cinfinite natLeq_Card_order]};
@@ -247,7 +247,7 @@
unfold_thms_tac ctxt simplified_set_simps THEN rtac refl 1;
val bd_ordIso_natLeq_tac =
- HEADGOAL (REPEAT_DETERM o resolve_tac
+ HEADGOAL (REPEAT_DETERM o resolve0_tac
(@{thm ordIso_refl[OF natLeq_Card_order]} :: csum_thms @ cprod_thms));
end;
--- a/src/HOL/Tools/BNF/bnf_def_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -58,7 +58,7 @@
fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1
else (rtac subsetI THEN'
rtac CollectI) 1 THEN
- REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN
+ REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN
REPEAT_DETERM_N (n - 1)
((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN
(etac subset_trans THEN' atac) 1;
@@ -102,21 +102,21 @@
else
EVERY' [rel_OO_Grps_tac, rtac @{thm antisym}, rtac @{thm predicate2I},
REPEAT_DETERM o
- eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
+ eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
hyp_subst_tac ctxt, rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0,
REPEAT_DETERM_N n o EVERY' [rtac @{thm Collect_split_Grp_eqD}, etac @{thm set_mp}, atac],
rtac CollectI,
CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
rtac @{thm image_subsetI}, rtac @{thm Collect_split_Grp_inD}, etac @{thm set_mp}, atac])
set_maps,
- rtac @{thm predicate2I}, REPEAT_DETERM o eresolve_tac [@{thm GrpE}, exE, conjE],
+ rtac @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [@{thm GrpE}, exE, conjE],
hyp_subst_tac ctxt,
rtac @{thm relcomppI}, rtac @{thm conversepI},
EVERY' (map2 (fn convol => fn map_id0 =>
EVERY' [rtac @{thm GrpI},
rtac (@{thm box_equals} OF [map_cong0, map_comp RS sym, map_id0]),
REPEAT_DETERM_N n o rtac (convol RS fun_cong),
- REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE],
rtac CollectI,
CONJ_WRAP' (fn thm =>
EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
@@ -145,11 +145,11 @@
Goal.conjunction_tac 1 THEN
unfold_thms_tac ctxt [rel_compp, rel_conversep, rel_Grp, @{thm vimage2p_Grp}] THEN
TRYALL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI},
- resolve_tac [map_id, refl], rtac CollectI,
+ resolve_tac ctxt [map_id, refl], rtac CollectI,
CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac,
- rtac @{thm conversepI}, rtac @{thm GrpI}, resolve_tac [map_id, refl], rtac CollectI,
+ rtac @{thm conversepI}, rtac @{thm GrpI}, resolve_tac ctxt [map_id, refl], rtac CollectI,
CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks,
- REPEAT_DETERM o eresolve_tac @{thms relcomppE conversepE GrpE},
+ REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE conversepE GrpE},
dtac (trans OF [sym, map_id]), hyp_subst_tac ctxt, atac])
end;
@@ -175,7 +175,7 @@
else
EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I},
REPEAT_DETERM o
- eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
+ eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
hyp_subst_tac ctxt, rtac @{thm conversepI}, rtac @{thm relcomppI}, rtac @{thm conversepI},
EVERY' (map (fn thm => EVERY' [rtac @{thm GrpI}, rtac sym, rtac trans,
rtac map_cong0, REPEAT_DETERM_N n o rtac thm,
@@ -204,7 +204,7 @@
else
EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I},
REPEAT_DETERM o
- eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
+ eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
hyp_subst_tac ctxt,
rtac @{thm relcomppI}, rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI},
rtac trans, rtac map_comp, rtac sym, rtac map_cong0,
@@ -229,7 +229,7 @@
if null set_maps then atac 1
else
unfold_tac ctxt [in_rel] THEN
- REPEAT_DETERM (eresolve_tac [exE, CollectE, conjE] 1) THEN
+ REPEAT_DETERM (eresolve_tac ctxt [exE, CollectE, conjE] 1) THEN
hyp_subst_tac ctxt 1 THEN
EVERY' [rtac exI, rtac @{thm conjI[OF CollectI conjI[OF refl refl]]},
CONJ_WRAP' (fn thm =>
@@ -245,7 +245,7 @@
in
REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN
(HEADGOAL (hyp_subst_tac ctxt THEN' rtac refl) ORELSE
- REPEAT_DETERM (HEADGOAL (eresolve_tac (Tactic.make_elim (in_rel RS iffD1) ::
+ REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt (Tactic.make_elim (in_rel RS iffD1) ::
@{thms exE conjE CollectE}))) THEN
HEADGOAL (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt rel_map THEN HEADGOAL (rtac iffI) THEN
last_tac iffD1 THEN last_tac iffD2)
@@ -264,7 +264,7 @@
unfold_thms_tac ctxt @{thms rel_fun_iff_leq_vimage2p} THEN
HEADGOAL (EVERY' [rtac @{thm order_trans}, rtac rel_mono, REPEAT_DETERM_N n o atac,
rtac @{thm predicate2I}, dtac (in_rel RS iffD1),
- REPEAT_DETERM o eresolve_tac [exE, CollectE, conjE], hyp_subst_tac ctxt,
+ REPEAT_DETERM o eresolve_tac ctxt [exE, CollectE, conjE], hyp_subst_tac ctxt,
rtac @{thm vimage2pI}, rtac (in_rel RS iffD2), rtac exI, rtac conjI, in_tac,
rtac conjI,
EVERY' (map (fn convol =>
@@ -311,7 +311,7 @@
unfold_thms_tac ctxt @{thms cprod_def Field_card_of} THEN
EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac surj_imp_ordLeq_inst, rtac subsetI,
Method.insert_tac inserts, REPEAT_DETERM o dtac meta_spec,
- REPEAT_DETERM o eresolve_tac [exE, Tactic.make_elim conjunct1], etac CollectE,
+ REPEAT_DETERM o eresolve_tac ctxt [exE, Tactic.make_elim conjunct1], etac CollectE,
if live = 1 then K all_tac
else REPEAT_DETERM_N (live - 2) o (etac conjE THEN' rotate_tac ~1) THEN' etac conjE,
rtac (Drule.rotate_prems 1 @{thm image_eqI}), rtac @{thm SigmaI}, rtac @{thm UNIV_I},
@@ -325,7 +325,7 @@
REPEAT_DETERM_N live o rtac (@{thm prod.case} RS trans),
rtac refl,
rtac @{thm surj_imp_ordLeq}, rtac subsetI, rtac (Drule.rotate_prems 1 @{thm image_eqI}),
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], rtac CollectI,
CONJ_WRAP' (fn thm =>
rtac (thm RS ord_eq_le_trans) THEN' etac @{thm subset_trans[OF image_mono Un_upper1]})
set_maps,
@@ -341,12 +341,12 @@
fun mk_set_transfer_tac ctxt in_rel set_maps =
Goal.conjunction_tac 1 THEN
EVERY (map (fn set_map => HEADGOAL (rtac rel_funI) THEN
- REPEAT_DETERM (HEADGOAL (eresolve_tac (Tactic.make_elim (in_rel RS iffD1) ::
+ REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt (Tactic.make_elim (in_rel RS iffD1) ::
@{thms exE conjE CollectE}))) THEN
HEADGOAL (hyp_subst_tac ctxt THEN' rtac (@{thm iffD2[OF arg_cong2]} OF [set_map, set_map]) THEN'
rtac @{thm rel_setI}) THEN
REPEAT (HEADGOAL (etac imageE THEN' dtac @{thm set_mp} THEN' atac THEN'
- REPEAT_DETERM o (eresolve_tac @{thms CollectE case_prodE}) THEN' hyp_subst_tac ctxt THEN'
+ REPEAT_DETERM o (eresolve_tac ctxt @{thms CollectE case_prodE}) THEN' hyp_subst_tac ctxt THEN'
rtac bexI THEN' etac @{thm subst_Pair[OF _ refl]} THEN' etac imageI))) set_maps);
end;
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -120,16 +120,16 @@
fun mk_ctr_transfer_tac ctxt rel_intros rel_eqs =
HEADGOAL Goal.conjunction_tac THEN
- ALLGOALS (REPEAT o (resolve_tac (rel_funI :: rel_intros) THEN'
+ ALLGOALS (REPEAT o (resolve_tac ctxt (rel_funI :: rel_intros) THEN'
TRY o (REPEAT_DETERM1 o (atac ORELSE'
K (unfold_thms_tac ctxt rel_eqs) THEN' hyp_subst_tac ctxt THEN' rtac refl))));
-fun mk_disc_transfer_tac rel_sel exhaust_disc distinct_disc=
+fun mk_disc_transfer_tac rel_sel exhaust_disc distinct_disc =
let
fun last_disc_tac iffD =
HEADGOAL (rtac (rotate_prems ~1 exhaust_disc) THEN' atac THEN'
REPEAT_DETERM o (rotate_tac ~1 THEN' dtac (rotate_prems 1 iffD) THEN' atac THEN'
- rotate_tac ~1 THEN' etac (rotate_prems 1 notE) THEN' eresolve_tac distinct_disc));
+ rotate_tac ~1 THEN' etac (rotate_prems 1 notE) THEN' eresolve0_tac distinct_disc));
in
HEADGOAL Goal.conjunction_tac THEN
REPEAT_DETERM (HEADGOAL (rtac rel_funI THEN' dtac (rel_sel RS iffD1) THEN'
@@ -169,7 +169,7 @@
case_prod_map_prod id_bnf_def map_prod_simp map_sum_if_distrib_then map_sum_if_distrib_else
if_distrib[THEN sym]};
in
- HEADGOAL (subst_tac @{context} (SOME [1, 2]) [co_rec_def] THEN'
+ HEADGOAL (subst_tac ctxt (SOME [1, 2]) [co_rec_def] THEN'
rtac (xtor_co_rec_o_map RS trans) THEN'
CONVERSION Thm.eta_long_conversion THEN'
asm_simp_tac (ss_only (pre_map_defs @ distinct Thm.eq_thm_prop (map_ident0s @ abs_inverses) @
@@ -208,11 +208,11 @@
HEADGOAL (rtac @{thm vimage2p_rel_fun}) THEN
unfold_thms_tac ctxt rel_eqs THEN
EVERY (@{map 2} (fn n => fn xss =>
- REPEAT_DETERM (HEADGOAL (resolve_tac
+ REPEAT_DETERM (HEADGOAL (resolve_tac ctxt
[mk_rel_funDN 2 case_sum_transfer_eq, mk_rel_funDN 2 case_sum_transfer])) THEN
HEADGOAL (select_prem_tac total_n (dtac asm_rl) (acc + n)) THEN
HEADGOAL (SELECT_GOAL (HEADGOAL
- (REPEAT_DETERM o (atac ORELSE' resolve_tac
+ (REPEAT_DETERM o (atac ORELSE' resolve_tac ctxt
[mk_rel_funDN 1 case_prod_transfer_eq,
mk_rel_funDN 1 case_prod_transfer,
rel_funI]) THEN_ALL_NEW
@@ -251,7 +251,7 @@
val num_pgs = length pgs;
fun prem_no_of x = 1 + find_index (curry (op =) x) pgs;
- val Inl_Inr_Pair_tac = REPEAT_DETERM o (resolve_tac
+ val Inl_Inr_Pair_tac = REPEAT_DETERM o (resolve_tac ctxt
[mk_rel_funDN 1 @{thm Inl_transfer},
mk_rel_funDN 1 @{thm Inl_transfer[of "op =" "op =", unfolded sum.rel_eq]},
mk_rel_funDN 1 @{thm Inr_transfer},
@@ -314,8 +314,8 @@
end;
fun solve_prem_prem_tac ctxt =
- REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
- hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
+ REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
+ hyp_subst_tac ctxt ORELSE' resolve_tac ctxt @{thms disjI1 disjI2}) THEN'
(rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
@@ -361,7 +361,7 @@
(atac ORELSE' REPEAT o etac conjE THEN'
full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
- REPEAT o (resolve_tac [refl, conjI] ORELSE' atac));
+ REPEAT o (resolve_tac ctxt [refl, conjI] ORELSE' atac));
fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
let
@@ -424,7 +424,7 @@
map (fn thm => thm RS eqTrueI) rel_injects) THEN
TRYALL (atac ORELSE' etac FalseE ORELSE'
(REPEAT_DETERM o dtac @{thm meta_spec} THEN'
- TRY o filter_prems_tac
+ TRY o filter_prems_tac ctxt
(forall (curry (op <>) (HOLogic.mk_Trueprop @{term False})) o Logic.strip_imp_prems) THEN'
REPEAT_DETERM o (dtac @{thm meta_mp} THEN' rtac refl) THEN' Goal.assume_rule_tac ctxt));
@@ -454,7 +454,7 @@
fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse =>
HEADGOAL (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW
(rtac (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2)
- THEN' atac THEN' atac THEN' TRY o resolve_tac assms))) THEN
+ THEN' atac THEN' atac THEN' TRY o resolve_tac ctxt assms))) THEN
unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @
@{thms id_bnf_def vimage2p_def}) THEN
TRYALL (hyp_subst_tac ctxt) THEN
@@ -470,7 +470,7 @@
unfold_thms_tac ctxt (sels @ rel_injects @ rel_eqs @
@{thms simp_thms(6,7,8,11,12,15,16,21,22,24)} @ ((discs @ distincts) RL [eqTrueI, eqFalseI]) @
(rel_injects RL [eqTrueI]) @ (rel_distincts RL [eqFalseI])) THEN
- TRYALL (resolve_tac [TrueI, refl]);
+ TRYALL (resolve_tac ctxt [TrueI, refl]);
fun mk_sel_transfer_tac ctxt n sel_defs case_transfer =
TRYALL Goal.conjunction_tac THEN
@@ -486,8 +486,8 @@
@{thms not_True_eq_False not_False_eq_True}) THEN
TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN
unfold_thms_tac ctxt (sels @ sets) THEN
- ALLGOALS (REPEAT o (resolve_tac @{thms UnI1 UnI2 imageI} ORELSE'
- eresolve_tac @{thms UN_I UN_I[rotated] imageE} ORELSE'
+ ALLGOALS (REPEAT o (resolve_tac ctxt @{thms UnI1 UnI2 imageI} ORELSE'
+ eresolve_tac ctxt @{thms UN_I UN_I[rotated] imageE} ORELSE'
hyp_subst_tac ctxt) THEN'
(rtac @{thm singletonI} ORELSE' atac));
@@ -495,16 +495,16 @@
HEADGOAL (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN
unfold_thms_tac ctxt sets THEN
REPEAT_DETERM (HEADGOAL
- (eresolve_tac @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE'
+ (eresolve_tac ctxt @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE'
hyp_subst_tac ctxt ORELSE'
- SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac assms THEN' REPEAT_DETERM o atac)))));
+ SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o atac)))));
fun mk_set_intros_tac ctxt sets =
TRYALL Goal.conjunction_tac THEN
unfold_thms_tac ctxt sets THEN
TRYALL (REPEAT o
- (resolve_tac @{thms UnI1 UnI2} ORELSE'
- eresolve_tac @{thms UN_I UN_I[rotated]}) THEN' (rtac @{thm singletonI} ORELSE' atac));
+ (resolve_tac ctxt @{thms UnI1 UnI2} ORELSE'
+ eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN' (rtac @{thm singletonI} ORELSE' atac));
fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors
Abs_pre_inverses =
@@ -518,15 +518,15 @@
val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
|> map2 (fn ct => cterm_instantiate_pos [NONE, SOME ct]) cts;
in
- ALLGOALS (resolve_tac dtor_set_inducts) THEN
- TRYALL (resolve_tac exhausts' THEN_ALL_NEW
- (resolve_tac (map (fn ct => refl RS
+ ALLGOALS (resolve_tac ctxt dtor_set_inducts) THEN
+ TRYALL (resolve_tac ctxt exhausts' THEN_ALL_NEW
+ (resolve_tac ctxt (map (fn ct => refl RS
cterm_instantiate_pos (replicate 4 NONE @ [SOME ct]) @{thm arg_cong2} RS iffD2) cts)
THEN' atac THEN' hyp_subst_tac ctxt)) THEN
unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @
@{thms id_bnf_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert Un_empty_left
Un_empty_right empty_iff singleton_iff}) THEN
- REPEAT (HEADGOAL (hyp_subst_tac ctxt ORELSE' eresolve_tac @{thms UN_E UnE singletonE} ORELSE'
+ REPEAT (HEADGOAL (hyp_subst_tac ctxt ORELSE' eresolve_tac ctxt @{thms UN_E UnE singletonE} ORELSE'
assms_tac))
end;
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -42,10 +42,10 @@
unfold_thms_tac ctxt vimage2p_unfolds THEN
HEADGOAL (CONJ_WRAP_GEN' (rtac @{thm context_conjI})
(fn thm => rtac thm THEN_ALL_NEW (rotate_tac ~1 THEN'
- REPEAT_ALL_NEW (FIRST' [eresolve_tac C_IHs, eresolve_tac C_IH_monos,
+ REPEAT_ALL_NEW (FIRST' [eresolve_tac ctxt C_IHs, eresolve_tac ctxt C_IH_monos,
SELECT_GOAL (unfold_thms_tac ctxt nesting_rel_eqs) THEN' rtac @{thm order_refl},
- assume_tac ctxt, resolve_tac co_inducts,
- resolve_tac C_IH_monos THEN' REPEAT_ALL_NEW (eresolve_tac nesting_rel_monos)])))
+ assume_tac ctxt, resolve_tac ctxt co_inducts,
+ resolve_tac ctxt C_IH_monos THEN' REPEAT_ALL_NEW (eresolve_tac ctxt nesting_rel_monos)])))
co_inducts)
end;
--- a/src/HOL/Tools/BNF/bnf_gfp.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML Tue Feb 10 14:48:26 2015 +0100
@@ -440,7 +440,7 @@
HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
in
Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
- (K (mk_mor_comp_tac mor_def mor_image'_thms morE_thms map_comp_id_thms))
+ (K (mk_mor_comp_tac lthy mor_def mor_image'_thms morE_thms map_comp_id_thms))
|> Thm.close_derivation
|> singleton (Proof_Context.export names_lthy lthy)
end;
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -52,7 +52,7 @@
val exhaust_inst_as_projs_tac = PRIMITIVE oo exhaust_inst_as_projs;
fun distinct_in_prems_tac distincts =
- eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
+ eresolve0_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
fun mk_primcorec_nchotomy_tac ctxt exhaust_discs =
HEADGOAL (Method.insert_tac exhaust_discs THEN' clean_blast_tac ctxt);
@@ -74,9 +74,9 @@
SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True
not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
- eresolve_tac falseEs ORELSE'
- resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
- dresolve_tac discIs THEN' atac ORELSE'
+ eresolve_tac ctxt falseEs ORELSE'
+ resolve_tac ctxt @{thms TrueI conjI disjI1 disjI2} ORELSE'
+ dresolve_tac ctxt discIs THEN' atac ORELSE'
etac notE THEN' atac ORELSE'
etac disjE))));
@@ -121,21 +121,21 @@
rtac FalseE THEN'
(rotate_tac 1 THEN' dtac fun_disc' THEN' REPEAT_DETERM o atac ORELSE'
cut_tac fun_disc') THEN'
- dresolve_tac distinct_discs THEN' etac notE THEN' atac)
+ dresolve_tac ctxt distinct_discs THEN' etac notE THEN' atac)
fun_discss) THEN'
(etac FalseE ORELSE'
- resolve_tac (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
+ resolve_tac ctxt (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms mapsx map_ident0s map_comps fun_sel k
m excludesss =
prelude_tac ctxt defs (fun_sel RS trans) THEN
cases_tac ctxt k m excludesss THEN
HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE'
- eresolve_tac falseEs ORELSE'
- resolve_tac split_connectI ORELSE'
+ eresolve_tac ctxt falseEs ORELSE'
+ resolve_tac ctxt split_connectI ORELSE'
Splitter.split_asm_tac ctxt (split_if_asm :: split_asms) ORELSE'
Splitter.split_tac ctxt (split_if :: splits) ORELSE'
- eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
+ eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
etac notE THEN' atac ORELSE'
(CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def
sum.case sum.sel sum.distinct[THEN eq_False[THEN iffD2]]} @
@@ -165,17 +165,17 @@
val splits' =
map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits);
in
- HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
+ HEADGOAL (REPEAT o (resolve_tac ctxt (splits' @ split_connectI))) THEN
prelude_tac ctxt [] (fun_ctr RS trans) THEN
HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
(rtac refl ORELSE' atac ORELSE'
- resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
+ resolve_tac ctxt (@{thm Code.abort_def} :: split_connectI) ORELSE'
Splitter.split_tac ctxt (split_if :: splits) ORELSE'
Splitter.split_asm_tac ctxt (split_if_asm :: split_asms) ORELSE'
mk_primcorec_assumption_tac ctxt discIs ORELSE'
distinct_in_prems_tac distincts ORELSE'
- (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))))
+ (TRY o dresolve_tac ctxt discIs) THEN' etac notE THEN' atac)))))
end;
fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'});
@@ -196,14 +196,14 @@
in
EVERY (map2 (raw_code_single_tac ctxt distincts discIs splits split_asms) ms' fun_ctrs') THEN
IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
- HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)))
+ HEADGOAL (REPEAT_DETERM o resolve_tac ctxt (refl :: split_connectI)))
end;
fun mk_primcorec_code_tac ctxt distincts splits raw =
HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o
(rtac refl ORELSE' atac ORELSE'
- resolve_tac split_connectI ORELSE'
+ resolve_tac ctxt split_connectI ORELSE'
Splitter.split_tac ctxt (split_if :: splits) ORELSE'
distinct_in_prems_tac distincts ORELSE'
rtac sym THEN' atac ORELSE'
--- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -56,7 +56,7 @@
thm list -> thm list -> thm list list -> thm list list list -> thm list list list ->
thm list list -> thm list -> thm list -> tactic
val mk_mor_case_sum_tac: 'a list -> thm -> tactic
- val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
+ val mk_mor_comp_tac: Proof.context -> thm -> thm list -> thm list -> thm list -> tactic
val mk_mor_elim_tac: thm -> tactic
val mk_mor_col_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
thm list -> thm list list -> thm list list -> tactic
@@ -116,7 +116,7 @@
dtac (coalg_def RS iffD1) 1 THEN
REPEAT_DETERM (etac conjE 1) THEN
EVERY' [dtac rev_bspec, atac] 1 THEN
- REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN atac 1;
+ REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN atac 1;
fun mk_mor_elim_tac mor_def =
(dtac (mor_def RS iffD1) THEN'
@@ -133,12 +133,12 @@
CONJ_WRAP' (fn thm =>
(EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (id_apply RS arg_cong)])) map_ids) 1;
-fun mk_mor_comp_tac mor_def mor_images morEs map_comp_ids =
+fun mk_mor_comp_tac ctxt mor_def mor_images morEs map_comp_ids =
let
fun fbetw_tac image = EVERY' [rtac ballI, rtac (o_apply RS @{thm ssubst_mem}), etac image,
etac image, atac];
fun mor_tac ((mor_image, morE), map_comp_id) =
- EVERY' [rtac ballI, stac o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
+ EVERY' [rtac ballI, stac ctxt o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac];
in
(rtac (mor_def RS iffD2) THEN' rtac conjI THEN'
@@ -183,14 +183,14 @@
if m = 0 then K all_tac
else (rtac @{thm Un_least} THEN' Goal.assume_rule_tac ctxt),
CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
- (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac [allE, conjE],
+ (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac ctxt [allE, conjE],
rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
rec_Sucs] 1;
fun mk_Jset_minimal_tac ctxt n col_minimal =
(CONJ_WRAP' (K (EVERY' [rtac @{thm UN_least}, rtac rev_mp, rtac col_minimal,
EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI,
- REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) (1 upto n)) 1
+ REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], atac])) (1 upto n)) 1
fun mk_mor_col_tac m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss =
EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
@@ -222,7 +222,7 @@
fun mk_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) =
EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i),
etac allE, etac allE, etac impE, atac, etac bexE,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ REPEAT_DETERM o eresolve0_tac [CollectE, conjE],
rtac (in_rel RS iffD2), rtac exI, rtac (Drule.rotate_prems 1 conjI),
CONJ_WRAP' (fn thm => EVERY' [rtac trans, rtac trans, rtac map_comp0, rtac map_cong0,
REPEAT_DETERM_N m o rtac thm, REPEAT_DETERM_N n o rtac (@{thm comp_id} RS fun_cong),
@@ -242,7 +242,7 @@
EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI,
etac allE, etac allE, etac impE, atac,
dtac (in_rel RS @{thm iffD1}),
- REPEAT_DETERM o eresolve_tac ([CollectE, conjE, exE] @
+ REPEAT_DETERM o eresolve0_tac ([CollectE, conjE, exE] @
@{thms CollectE Collect_split_in_rel_leE}),
rtac bexI, rtac conjI, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
REPEAT_DETERM_N m o rtac (@{thm id_comp} RS fun_cong),
@@ -327,7 +327,7 @@
val n = length lsbis_defs;
in
EVERY' [rtac (Thm.permute_prems 0 1 bis_cong), EVERY' (map rtac lsbis_defs),
- rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE, exE],
+ rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, exE],
hyp_subst_tac ctxt, etac bis_cong, EVERY' (map (rtac o mk_nth_conv n) (1 upto n))] 1
end;
@@ -357,7 +357,7 @@
val n = length strT_defs;
val ks = 1 upto n;
fun coalg_tac (i, (active_sets, def)) =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE],
hyp_subst_tac ctxt, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans),
rtac (mk_sum_caseN n i), rtac CollectI,
REPEAT_DETERM_N m o EVERY' [rtac conjI, rtac subset_UNIV],
@@ -370,7 +370,7 @@
CONJ_WRAP' (fn i => EVERY' [rtac ballI, etac CollectE, dtac @{thm ShiftD},
dtac bspec, etac thin_rl, atac, dtac (mk_conjunctN n i),
dtac bspec, rtac CollectI, etac @{thm set_mp[OF equalityD1[OF Succ_Shift]]},
- REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
+ REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac exI,
rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
rtac (@{thm append_Cons} RS sym RS arg_cong RS trans), atac,
CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
@@ -378,7 +378,7 @@
rtac (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks,
dtac bspec, atac, dtac (mk_conjunctN n i), dtac bspec,
etac @{thm set_mp[OF equalityD1]}, atac,
- REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
+ REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac exI,
rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
etac (@{thm append_Nil} RS sym RS arg_cong RS trans),
REPEAT_DETERM_N m o (rtac conjI THEN' atac),
@@ -405,7 +405,7 @@
EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
(fn i =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+ EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt,
rtac trans, rtac @{thm length_Cons}, rtac @{thm arg_cong[of _ _ Suc]},
REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), etac mp, atac]) ks])
Lev_Sucs] 1
@@ -467,7 +467,7 @@
EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
(fn (i, from_to_sbd) =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+ EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt,
CONJ_WRAP' (fn i' => rtac impI THEN'
CONJ_WRAP' (fn i'' =>
EVERY' [rtac impI, rtac (Lev_Suc RS equalityD2 RS set_mp),
@@ -502,7 +502,7 @@
dtac (Lev_Suc RS equalityD1 RS set_mp), dtac (mk_InN_inject n i),
hyp_subst_tac ctxt,
CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn k => REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
+ (fn k => REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN'
dtac list_inject_iffD1 THEN' etac conjE THEN'
(if k = i'
then EVERY' [dtac (mk_InN_inject n k), hyp_subst_tac ctxt, etac imageI]
@@ -517,12 +517,12 @@
EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
(fn (i, (from_to_sbd, to_sbd_inj)) =>
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN' hyp_subst_tac ctxt THEN'
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN' hyp_subst_tac ctxt THEN'
CONJ_WRAP' (fn i' => rtac impI THEN'
dtac @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN'
dtac (Lev_Suc RS equalityD1 RS set_mp) THEN'
CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn k =>
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE] THEN'
dtac list_inject_iffD1 THEN' etac conjE THEN'
(if k = i
then EVERY' [dtac (mk_InN_inject n i),
@@ -563,7 +563,7 @@
CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_map0s,
(set_Levs, set_image_Levs))))) =>
EVERY' [rtac ballI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}],
rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS iffD2),
rtac exI, rtac conjI,
if n = 1 then rtac refl
@@ -572,7 +572,7 @@
EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
if n = 1 then rtac refl else atac, atac, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}],
REPEAT_DETERM_N 4 o etac thin_rl,
rtac set_image_Lev,
atac, dtac length_Lev, hyp_subst_tac ctxt, dtac length_Lev',
@@ -591,7 +591,7 @@
rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, rtac set_Lev,
rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI}, rtac rv_Nil,
atac, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}],
rtac set_image_Lev, rtac (Lev_0 RS equalityD2 RS set_mp),
rtac @{thm singletonI}, dtac length_Lev',
etac @{thm set_mp[OF equalityD1[OF arg_cong[OF
@@ -612,11 +612,11 @@
DETERM o EVERY' [rtac trans, rtac o_apply, rtac prod_injectI, rtac conjI,
rtac trans, rtac @{thm Shift_def},
rtac equalityI, rtac subsetI, etac thin_rl,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl,
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl,
etac thin_rl, dtac @{thm set_rev_mp[OF _ equalityD1]},
rtac (@{thm length_Cons} RS arg_cong RS trans), rtac Lev_Suc,
CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE],
dtac list_inject_iffD1, etac conjE,
if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
@@ -648,7 +648,7 @@
fun mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBISs =
EVERY' [rtac @{thm congruentI}, dtac lsbisE,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE], rtac (o_apply RS trans),
+ REPEAT_DETERM o eresolve0_tac [CollectE, conjE, bexE], rtac (o_apply RS trans),
etac (sym RS arg_cong RS trans), rtac (map_comp_id RS trans),
rtac (map_cong0 RS trans), REPEAT_DETERM_N m o rtac refl,
EVERY' (map (fn equiv_LSBIS =>
@@ -818,8 +818,8 @@
EVERY' ([rtac rev_mp, coinduct_tac] @
maps (fn ((((((map_comp_trans, dtor_maps_trans), map_cong0), set_map0s), set_Jsets),
set_Jset_Jsetss), in_rel) =>
- [REPEAT_DETERM o resolve_tac [allI, impI, in_rel RS iffD2],
- REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac ctxt, rtac exI,
+ [REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2],
+ REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], hyp_subst_tac ctxt, rtac exI,
rtac (Drule.rotate_prems 1 conjI),
rtac conjI, rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0,
REPEAT_DETERM_N m o (rtac o_apply_trans_sym THEN' rtac @{thm fst_conv}),
@@ -910,17 +910,17 @@
rel_Jrels le_rel_OOs) 1;
fun mk_wit_tac ctxt n dtor_ctors dtor_set wit coind_wits =
- ALLGOALS (TRY o (eresolve_tac coind_wits THEN' rtac refl)) THEN
+ ALLGOALS (TRY o (eresolve_tac ctxt coind_wits THEN' rtac refl)) THEN
REPEAT_DETERM (atac 1 ORELSE
- EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac dtor_set,
+ EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt dtor_set,
K (unfold_thms_tac ctxt dtor_ctors),
REPEAT_DETERM_N n o etac UnE,
REPEAT_DETERM o
(TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
- (eresolve_tac wit ORELSE'
- (dresolve_tac wit THEN'
+ (eresolve_tac ctxt wit ORELSE'
+ (dresolve_tac ctxt wit THEN'
(etac FalseE ORELSE'
- EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac dtor_set,
+ EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt dtor_set,
K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac UnE]))))] 1);
fun mk_coind_wit_tac ctxt induct unfolds set_nats wits =
@@ -928,7 +928,7 @@
unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN
ALLGOALS (REPEAT_DETERM o etac imageE THEN' TRY o hyp_subst_tac ctxt) THEN
ALLGOALS (TRY o
- FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac wits THEN' etac FalseE]);
+ FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac ctxt wits THEN' etac FalseE]);
fun mk_dtor_corec_transfer_tac ctxt n m dtor_corec_defs dtor_unfold_transfer pre_T_map_transfers
dtor_rels =
@@ -955,7 +955,7 @@
val (passive_set_map0s, active_set_map0s) = chop m set_map0s;
val in_Jrel = nth in_Jrels (i - 1);
val if_tac =
- EVERY' [dtac (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ EVERY' [dtac (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE],
rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
EVERY' (map2 (fn set_map0 => fn set_incl =>
EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0,
@@ -975,7 +975,7 @@
rtac trans, rtac sym, rtac dtor_map, rtac (dtor_inject RS iffD2), atac])
@{thms fst_conv snd_conv}];
val only_if_tac =
- EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE],
rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
CONJ_WRAP' (fn (dtor_set, passive_set_map0) =>
EVERY' [rtac ord_eq_le_trans, rtac dtor_set, rtac @{thm Un_least},
@@ -987,12 +987,12 @@
rtac active_set_map0, rtac (dtor_ctor RS sym RS arg_cong), rtac @{thm UN_least},
dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
dtac @{thm ssubst_mem[OF pair_collapse]},
- REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
+ REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE ::
@{thms case_prodE iffD1[OF prod.inject, elim_format]}),
hyp_subst_tac ctxt,
dtac (in_Jrel RS iffD1),
dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], atac])
(rev (active_set_map0s ~~ in_Jrels))])
(dtor_sets ~~ passive_set_map0s),
rtac conjI,
@@ -1001,7 +1001,7 @@
rtac map_cong0, REPEAT_DETERM_N m o rtac @{thm fun_cong[OF comp_id]},
EVERY' (map (fn in_Jrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
dtac @{thm ssubst_mem[OF pair_collapse]},
- REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
+ REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE ::
@{thms case_prodE iffD1[OF prod.inject, elim_format]}),
hyp_subst_tac ctxt, dtac (in_Jrel RS iffD1),
dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jrels),
@@ -1020,10 +1020,10 @@
EVERY' [rtac coinduct,
EVERY' (@{map 8} (fn i => fn map_comp0 => fn map_cong => fn map_arg_cong => fn set_map0s =>
fn dtor_unfold => fn dtor_map => fn in_rel =>
- EVERY' [REPEAT_DETERM o resolve_tac [allI, impI, in_rel RS iffD2],
- REPEAT_DETERM o eresolve_tac [exE, conjE],
+ EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2],
+ REPEAT_DETERM o eresolve_tac ctxt [exE, conjE],
select_prem_tac (length ks) (dtac @{thm spec2}) i, dtac mp, atac,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], hyp_subst_tac ctxt,
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt,
rtac exI, rtac (Drule.rotate_prems 1 conjI), rtac conjI,
rtac (map_comp0 RS trans), rtac (dtor_map RS trans RS sym),
rtac (dtor_unfold RS map_arg_cong RS trans), rtac (trans OF [map_comp0, map_cong]),
@@ -1051,20 +1051,20 @@
in
rtac set_induct 1 THEN
EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i =>
- EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], etac conjE,
+ EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac conjE,
select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE, Collect_splitD_set_mp, set_rev_mp],
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE, Collect_splitD_set_mp, set_rev_mp],
hyp_subst_tac ctxt,
SELECT_GOAL (unfold_thms_tac ctxt ([unfold, nth set_map0s (j - 1)] @ split_id_unfolds)),
rtac subset_refl])
unfolds set_map0ss ks) 1 THEN
EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i =>
EVERY' (map (fn set_map0 =>
- EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], etac conjE,
+ EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac conjE,
select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], hyp_subst_tac ctxt,
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt,
SELECT_GOAL (unfold_thms_tac ctxt ([unfold, set_map0] @ split_id_unfolds)),
- etac imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac [allE, mp],
+ etac imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp],
rtac conjI, etac Collect_splitD_set_mp, atac, rtac (@{thm surjective_pairing} RS arg_cong)])
(drop m set_map0s)))
unfolds set_map0ss ks) 1
@@ -1083,8 +1083,8 @@
CONJ_WRAP' (fn helper_ind =>
EVERY' [rtac (helper_ind RS rev_mp), REPEAT_DETERM_N n o atac,
REPEAT_DETERM_N n o etac thin_rl, rtac impI,
- REPEAT_DETERM o resolve_tac [subsetI, CollectI, @{thm iffD2[OF split_beta]}],
- dtac bspec, atac, REPEAT_DETERM o eresolve_tac [allE, mp, conjE],
+ REPEAT_DETERM o resolve_tac ctxt [subsetI, CollectI, @{thm iffD2[OF split_beta]}],
+ dtac bspec, atac, REPEAT_DETERM o eresolve_tac ctxt [allE, mp, conjE],
etac (refl RSN (2, conjI))])
helper_inds,
rtac conjI,
@@ -1103,9 +1103,9 @@
@{thms rel_fun_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN
unfold_thms_tac ctxt @{thms rel_fun_iff_geq_image2p} THEN
HEADGOAL (EVERY'
- [REPEAT_DETERM o resolve_tac [allI, impI], rtac ctor_rel_coinduct,
+ [REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctor_rel_coinduct,
EVERY' (map (fn map_transfer => EVERY'
- [REPEAT_DETERM o resolve_tac [allI, impI], etac @{thm image2pE}, hyp_subst_tac ctxt,
+ [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac @{thm image2pE}, hyp_subst_tac ctxt,
SELECT_GOAL (unfold_thms_tac ctxt unfolds),
rtac (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer),
REPEAT_DETERM_N m o rtac @{thm id_transfer},
--- a/src/HOL/Tools/BNF/bnf_lfp.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML Tue Feb 10 14:48:26 2015 +0100
@@ -368,7 +368,7 @@
val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
in
Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
- (K (mk_mor_incl_tac mor_def map_ids))
+ (fn {context = ctxt, ...} => mk_mor_incl_tac ctxt mor_def map_ids)
|> Thm.close_derivation
|> singleton (Proof_Context.export names_lthy lthy)
end;
@@ -860,7 +860,8 @@
lthy
|> @{fold_map 3} (fn b => fn mx => fn car_init =>
typedef (b, params, mx) car_init NONE
- (fn _ => EVERY' [rtac iffD2, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
+ (fn ctxt =>
+ EVERY' [rtac iffD2, rtac @{thm ex_in_conv}, resolve_tac ctxt alg_not_empty_thms,
rtac alg_init_thm] 1)) bs mixfixes car_inits
|>> apsnd split_list o split_list;
@@ -1024,7 +1025,8 @@
in
Goal.prove_sorry lthy [] []
(HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)))
- (K (mk_mor_fold_tac cT ct fold_defs init_ex_mor_thm mor_cong))
+ (fn {context = ctxt, ...} =>
+ mk_mor_fold_tac ctxt cT ct fold_defs init_ex_mor_thm mor_cong)
|> Thm.close_derivation
|> singleton (Proof_Context.export names_lthy lthy)
end;
--- a/src/HOL/Tools/BNF/bnf_lfp_size.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML Tue Feb 10 14:48:26 2015 +0100
@@ -221,7 +221,7 @@
(Class.instantiation (T_names, map dest_TFree As, [HOLogic.class_size])
#> @{fold_map 3} define_overloaded_size overloaded_size_def_bs overloaded_size_consts
overloaded_size_rhss
- ##> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
+ ##> Class.prove_instantiation_instance (fn ctxt => Class.intro_classes_tac ctxt [])
##> Local_Theory.exit_global);
val size_defs' =
--- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -56,8 +56,8 @@
val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic
val mk_mor_convol_tac: 'a list -> thm -> tactic
val mk_mor_elim_tac: thm -> tactic
- val mk_mor_incl_tac: thm -> thm list -> tactic
- val mk_mor_fold_tac: ctyp -> cterm -> thm list -> thm -> thm -> tactic
+ val mk_mor_incl_tac: Proof.context -> thm -> thm list -> tactic
+ val mk_mor_fold_tac: Proof.context -> ctyp -> cterm -> thm list -> thm -> thm -> tactic
val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list ->
thm list -> tactic
val mk_mor_str_tac: 'a list -> thm -> tactic
@@ -96,9 +96,9 @@
fun mk_alg_not_empty_tac ctxt alg_set alg_sets wits =
(EVERY' [rtac notI, hyp_subst_tac ctxt, ftac alg_set] THEN'
REPEAT_DETERM o FIRST'
- [EVERY' [rtac @{thm subset_emptyI}, eresolve_tac wits],
- EVERY' [rtac subsetI, rtac FalseE, eresolve_tac wits],
- EVERY' [rtac subsetI, dresolve_tac wits, hyp_subst_tac ctxt,
+ [EVERY' [rtac @{thm subset_emptyI}, eresolve_tac ctxt wits],
+ EVERY' [rtac subsetI, rtac FalseE, eresolve_tac ctxt wits],
+ EVERY' [rtac subsetI, dresolve_tac ctxt wits, hyp_subst_tac ctxt,
FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN'
etac @{thm emptyE}) 1;
@@ -109,13 +109,13 @@
etac bspec THEN'
atac) 1;
-fun mk_mor_incl_tac mor_def map_ids =
+fun mk_mor_incl_tac ctxt mor_def map_ids =
(rtac (mor_def RS iffD2) THEN'
rtac conjI THEN'
CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, etac (id_apply RS @{thm ssubst_mem})]))
map_ids THEN'
CONJ_WRAP' (fn thm =>
- (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_ids) 1;
+ (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac ctxt thm, rtac refl])) map_ids) 1;
fun mk_mor_comp_tac mor_def set_maps map_comp_ids =
let
@@ -124,7 +124,7 @@
fun mor_tac (set_map, map_comp_id) =
EVERY' [rtac ballI, rtac (o_apply RS trans), rtac trans,
rtac trans, dtac rev_bspec, atac, etac arg_cong,
- REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
+ REPEAT o eresolve0_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
CONJ_WRAP' (fn thm =>
FIRST' [rtac subset_UNIV,
(EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
@@ -170,14 +170,14 @@
rtac equalityD1, etac @{thm bij_betw_imageE}];
val alg_tac =
CONJ_WRAP' (fn (set_maps, alg_set) =>
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac set_mp,
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac set_mp,
rtac equalityD1, etac @{thm bij_betw_imageE[OF bij_betw_the_inv_into]},
rtac imageI, etac alg_set, EVERY' (map set_tac (drop m set_maps))])
(set_mapss ~~ alg_sets);
val mor_tac = rtac conjI THEN' CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets THEN'
CONJ_WRAP' (fn (set_maps, alg_set) =>
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE],
etac @{thm f_the_inv_into_f_bij_betw}, etac alg_set,
EVERY' (map set_tac (drop m set_maps))])
(set_mapss ~~ alg_sets);
@@ -251,7 +251,7 @@
rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite},
rtac Asuc_Cinfinite, rtac bd_Card_order,
rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq},
- resolve_tac @{thms Card_order_csum Card_order_ctwo}, rtac suc_Cinfinite,
+ resolve0_tac @{thms Card_order_csum Card_order_ctwo}, rtac suc_Cinfinite,
rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite];
in
(rtac induct THEN'
@@ -269,7 +269,7 @@
fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg,
rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI},
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac alg_set,
+ REPEAT_DETERM o eresolve0_tac [CollectE, conjE], etac alg_set,
REPEAT_DETERM o (etac subset_trans THEN' minG_tac)];
in
(rtac induct THEN'
@@ -283,9 +283,9 @@
val n = length min_algs;
fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY'
[rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono,
- etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve_tac set_bds];
+ etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve0_tac set_bds];
fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE],
EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE,
rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac,
rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}),
@@ -314,7 +314,8 @@
REPEAT_DETERM o etac conjE, atac] 1;
fun mk_alg_select_tac ctxt Abs_inverse =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt] 1 THEN
+ EVERY' [rtac ballI,
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt] 1 THEN
unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1;
fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select alg_sets
@@ -326,9 +327,9 @@
val mor_tac =
CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
fun alg_epi_tac ((alg_set, str_init_def), set_map) =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac CollectI,
rtac ballI, ftac (alg_select RS bspec), rtac (str_init_def RS @{thm ssubst_mem}),
- etac alg_set, REPEAT_DETERM o EVERY' [rtac ord_eq_le_trans, resolve_tac set_map,
+ etac alg_set, REPEAT_DETERM o EVERY' [rtac ord_eq_le_trans, resolve0_tac set_map,
rtac subset_trans, etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec,
atac]];
in
@@ -344,7 +345,7 @@
in
EVERY' [Method.insert_tac (map (fn thm => thm RS @{thm ex_bij_betw}) card_of_min_algs),
REPEAT_DETERM o etac exE, rtac rev_mp, rtac copy, REPEAT_DETERM_N n o atac,
- rtac impI, REPEAT_DETERM o eresolve_tac [exE, conjE], REPEAT_DETERM_N n o rtac exI,
+ rtac impI, REPEAT_DETERM o eresolve0_tac [exE, conjE], REPEAT_DETERM_N n o rtac exI,
rtac mor_comp, rtac mor_Rep, rtac mor_select, rtac CollectI, REPEAT_DETERM o rtac exI,
rtac conjI, rtac refl, atac,
SELECT_GOAL (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)),
@@ -367,7 +368,7 @@
fun mk_alg_tac (ct, (alg_set, (in_mono, (morE, map_cong0)))) =
EVERY' [rtac ballI, rtac CollectI,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
+ REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
rtac trans, mor_tac morE in_mono,
rtac trans, cong_tac ct map_cong0,
@@ -386,7 +387,7 @@
val n = length least_min_algs;
fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
+ REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
rtac mp, etac bspec, rtac CollectI,
REPEAT_DETERM_N m o (rtac conjI THEN' atac),
@@ -417,15 +418,15 @@
EVERY' [rtac conjI,
CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses,
CONJ_WRAP' (fn (ct, thm) =>
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE],
rtac (thm RS (cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong) RS sym),
EVERY' (map (fn Abs_inverse =>
EVERY' [rtac (o_apply RS trans RS ballI), etac (set_mp RS Abs_inverse), atac])
Abs_inverses)])
(cts ~~ map2 mk_trans map_comp_ids map_congLs)] 1;;
-fun mk_mor_fold_tac cT ct fold_defs ex_mor mor =
- (EVERY' (map stac fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
+fun mk_mor_fold_tac ctxt cT ct fold_defs ex_mor mor =
+ (EVERY' (map (stac ctxt) fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
REPEAT_DETERM_N (length fold_defs) o etac exE THEN'
rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
@@ -433,7 +434,7 @@
let
fun mk_unique type_def =
EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}),
- rtac ballI, resolve_tac init_unique_mors,
+ rtac ballI, resolve0_tac init_unique_mors,
EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps),
rtac mor_comp, rtac mor_Abs, atac,
rtac mor_comp, rtac mor_Abs, rtac mor_fold];
@@ -472,7 +473,7 @@
fun mk_closed_tac (k, (morE, set_maps)) =
EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
rtac (mor_Abs RS morE RS arg_cong RS iffD2), atac,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], dtac @{thm meta_spec},
EVERY' (@{map 3} mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), atac];
fun mk_induct_tac (Rep, Rep_inv) =
@@ -498,7 +499,7 @@
in
EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct),
EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI,
- REPEAT_DETERM o eresolve_tac [conjE, allE],
+ REPEAT_DETERM o eresolve_tac ctxt [conjE, allE],
CONJ_WRAP' (K atac) ks] 1
end;
@@ -626,13 +627,13 @@
fun mk_wit_tac ctxt n ctor_set wit =
REPEAT_DETERM (atac 1 ORELSE
- EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
+ EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt ctor_set,
REPEAT_DETERM o
(TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
- (eresolve_tac wit ORELSE'
- (dresolve_tac wit THEN'
+ (eresolve_tac ctxt wit ORELSE'
+ (dresolve_tac ctxt wit THEN'
(etac FalseE ORELSE'
- EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
+ EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctxt ctor_set,
REPEAT_DETERM_N n o etac UnE]))))] 1);
fun mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets ctor_inject
@@ -646,7 +647,7 @@
val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans;
val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans;
val if_tac =
- EVERY' [dtac (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ EVERY' [dtac (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE],
rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
EVERY' (map2 (fn set_map0 => fn ctor_set_incl =>
EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0,
@@ -669,7 +670,7 @@
etac eq_arg_cong_ctor_dtor])
fst_snd_convs];
val only_if_tac =
- EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE],
rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
CONJ_WRAP' (fn (ctor_set, passive_set_map0) =>
EVERY' [rtac ord_eq_le_trans, rtac ctor_set, rtac @{thm Un_least},
@@ -680,11 +681,11 @@
rtac @{thm SUP_cong[OF _ refl]}, rtac active_set_map0, rtac @{thm UN_least},
dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
dtac @{thm ssubst_mem[OF pair_collapse]},
- REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
+ REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE ::
@{thms case_prodE iffD1[OF prod.inject, elim_format]}),
hyp_subst_tac ctxt,
dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], atac])
(rev (active_set_map0s ~~ in_Irels))])
(ctor_sets ~~ passive_set_map0s),
rtac conjI,
@@ -693,7 +694,7 @@
REPEAT_DETERM_N m o rtac @{thm fun_cong[OF comp_id]},
EVERY' (map (fn in_Irel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
dtac @{thm ssubst_mem[OF pair_collapse]},
- REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
+ REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE ::
@{thms case_prodE iffD1[OF prod.inject, elim_format]}),
hyp_subst_tac ctxt,
dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac])
@@ -743,9 +744,9 @@
@{thms rel_fun_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN
unfold_thms_tac ctxt @{thms rel_fun_iff_leq_vimage2p} THEN
HEADGOAL (EVERY'
- [REPEAT_DETERM o resolve_tac [allI, impI], rtac ctor_rel_induct,
+ [REPEAT_DETERM o resolve_tac ctxt [allI, impI], rtac ctor_rel_induct,
EVERY' (map (fn map_transfer => EVERY'
- [REPEAT_DETERM o resolve_tac [allI, impI, @{thm vimage2pI}],
+ [REPEAT_DETERM o resolve_tac ctxt [allI, impI, @{thm vimage2pI}],
SELECT_GOAL (unfold_thms_tac ctxt folds),
etac @{thm predicate2D_vimage2p},
rtac (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer),
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Tue Feb 10 14:48:26 2015 +0100
@@ -93,7 +93,8 @@
(def, lthy')
end;
- fun tac ctxt thms = Class.intro_classes_tac [] THEN ALLGOALS (Proof_Context.fact_tac ctxt thms);
+ fun tac ctxt thms =
+ Class.intro_classes_tac ctxt [] THEN ALLGOALS (Proof_Context.fact_tac ctxt thms);
val qualify =
Binding.qualify true (Long_Name.base_name fcT_name) o Binding.qualify true eqN o Binding.name;
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -85,7 +85,7 @@
fun mk_exhaust_sel_tac n exhaust_disc collapses =
mk_exhaust_disc_or_sel_tac n exhaust_disc collapses ORELSE
- HEADGOAL (etac meta_mp THEN' resolve_tac collapses);
+ HEADGOAL (etac meta_mp THEN' resolve0_tac collapses);
fun mk_collapse_tac ctxt m discD sels =
HEADGOAL (dtac discD THEN'
@@ -101,7 +101,7 @@
(hyp_subst_tac ctxt THEN'
SELECT_GOAL (unfold_thms_tac ctxt (@{thms not_True_eq_False not_False_eq_True} @ cases @
((refl :: discs @ distincts) RL [eqTrueI, eqFalseI]))) THEN'
- resolve_tac @{thms TrueI True_not_False False_not_True}));
+ resolve_tac ctxt @{thms TrueI True_not_False False_not_True}));
fun mk_expand_tac ctxt n ms uexhaust_disc vexhaust_disc uncollapses distinct_discsss
distinct_discsss' =
--- a/src/HOL/Tools/Function/function_common.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/function_common.ML Tue Feb 10 14:48:26 2015 +0100
@@ -279,7 +279,8 @@
preproc1)
)
-val termination_rule_tac = resolve_tac o #1 o Data.get o Context.Proof
+fun termination_rule_tac ctxt = resolve_tac ctxt (#1 (Data.get (Context.Proof ctxt)))
+
val store_termination_rule = Data.map o @{apply 4(1)} o cons
val get_functions = #2 o Data.get o Context.Proof
--- a/src/HOL/Tools/Function/function_core.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/function_core.ML Tue Feb 10 14:48:26 2015 +0100
@@ -707,8 +707,8 @@
|> cterm_of thy
in
Goal.init goal
- |> (SINGLE (resolve_tac [accI] 1)) |> the
- |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
+ |> (SINGLE (resolve_tac ctxt [accI] 1)) |> the
+ |> (SINGLE (eresolve_tac ctxt [Thm.forall_elim_vars 0 R_cases] 1)) |> the
|> (SINGLE (auto_tac ctxt)) |> the
|> Goal.conclude
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
--- a/src/HOL/Tools/Function/function_elims.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/function_elims.ML Tue Feb 10 14:48:26 2015 +0100
@@ -60,8 +60,8 @@
fun bool_subst_tac ctxt i =
REPEAT (EqSubst.eqsubst_asm_tac ctxt [1] eq_bool i)
- THEN REPEAT (dresolve_tac boolD i)
- THEN REPEAT (eresolve_tac boolE i)
+ THEN REPEAT (dresolve_tac ctxt boolD i)
+ THEN REPEAT (eresolve_tac ctxt boolE i)
fun mk_bool_elims ctxt elim =
let
@@ -69,7 +69,7 @@
fun mk_bool_elim b =
elim
|> Thm.forall_elim b
- |> Tactic.rule_by_tactic ctxt (TRY (resolve_tac eq_boolI 1))
+ |> Tactic.rule_by_tactic ctxt (TRY (resolve_tac ctxt eq_boolI 1))
|> Tactic.rule_by_tactic ctxt tac;
in
map mk_bool_elim [@{cterm True}, @{cterm False}]
@@ -125,7 +125,7 @@
val asms_thms = map Thm.assume asms;
fun prep_subgoal_tac i =
- REPEAT (eresolve_tac @{thms Pair_inject} i)
+ REPEAT (eresolve_tac ctxt @{thms Pair_inject} i)
THEN Method.insert_tac (case asms_thms of thm :: thms => (thm RS sym) :: thms) i
THEN propagate_tac ctxt i
THEN TRY ((EqSubst.eqsubst_asm_tac ctxt [1] psimps i) THEN assume_tac ctxt i)
--- a/src/HOL/Tools/Function/function_lib.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/function_lib.ML Tue Feb 10 14:48:26 2015 +0100
@@ -113,7 +113,7 @@
then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
(K (rewrite_goals_tac ctxt ac
- THEN resolve_tac [Drule.reflexive_thm] 1))
+ THEN resolve_tac ctxt [Drule.reflexive_thm] 1))
end
(* instance for unions *)
--- a/src/HOL/Tools/Function/measure_functions.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/measure_functions.ML Tue Feb 10 14:48:26 2015 +0100
@@ -18,7 +18,7 @@
Const (@{const_name is_measure}, fastype_of t --> HOLogic.boolT) $ t
fun find_measures ctxt T =
- DEPTH_SOLVE (resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems measure_function})) 1)
+ DEPTH_SOLVE (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems measure_function})) 1)
(HOLogic.mk_Trueprop (mk_is_measure (Var (("f",0), T --> HOLogic.natT)))
|> cterm_of (Proof_Context.theory_of ctxt) |> Goal.init)
|> Seq.map (prop_of #> (fn _ $ (_ $ (_ $ f)) => f))
--- a/src/HOL/Tools/Function/mutual.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/mutual.ML Tue Feb 10 14:48:26 2015 +0100
@@ -267,8 +267,8 @@
@{thms HOL.notE[OF Sum_Type.sum.distinct(1)] HOL.notE[OF Sum_Type.sum.distinct(2)]}
fun prep_subgoal i =
- REPEAT (eresolve_tac @{thms Pair_inject Inl_inject[elim_format] Inr_inject[elim_format]} i)
- THEN REPEAT (Tactic.eresolve_tac sum_elims i)
+ REPEAT (eresolve_tac ctxt @{thms Pair_inject Inl_inject[elim_format] Inr_inject[elim_format]} i)
+ THEN REPEAT (eresolve_tac ctxt sum_elims i)
in
cases_rule
|> Thm.forall_elim P
--- a/src/HOL/Tools/Function/partial_function.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/partial_function.ML Tue Feb 10 14:48:26 2015 +0100
@@ -80,7 +80,7 @@
| _ => NONE;
(*split on case expressions*)
-val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context=ctxt, ...} =>
+val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} =>
SUBGOAL (fn (t, i) => case t of
_ $ (_ $ Abs (_, _, body)) =>
(case dest_case ctxt body of
@@ -90,7 +90,7 @@
if Term.is_open arg then no_tac
else ((DETERM o strip_cases o Induct.cases_tac ctxt false [[SOME arg]] NONE [])
THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0)
- THEN_ALL_NEW eresolve_tac @{thms thin_rl}
+ THEN_ALL_NEW eresolve_tac ctxt @{thms thin_rl}
THEN_ALL_NEW (CONVERSION
(params_conv ~1 (fn ctxt' =>
arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i
@@ -101,7 +101,7 @@
fun mono_tac ctxt =
K (Local_Defs.unfold_tac ctxt [@{thm curry_def}])
THEN' (TRY o REPEAT_ALL_NEW
- (resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems partial_function_mono}))
+ (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems partial_function_mono}))
ORELSE' split_cases_tac ctxt));
@@ -290,7 +290,7 @@
val rec_rule = let open Conv in
Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>
CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1
- THEN resolve_tac @{thms refl} 1) end;
+ THEN resolve_tac lthy' @{thms refl} 1) end;
in
lthy'
|> Local_Theory.note (eq_abinding, [rec_rule])
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Tue Feb 10 14:48:26 2015 +0100
@@ -282,7 +282,7 @@
THEN PRIMITIVE (instantiate' [] [SOME level_mapping])
THEN unfold_tac ctxt @{thms rp_inv_image_def}
THEN Local_Defs.unfold_tac ctxt @{thms split_conv fst_conv snd_conv}
- THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}]))
+ THEN REPEAT (SOMEGOAL (resolve_tac ctxt [@{thm Un_least}, @{thm empty_subsetI}]))
THEN EVERY (map (prove_lev true) sl)
THEN EVERY (map (prove_lev false) (subtract (op =) sl (0 upto length cs - 1))))
end
--- a/src/HOL/Tools/Lifting/lifting_def.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_def.ML Tue Feb 10 14:48:26 2015 +0100
@@ -33,8 +33,8 @@
val refl_rules = Lifting_Info.get_reflexivity_rules ctxt
val transfer_rules = Transfer.get_transfer_raw ctxt
- fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac refl_rules) THEN_ALL_NEW
- (REPEAT_ALL_NEW (DETERM o resolve_tac transfer_rules))) i
+ fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW
+ (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i
in
SOME (Goal.prove ctxt [] [] prop (K (main_tac 1)))
handle ERROR _ => NONE
@@ -545,7 +545,7 @@
fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
val eq_onp_assms_tac_rules = @{thm left_unique_OO} ::
eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw lthy)
- val eq_onp_assms_tac = (TRY o REPEAT_ALL_NEW (resolve_tac eq_onp_assms_tac_rules)
+ val eq_onp_assms_tac = (TRY o REPEAT_ALL_NEW (resolve_tac lthy eq_onp_assms_tac_rules)
THEN_ALL_NEW (DETERM o Transfer.eq_tac lthy)) 1
val relator_eq_onp_conv = Conv.bottom_conv
(K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac
--- a/src/HOL/Tools/Lifting/lifting_setup.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_setup.ML Tue Feb 10 14:48:26 2015 +0100
@@ -310,7 +310,8 @@
val init_goal = Goal.init (cterm_of thy fixed_goal)
val rules = Transfer.get_transfer_raw ctxt
val rules = constraint :: OO_rules @ rules
- val tac = K (Local_Defs.unfold_tac ctxt [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac rules)
+ val tac =
+ K (Local_Defs.unfold_tac ctxt [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac ctxt rules)
in
(singleton (Variable.export ctxt orig_ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal))
end
@@ -414,7 +415,8 @@
let
val goal = thm |> prems_of |> hd
val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var
- val reduced_assm = reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac rules) 1)) ctxt
+ val reduced_assm =
+ reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt
in
reduced_assm RS thm
end
@@ -424,7 +426,8 @@
fun reduce_first_assm ctxt rules thm =
let
val goal = thm |> prems_of |> hd
- val reduced_assm = reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac rules) 1)) ctxt
+ val reduced_assm =
+ reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt
in
reduced_assm RS thm
end
--- a/src/HOL/Tools/Lifting/lifting_term.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_term.ML Tue Feb 10 14:48:26 2015 +0100
@@ -473,7 +473,7 @@
fun prove_extra_assms ctxt ctm distr_rule =
let
fun prove_assm assm = try (Goal.prove ctxt [] [] (term_of assm))
- (fn _ => SOLVED' (REPEAT_ALL_NEW (resolve_tac (Transfer.get_transfer_raw ctxt))) 1)
+ (fn _ => SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt (Transfer.get_transfer_raw ctxt))) 1)
fun is_POS_or_NEG ctm =
case (head_of o term_of o Thm.dest_arg) ctm of
--- a/src/HOL/Tools/Meson/meson.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Meson/meson.ML Tue Feb 10 14:48:26 2015 +0100
@@ -170,14 +170,14 @@
property of the form "... c ... c ... c ..." will lead to a huge unification
problem, due to the (spurious) choices between projection and imitation. The
workaround is to instantiate "?P := (%c. ... c ... c ... c ...)" manually. *)
-fun quant_resolve_tac th i st =
+fun quant_resolve_tac ctxt th i st =
case (concl_of st, prop_of th) of
(@{const Trueprop} $ (Var _ $ (c as Free _)), @{const Trueprop} $ _) =>
let
val cc = cterm_of (theory_of_thm th) c
val ct = Thm.dest_arg (cprop_of th)
- in resolve_tac [th] i (Drule.instantiate' [] [SOME (Thm.lambda cc ct)] st) end
- | _ => resolve_tac [th] i st
+ in resolve_tac ctxt [th] i (Drule.instantiate' [] [SOME (Thm.lambda cc ct)] st) end
+ | _ => resolve_tac ctxt [th] i st
(*Permits forward proof from rules that discharge assumptions. The supplied proof state st,
e.g. from conj_forward, should have the form
@@ -185,7 +185,7 @@
and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*)
fun forward_res ctxt nf st =
let
- fun tacf [prem] = quant_resolve_tac (nf prem) 1
+ fun tacf [prem] = quant_resolve_tac ctxt (nf prem) 1
| tacf prems =
error (cat_lines
("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" ::
@@ -287,7 +287,7 @@
case Seq.pull
(REPEAT
(Misc_Legacy.METAHYPS ctxt
- (fn major::minors => resolve_tac [nf (minors @ hyps) major] 1) 1)
+ (fn major::minors => resolve_tac ctxt [nf (minors @ hyps) major] 1) 1)
st)
of SOME(th,_) => th
| NONE => raise THM("forward_res2", 0, [st]);
@@ -390,7 +390,7 @@
cnf_nil. The normal form is given to resolve_tac, instantiate a Boolean
variable created by resolution with disj_forward. Since (cnf_nil prem)
returns a LIST of theorems, we can backtrack to get all combinations.*)
- let val tac = Misc_Legacy.METAHYPS ctxt (fn [prem] => resolve_tac (cnf_nil prem) 1) 1
+ let val tac = Misc_Legacy.METAHYPS ctxt (fn [prem] => resolve_tac ctxt (cnf_nil prem) 1) 1
in Seq.list_of ((tac THEN tac) (th RS disj_forward)) @ ths end
| _ => nodups ctxt th :: ths (*no work to do*)
and cnf_nil th = cnf_aux (th, [])
@@ -500,7 +500,7 @@
(* resolve_from_net_tac actually made it slower... *)
fun prolog_step_tac ctxt horns i =
- (assume_tac ctxt i APPEND resolve_tac horns i) THEN check_tac THEN
+ (assume_tac ctxt i APPEND resolve_tac ctxt horns i) THEN check_tac THEN
TRYALL_eq_assume_tac;
(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*)
@@ -698,14 +698,14 @@
fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls));
fun skolemize_prems_tac ctxt prems =
- cut_facts_tac (maps (try_skolemize_etc ctxt) prems) THEN' REPEAT o eresolve_tac [exE]
+ cut_facts_tac (maps (try_skolemize_etc ctxt) prems) THEN' REPEAT o eresolve_tac ctxt [exE]
(*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions.
Function mkcl converts theorems to clauses.*)
fun MESON preskolem_tac mkcl cltac ctxt i st =
SELECT_GOAL
(EVERY [Object_Logic.atomize_prems_tac ctxt 1,
- resolve_tac @{thms ccontr} 1,
+ resolve_tac ctxt @{thms ccontr} 1,
preskolem_tac,
Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} =>
EVERY1 [skolemize_prems_tac ctxt negs,
@@ -719,7 +719,7 @@
fun best_meson_tac sizef ctxt =
MESON all_tac (make_clauses ctxt)
(fn cls =>
- THEN_BEST_FIRST (resolve_tac (gocls cls) 1)
+ THEN_BEST_FIRST (resolve_tac ctxt (gocls cls) 1)
(has_fewer_prems 1, sizef)
(prolog_step_tac ctxt (make_horns cls) 1))
ctxt
@@ -732,7 +732,7 @@
fun depth_meson_tac ctxt =
MESON all_tac (make_clauses ctxt)
- (fn cls => EVERY [resolve_tac (gocls cls) 1, depth_prolog_tac ctxt (make_horns cls)])
+ (fn cls => EVERY [resolve_tac ctxt (gocls cls) 1, depth_prolog_tac ctxt (make_horns cls)])
ctxt
(** Iterative deepening version **)
@@ -764,7 +764,7 @@
["clauses:"] @ map (Display.string_of_thm ctxt) horns))
in
THEN_ITER_DEEPEN iter_deepen_limit
- (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' ctxt horns)
+ (resolve_tac ctxt goes 1) (has_fewer_prems 1) (prolog_step_tac' ctxt horns)
end));
fun meson_tac ctxt ths =
--- a/src/HOL/Tools/Meson/meson_clausify.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Meson/meson_clausify.ML Tue Feb 10 14:48:26 2015 +0100
@@ -208,7 +208,7 @@
|> Drule.beta_conv cabs |> Thm.apply cTrueprop
fun tacf [prem] =
rewrite_goals_tac ctxt @{thms skolem_def [abs_def]}
- THEN resolve_tac [(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]})
+ THEN resolve_tac ctxt [(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]})
RS Global_Theory.get_thm thy "Hilbert_Choice.someI_ex"] 1
in
Goal.prove_internal ctxt [ex_tm] conc tacf
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML Tue Feb 10 14:48:26 2015 +0100
@@ -530,17 +530,19 @@
val copy_prem = @{lemma "P ==> (P ==> P ==> Q) ==> Q" by fast}
-fun copy_prems_tac [] ns i =
- if forall (curry (op =) 1) ns then all_tac else copy_prems_tac (rev ns) [] i
- | copy_prems_tac (1 :: ms) ns i = rotate_tac 1 i THEN copy_prems_tac ms (1 :: ns) i
- | copy_prems_tac (m :: ms) ns i =
- eresolve_tac [copy_prem] i THEN copy_prems_tac ms (m div 2 :: (m + 1) div 2 :: ns) i
+fun copy_prems_tac ctxt [] ns i =
+ if forall (curry (op =) 1) ns then all_tac else copy_prems_tac ctxt (rev ns) [] i
+ | copy_prems_tac ctxt (1 :: ms) ns i = rotate_tac 1 i THEN copy_prems_tac ctxt ms (1 :: ns) i
+ | copy_prems_tac ctxt (m :: ms) ns i =
+ eresolve_tac ctxt [copy_prem] i THEN copy_prems_tac ctxt ms (m div 2 :: (m + 1) div 2 :: ns) i
(* Metis generates variables of the form _nnn. *)
val is_metis_fresh_variable = String.isPrefix "_"
-fun instantiate_forall_tac thy t i st =
+fun instantiate_forall_tac ctxt t i st =
let
+ val thy = Proof_Context.theory_of ctxt
+
val params = Logic.strip_params (Logic.get_goal (prop_of st) i) |> rev
fun repair (t as (Var ((s, _), _))) =
@@ -581,16 +583,16 @@
end
| _ => raise Fail "expected a single non-zapped, non-Metis Var")
in
- (DETERM (eresolve_tac @{thms allE} i THEN rotate_tac ~1 i) THEN PRIMITIVE do_instantiate) st
+ (DETERM (eresolve_tac ctxt @{thms allE} i THEN rotate_tac ~1 i) THEN PRIMITIVE do_instantiate) st
end
-fun fix_exists_tac t = eresolve_tac [exE] THEN' rename_tac [t |> dest_Var |> fst |> fst]
+fun fix_exists_tac ctxt t = eresolve_tac ctxt [exE] THEN' rename_tac [t |> dest_Var |> fst |> fst]
-fun release_quantifier_tac thy (skolem, t) =
- (if skolem then fix_exists_tac else instantiate_forall_tac thy) t
+fun release_quantifier_tac ctxt (skolem, t) =
+ (if skolem then fix_exists_tac ctxt else instantiate_forall_tac ctxt) t
fun release_clusters_tac _ _ _ [] = K all_tac
- | release_clusters_tac thy ax_counts substs ((ax_no, cluster_no) :: clusters) =
+ | release_clusters_tac ctxt ax_counts substs ((ax_no, cluster_no) :: clusters) =
let
val cluster_of_var = Meson_Clausify.cluster_of_zapped_var_name o fst o fst o dest_Var
fun in_right_cluster ((_, (cluster_no', _)), _) = cluster_no' = cluster_no
@@ -605,13 +607,13 @@
else
NONE)
fun do_cluster_subst cluster_subst =
- map (release_quantifier_tac thy) cluster_subst @ [rotate_tac 1]
+ map (release_quantifier_tac ctxt) cluster_subst @ [rotate_tac 1]
val first_prem = find_index (fn (ax_no', _) => ax_no' = ax_no) substs
in
rotate_tac first_prem
THEN' (EVERY' (maps do_cluster_subst cluster_substs))
THEN' rotate_tac (~ first_prem - length cluster_substs)
- THEN' release_clusters_tac thy ax_counts substs clusters
+ THEN' release_clusters_tac ctxt ax_counts substs clusters
end
fun cluster_key ((ax_no, (cluster_no, index_no)), skolem) =
@@ -731,21 +733,20 @@
val _ = tracing ("SUBSTS (" ^ string_of_int (length substs) ^ "):\n" ^
cat_lines (map string_of_subst_info substs))
*)
+ val ctxt' = fold Thm.declare_hyps (#hyps (Thm.crep_thm prems_imp_false)) ctxt
fun cut_and_ex_tac axiom =
- cut_tac axiom 1 THEN TRY (REPEAT_ALL_NEW (eresolve_tac @{thms exE}) 1)
+ cut_tac axiom 1 THEN TRY (REPEAT_ALL_NEW (eresolve_tac ctxt' @{thms exE}) 1)
fun rotation_of_subgoal i =
find_index (fn (_, (subgoal_no, _)) => subgoal_no = i) substs
-
- val ctxt' = fold Thm.declare_hyps (#hyps (Thm.crep_thm prems_imp_false)) ctxt
in
Goal.prove ctxt' [] [] @{prop False}
(K (DETERM (EVERY (map (cut_and_ex_tac o fst o the o nth axioms o fst o fst) ax_counts)
THEN rename_tac outer_param_names 1
- THEN copy_prems_tac (map snd ax_counts) [] 1)
- THEN release_clusters_tac thy ax_counts substs ordered_clusters 1
+ THEN copy_prems_tac ctxt' (map snd ax_counts) [] 1)
+ THEN release_clusters_tac ctxt' ax_counts substs ordered_clusters 1
THEN match_tac ctxt' [prems_imp_false] 1
- THEN ALLGOALS (fn i => resolve_tac @{thms Meson.skolem_COMBK_I} i
+ THEN ALLGOALS (fn i => resolve_tac ctxt' @{thms Meson.skolem_COMBK_I} i
THEN rotate_tac (rotation_of_subgoal i) i
THEN PRIMITIVE (unify_first_prem_with_concl thy i)
THEN assume_tac ctxt' i
--- a/src/HOL/Tools/Metis/metis_tactic.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Metis/metis_tactic.ML Tue Feb 10 14:48:26 2015 +0100
@@ -61,7 +61,7 @@
fun lam_lifted_of_metis ctxt type_enc sym_tab concealed mth =
let
val thy = Proof_Context.theory_of ctxt
- val tac = rewrite_goals_tac ctxt @{thms lambda_def [abs_def]} THEN resolve_tac [refl] 1
+ val tac = rewrite_goals_tac ctxt @{thms lambda_def [abs_def]} THEN resolve_tac ctxt [refl] 1
val t = hol_clause_of_metis ctxt type_enc sym_tab concealed mth
val ct = cterm_of thy (HOLogic.mk_Trueprop t)
in Goal.prove_internal ctxt [] ct (K tac) |> Meson.make_meta_clause end
@@ -102,7 +102,7 @@
so that "Thm.equal_elim" works below. *)
val t0 $ _ $ t2 = prop_of eq_th
val eq_ct = t0 $ prop_of th $ t2 |> cterm_of thy
- val eq_th' = Goal.prove_internal ctxt [] eq_ct (K (resolve_tac [eq_th] 1))
+ val eq_th' = Goal.prove_internal ctxt [] eq_ct (K (resolve_tac ctxt [eq_th] 1))
in Thm.equal_elim eq_th' th end
fun clause_params ordering =
@@ -257,7 +257,7 @@
"Metis called with theorems\n" ^ cat_lines (map (Display.string_of_thm ctxt) ths))
val type_encs = type_encs |> maps unalias_type_enc
val combs = (lam_trans = combsN)
- fun tac clause = resolve_tac (FOL_SOLVE unused type_encs lam_trans ctxt clause ths) 1
+ fun tac clause = resolve_tac ctxt (FOL_SOLVE unused type_encs lam_trans ctxt clause ths) 1
val seq = Meson.MESON (preskolem_tac ctxt) (maps (neg_clausify ctxt combs)) tac ctxt i st0
in
(!unused, seq)
--- a/src/HOL/Tools/Old_Datatype/old_datatype.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML Tue Feb 10 14:48:26 2015 +0100
@@ -188,9 +188,9 @@
(fn (((name, mx), tvs), c) =>
Typedef.add_typedef_global false (name, tvs, mx)
(Collect $ Const (c, UnivT')) NONE
- (fn _ => rtac exI 1 THEN rtac CollectI 1 THEN
+ (fn ctxt => rtac exI 1 THEN rtac CollectI 1 THEN
QUIET_BREADTH_FIRST (has_fewer_prems 1)
- (resolve_tac rep_intrs 1)))
+ (resolve_tac ctxt rep_intrs 1)))
(types_syntax ~~ tyvars ~~ take (length newTs) rep_set_names)
||> Sign.add_path big_name;
@@ -425,15 +425,15 @@
REPEAT (EVERY
[hyp_subst_tac ctxt 1,
rewrite_goals_tac ctxt rewrites,
- REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
- (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
+ REPEAT (dresolve_tac ctxt [In0_inject, In1_inject] 1),
+ (eresolve_tac ctxt [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
ORELSE (EVERY
- [REPEAT (eresolve_tac (Scons_inject ::
+ [REPEAT (eresolve_tac ctxt (Scons_inject ::
map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
REPEAT (cong_tac ctxt 1), rtac refl 1,
REPEAT (assume_tac ctxt 1 ORELSE (EVERY
[REPEAT (rtac @{thm ext} 1),
- REPEAT (eresolve_tac (mp :: allE ::
+ REPEAT (eresolve_tac ctxt (mp :: allE ::
map make_elim (Suml_inject :: Sumr_inject ::
Lim_inject :: inj_thms') @ fun_congs) 1),
assume_tac ctxt 1]))])])])]);
@@ -447,7 +447,7 @@
EVERY [
(Old_Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1,
rewrite_goals_tac ctxt rewrites,
- REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
+ REPEAT ((resolve_tac ctxt rep_intrs THEN_ALL_NEW
((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
in (inj_thms'' @ inj_thms, elem_thms @ Old_Datatype_Aux.split_conj_thm elem_thm) end;
@@ -488,11 +488,11 @@
Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
rewrite_goals_tac ctxt (map Thm.symmetric range_eqs),
REPEAT (EVERY
- [REPEAT (eresolve_tac ([rangeE, @{thm ex1_implies_ex} RS exE] @
+ [REPEAT (eresolve_tac ctxt ([rangeE, @{thm ex1_implies_ex} RS exE] @
maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
TRY (hyp_subst_tac ctxt 1),
rtac (sym RS range_eqI) 1,
- resolve_tac iso_char_thms 1])])));
+ resolve_tac ctxt iso_char_thms 1])])));
val Abs_inverse_thms' =
map #1 newT_iso_axms @
@@ -514,11 +514,11 @@
in
Goal.prove_sorry_global thy5 [] [] eqn
(fn {context = ctxt, ...} => EVERY
- [resolve_tac inj_thms 1,
+ [resolve_tac ctxt inj_thms 1,
rewrite_goals_tac ctxt rewrites,
rtac refl 3,
- resolve_tac rep_intrs 2,
- REPEAT (resolve_tac iso_elem_thms 1)])
+ resolve_tac ctxt rep_intrs 2,
+ REPEAT (resolve_tac ctxt iso_elem_thms 1)])
end;
(*--------------------------------------------------------------*)
@@ -560,11 +560,11 @@
(fn {context = ctxt, ...} => EVERY
[rtac iffI 1,
REPEAT (etac conjE 2), hyp_subst_tac ctxt 2, rtac refl 2,
- dresolve_tac rep_congs 1, dtac @{thm box_equals} 1,
- REPEAT (resolve_tac rep_thms 1),
- REPEAT (eresolve_tac inj_thms 1),
+ dresolve_tac ctxt rep_congs 1, dtac @{thm box_equals} 1,
+ REPEAT (resolve_tac ctxt rep_thms 1),
+ REPEAT (eresolve_tac ctxt inj_thms 1),
REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac @{thm ext} 1),
- REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
+ REPEAT (eresolve_tac ctxt (make_elim fun_cong :: inj_thms) 1),
assume_tac ctxt 1]))])
end;
@@ -616,12 +616,12 @@
(Logic.mk_implies
(HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems),
HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls)))
- (fn _ =>
+ (fn {context = ctxt, ...} =>
EVERY
[REPEAT (etac conjE 1),
REPEAT (EVERY
- [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
- etac mp 1, resolve_tac iso_elem_thms 1])]);
+ [TRY (rtac conjI 1), resolve_tac ctxt Rep_inverse_thms 1,
+ etac mp 1, resolve_tac ctxt iso_elem_thms 1])]);
val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
val frees =
@@ -640,7 +640,7 @@
(Old_Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW
Object_Logic.atomize_prems_tac ctxt) 1,
EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms 1),
+ [REPEAT (eresolve_tac ctxt Abs_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)]))
--- a/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Tue Feb 10 14:48:26 2015 +0100
@@ -150,7 +150,7 @@
NONE => NONE
| SOME u' => SOME (t |> getP |> snd |> head_of |> cert, cert u'))) (ts ~~ ts');
val indrule' = cterm_instantiate insts indrule;
- in resolve_tac [indrule'] i end);
+ in resolve0_tac [indrule'] i end);
(* perform exhaustive case analysis on last parameter of subgoal i *)
--- a/src/HOL/Tools/Old_Datatype/old_primrec.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_primrec.ML Tue Feb 10 14:48:26 2015 +0100
@@ -248,7 +248,7 @@
in
map (fn eq => Goal.prove ctxt frees [] eq
(fn {context = ctxt', ...} =>
- EVERY [rewrite_goals_tac ctxt' rewrites, resolve_tac [refl] 1])) eqs
+ EVERY [rewrite_goals_tac ctxt' rewrites, resolve_tac ctxt' [refl] 1])) eqs
end;
in ((prefix, (fs, defs)), prove) end
handle PrimrecError (msg, some_eqn) =>
--- a/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Tue Feb 10 14:48:26 2015 +0100
@@ -55,12 +55,12 @@
Goal.prove_sorry_global thy []
(Logic.strip_imp_prems t)
(Logic.strip_imp_concl t)
- (fn {prems, ...} =>
+ (fn {context = ctxt, prems, ...} =>
EVERY
- [resolve_tac [induct'] 1,
- REPEAT (resolve_tac [TrueI] 1),
- REPEAT ((resolve_tac [impI] 1) THEN (eresolve_tac prems 1)),
- REPEAT (resolve_tac [TrueI] 1)])
+ [resolve_tac ctxt [induct'] 1,
+ REPEAT (resolve_tac ctxt [TrueI] 1),
+ REPEAT ((resolve_tac ctxt [impI] 1) THEN (eresolve_tac ctxt prems 1)),
+ REPEAT (resolve_tac ctxt [TrueI] 1)])
end;
val casedist_thms =
@@ -176,16 +176,17 @@
in
(EVERY
[DETERM tac,
- REPEAT (eresolve_tac @{thms ex1E} 1), resolve_tac @{thms ex1I} 1,
+ REPEAT (eresolve_tac ctxt @{thms ex1E} 1), resolve_tac ctxt @{thms ex1I} 1,
DEPTH_SOLVE_1 (ares_tac [intr] 1),
- REPEAT_DETERM_N k (eresolve_tac [thin_rl] 1 THEN rotate_tac 1 1),
- eresolve_tac [elim] 1,
+ REPEAT_DETERM_N k (eresolve_tac ctxt [thin_rl] 1 THEN rotate_tac 1 1),
+ eresolve_tac ctxt [elim] 1,
REPEAT_DETERM_N j distinct_tac,
- TRY (dresolve_tac inject 1),
- REPEAT (eresolve_tac [conjE] 1), hyp_subst_tac ctxt 1,
- REPEAT (EVERY [eresolve_tac [allE] 1, dresolve_tac [mp] 1, assume_tac ctxt 1]),
+ TRY (dresolve_tac ctxt inject 1),
+ REPEAT (eresolve_tac ctxt [conjE] 1), hyp_subst_tac ctxt 1,
+ REPEAT
+ (EVERY [eresolve_tac ctxt [allE] 1, dresolve_tac ctxt [mp] 1, assume_tac ctxt 1]),
TRY (hyp_subst_tac ctxt 1),
- resolve_tac [refl] 1,
+ resolve_tac ctxt [refl] 1,
REPEAT_DETERM_N (n - j - 1) distinct_tac],
intrs, j + 1)
end;
@@ -211,7 +212,7 @@
(HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj rec_unique_ts))
(fn {context = ctxt, ...} =>
#1 (fold (mk_unique_tac ctxt) (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
- (((resolve_tac [induct'] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1 THEN
+ (((resolve_tac ctxt [induct'] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1 THEN
rewrite_goals_tac ctxt [mk_meta_eq @{thm choice_eq}], rec_intrs)))))
end;
@@ -254,10 +255,10 @@
Goal.prove_sorry_global thy2 [] [] t
(fn {context = ctxt, ...} => EVERY
[rewrite_goals_tac ctxt reccomb_defs,
- resolve_tac @{thms the1_equality} 1,
- resolve_tac rec_unique_thms 1,
- resolve_tac rec_intrs 1,
- REPEAT (resolve_tac [allI] 1 ORELSE resolve_tac rec_total_thms 1)]))
+ resolve_tac ctxt @{thms the1_equality} 1,
+ resolve_tac ctxt rec_unique_thms 1,
+ resolve_tac ctxt rec_intrs 1,
+ REPEAT (resolve_tac ctxt [allI] 1 ORELSE resolve_tac ctxt rec_total_thms 1)]))
(Old_Datatype_Prop.make_primrecs reccomb_names descr thy2);
in
thy2
@@ -339,7 +340,7 @@
fun prove_case t =
Goal.prove_sorry_global thy2 [] [] t (fn {context = ctxt, ...} =>
EVERY [rewrite_goals_tac ctxt (case_defs @ map mk_meta_eq primrec_thms),
- resolve_tac [refl] 1]);
+ resolve_tac ctxt [refl] 1]);
fun prove_cases (Type (Tcon, _)) ts =
(case Ctr_Sugar.ctr_sugar_of ctxt Tcon of
@@ -380,7 +381,7 @@
val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion;
fun tac ctxt =
- EVERY [resolve_tac [exhaustion'] 1,
+ EVERY [resolve_tac ctxt [exhaustion'] 1,
ALLGOALS (asm_simp_tac
(put_simpset HOL_ss ctxt addsimps (dist_rewrites' @ inject @ case_thms')))];
in
@@ -406,7 +407,8 @@
let
fun prove_case_cong_weak t =
Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} => EVERY [resolve_tac [hd prems RS arg_cong] 1]);
+ (fn {context = ctxt, prems, ...} =>
+ EVERY [resolve_tac ctxt [hd prems RS arg_cong] 1]);
val case_cong_weaks =
map prove_case_cong_weak (Old_Datatype_Prop.make_case_cong_weaks case_names descr thy);
@@ -424,13 +426,13 @@
let
(* For goal i, select the correct disjunct to attack, then prove it *)
fun tac ctxt i 0 =
- EVERY [TRY (resolve_tac [disjI1] i), hyp_subst_tac ctxt i,
- REPEAT (resolve_tac [exI] i), resolve_tac [refl] i]
- | tac ctxt i n = resolve_tac [disjI2] i THEN tac ctxt i (n - 1);
+ EVERY [TRY (resolve_tac ctxt [disjI1] i), hyp_subst_tac ctxt i,
+ REPEAT (resolve_tac ctxt [exI] i), resolve_tac ctxt [refl] i]
+ | tac ctxt i n = resolve_tac ctxt [disjI2] i THEN tac ctxt i (n - 1);
in
Goal.prove_sorry_global thy [] [] t
(fn {context = ctxt, ...} =>
- EVERY [resolve_tac [allI] 1,
+ EVERY [resolve_tac ctxt [allI] 1,
Old_Datatype_Aux.exh_tac ctxt (K exhaustion) 1,
ALLGOALS (fn i => tac ctxt i (i - 1))])
end;
@@ -459,8 +461,9 @@
EVERY [
simp_tac (put_simpset HOL_ss ctxt addsimps [hd prems]) 1,
cut_tac nchotomy'' 1,
- REPEAT (eresolve_tac [disjE] 1 THEN REPEAT (eresolve_tac [exE] 1) THEN simplify 1),
- REPEAT (eresolve_tac [exE] 1) THEN simplify 1 (* Get last disjunct *)]
+ REPEAT (eresolve_tac ctxt [disjE] 1 THEN
+ REPEAT (eresolve_tac ctxt [exE] 1) THEN simplify 1),
+ REPEAT (eresolve_tac ctxt [exE] 1) THEN simplify 1 (* Get last disjunct *)]
end)
end;
--- a/src/HOL/Tools/Old_Datatype/old_size.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_size.ML Tue Feb 10 14:48:26 2015 +0100
@@ -131,7 +131,7 @@
||> Class.instantiation (tycos, map dest_TFree paramTs, [HOLogic.class_size])
||>> fold_map define_overloaded
(def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1))
- ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
+ ||> Class.prove_instantiation_instance (fn ctxt => Class.intro_classes_tac ctxt [])
||> Local_Theory.exit_global;
val ctxt = Proof_Context.init_global thy';
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Tue Feb 10 14:48:26 2015 +0100
@@ -1216,7 +1216,7 @@
HOLogic.mk_Trueprop (list_comb (const, map Free vs')))
val intro =
Goal.prove (Proof_Context.init_global thy1) (map fst vs') [] t
- (fn _ => ALLGOALS Skip_Proof.cheat_tac)
+ (fn {context = ctxt, ...} => ALLGOALS (Skip_Proof.cheat_tac ctxt))
in
((((full_constname, constT), vs'), intro), thy1)
end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Tue Feb 10 14:48:26 2015 +0100
@@ -151,7 +151,7 @@
val t' = Pattern.rewrite_term thy rewr [] t
val th'' =
Goal.prove ctxt (Term.add_free_names t' []) [] t'
- (fn _ => ALLGOALS Skip_Proof.cheat_tac)
+ (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt))
val th''' = Local_Defs.unfold ctxt [@{thm split_conv}, @{thm fst_conv}, @{thm snd_conv}] th''
in
th'''
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Tue Feb 10 14:48:26 2015 +0100
@@ -135,7 +135,7 @@
(flatten constname) (map prop_of intros) ([], thy)
val ctxt'' = Proof_Context.transfer thy' ctxt'
val intros'' =
- map (fn t => Goal.prove ctxt'' [] [] t (fn _ => ALLGOALS Skip_Proof.cheat_tac)) intros'
+ map (fn t => Goal.prove ctxt'' [] [] t (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt''))) intros'
|> Variable.export ctxt'' ctxt
in
(intros'', (local_defs, thy'))
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Tue Feb 10 14:48:26 2015 +0100
@@ -502,7 +502,7 @@
THEN trace_tac ctxt options "proved one direction"
THEN prove_other_direction options ctxt pred mode moded_clauses
THEN trace_tac ctxt options "proved other direction")
- else (fn _ => ALLGOALS Skip_Proof.cheat_tac))
+ else (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt)))
end
end
\ No newline at end of file
--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML Tue Feb 10 14:48:26 2015 +0100
@@ -211,7 +211,7 @@
|> Quickcheck_Common.define_functions
(fn functerms => mk_equations functerms (descr, vs, Ts @ Us), NONE)
prfx argnames fullnames (map mk_T (Ts @ Us))
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
end
else
(Old_Datatype_Aux.message config
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML Tue Feb 10 14:48:26 2015 +0100
@@ -54,7 +54,7 @@
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
|> snd
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
end;
fun ensure_partial_term_of (tyco, (raw_vs, _)) thy =
@@ -187,7 +187,7 @@
|> Quickcheck_Common.define_functions
(fn narrowings => mk_equations descr vs narrowings, NONE)
prfx [] narrowingsN (map narrowingT (Ts @ Us))
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
else
thy
end;
--- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML Tue Feb 10 14:48:26 2015 +0100
@@ -379,7 +379,7 @@
let
val eqs_t = mk_equations consts
val eqs = map (fn eq => Goal.prove lthy argnames [] eq
- (fn _ => ALLGOALS Skip_Proof.cheat_tac)) eqs_t
+ (fn {context = ctxt, ...} => ALLGOALS (Skip_Proof.cheat_tac ctxt))) eqs_t
in
fold (fn (name, eq) => Local_Theory.note
((Binding.qualify true prfx
--- a/src/HOL/Tools/Quickcheck/random_generators.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML Tue Feb 10 14:48:26 2015 +0100
@@ -263,7 +263,7 @@
Specification.definition (NONE, (apfst Binding.conceal
Attrib.empty_binding, random_def))) random_defs')
|> snd
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
end;
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Tue Feb 10 14:48:26 2015 +0100
@@ -55,14 +55,14 @@
(** solvers for equivp and quotient assumptions **)
fun equiv_tac ctxt =
- REPEAT_ALL_NEW (resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv})))
+ REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv})))
val equiv_solver = mk_solver "Equivalence goal solver" equiv_tac
fun quotient_tac ctxt =
(REPEAT_ALL_NEW (FIRST'
[rtac @{thm identity_quotient3},
- resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems quot_thm}))]))
+ resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems quot_thm}))]))
val quotient_solver = mk_solver "Quotient goal solver" quotient_tac
@@ -167,10 +167,10 @@
in
simp_tac simpset THEN'
TRY o REPEAT_ALL_NEW (CHANGED o FIRST'
- [resolve_tac @{thms ball_reg_right bex_reg_left bex1_bexeq_reg},
- resolve_tac (Inductive.get_monos ctxt),
- resolve_tac @{thms ball_all_comm bex_ex_comm},
- resolve_tac eq_eqvs,
+ [resolve_tac ctxt @{thms ball_reg_right bex_reg_left bex1_bexeq_reg},
+ resolve_tac ctxt (Inductive.get_monos ctxt),
+ resolve_tac ctxt @{thms ball_all_comm bex_ex_comm},
+ resolve_tac ctxt eq_eqvs,
simp_tac simpset])
end
@@ -371,7 +371,7 @@
(* respectfulness of constants; in particular of a simple relation *)
| _ $ (Const _) $ (Const _) (* rel_fun, list_rel, etc but not equality *)
- => resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems quot_respect}))
+ => resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems quot_respect}))
THEN_ALL_NEW quotient_tac ctxt
(* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
@@ -403,7 +403,7 @@
(* reflexivity of the basic relations *)
(* R ... ... *)
- resolve_tac rel_refl]
+ resolve_tac ctxt rel_refl]
fun injection_tac ctxt =
let
--- a/src/HOL/Tools/SMT/z3_replay.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/SMT/z3_replay.ML Tue Feb 10 14:48:26 2015 +0100
@@ -147,10 +147,10 @@
by (metis someI_ex)+}
in
-fun discharge_sk_tac i st =
+fun discharge_sk_tac ctxt i st =
(rtac @{thm trans} i
- THEN resolve_tac sk_rules i
- THEN (rtac @{thm refl} ORELSE' discharge_sk_tac) (i+1)
+ THEN resolve_tac ctxt sk_rules i
+ THEN (rtac @{thm refl} ORELSE' discharge_sk_tac ctxt) (i+1)
THEN rtac @{thm refl} i) st
end
@@ -163,14 +163,16 @@
"(P | ~ P) & (~ P | P)"
by fast+}
-fun discharge_assms_tac rules =
- REPEAT (HEADGOAL (resolve_tac (intro_def_rules @ rules) ORELSE' SOLVED' discharge_sk_tac))
+fun discharge_assms_tac ctxt rules =
+ REPEAT
+ (HEADGOAL (resolve_tac ctxt (intro_def_rules @ rules) ORELSE'
+ SOLVED' (discharge_sk_tac ctxt)))
fun discharge_assms ctxt rules thm =
(if Thm.nprems_of thm = 0 then
thm
else
- (case Seq.pull (discharge_assms_tac rules thm) of
+ (case Seq.pull (discharge_assms_tac ctxt rules thm) of
SOME (thm', _) => thm'
| NONE => raise THM ("failed to discharge premise", 1, [thm])))
|> Goal.norm_result ctxt
--- a/src/HOL/Tools/SMT/z3_replay_methods.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/SMT/z3_replay_methods.ML Tue Feb 10 14:48:26 2015 +0100
@@ -309,7 +309,7 @@
(* congruence *)
fun ctac ctxt prems i st = st |> (
- resolve_tac (@{thm refl} :: prems) i
+ resolve_tac ctxt (@{thm refl} :: prems) i
ORELSE (cong_tac ctxt i THEN ctac ctxt prems (i + 1) THEN ctac ctxt prems i))
fun cong_basic ctxt thms t =
@@ -328,7 +328,7 @@
fun cong_full ctxt thms t = prove ctxt t (fn ctxt' =>
Method.insert_tac thms
THEN' (Classical.fast_tac ctxt'
- ORELSE' dresolve_tac cong_dest_rules
+ ORELSE' dresolve_tac ctxt cong_dest_rules
THEN' Classical.fast_tac ctxt'))
fun cong ctxt thms = try_provers ctxt Z3_Proof.Monotonicity [
@@ -346,7 +346,7 @@
by fast+}
fun quant_intro ctxt [thm] t =
- prove ctxt t (K (REPEAT_ALL_NEW (resolve_tac (thm :: quant_intro_rules))))
+ prove ctxt t (K (REPEAT_ALL_NEW (resolve_tac ctxt (thm :: quant_intro_rules))))
| quant_intro ctxt thms t = replay_rule_error ctxt Z3_Proof.Quant_Intro thms t
--- a/src/HOL/Tools/TFL/post.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/TFL/post.ML Tue Feb 10 14:48:26 2015 +0100
@@ -125,7 +125,7 @@
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
fun meta_outer ctxt =
curry_rule ctxt o Drule.export_without_context o
- rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE)));
+ rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' etac conjE)));
(*Strip off the outer !P*)
val spec'= Rule_Insts.read_instantiate @{context} [(("x", 0), "P::'b=>bool")] [] spec;
--- a/src/HOL/Tools/Transfer/transfer.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Transfer/transfer.ML Tue Feb 10 14:48:26 2015 +0100
@@ -611,12 +611,12 @@
|> Thm.instantiate (map tinst binsts, map inst binsts)
end
-fun eq_rules_tac eq_rules = TRY o REPEAT_ALL_NEW (resolve_tac eq_rules)
+fun eq_rules_tac eq_rules = TRY o REPEAT_ALL_NEW (resolve0_tac eq_rules)
THEN_ALL_NEW rtac @{thm is_equality_eq}
fun eq_tac ctxt = eq_rules_tac (get_relator_eq_raw ctxt)
-fun transfer_step_tac ctxt = (REPEAT_ALL_NEW (resolve_tac (get_transfer_raw ctxt))
+fun transfer_step_tac ctxt = (REPEAT_ALL_NEW (resolve_tac ctxt (get_transfer_raw ctxt))
THEN_ALL_NEW (DETERM o eq_rules_tac (get_relator_eq_raw ctxt)))
fun transfer_tac equiv ctxt i =
@@ -633,7 +633,9 @@
rtac start_rule i THEN
(rtac (transfer_rule_of_term ctxt equiv (HOLogic.dest_Trueprop t))
THEN_ALL_NEW
- (SOLVED' (REPEAT_ALL_NEW (resolve_tac rules) THEN_ALL_NEW (DETERM o eq_rules_tac eq_rules))
+ (SOLVED'
+ (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW
+ (DETERM o eq_rules_tac eq_rules))
ORELSE' end_tac)) (i + 1)
handle TERM (_, ts) => raise TERM (err_msg, ts)
in
@@ -658,7 +660,8 @@
rtac @{thm transfer_prover_start} i,
((rtac rule1 ORELSE' (CONVERSION expand_eq_in_rel THEN' rtac rule1))
THEN_ALL_NEW
- (REPEAT_ALL_NEW (resolve_tac rules) THEN_ALL_NEW (DETERM o eq_rules_tac eq_rules))) (i+1),
+ (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW
+ (DETERM o eq_rules_tac eq_rules))) (i + 1),
rtac @{thm refl} i]
end)
@@ -682,10 +685,10 @@
val t = HOLogic.dest_Trueprop (Thm.concl_of thm2)
val rule = transfer_rule_of_lhs ctxt' t
val tac =
- resolve_tac [thm2 RS start_rule', thm2 RS start_rule] 1 THEN
+ resolve_tac ctxt' [thm2 RS start_rule', thm2 RS start_rule] 1 THEN
(rtac rule
THEN_ALL_NEW
- (SOLVED' (REPEAT_ALL_NEW (resolve_tac rules)
+ (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules)
THEN_ALL_NEW (DETERM o eq_rules_tac eq_rules)))) 1
handle TERM (_, ts) => raise TERM (err_msg, ts)
val thm3 = Goal.prove_internal ctxt' [] @{cpat "Trueprop ?P"} (K tac)
@@ -721,7 +724,7 @@
rtac (thm2 RS start_rule) 1 THEN
(rtac rule
THEN_ALL_NEW
- (SOLVED' (REPEAT_ALL_NEW (resolve_tac rules)
+ (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules)
THEN_ALL_NEW (DETERM o eq_rules_tac eq_rules)))) 1
handle TERM (_, ts) => raise TERM (err_msg, ts)
val thm3 = Goal.prove_internal ctxt' [] @{cpat "Trueprop ?P"} (K tac)
--- a/src/HOL/Tools/Transfer/transfer_bnf.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/Transfer/transfer_bnf.ML Tue Feb 10 14:48:26 2015 +0100
@@ -201,19 +201,20 @@
in
EVERY' [rtac ext, SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm Domainp.simps},
in_rel_of_bnf bnf, pred_def]), rtac iffI,
- REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE], hyp_subst_tac ctxt,
+ REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], hyp_subst_tac ctxt,
CONJ_WRAP' (fn set_map => EVERY' [rtac ballI, dtac (set_map RS equalityD1 RS set_mp),
etac imageE, dtac set_rev_mp, assume_tac ctxt,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm case_prodE}],
+ REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm case_prodE}],
hyp_subst_tac ctxt, rtac @{thm iffD2[OF arg_cong2[of _ _ _ _ Domainp, OF refl fst_conv]]},
etac @{thm DomainPI}]) set_map's,
- REPEAT_DETERM o etac conjE, REPEAT_DETERM o resolve_tac [exI, (refl RS conjI), rotate_prems 1 conjI],
+ REPEAT_DETERM o etac conjE,
+ REPEAT_DETERM o resolve_tac ctxt [exI, (refl RS conjI), rotate_prems 1 conjI],
rtac refl, rtac (box_equals OF [map_cong0_of_bnf bnf, map_comp_of_bnf bnf RS sym,
map_id_of_bnf bnf]),
REPEAT_DETERM_N n o (EVERY' [rtac @{thm box_equals[OF _ sym[OF o_apply] sym[OF id_apply]]},
rtac @{thm fst_conv}]), rtac CollectI,
CONJ_WRAP' (fn set_map => EVERY' [rtac (set_map RS @{thm ord_eq_le_trans}),
- REPEAT_DETERM o resolve_tac [@{thm image_subsetI}, CollectI, @{thm case_prodI}],
+ REPEAT_DETERM o resolve_tac ctxt [@{thm image_subsetI}, CollectI, @{thm case_prodI}],
dtac (rotate_prems 1 bspec), assume_tac ctxt,
etac @{thm DomainpE}, etac @{thm someI}]) set_map's
] i
--- a/src/HOL/Tools/cnf.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/cnf.ML Tue Feb 10 14:48:26 2015 +0100
@@ -39,7 +39,7 @@
val is_clause: term -> bool
val clause_is_trivial: term -> bool
- val clause2raw_thm: thm -> thm
+ val clause2raw_thm: Proof.context -> thm -> thm
val make_nnf_thm: theory -> term -> thm
val weakening_tac: Proof.context -> int -> tactic (* removes the first hypothesis of a subgoal *)
@@ -132,7 +132,7 @@
(* where each xi' is the negation normal form of ~xi *)
(* ------------------------------------------------------------------------- *)
-fun clause2raw_thm clause =
+fun clause2raw_thm ctxt clause =
let
(* eliminates negated disjunctions from the i-th premise, possibly *)
(* adding new premises, then continues with the (i+1)-th premise *)
@@ -141,7 +141,8 @@
if i > nprems_of thm then
thm
else
- not_disj_to_prem (i+1) (Seq.hd (REPEAT_DETERM (resolve_tac [clause2raw_not_disj] i) thm))
+ not_disj_to_prem (i+1)
+ (Seq.hd (REPEAT_DETERM (resolve_tac ctxt [clause2raw_not_disj] i) thm))
(* moves all premises to hyps, i.e. "[...] |- A1 ==> ... ==> An ==> B" *)
(* becomes "[..., A1, ..., An] |- B" *)
(* Thm.thm -> Thm.thm *)
@@ -154,7 +155,7 @@
(* [...] |- ~x1 ==> ... ==> ~xn ==> False *)
|> not_disj_to_prem 1
(* [...] |- x1' ==> ... ==> xn' ==> False *)
- |> Seq.hd o TRYALL (resolve_tac [clause2raw_not_not])
+ |> Seq.hd o TRYALL (resolve_tac ctxt [clause2raw_not_not])
(* [..., x1', ..., xn'] |- False *)
|> prems_to_hyps
end;
@@ -529,7 +530,7 @@
(* ------------------------------------------------------------------------- *)
fun weakening_tac ctxt i =
- dresolve_tac [weakening_thm] i THEN assume_tac ctxt (i+1);
+ dresolve_tac ctxt [weakening_thm] i THEN assume_tac ctxt (i+1);
(* ------------------------------------------------------------------------- *)
(* cnf_rewrite_tac: converts all premises of the 'i'-th subgoal to CNF *)
--- a/src/HOL/Tools/code_evaluation.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/code_evaluation.ML Tue Feb 10 14:48:26 2015 +0100
@@ -45,7 +45,7 @@
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
|> snd
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
end;
fun ensure_term_of (tyco, (raw_vs, _)) thy =
--- a/src/HOL/Tools/coinduction.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/coinduction.ML Tue Feb 10 14:48:26 2015 +0100
@@ -34,7 +34,7 @@
let
val n = nth (prems_of st) (i - 1) |> Logic.strip_assums_hyp |> length;
in
- (THEN_ALL_NEW_SKIP skip tac (REPEAT_DETERM_N n o eresolve_tac [thin_rl])) i st
+ (THEN_ALL_NEW_SKIP skip tac (REPEAT_DETERM_N n o eresolve0_tac [thin_rl])) i st
end;
fun coinduction_tac ctxt raw_vars opt_raw_thm prems = HEADGOAL (SUBGOAL_CASES (fn (goal, _, _) =>
@@ -84,17 +84,17 @@
val e = length eqs;
val p = length prems;
in
- HEADGOAL (EVERY' [resolve_tac [thm],
+ HEADGOAL (EVERY' [resolve_tac ctxt [thm],
EVERY' (map (fn var =>
- resolve_tac
+ resolve_tac ctxt
[Ctr_Sugar_Util.cterm_instantiate_pos
[NONE, SOME (Ctr_Sugar_Util.certify ctxt var)] exI]) vars),
- if p = 0 then Ctr_Sugar_Util.CONJ_WRAP' (K (resolve_tac [refl])) eqs
+ if p = 0 then Ctr_Sugar_Util.CONJ_WRAP' (K (resolve_tac ctxt [refl])) eqs
else
- REPEAT_DETERM_N e o (resolve_tac [conjI] THEN' resolve_tac [refl]) THEN'
- Ctr_Sugar_Util.CONJ_WRAP' (resolve_tac o single) prems,
+ REPEAT_DETERM_N e o (resolve_tac ctxt [conjI] THEN' resolve_tac ctxt [refl]) THEN'
+ Ctr_Sugar_Util.CONJ_WRAP' (resolve_tac ctxt o single) prems,
K (ALLGOALS_SKIP skip
- (REPEAT_DETERM_N (length vars) o (eresolve_tac [exE] THEN' rotate_tac ~1) THEN'
+ (REPEAT_DETERM_N (length vars) o (eresolve_tac ctxt [exE] THEN' rotate_tac ~1) THEN'
DELETE_PREMS_AFTER 0 (Subgoal.FOCUS (fn {prems, params, context = ctxt, ...} =>
(case prems of
[] => all_tac
--- a/src/HOL/Tools/datatype_realizer.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/datatype_realizer.ML Tue Feb 10 14:48:26 2015 +0100
@@ -117,7 +117,7 @@
rtac (cterm_instantiate inst induct) 1,
ALLGOALS (Object_Logic.atomize_prems_tac ctxt),
rewrite_goals_tac ctxt (@{thm o_def} :: map mk_meta_eq rec_rewrites),
- REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
+ REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW (fn i =>
REPEAT (etac allE i) THEN assume_tac ctxt i)) 1)])
|> Drule.export_without_context;
@@ -194,7 +194,7 @@
rtac (cterm_instantiate [(cert y, cert y')] exhaust) 1,
ALLGOALS (EVERY'
[asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites),
- resolve_tac prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])])
+ resolve_tac ctxt prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])])
|> Drule.export_without_context;
val exh_name = Thm.derivation_name exhaust;
--- a/src/HOL/Tools/inductive.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/inductive.ML Tue Feb 10 14:48:26 2015 +0100
@@ -169,8 +169,8 @@
| mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n);
fun select_disj 1 1 = []
- | select_disj _ 1 = [resolve_tac [disjI1]]
- | select_disj n i = resolve_tac [disjI2] :: select_disj (n - 1) (i - 1);
+ | select_disj _ 1 = [resolve0_tac [disjI1]]
+ | select_disj n i = resolve0_tac [disjI2] :: select_disj (n - 1) (i - 1);
@@ -261,7 +261,7 @@
| _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => eq_to_mono thm
| _ $ (Const (@{const_name Orderings.less_eq}, _) $ _ $ _) =>
dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
- (resolve_tac [@{thm le_funI}, @{thm le_boolI'}])) thm))
+ (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}])) thm))
| _ => thm)
end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm ctxt thm);
@@ -378,13 +378,13 @@
[] []
(HOLogic.mk_Trueprop
(Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
- (fn _ => EVERY [resolve_tac @{thms monoI} 1,
- REPEAT (resolve_tac [@{thm le_funI}, @{thm le_boolI'}] 1),
+ (fn _ => EVERY [resolve_tac ctxt @{thms monoI} 1,
+ REPEAT (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}] 1),
REPEAT (FIRST
[assume_tac ctxt 1,
- resolve_tac (map (mk_mono ctxt) monos @ get_monos ctxt) 1,
- eresolve_tac @{thms le_funE} 1,
- dresolve_tac @{thms le_boolD} 1])]));
+ resolve_tac ctxt (map (mk_mono ctxt) monos @ get_monos ctxt) 1,
+ eresolve_tac ctxt @{thms le_funE} 1,
+ dresolve_tac ctxt @{thms le_boolD} 1])]));
(* prove introduction rules *)
@@ -402,11 +402,11 @@
val intrs = map_index (fn (i, intr) =>
Goal.prove_sorry ctxt [] [] intr (fn _ => EVERY
[rewrite_goals_tac ctxt rec_preds_defs,
- resolve_tac [unfold RS iffD2] 1,
+ resolve_tac ctxt [unfold RS iffD2] 1,
EVERY1 (select_disj (length intr_ts) (i + 1)),
(*Not ares_tac, since refl must be tried before any equality assumptions;
backtracking may occur if the premises have extra variables!*)
- DEPTH_SOLVE_1 (resolve_tac rules 1 APPEND assume_tac ctxt 1)])
+ DEPTH_SOLVE_1 (resolve_tac ctxt rules 1 APPEND assume_tac ctxt 1)])
|> singleton (Proof_Context.export ctxt ctxt')) intr_ts
in (intrs, unfold) end;
@@ -448,9 +448,9 @@
(fn {context = ctxt4, prems} => EVERY
[cut_tac (hd prems) 1,
rewrite_goals_tac ctxt4 rec_preds_defs,
- dresolve_tac [unfold RS iffD1] 1,
- REPEAT (FIRSTGOAL (eresolve_tac rules1)),
- REPEAT (FIRSTGOAL (eresolve_tac rules2)),
+ dresolve_tac ctxt4 [unfold RS iffD1] 1,
+ REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules1)),
+ REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules2)),
EVERY (map (fn prem =>
DEPTH_SOLVE_1 (ares_tac [rewrite_rule ctxt4 rec_preds_defs prem, conjI] 1))
(tl prems))])
@@ -493,42 +493,42 @@
if null c_intrs then @{term False}
else foldr1 HOLogic.mk_disj (map mk_intr_conj c_intrs);
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
- fun prove_intr1 (i, _) = Subgoal.FOCUS_PREMS (fn {params, prems, ...} =>
+ fun prove_intr1 (i, _) = Subgoal.FOCUS_PREMS (fn {context = ctxt'', params, prems, ...} =>
EVERY1 (select_disj (length c_intrs) (i + 1)) THEN
- EVERY (replicate (length params) (resolve_tac @{thms exI} 1)) THEN
- (if null prems then resolve_tac @{thms TrueI} 1
+ EVERY (replicate (length params) (resolve_tac ctxt'' @{thms exI} 1)) THEN
+ (if null prems then resolve_tac ctxt'' @{thms TrueI} 1
else
let
val (prems', last_prem) = split_last prems;
in
EVERY (map (fn prem =>
- (resolve_tac @{thms conjI} 1 THEN resolve_tac [prem] 1)) prems') THEN
- resolve_tac [last_prem] 1
+ (resolve_tac ctxt'' @{thms conjI} 1 THEN resolve_tac ctxt'' [prem] 1)) prems')
+ THEN resolve_tac ctxt'' [last_prem] 1
end)) ctxt' 1;
fun prove_intr2 (((_, _, us, _), ts, params'), intr) =
- EVERY (replicate (length params') (eresolve_tac @{thms exE} 1)) THEN
- (if null ts andalso null us then resolve_tac [intr] 1
+ EVERY (replicate (length params') (eresolve_tac ctxt' @{thms exE} 1)) THEN
+ (if null ts andalso null us then resolve_tac ctxt' [intr] 1
else
- EVERY (replicate (length ts + length us - 1) (eresolve_tac @{thms conjE} 1)) THEN
+ EVERY (replicate (length ts + length us - 1) (eresolve_tac ctxt' @{thms conjE} 1)) THEN
Subgoal.FOCUS_PREMS (fn {context = ctxt'', prems, ...} =>
let
val (eqs, prems') = chop (length us) prems;
val rew_thms = map (fn th => th RS @{thm eq_reflection}) eqs;
in
rewrite_goal_tac ctxt'' rew_thms 1 THEN
- resolve_tac [intr] 1 THEN
- EVERY (map (fn p => resolve_tac [p] 1) prems')
+ resolve_tac ctxt'' [intr] 1 THEN
+ EVERY (map (fn p => resolve_tac ctxt'' [p] 1) prems')
end) ctxt' 1);
in
Goal.prove_sorry ctxt' [] [] eq (fn _ =>
- resolve_tac @{thms iffI} 1 THEN
- eresolve_tac [#1 elim] 1 THEN
+ resolve_tac ctxt' @{thms iffI} 1 THEN
+ eresolve_tac ctxt' [#1 elim] 1 THEN
EVERY (map_index prove_intr1 c_intrs) THEN
- (if null c_intrs then eresolve_tac @{thms FalseE} 1
+ (if null c_intrs then eresolve_tac ctxt' @{thms FalseE} 1
else
let val (c_intrs', last_c_intr) = split_last c_intrs in
- EVERY (map (fn ci => eresolve_tac @{thms disjE} 1 THEN prove_intr2 ci) c_intrs') THEN
- prove_intr2 last_c_intr
+ EVERY (map (fn ci => eresolve_tac ctxt' @{thms disjE} 1 THEN prove_intr2 ci) c_intrs')
+ THEN prove_intr2 last_c_intr
end))
|> rulify ctxt'
|> singleton (Proof_Context.export ctxt' ctxt'')
@@ -546,10 +546,13 @@
val refl_thin = Goal.prove_global @{theory HOL} [] [] @{prop "!!P. a = a ==> P ==> P"}
(fn {context = ctxt, ...} => assume_tac ctxt 1);
val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE];
-val elim_tac = REPEAT o eresolve_tac elim_rls;
+fun elim_tac ctxt = REPEAT o eresolve_tac ctxt elim_rls;
fun simp_case_tac ctxt i =
- EVERY' [elim_tac, asm_full_simp_tac ctxt, elim_tac, REPEAT o bound_hyp_subst_tac ctxt] i;
+ EVERY' [elim_tac ctxt,
+ asm_full_simp_tac ctxt,
+ elim_tac ctxt,
+ REPEAT o bound_hyp_subst_tac ctxt] i;
in
@@ -732,16 +735,17 @@
val induct = Goal.prove_sorry ctxt'' [] ind_prems ind_concl
(fn {context = ctxt3, prems} => EVERY
[rewrite_goals_tac ctxt3 [inductive_conj_def],
- DETERM (resolve_tac [raw_fp_induct] 1),
- REPEAT (resolve_tac [@{thm le_funI}, @{thm le_boolI}] 1),
+ DETERM (resolve_tac ctxt3 [raw_fp_induct] 1),
+ REPEAT (resolve_tac ctxt3 [@{thm le_funI}, @{thm le_boolI}] 1),
rewrite_goals_tac ctxt3 simp_thms2,
(*This disjE separates out the introduction rules*)
- REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
+ REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [disjE, exE, FalseE])),
(*Now break down the individual cases. No disjE here in case
some premise involves disjunction.*)
- REPEAT (FIRSTGOAL (eresolve_tac [conjE] ORELSE' bound_hyp_subst_tac ctxt3)),
+ REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [conjE] ORELSE' bound_hyp_subst_tac ctxt3)),
REPEAT (FIRSTGOAL
- (resolve_tac [conjI, impI] ORELSE' (eresolve_tac [notE] THEN' assume_tac ctxt3))),
+ (resolve_tac ctxt3 [conjI, impI] ORELSE'
+ (eresolve_tac ctxt3 [notE] THEN' assume_tac ctxt3))),
EVERY (map (fn prem => DEPTH_SOLVE_1 (ares_tac [rewrite_rule ctxt3
(inductive_conj_def :: rec_preds_defs @ simp_thms2) prem,
conjI, refl] 1)) prems)]);
@@ -750,8 +754,8 @@
(Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn {context = ctxt3, ...} => EVERY
[rewrite_goals_tac ctxt3 rec_preds_defs,
REPEAT (EVERY
- [REPEAT (resolve_tac [conjI, impI] 1),
- REPEAT (eresolve_tac [@{thm le_funE}, @{thm le_boolE}] 1),
+ [REPEAT (resolve_tac ctxt3 [conjI, impI] 1),
+ REPEAT (eresolve_tac ctxt3 [@{thm le_funE}, @{thm le_boolE}] 1),
assume_tac ctxt3 1,
rewrite_goals_tac ctxt3 simp_thms1,
assume_tac ctxt3 1])]);
--- a/src/HOL/Tools/inductive_realizer.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML Tue Feb 10 14:48:26 2015 +0100
@@ -395,7 +395,7 @@
(fn {context = ctxt, prems} => EVERY
[rtac (#raw_induct ind_info) 1,
rewrite_goals_tac ctxt rews,
- REPEAT ((resolve_tac prems THEN_ALL_NEW EVERY'
+ REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW EVERY'
[K (rewrite_goals_tac ctxt rews), Object_Logic.atomize_prems_tac ctxt,
DEPTH_SOLVE_1 o FIRST' [assume_tac ctxt, etac allE, etac impE]]) 1)]);
val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"
@@ -459,7 +459,7 @@
etac elimR 1,
ALLGOALS (asm_simp_tac (put_simpset HOL_basic_ss ctxt)),
rewrite_goals_tac ctxt rews,
- REPEAT ((resolve_tac prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac ctxt THEN'
+ REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac ctxt THEN'
DEPTH_SOLVE_1 o FIRST' [assume_tac ctxt, 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
--- a/src/HOL/Tools/inductive_set.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/inductive_set.ML Tue Feb 10 14:48:26 2015 +0100
@@ -75,14 +75,15 @@
SOME (close (Goal.prove ctxt [] [])
(Logic.mk_equals (t, fold_rev Term.abs xs (m $ p $ (bop $ S $ S'))))
(K (EVERY
- [resolve_tac [eq_reflection] 1, REPEAT (resolve_tac @{thms ext} 1),
- resolve_tac [iffI] 1,
- EVERY [eresolve_tac [conjE] 1, resolve_tac [IntI] 1, simp, simp,
- eresolve_tac [IntE] 1, resolve_tac [conjI] 1, simp, simp] ORELSE
- EVERY [eresolve_tac [disjE] 1, resolve_tac [UnI1] 1, simp,
- resolve_tac [UnI2] 1, simp,
- eresolve_tac [UnE] 1, resolve_tac [disjI1] 1, simp,
- resolve_tac [disjI2] 1, simp]])))
+ [resolve_tac ctxt [eq_reflection] 1,
+ REPEAT (resolve_tac ctxt @{thms ext} 1),
+ resolve_tac ctxt [iffI] 1,
+ EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [IntI] 1, simp, simp,
+ eresolve_tac ctxt [IntE] 1, resolve_tac ctxt [conjI] 1, simp, simp] ORELSE
+ EVERY [eresolve_tac ctxt [disjE] 1, resolve_tac ctxt [UnI1] 1, simp,
+ resolve_tac ctxt [UnI2] 1, simp,
+ eresolve_tac ctxt [UnE] 1, resolve_tac ctxt [disjI1] 1, simp,
+ resolve_tac ctxt [disjI2] 1, simp]])))
handle ERROR _ => NONE))
in
case strip_comb t of
@@ -505,7 +506,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 (resolve_tac @{thms ext} 1) THEN
+ (K (REPEAT (resolve_tac lthy @{thms ext} 1) THEN
simp_tac (put_simpset HOL_basic_ss lthy addsimps
[def, mem_Collect_eq, @{thm split_conv}]) 1))
in
--- a/src/HOL/Tools/lin_arith.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/lin_arith.ML Tue Feb 10 14:48:26 2015 +0100
@@ -732,11 +732,12 @@
end)
in
EVERY' [
- REPEAT_DETERM o eresolve_tac [rev_mp],
+ REPEAT_DETERM o eresolve_tac ctxt [rev_mp],
cond_split_tac,
- resolve_tac @{thms ccontr},
+ resolve_tac ctxt @{thms ccontr},
prem_nnf_tac ctxt,
- TRY o REPEAT_ALL_NEW (DETERM o (eresolve_tac [conjE, exE] ORELSE' eresolve_tac [disjE]))
+ TRY o REPEAT_ALL_NEW
+ (DETERM o (eresolve_tac ctxt [conjE, exE] ORELSE' eresolve_tac ctxt [disjE]))
]
end;
@@ -753,13 +754,13 @@
fun is_relevant t = is_some (decomp ctxt t)
in
DETERM (
- TRY (filter_prems_tac is_relevant i)
+ TRY (filter_prems_tac ctxt is_relevant i)
THEN (
(TRY o REPEAT_ALL_NEW (split_once_tac ctxt split_thms))
THEN_ALL_NEW
(CONVERSION Drule.beta_eta_conversion
THEN'
- (TRY o (eresolve_tac [notE] THEN' eq_assume_tac)))
+ (TRY o (eresolve_tac ctxt [notE] THEN' eq_assume_tac)))
) i
)
end;
@@ -834,14 +835,14 @@
fun refute_tac ctxt test prep_tac ref_tac =
let val refute_prems_tac =
REPEAT_DETERM
- (eresolve_tac [@{thm conjE}, @{thm exE}] 1 ORELSE
- filter_prems_tac test 1 ORELSE
- eresolve_tac @{thms disjE} 1) THEN
- (DETERM (eresolve_tac @{thms notE} 1 THEN eq_assume_tac 1) ORELSE
+ (eresolve_tac ctxt [@{thm conjE}, @{thm exE}] 1 ORELSE
+ filter_prems_tac ctxt test 1 ORELSE
+ eresolve_tac ctxt @{thms disjE} 1) THEN
+ (DETERM (eresolve_tac ctxt @{thms notE} 1 THEN eq_assume_tac 1) ORELSE
ref_tac 1);
- in EVERY'[TRY o filter_prems_tac test,
- REPEAT_DETERM o eresolve_tac @{thms rev_mp}, prep_tac,
- resolve_tac @{thms ccontr}, prem_nnf_tac ctxt,
+ in EVERY'[TRY o filter_prems_tac ctxt test,
+ REPEAT_DETERM o eresolve_tac ctxt @{thms rev_mp}, prep_tac,
+ resolve_tac ctxt @{thms ccontr}, prem_nnf_tac ctxt,
SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)]
end;
@@ -875,7 +876,7 @@
fun gen_tac ex ctxt =
FIRST' [simple_tac ctxt,
Object_Logic.full_atomize_tac ctxt THEN'
- (REPEAT_DETERM o resolve_tac [impI]) THEN' raw_tac ctxt ex];
+ (REPEAT_DETERM o resolve_tac ctxt [impI]) THEN' raw_tac ctxt ex];
val tac = gen_tac true;
--- a/src/HOL/Tools/record.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/record.ML Tue Feb 10 14:48:26 2015 +0100
@@ -1128,7 +1128,7 @@
(fn {context = ctxt, ...} =>
simp_tac (put_simpset ss ctxt) 1 THEN
REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1) THEN
- TRY (resolve_tac [updacc_cong_idI] 1))
+ TRY (resolve_tac ctxt [updacc_cong_idI] 1))
end;
@@ -1634,7 +1634,7 @@
Goal.prove_sorry_global defs_thy [] [assm] concl
(fn {context = ctxt, prems, ...} =>
cut_tac (split_meta RS Drule.equal_elim_rule2) 1 THEN
- resolve_tac prems 2 THEN
+ resolve_tac ctxt prems 2 THEN
asm_simp_tac (put_simpset HOL_ss ctxt) 1)
end);
@@ -1724,7 +1724,7 @@
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition (NONE, (apfst Binding.conceal Attrib.empty_binding, eq)))
|> snd
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
end;
fun ensure_sort_record (sort, mk_eq) ext_tyco vs extN Ts thy =
@@ -1749,7 +1749,7 @@
(Const (@{const_name HOL.equal}, extT --> extT --> HOLogic.boolT),
Const (@{const_name HOL.eq}, extT --> extT --> HOLogic.boolT)));
fun tac ctxt eq_def =
- Class.intro_classes_tac []
+ Class.intro_classes_tac ctxt []
THEN rewrite_goals_tac ctxt [Simpdata.mk_eq eq_def]
THEN ALLGOALS (rtac @{thm refl});
fun mk_eq thy eq_def =
@@ -2128,7 +2128,7 @@
(fn {context = ctxt, prems, ...} =>
try_param_tac ctxt rN induct_scheme 1
THEN try_param_tac ctxt "more" @{thm unit.induct} 1
- THEN resolve_tac prems 1));
+ THEN resolve_tac ctxt prems 1));
val surjective = timeit_msg ctxt "record surjective proof:" (fn () =>
Goal.prove_sorry_global defs_thy [] [] surjective_prop
@@ -2144,8 +2144,8 @@
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)
- (fn {prems, ...} =>
- resolve_tac prems 1 THEN
+ (fn {context = ctxt, prems, ...} =>
+ resolve_tac ctxt prems 1 THEN
rtac surjective 1));
val cases = timeit_msg ctxt "record cases proof:" (fn () =>
--- a/src/HOL/Tools/sat.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/sat.ML Tue Feb 10 14:48:26 2015 +0100
@@ -226,7 +226,7 @@
(* convert the original clause *)
let
val _ = cond_tracing ctxt (fn () => "Using original clause #" ^ string_of_int id)
- val raw = CNF.clause2raw_thm thm
+ val raw = CNF.clause2raw_thm ctxt thm
val hyps = sort (lit_ord o apply2 fst) (map_filter (fn chyp =>
Option.map (rpair chyp) (index_of_literal chyp)) (#hyps (Thm.crep_thm raw)))
val clause = (raw, hyps)
@@ -409,7 +409,7 @@
fun rawsat_tac ctxt i =
Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
- resolve_tac [rawsat_thm ctxt' (map cprop_of prems)] 1) ctxt i;
+ resolve_tac ctxt' [rawsat_thm ctxt' (map cprop_of prems)] 1) ctxt i;
(* ------------------------------------------------------------------------- *)
(* pre_cnf_tac: converts the i-th subgoal *)
@@ -424,7 +424,7 @@
(* ------------------------------------------------------------------------- *)
fun pre_cnf_tac ctxt =
- resolve_tac @{thms ccontr} THEN'
+ resolve_tac ctxt @{thms ccontr} THEN'
Object_Logic.atomize_prems_tac ctxt THEN'
CONVERSION Drule.beta_eta_conversion;
@@ -436,7 +436,8 @@
(* ------------------------------------------------------------------------- *)
fun cnfsat_tac ctxt i =
- (eresolve_tac [FalseE] i) ORELSE (REPEAT_DETERM (eresolve_tac [conjE] i) THEN rawsat_tac ctxt i);
+ (eresolve_tac ctxt [FalseE] i) ORELSE
+ (REPEAT_DETERM (eresolve_tac ctxt [conjE] i) THEN rawsat_tac ctxt i);
(* ------------------------------------------------------------------------- *)
(* cnfxsat_tac: checks if the empty clause "False" occurs among the *)
@@ -446,8 +447,9 @@
(* ------------------------------------------------------------------------- *)
fun cnfxsat_tac ctxt i =
- (eresolve_tac [FalseE] i) ORELSE
- (REPEAT_DETERM (eresolve_tac [conjE] i ORELSE eresolve_tac [exE] i) THEN rawsat_tac ctxt i);
+ (eresolve_tac ctxt [FalseE] i) ORELSE
+ (REPEAT_DETERM (eresolve_tac ctxt [conjE] i ORELSE eresolve_tac ctxt [exE] i) THEN
+ rawsat_tac ctxt i);
(* ------------------------------------------------------------------------- *)
(* sat_tac: tactic for calling an external SAT solver, taking as input an *)
--- a/src/HOL/Tools/set_comprehension_pointfree.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/set_comprehension_pointfree.ML Tue Feb 10 14:48:26 2015 +0100
@@ -314,96 +314,99 @@
val collectI' = @{lemma "\<not> P a ==> a \<notin> {x. P x}" by auto}
val collectE' = @{lemma "a \<notin> {x. P x} ==> (\<not> P a ==> Q) ==> Q" by auto}
-fun elim_Collect_tac ctxt = dresolve_tac @{thms iffD1 [OF mem_Collect_eq]}
- THEN' (REPEAT_DETERM o (eresolve_tac @{thms exE}))
- THEN' REPEAT_DETERM o eresolve_tac @{thms conjE}
+fun elim_Collect_tac ctxt =
+ dresolve_tac ctxt @{thms iffD1 [OF mem_Collect_eq]}
+ THEN' (REPEAT_DETERM o (eresolve_tac ctxt @{thms exE}))
+ THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}
THEN' TRY o hyp_subst_tac ctxt;
-fun intro_image_tac ctxt = resolve_tac @{thms image_eqI}
- THEN' (REPEAT_DETERM1 o
- (resolve_tac @{thms refl}
- ORELSE' resolve_tac @{thms arg_cong2 [OF refl, where f = "op =", OF prod.case, THEN iffD2]}
+fun intro_image_tac ctxt =
+ resolve_tac ctxt @{thms image_eqI}
+ THEN' (REPEAT_DETERM1 o
+ (resolve_tac ctxt @{thms refl}
+ ORELSE' resolve_tac ctxt @{thms arg_cong2 [OF refl, where f = "op =", OF prod.case, THEN iffD2]}
ORELSE' CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1
(HOLogic.Trueprop_conv
(HOLogic.eq_conv Conv.all_conv (Conv.rewr_conv (mk_meta_eq case_prod_distrib)))))) ctxt)))
-fun elim_image_tac ctxt = eresolve_tac @{thms imageE}
+fun elim_image_tac ctxt =
+ eresolve_tac ctxt @{thms imageE}
THEN' REPEAT_DETERM o CHANGED o
(TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.case})
- THEN' REPEAT_DETERM o eresolve_tac @{thms Pair_inject}
+ THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
THEN' TRY o hyp_subst_tac ctxt)
fun tac1_of_formula ctxt (Int (fm1, fm2)) =
- TRY o eresolve_tac @{thms conjE}
- THEN' resolve_tac @{thms IntI}
+ TRY o eresolve_tac ctxt @{thms conjE}
+ THEN' resolve_tac ctxt @{thms IntI}
THEN' (fn i => tac1_of_formula ctxt fm2 (i + 1))
THEN' tac1_of_formula ctxt fm1
| tac1_of_formula ctxt (Un (fm1, fm2)) =
- eresolve_tac @{thms disjE} THEN' resolve_tac @{thms UnI1}
+ eresolve_tac ctxt @{thms disjE} THEN' resolve_tac ctxt @{thms UnI1}
THEN' tac1_of_formula ctxt fm1
- THEN' resolve_tac @{thms UnI2}
+ THEN' resolve_tac ctxt @{thms UnI2}
THEN' tac1_of_formula ctxt fm2
| tac1_of_formula ctxt (Atom _) =
REPEAT_DETERM1 o (assume_tac ctxt
- ORELSE' resolve_tac @{thms SigmaI}
- ORELSE' ((resolve_tac @{thms CollectI} ORELSE' resolve_tac [collectI']) THEN'
+ ORELSE' resolve_tac ctxt @{thms SigmaI}
+ ORELSE' ((resolve_tac ctxt @{thms CollectI} ORELSE' resolve_tac ctxt [collectI']) THEN'
TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]))
- ORELSE' ((resolve_tac @{thms vimageI2} ORELSE' resolve_tac [vimageI2']) THEN'
+ ORELSE' ((resolve_tac ctxt @{thms vimageI2} ORELSE' resolve_tac ctxt [vimageI2']) THEN'
TRY o simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]))
- ORELSE' (resolve_tac @{thms image_eqI} THEN'
+ ORELSE' (resolve_tac ctxt @{thms image_eqI} THEN'
(REPEAT_DETERM o
- (resolve_tac @{thms refl}
- ORELSE' resolve_tac @{thms arg_cong2[OF refl, where f = "op =", OF prod.case, THEN iffD2]})))
- ORELSE' resolve_tac @{thms UNIV_I}
- ORELSE' resolve_tac @{thms iffD2[OF Compl_iff]}
+ (resolve_tac ctxt @{thms refl}
+ ORELSE' resolve_tac ctxt @{thms arg_cong2[OF refl, where f = "op =", OF prod.case, THEN iffD2]})))
+ ORELSE' resolve_tac ctxt @{thms UNIV_I}
+ ORELSE' resolve_tac ctxt @{thms iffD2[OF Compl_iff]}
ORELSE' assume_tac ctxt)
fun tac2_of_formula ctxt (Int (fm1, fm2)) =
- TRY o eresolve_tac @{thms IntE}
- THEN' TRY o resolve_tac @{thms conjI}
+ TRY o eresolve_tac ctxt @{thms IntE}
+ THEN' TRY o resolve_tac ctxt @{thms conjI}
THEN' (fn i => tac2_of_formula ctxt fm2 (i + 1))
THEN' tac2_of_formula ctxt fm1
| tac2_of_formula ctxt (Un (fm1, fm2)) =
- eresolve_tac @{thms UnE} THEN' resolve_tac @{thms disjI1}
+ eresolve_tac ctxt @{thms UnE} THEN' resolve_tac ctxt @{thms disjI1}
THEN' tac2_of_formula ctxt fm1
- THEN' resolve_tac @{thms disjI2}
+ THEN' resolve_tac ctxt @{thms disjI2}
THEN' tac2_of_formula ctxt fm2
| tac2_of_formula ctxt (Atom _) =
REPEAT_DETERM o
(assume_tac ctxt
- ORELSE' dresolve_tac @{thms iffD1[OF mem_Sigma_iff]}
- ORELSE' eresolve_tac @{thms conjE}
- ORELSE' ((eresolve_tac @{thms CollectE} ORELSE' eresolve_tac [collectE']) THEN'
+ ORELSE' dresolve_tac ctxt @{thms iffD1[OF mem_Sigma_iff]}
+ ORELSE' eresolve_tac ctxt @{thms conjE}
+ ORELSE' ((eresolve_tac ctxt @{thms CollectE} ORELSE' eresolve_tac ctxt [collectE']) THEN'
TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}]) THEN'
- REPEAT_DETERM o eresolve_tac @{thms Pair_inject} THEN' TRY o hyp_subst_tac ctxt THEN'
- TRY o resolve_tac @{thms refl})
- ORELSE' (eresolve_tac @{thms imageE}
+ REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject} THEN' TRY o hyp_subst_tac ctxt THEN'
+ TRY o resolve_tac ctxt @{thms refl})
+ ORELSE' (eresolve_tac ctxt @{thms imageE}
THEN' (REPEAT_DETERM o CHANGED o
(TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_paired_all prod.case})
- THEN' REPEAT_DETERM o eresolve_tac @{thms Pair_inject}
- THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac @{thms refl})))
- ORELSE' eresolve_tac @{thms ComplE}
- ORELSE' ((eresolve_tac @{thms vimageE} ORELSE' eresolve_tac [vimageE'])
+ THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms Pair_inject}
+ THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac ctxt @{thms refl})))
+ ORELSE' eresolve_tac ctxt @{thms ComplE}
+ ORELSE' ((eresolve_tac ctxt @{thms vimageE} ORELSE' eresolve_tac ctxt [vimageE'])
THEN' TRY o full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm prod.case}])
- THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac @{thms refl}))
+ THEN' TRY o hyp_subst_tac ctxt THEN' TRY o resolve_tac ctxt @{thms refl}))
fun tac ctxt fm =
let
- val subset_tac1 = resolve_tac @{thms subsetI}
+ val subset_tac1 = resolve_tac ctxt @{thms subsetI}
THEN' elim_Collect_tac ctxt
THEN' intro_image_tac ctxt
THEN' tac1_of_formula ctxt fm
- val subset_tac2 = resolve_tac @{thms subsetI}
+ val subset_tac2 = resolve_tac ctxt @{thms subsetI}
THEN' elim_image_tac ctxt
- THEN' resolve_tac @{thms iffD2[OF mem_Collect_eq]}
- THEN' REPEAT_DETERM o resolve_tac @{thms exI}
- THEN' (TRY o REPEAT_ALL_NEW (resolve_tac @{thms conjI}))
- THEN' (K (TRY (FIRSTGOAL ((TRY o hyp_subst_tac ctxt) THEN' resolve_tac @{thms refl}))))
+ THEN' resolve_tac ctxt @{thms iffD2[OF mem_Collect_eq]}
+ THEN' REPEAT_DETERM o resolve_tac ctxt @{thms exI}
+ THEN' (TRY o REPEAT_ALL_NEW (resolve_tac ctxt @{thms conjI}))
+ THEN' (K (TRY (FIRSTGOAL ((TRY o hyp_subst_tac ctxt) THEN' resolve_tac ctxt @{thms refl}))))
THEN' (fn i => EVERY (rev (map_index (fn (j, f) =>
- REPEAT_DETERM (eresolve_tac @{thms IntE} (i + j)) THEN
+ REPEAT_DETERM (eresolve_tac ctxt @{thms IntE} (i + j)) THEN
tac2_of_formula ctxt f (i + j)) (strip_Int fm))))
in
- resolve_tac @{thms subset_antisym} THEN' subset_tac1 THEN' subset_tac2
+ resolve_tac ctxt @{thms subset_antisym} THEN' subset_tac1 THEN' subset_tac2
end;
@@ -430,14 +433,14 @@
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.case}
fun tac ctxt =
- resolve_tac @{thms set_eqI}
+ resolve_tac ctxt @{thms set_eqI}
THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps unfold_thms)
- THEN' resolve_tac @{thms iffI}
- THEN' REPEAT_DETERM o resolve_tac @{thms exI}
- THEN' resolve_tac @{thms conjI} THEN' resolve_tac @{thms refl} THEN' assume_tac ctxt
- THEN' REPEAT_DETERM o eresolve_tac @{thms exE}
- THEN' eresolve_tac @{thms conjE}
- THEN' REPEAT_DETERM o eresolve_tac @{thms Pair_inject}
+ THEN' resolve_tac ctxt @{thms iffI}
+ THEN' REPEAT_DETERM o resolve_tac ctxt @{thms exI}
+ THEN' resolve_tac ctxt @{thms conjI} THEN' resolve_tac ctxt @{thms refl} THEN' assume_tac ctxt
+ THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms exE}
+ THEN' eresolve_tac ctxt @{thms conjE}
+ THEN' REPEAT_DETERM o eresolve_tac ctxt @{thms 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
--- a/src/HOL/Tools/simpdata.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Tools/simpdata.ML Tue Feb 10 14:48:26 2015 +0100
@@ -79,13 +79,15 @@
end;
(*Congruence rules for = (instead of ==)*)
-fun mk_meta_cong (_: Proof.context) rl = zero_var_indexes
- (let val rl' = Seq.hd (TRYALL (fn i => fn st =>
- resolve_tac [lift_meta_eq_to_obj_eq i st] i st) rl)
- in mk_meta_eq rl' handle THM _ =>
- if can Logic.dest_equals (concl_of rl') then rl'
- else error "Conclusion of congruence rules must be =-equality"
- end);
+fun mk_meta_cong ctxt rl =
+ let
+ val rl' = Seq.hd (TRYALL (fn i => fn st =>
+ resolve_tac ctxt [lift_meta_eq_to_obj_eq i st] i st) rl)
+ in
+ mk_meta_eq rl' handle THM _ =>
+ if can Logic.dest_equals (concl_of rl') then rl'
+ else error "Conclusion of congruence rules must be =-equality"
+ end |> zero_var_indexes;
fun mk_atomize ctxt pairs =
let
@@ -120,7 +122,10 @@
val sol_thms =
reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt;
fun sol_tac i =
- FIRST [resolve_tac sol_thms i, assume_tac ctxt i , eresolve_tac @{thms FalseE} i] ORELSE
+ FIRST
+ [resolve_tac ctxt sol_thms i,
+ assume_tac ctxt i,
+ eresolve_tac ctxt @{thms FalseE} i] ORELSE
(match_tac ctxt intros THEN_ALL_NEW sol_tac) i
in
(fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' sol_tac
--- a/src/HOL/Typerep.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Typerep.thy Tue Feb 10 14:48:26 2015 +0100
@@ -60,7 +60,7 @@
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
|> snd
- |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
end;
fun ensure_typerep tyco thy =
--- a/src/HOL/UNITY/Simple/NSP_Bad.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/UNITY/Simple/NSP_Bad.thy Tue Feb 10 14:48:26 2015 +0100
@@ -105,7 +105,7 @@
SELECT_GOAL
(EVERY
[REPEAT (etac @{thm Always_ConstrainsI} 1),
- REPEAT (resolve_tac [@{thm StableI}, @{thm stableI}, @{thm constrains_imp_Constrains}] 1),
+ REPEAT (resolve_tac ctxt [@{thm StableI}, @{thm stableI}, @{thm constrains_imp_Constrains}] 1),
rtac @{thm ns_constrainsI} 1,
full_simp_tac ctxt 1,
REPEAT (FIRSTGOAL (etac disjE)),
--- a/src/HOL/UNITY/UNITY_tactics.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/UNITY/UNITY_tactics.ML Tue Feb 10 14:48:26 2015 +0100
@@ -21,7 +21,7 @@
(*reduce the fancy safety properties to "constrains"*)
REPEAT (etac @{thm Always_ConstrainsI} 1
ORELSE
- resolve_tac [@{thm StableI}, @{thm stableI},
+ resolve_tac ctxt [@{thm StableI}, @{thm stableI},
@{thm constrains_imp_Constrains}] 1),
(*for safety, the totalize operator can be ignored*)
simp_tac (put_simpset HOL_ss ctxt
--- a/src/HOL/Word/Word.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Word/Word.thy Tue Feb 10 14:48:26 2015 +0100
@@ -1561,7 +1561,7 @@
|> fold Splitter.add_split @{thms uint_splits}
|> fold Simplifier.add_cong @{thms power_False_cong})),
rewrite_goals_tac ctxt @{thms word_size},
- ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN
+ ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN
REPEAT (etac conjE n) THEN
REPEAT (dtac @{thm word_of_int_inverse} n
THEN assume_tac ctxt n
@@ -2063,7 +2063,7 @@
|> fold Splitter.add_split @{thms unat_splits}
|> fold Simplifier.add_cong @{thms power_False_cong})),
rewrite_goals_tac ctxt @{thms word_size},
- ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN
+ ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN
REPEAT (etac conjE n) THEN
REPEAT (dtac @{thm of_nat_inverse} n THEN atac n)),
TRYALL arith_tac' ]
--- a/src/HOL/Word/WordBitwise.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/HOL/Word/WordBitwise.thy Tue Feb 10 14:48:26 2015 +0100
@@ -518,7 +518,7 @@
in
try (fn () =>
Goal.prove_internal ctxt [] prop
- (K (REPEAT_DETERM (resolve_tac @{thms upt_eq_list_intros} 1
+ (K (REPEAT_DETERM (resolve_tac ctxt @{thms upt_eq_list_intros} 1
ORELSE simp_tac (put_simpset word_ss ctxt) 1))) |> mk_meta_eq) ()
end
| _ => NONE;
--- a/src/LCF/LCF.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/LCF/LCF.thy Tue Feb 10 14:48:26 2015 +0100
@@ -322,7 +322,7 @@
Scan.lift Args.name_inner_syntax >> (fn v => fn ctxt =>
SIMPLE_METHOD' (fn i =>
res_inst_tac ctxt [(("f", 0), v)] @{thm induct} i THEN
- REPEAT (resolve_tac @{thms adm_lemmas} i)))
+ REPEAT (resolve_tac ctxt @{thms adm_lemmas} i)))
*}
lemma least_FIX: "f(p) = p \<Longrightarrow> FIX(f) << p"
@@ -381,7 +381,7 @@
ML {*
fun induct2_tac ctxt (f, g) i =
res_inst_tac ctxt [(("f", 0), f), (("g", 0), g)] @{thm induct2} i THEN
- REPEAT(resolve_tac @{thms adm_lemmas} i)
+ REPEAT(resolve_tac ctxt @{thms adm_lemmas} i)
*}
end
--- a/src/Provers/Arith/cancel_numerals.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/Arith/cancel_numerals.ML Tue Feb 10 14:48:26 2015 +0100
@@ -92,12 +92,12 @@
Option.map (export o Data.simplify_meta_eq ctxt)
(if n2 <= n1 then
Data.prove_conv
- [Data.trans_tac reshape, resolve_tac [Data.bal_add1] 1,
+ [Data.trans_tac reshape, resolve_tac ctxt [Data.bal_add1] 1,
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, resolve_tac [Data.bal_add2] 1,
+ [Data.trans_tac reshape, resolve_tac ctxt [Data.bal_add2] 1,
Data.numeral_simp_tac ctxt] ctxt prems
(t', Data.mk_bal (Data.mk_sum T terms1', newshape(n2-n1,terms2'))))
end
--- a/src/Provers/Arith/combine_numerals.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/Arith/combine_numerals.ML Tue Feb 10 14:48:26 2015 +0100
@@ -83,7 +83,7 @@
in
Option.map (export o Data.simplify_meta_eq ctxt)
(Data.prove_conv
- [Data.trans_tac reshape, resolve_tac [Data.left_distrib] 1,
+ [Data.trans_tac reshape, resolve_tac ctxt [Data.left_distrib] 1,
Data.numeral_simp_tac ctxt] ctxt
(t', Data.mk_sum T (Data.mk_coeff(Data.add(m,n), u) :: terms)))
end
--- a/src/Provers/Arith/fast_lin_arith.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML Tue Feb 10 14:48:26 2015 +0100
@@ -810,15 +810,15 @@
fun just1 j =
(* eliminate inequalities *)
(if split_neq then
- REPEAT_DETERM (eresolve_tac neqE i)
+ REPEAT_DETERM (eresolve_tac ctxt neqE i)
else
all_tac) THEN
PRIMITIVE (trace_thm ctxt ["State after neqE:"]) THEN
(* use theorems generated from the actual justifications *)
- Subgoal.FOCUS (fn {prems, ...} => resolve_tac [mkthm ctxt prems j] 1) ctxt i
+ Subgoal.FOCUS (fn {prems, ...} => resolve_tac ctxt [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
+ DETERM (resolve_tac ctxt [LA_Logic.notI, LA_Logic.ccontr] i) THEN
(* user-defined preprocessing of the subgoal *)
DETERM (LA_Data.pre_tac ctxt i) THEN
PRIMITIVE (trace_thm ctxt ["State after pre_tac:"]) THEN
--- a/src/Provers/blast.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/blast.ML Tue Feb 10 14:48:26 2015 +0100
@@ -488,8 +488,11 @@
(*Resolution/matching tactics: if upd then the proof state may be updated.
Matching makes the tactics more deterministic in the presence of Vars.*)
-fun emtac ctxt upd rl = TRACE ctxt rl (if upd then eresolve_tac [rl] else ematch_tac ctxt [rl]);
-fun rmtac ctxt upd rl = TRACE ctxt rl (if upd then resolve_tac [rl] else match_tac ctxt [rl]);
+fun emtac ctxt upd rl =
+ TRACE ctxt rl (if upd then eresolve_tac ctxt [rl] else ematch_tac ctxt [rl]);
+
+fun rmtac ctxt upd rl =
+ TRACE ctxt rl (if upd then resolve_tac ctxt [rl] else match_tac ctxt [rl]);
(*Tableau rule from elimination rule.
Flag "upd" says that the inference updated the branch.
@@ -624,7 +627,7 @@
(*Tactic. Convert *Goal* to negated assumption in FIRST position*)
fun negOfGoal_tac ctxt i =
- TRACE ctxt Data.ccontr (resolve_tac [Data.ccontr]) i THEN rotate_tac ~1 i;
+ TRACE ctxt Data.ccontr (resolve_tac ctxt [Data.ccontr]) i THEN rotate_tac ~1 i;
fun traceTerm ctxt t =
let val thy = Proof_Context.theory_of ctxt
--- a/src/Provers/classical.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/classical.ML Tue Feb 10 14:48:26 2015 +0100
@@ -157,7 +157,7 @@
val rule'' =
rule' |> ALLGOALS (SUBGOAL (fn (goal, i) =>
if i = 1 orelse redundant_hyp goal
- then eresolve_tac [thin_rl] i
+ then eresolve0_tac [thin_rl] i
else all_tac))
|> Seq.hd
|> Drule.zero_var_indexes;
@@ -180,11 +180,12 @@
No backtracking if it finds an equal assumption. Perhaps should call
ematch_tac instead of eresolve_tac, but then cannot prove ZF/cantor.*)
fun contr_tac ctxt =
- eresolve_tac [Data.not_elim] THEN' (eq_assume_tac ORELSE' assume_tac ctxt);
+ eresolve_tac ctxt [Data.not_elim] THEN' (eq_assume_tac ORELSE' assume_tac ctxt);
(*Finds P-->Q and P in the assumptions, replaces implication by Q.
Could do the same thing for P<->Q and P... *)
-fun mp_tac ctxt i = eresolve_tac [Data.not_elim, Data.imp_elim] i THEN assume_tac ctxt i;
+fun mp_tac ctxt i =
+ eresolve_tac ctxt [Data.not_elim, Data.imp_elim] i THEN assume_tac ctxt i;
(*Like mp_tac but instantiates no variables*)
fun eq_mp_tac ctxt i = ematch_tac ctxt [Data.not_elim, Data.imp_elim] i THEN eq_assume_tac i;
@@ -199,7 +200,7 @@
let fun addrl rl brls = (false, rl) :: (true, rl RSN (2, Data.swap)) :: brls in
assume_tac ctxt ORELSE'
contr_tac ctxt ORELSE'
- biresolve_tac (fold_rev addrl rls [])
+ biresolve_tac ctxt (fold_rev addrl rls [])
end;
(*Duplication of hazardous rules, for complete provers*)
@@ -212,7 +213,7 @@
SOME c => Context.proof_of c
| NONE => Proof_Context.init_global (Thm.theory_of_thm th));
val rl = (th RSN (2, revcut_rl)) |> Thm.assumption (SOME ctxt) 2 |> Seq.hd;
- in rule_by_tactic ctxt (TRYALL (eresolve_tac [revcut_rl])) rl end;
+ in rule_by_tactic ctxt (TRYALL (eresolve_tac ctxt [revcut_rl])) rl end;
(**** Classical rule sets ****)
@@ -679,9 +680,9 @@
ctxt addWrapper (name, fn ctxt => fn tac1 => tac1 APPEND' tac2 ctxt);
fun ctxt addD2 (name, thm) =
- ctxt addafter (name, fn ctxt' => dresolve_tac [thm] THEN' assume_tac ctxt');
+ ctxt addafter (name, fn ctxt' => dresolve_tac ctxt' [thm] THEN' assume_tac ctxt');
fun ctxt addE2 (name, thm) =
- ctxt addafter (name, fn ctxt' => eresolve_tac [thm] THEN' assume_tac ctxt');
+ ctxt addafter (name, fn ctxt' => eresolve_tac ctxt' [thm] THEN' assume_tac ctxt');
fun ctxt addSD2 (name, thm) =
ctxt addSafter (name, fn ctxt' => dmatch_tac ctxt' [thm] THEN' eq_assume_tac);
fun ctxt addSE2 (name, thm) =
@@ -902,7 +903,7 @@
val ruleq = Drule.multi_resolves (SOME ctxt) facts rules;
val _ = Method.trace ctxt rules;
in
- fn st => Seq.maps (fn rule => resolve_tac [rule] i st) ruleq
+ fn st => Seq.maps (fn rule => resolve_tac ctxt [rule] i st) ruleq
end)
THEN_ALL_NEW Goal.norm_hhf_tac ctxt;
--- a/src/Provers/hypsubst.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/hypsubst.ML Tue Feb 10 14:48:26 2015 +0100
@@ -51,7 +51,7 @@
val hyp_subst_thin : bool Config.T
val hyp_subst_tac : Proof.context -> int -> tactic
val blast_hyp_subst_tac : Proof.context -> bool -> int -> tactic
- val stac : thm -> int -> tactic
+ val stac : Proof.context -> thm -> int -> tactic
end;
functor Hypsubst(Data: HYPSUBST_DATA): HYPSUBST =
@@ -122,13 +122,14 @@
in eq_var_aux 0 t [] end;
-val thin_free_eq_tac = SUBGOAL (fn (t, i) => let
- val (k, _) = eq_var false false false t
- val ok = (eq_var false false true t |> fst) > k
- handle EQ_VAR => true
- in if ok then EVERY [rotate_tac k i, eresolve_tac [thin_rl] i, rotate_tac (~k) i]
+fun thin_free_eq_tac ctxt = SUBGOAL (fn (t, i) =>
+ let
+ val (k, _) = eq_var false false false t
+ val ok = (eq_var false false true t |> fst) > k handle EQ_VAR => true
+ in
+ if ok then EVERY [rotate_tac k i, eresolve_tac ctxt [thin_rl] i, rotate_tac (~k) i]
else no_tac
- end handle EQ_VAR => no_tac)
+ end handle EQ_VAR => no_tac)
(*For the simpset. Adds ALL suitable equalities, even if not first!
No vars are allowed here, as simpsets are built from meta-assumptions*)
@@ -151,11 +152,11 @@
val (k, (orient, is_free)) = eq_var bnd true true Bi
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,
- if not is_free then eresolve_tac [thin_rl] i
- else if orient then eresolve_tac [Data.rev_mp] i
- else eresolve_tac [Data.sym RS Data.rev_mp] i,
+ if not is_free then eresolve_tac ctxt [thin_rl] i
+ else if orient then eresolve_tac ctxt [Data.rev_mp] i
+ else eresolve_tac ctxt [Data.sym RS Data.rev_mp] i,
rotate_tac (~k) i,
- if is_free then resolve_tac [Data.imp_intr] i else all_tac]
+ if is_free then resolve_tac ctxt [Data.imp_intr] i else all_tac]
end handle THM _ => no_tac | EQ_VAR => no_tac)
end;
@@ -194,7 +195,7 @@
end
| NONE => no_tac);
-val imp_intr_tac = resolve_tac [Data.imp_intr];
+fun imp_intr_tac ctxt = resolve_tac ctxt [Data.imp_intr];
fun rev_dup_elim ctxt th = (th RSN (2, revcut_rl)) |> Thm.assumption (SOME ctxt) 2 |> Seq.hd;
fun dup_subst ctxt = rev_dup_elim ctxt ssubst
@@ -211,11 +212,11 @@
val rl = if orient then rl else Data.sym RS rl
in
DETERM
- (EVERY [REPEAT_DETERM_N k (eresolve_tac [Data.rev_mp] i),
+ (EVERY [REPEAT_DETERM_N k (eresolve_tac ctxt [Data.rev_mp] i),
rotate_tac 1 i,
- REPEAT_DETERM_N (n-k) (eresolve_tac [Data.rev_mp] i),
+ REPEAT_DETERM_N (n-k) (eresolve_tac ctxt [Data.rev_mp] i),
inst_subst_tac ctxt orient rl i,
- REPEAT_DETERM_N n (imp_intr_tac i THEN rotate_tac ~1 i)])
+ REPEAT_DETERM_N n (imp_intr_tac ctxt i THEN rotate_tac ~1 i)])
end
handle THM _ => no_tac | EQ_VAR => no_tac);
@@ -225,7 +226,7 @@
fun hyp_subst_tac_thin thin ctxt =
REPEAT_DETERM1 o FIRST' [ematch_tac ctxt [Data.thin_refl],
gen_hyp_subst_tac ctxt false, vars_gen_hyp_subst_tac ctxt false,
- if thin then thin_free_eq_tac else K no_tac];
+ if thin then thin_free_eq_tac ctxt else K no_tac];
val hyp_subst_thin = Attrib.setup_config_bool @{binding hypsubst_thin} (K false);
@@ -242,16 +243,16 @@
substitutions in the arguments of a function Var. **)
(*final re-reversal of the changed assumptions*)
-fun reverse_n_tac 0 i = all_tac
- | reverse_n_tac 1 i = rotate_tac ~1 i
- | reverse_n_tac n i =
- REPEAT_DETERM_N n (rotate_tac ~1 i THEN eresolve_tac [Data.rev_mp] i) THEN
- REPEAT_DETERM_N n (imp_intr_tac i THEN rotate_tac ~1 i);
+fun reverse_n_tac _ 0 i = all_tac
+ | reverse_n_tac _ 1 i = rotate_tac ~1 i
+ | reverse_n_tac ctxt n i =
+ REPEAT_DETERM_N n (rotate_tac ~1 i THEN eresolve_tac ctxt [Data.rev_mp] i) THEN
+ REPEAT_DETERM_N n (imp_intr_tac ctxt i THEN rotate_tac ~1 i);
(*Use imp_intr, comparing the old hyps with the new ones as they come out.*)
-fun all_imp_intr_tac hyps i =
+fun all_imp_intr_tac ctxt hyps i =
let
- fun imptac (r, []) st = reverse_n_tac r i st
+ fun imptac (r, []) st = reverse_n_tac ctxt r i st
| imptac (r, hyp::hyps) st =
let
val (hyp', _) =
@@ -260,8 +261,8 @@
|> Data.dest_Trueprop |> Data.dest_imp;
val (r', tac) =
if Envir.aeconv (hyp, hyp')
- then (r, imp_intr_tac i THEN rotate_tac ~1 i)
- else (*leave affected hyps at end*) (r + 1, imp_intr_tac i);
+ then (r, imp_intr_tac ctxt i THEN rotate_tac ~1 i)
+ else (*leave affected hyps at end*) (r + 1, imp_intr_tac ctxt i);
in
(case Seq.pull (tac st) of
NONE => Seq.single st
@@ -270,28 +271,29 @@
in imptac (0, rev hyps) end;
-fun blast_hyp_subst_tac ctxt trace = SUBGOAL(fn (Bi,i) =>
- let val (k, (symopt, _)) = eq_var false false false Bi
- val hyps0 = map Data.dest_Trueprop (Logic.strip_assums_hyp Bi)
- (*omit selected equality, returning other hyps*)
- val hyps = List.take(hyps0, k) @ List.drop(hyps0, k+1)
- val n = length hyps
- in
- if trace then tracing "Substituting an equality" else ();
- DETERM
- (EVERY [REPEAT_DETERM_N k (eresolve_tac [Data.rev_mp] i),
- rotate_tac 1 i,
- REPEAT_DETERM_N (n-k) (eresolve_tac [Data.rev_mp] i),
- inst_subst_tac ctxt symopt (if symopt then ssubst else Data.subst) i,
- all_imp_intr_tac hyps i])
- end
- handle THM _ => no_tac | EQ_VAR => no_tac);
+fun blast_hyp_subst_tac ctxt trace = SUBGOAL(fn (Bi, i) =>
+ let
+ val (k, (symopt, _)) = eq_var false false false Bi
+ val hyps0 = map Data.dest_Trueprop (Logic.strip_assums_hyp Bi)
+ (*omit selected equality, returning other hyps*)
+ val hyps = List.take(hyps0, k) @ List.drop(hyps0, k+1)
+ val n = length hyps
+ in
+ if trace then tracing "Substituting an equality" else ();
+ DETERM
+ (EVERY [REPEAT_DETERM_N k (eresolve_tac ctxt [Data.rev_mp] i),
+ rotate_tac 1 i,
+ REPEAT_DETERM_N (n-k) (eresolve_tac ctxt [Data.rev_mp] i),
+ inst_subst_tac ctxt symopt (if symopt then ssubst else Data.subst) i,
+ all_imp_intr_tac ctxt hyps i])
+ end
+ handle THM _ => no_tac | EQ_VAR => no_tac);
(*apply an equality or definition ONCE;
fails unless the substitution has an effect*)
-fun stac th =
+fun stac ctxt th =
let val th' = th RS Data.rev_eq_reflection handle THM _ => th
- in CHANGED_GOAL (resolve_tac [th' RS ssubst]) end;
+ in CHANGED_GOAL (resolve_tac ctxt [th' RS ssubst]) end;
(* method setup *)
@@ -305,7 +307,8 @@
(Scan.succeed (fn ctxt => SIMPLE_METHOD'
(CHANGED_PROP o hyp_subst_tac_thin true ctxt)))
"substitution using an assumption, eliminating assumptions" #>
- Method.setup @{binding simplesubst} (Attrib.thm >> (fn th => K (SIMPLE_METHOD' (stac th))))
+ Method.setup @{binding simplesubst}
+ (Attrib.thm >> (fn th => fn ctxt => SIMPLE_METHOD' (stac ctxt th)))
"simple substitution");
end;
--- a/src/Provers/order.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/order.ML Tue Feb 10 14:48:26 2015 +0100
@@ -1241,12 +1241,13 @@
val prfs = gen_solve mkconcl thy (lesss, C);
val (subgoals, prf) = mkconcl decomp less_thms thy C;
in
- Subgoal.FOCUS (fn {prems = asms, ...} =>
+ Subgoal.FOCUS (fn {context = ctxt', prems = asms, ...} =>
let val thms = map (prove (prems @ asms)) prfs
- in resolve_tac [prove thms prf] 1 end) ctxt n st
+ in resolve_tac ctxt' [prove thms prf] 1 end) ctxt n st
end
handle Contr p =>
- (Subgoal.FOCUS (fn {prems = asms, ...} => resolve_tac [prove asms p] 1) ctxt n st
+ (Subgoal.FOCUS (fn {context = ctxt', prems = asms, ...} =>
+ resolve_tac ctxt' [prove asms p] 1) ctxt n st
handle General.Subscript => Seq.empty)
| Cannot => Seq.empty
| General.Subscript => Seq.empty)
--- a/src/Provers/quantifier1.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/quantifier1.ML Tue Feb 10 14:48:26 2015 +0100
@@ -64,8 +64,8 @@
signature QUANTIFIER1 =
sig
- val prove_one_point_all_tac: tactic
- val prove_one_point_ex_tac: tactic
+ val prove_one_point_all_tac: Proof.context -> tactic
+ val prove_one_point_ex_tac: Proof.context -> tactic
val rearrange_all: Proof.context -> cterm -> thm option
val rearrange_ex: Proof.context -> cterm -> thm option
val rearrange_ball: (Proof.context -> tactic) -> Proof.context -> cterm -> thm option
@@ -126,10 +126,12 @@
yield_singleton (Variable.import_terms true) (Logic.mk_equals tu) ctxt;
val thm =
Goal.prove ctxt' [] [] goal
- (fn {context = ctxt'', ...} => resolve_tac [Data.iff_reflection] 1 THEN tac ctxt'');
+ (fn {context = ctxt'', ...} =>
+ resolve_tac ctxt'' [Data.iff_reflection] 1 THEN tac ctxt'');
in singleton (Variable.export ctxt' ctxt) thm end;
-fun qcomm_tac qcomm qI i = REPEAT_DETERM (resolve_tac [qcomm] i THEN resolve_tac [qI] i);
+fun qcomm_tac ctxt qcomm qI i =
+ REPEAT_DETERM (resolve_tac ctxt [qcomm] i THEN resolve_tac ctxt [qI] i);
(* Proves (? x0..xn. ... & x0 = t & ...) = (? x1..xn x0. x0 = t & ... & ...)
Better: instantiate exI
@@ -137,11 +139,11 @@
local
val excomm = Data.ex_comm RS Data.iff_trans;
in
- val prove_one_point_ex_tac =
- qcomm_tac excomm Data.iff_exI 1 THEN resolve_tac [Data.iffI] 1 THEN
+ fun prove_one_point_ex_tac ctxt =
+ qcomm_tac ctxt excomm Data.iff_exI 1 THEN resolve_tac ctxt [Data.iffI] 1 THEN
ALLGOALS
- (EVERY' [eresolve_tac [Data.exE], REPEAT_DETERM o eresolve_tac [Data.conjE],
- resolve_tac [Data.exI],
+ (EVERY' [eresolve_tac ctxt [Data.exE], REPEAT_DETERM o eresolve_tac ctxt [Data.conjE],
+ resolve_tac ctxt [Data.exI],
DEPTH_SOLVE_1 o ares_tac [Data.conjI]])
end;
@@ -149,19 +151,19 @@
(! x1..xn x0. x0 = t --> (... & ...) --> P x0)
*)
local
- val tac =
+ fun tac ctxt =
SELECT_GOAL
- (EVERY1 [REPEAT o dresolve_tac [Data.uncurry],
- REPEAT o resolve_tac [Data.impI],
- eresolve_tac [Data.mp],
- REPEAT o eresolve_tac [Data.conjE],
+ (EVERY1 [REPEAT o dresolve_tac ctxt [Data.uncurry],
+ REPEAT o resolve_tac ctxt [Data.impI],
+ eresolve_tac ctxt [Data.mp],
+ REPEAT o eresolve_tac ctxt [Data.conjE],
REPEAT o ares_tac [Data.conjI]]);
val allcomm = Data.all_comm RS Data.iff_trans;
in
- val prove_one_point_all_tac =
- EVERY1 [qcomm_tac allcomm Data.iff_allI,
- resolve_tac [Data.iff_allI],
- resolve_tac [Data.iffI], tac, tac];
+ fun prove_one_point_all_tac ctxt =
+ EVERY1 [qcomm_tac ctxt allcomm Data.iff_allI,
+ resolve_tac ctxt [Data.iff_allI],
+ resolve_tac ctxt [Data.iffI], tac ctxt, tac ctxt];
end
fun renumber l u (Bound i) =
@@ -185,7 +187,7 @@
NONE => NONE
| SOME (xs, eq, Q) =>
let val R = quantify all x T xs (Data.imp $ eq $ Q)
- in SOME (prove_conv ctxt (F, R) (K prove_one_point_all_tac)) end)
+ in SOME (prove_conv ctxt (F, R) prove_one_point_all_tac) end)
| _ => NONE);
fun rearrange_ball tac ctxt ct =
@@ -207,7 +209,7 @@
NONE => NONE
| SOME (xs, eq, Q) =>
let val R = quantify ex x T xs (Data.conj $ eq $ Q)
- in SOME (prove_conv ctxt (F, R) (K prove_one_point_ex_tac)) end)
+ in SOME (prove_conv ctxt (F, R) prove_one_point_ex_tac) end)
| _ => NONE);
fun rearrange_bex tac ctxt ct =
--- a/src/Provers/quasi.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/quasi.ML Tue Feb 10 14:48:26 2015 +0100
@@ -563,12 +563,14 @@
val (subgoal, prf) = mkconcl_trans thy C;
in
- Subgoal.FOCUS (fn {prems, ...} =>
+ Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
let val thms = map (prove prems) prfs
- in resolve_tac [prove thms prf] 1 end) ctxt n st
+ in resolve_tac ctxt' [prove thms prf] 1 end) ctxt n st
end
- handle Contr p => Subgoal.FOCUS (fn {prems, ...} => resolve_tac [prove prems p] 1) ctxt n st
- | Cannot => Seq.empty);
+ handle Contr p =>
+ Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
+ resolve_tac ctxt' [prove prems p] 1) ctxt n st
+ | Cannot => Seq.empty);
(* quasi_tac - solves quasi orders *)
@@ -583,13 +585,13 @@
val prfs = solveQuasiOrder thy (lesss, C);
val (subgoals, prf) = mkconcl_quasi thy C;
in
- Subgoal.FOCUS (fn {prems, ...} =>
+ Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
let val thms = map (prove prems) prfs
- in resolve_tac [prove thms prf] 1 end) ctxt n st
+ in resolve_tac ctxt' [prove thms prf] 1 end) ctxt n st
end
handle Contr p =>
- (Subgoal.FOCUS (fn {prems, ...} => resolve_tac [prove prems p] 1) ctxt n st
- handle General.Subscript => Seq.empty)
+ (Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
+ resolve_tac ctxt' [prove prems p] 1) ctxt n st handle General.Subscript => Seq.empty)
| Cannot => Seq.empty
| General.Subscript => Seq.empty);
--- a/src/Provers/splitter.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/splitter.ML Tue Feb 10 14:48:26 2015 +0100
@@ -99,7 +99,8 @@
val lift = Goal.prove_global Pure.thy ["P", "Q", "R"]
[Syntax.read_prop_global Pure.thy "!!x :: 'b. Q(x) == R(x) :: 'c"]
(Syntax.read_prop_global Pure.thy "P(%x. Q(x)) == P(%x. R(x))")
- (fn {context = ctxt, prems} => rewrite_goals_tac ctxt prems THEN resolve_tac [reflexive_thm] 1)
+ (fn {context = ctxt, prems} =>
+ rewrite_goals_tac ctxt prems THEN resolve_tac ctxt [reflexive_thm] 1)
val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = prop_of lift;
@@ -365,11 +366,11 @@
(case apsns of
[] => compose_tac ctxt (false, inst_split Ts t tt thm TB state i, 0) i state
| p::_ => EVERY [lift_tac Ts t p,
- resolve_tac [reflexive_thm] (i+1),
+ resolve_tac ctxt [reflexive_thm] (i+1),
lift_split_tac] state)
end
in COND (has_fewer_prems i) no_tac
- (resolve_tac [meta_iffD] i THEN lift_split_tac)
+ (resolve_tac ctxt [meta_iffD] i THEN lift_split_tac)
end;
in (split_tac, exported_split_posns) end; (* mk_case_split_tac *)
@@ -400,16 +401,16 @@
(* does not work properly if the split variable is bound by a quantifier *)
fun flat_prems_tac i = SUBGOAL (fn (t,i) =>
(if first_prem_is_disj t
- then EVERY[eresolve_tac [Data.disjE] i, rotate_tac ~1 i,
+ then EVERY[eresolve_tac ctxt [Data.disjE] i, rotate_tac ~1 i,
rotate_tac ~1 (i+1),
flat_prems_tac (i+1)]
else all_tac)
- THEN REPEAT (eresolve_tac [Data.conjE,Data.exE] i)
- THEN REPEAT (dresolve_tac [Data.notnotD] i)) i;
+ THEN REPEAT (eresolve_tac ctxt [Data.conjE,Data.exE] i)
+ THEN REPEAT (dresolve_tac ctxt [Data.notnotD] i)) i;
in if n<0 then no_tac else (DETERM (EVERY'
- [rotate_tac n, eresolve_tac [Data.contrapos2],
+ [rotate_tac n, eresolve_tac ctxt [Data.contrapos2],
split_tac ctxt splits,
- rotate_tac ~1, eresolve_tac [Data.contrapos], rotate_tac ~1,
+ rotate_tac ~1, eresolve_tac ctxt [Data.contrapos], rotate_tac ~1,
flat_prems_tac] i))
end;
in SUBGOAL tac
--- a/src/Provers/trancl.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/trancl.ML Tue Feb 10 14:48:26 2015 +0100
@@ -541,11 +541,11 @@
val prems = flat (map_index (mkasm_trancl rel o swap) Hs);
val prfs = solveTrancl (prems, C);
in
- Subgoal.FOCUS (fn {prems, concl, ...} =>
+ Subgoal.FOCUS (fn {context = ctxt', prems, concl, ...} =>
let
val SOME (_, _, rel', _) = decomp (term_of concl);
val thms = map (prove thy rel' prems) prfs
- in resolve_tac [prove thy rel' thms prf] 1 end) ctxt n st
+ in resolve_tac ctxt' [prove thy rel' thms prf] 1 end) ctxt n st
end
handle Cannot => Seq.empty);
@@ -560,11 +560,11 @@
val prems = flat (map_index (mkasm_rtrancl rel o swap) Hs);
val prfs = solveRtrancl (prems, C);
in
- Subgoal.FOCUS (fn {prems, concl, ...} =>
+ Subgoal.FOCUS (fn {context = ctxt', prems, concl, ...} =>
let
val SOME (_, _, rel', _) = decomp (term_of concl);
val thms = map (prove thy rel' prems) prfs
- in resolve_tac [prove thy rel' thms prf] 1 end) ctxt n st
+ in resolve_tac ctxt' [prove thy rel' thms prf] 1 end) ctxt n st
end
handle Cannot => Seq.empty | General.Subscript => Seq.empty);
--- a/src/Provers/typedsimp.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Provers/typedsimp.ML Tue Feb 10 14:48:26 2015 +0100
@@ -109,16 +109,16 @@
(*Single step for conditional rewriting: prove_cond_tac handles new subgoals.*)
fun cond_step_tac ctxt (prove_cond_tac, congr_rls, asms) =
asm_res_tac ctxt asms ORELSE' rewrite_res_tac ctxt ORELSE'
- (resolve_tac [trans, red_trans] THEN' prove_cond_tac) ORELSE'
+ (resolve_tac ctxt [trans, red_trans] THEN' prove_cond_tac) ORELSE'
subconv_res_tac ctxt congr_rls;
(*Unconditional normalization tactic*)
fun norm_tac ctxt arg = REPEAT_FIRST (step_tac ctxt arg) THEN
- TRYALL (resolve_tac [red_if_equal]);
+ TRYALL (resolve_tac ctxt [red_if_equal]);
(*Conditional normalization tactic*)
fun cond_norm_tac ctxt arg = REPEAT_FIRST (cond_step_tac ctxt arg) THEN
- TRYALL (resolve_tac [red_if_equal]);
+ TRYALL (resolve_tac ctxt [red_if_equal]);
end;
end;
--- a/src/Pure/Isar/class.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/class.ML Tue Feb 10 14:48:26 2015 +0100
@@ -47,7 +47,7 @@
-> morphism -> local_theory -> local_theory
(*tactics*)
- val intro_classes_tac: thm list -> tactic
+ val intro_classes_tac: Proof.context -> thm list -> tactic
val default_intro_tac: Proof.context -> thm list -> tactic
(*diagnostics*)
@@ -437,7 +437,7 @@
(fst o rules thy) sub];
val classrel =
Goal.prove_sorry_global thy [] [] (Logic.mk_classrel (sub, sup))
- (K (EVERY (map (TRYALL o resolve_tac o single) intros)));
+ (fn {context = ctxt, ...} => EVERY (map (TRYALL o resolve_tac ctxt o single) intros));
val diff_sort = Sign.complete_sort thy [sup]
|> subtract (op =) (Sign.complete_sort thy [sub])
|> filter (is_class thy);
@@ -702,7 +702,7 @@
(** tactics and methods **)
-fun intro_classes_tac facts st =
+fun intro_classes_tac ctxt facts st =
let
val thy = Thm.theory_of_thm st;
val classes = Sign.all_classes thy;
@@ -710,12 +710,12 @@
val class_intros = map_filter (try (#intro o Axclass.get_info thy)) classes;
val assm_intros = all_assm_intros thy;
in
- Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
+ Method.intros_tac ctxt (class_trivs @ class_intros @ assm_intros) facts st
end;
fun default_intro_tac ctxt [] =
COND Thm.no_prems no_tac
- (intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt [])
+ (intro_classes_tac ctxt [] ORELSE Locale.intro_locales_tac true ctxt [])
| default_intro_tac _ _ = no_tac;
fun default_tac rules ctxt facts =
@@ -723,7 +723,7 @@
default_intro_tac ctxt facts;
val _ = Theory.setup
- (Method.setup @{binding intro_classes} (Scan.succeed (K (METHOD intro_classes_tac)))
+ (Method.setup @{binding intro_classes} (Scan.succeed (METHOD o intro_classes_tac))
"back-chain introduction rules of classes" #>
Method.setup @{binding default} (Attrib.thms >> (METHOD oo default_tac))
"apply some intro/elim rule");
--- a/src/Pure/Isar/class_declaration.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/class_declaration.ML Tue Feb 10 14:48:26 2015 +0100
@@ -56,7 +56,7 @@
val loc_intro_tac =
(case Locale.intros_of thy class of
(_, NONE) => all_tac
- | (_, SOME intro) => ALLGOALS (resolve_tac [intro]));
+ | (_, SOME intro) => ALLGOALS (resolve_tac empty_ctxt [intro]));
val tac = loc_intro_tac
THEN ALLGOALS (Proof_Context.fact_tac empty_ctxt (sup_axioms @ the_list assm_axiom));
in Element.prove_witness empty_ctxt prop tac end) some_prop;
--- a/src/Pure/Isar/element.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/element.ML Tue Feb 10 14:48:26 2015 +0100
@@ -270,14 +270,14 @@
Witness (t,
Thm.close_derivation
(Goal.prove ctxt [] [] (mark_witness t)
- (fn _ => resolve_tac [Drule.protectI] 1 THEN tac)));
+ (fn _ => resolve_tac ctxt [Drule.protectI] 1 THEN tac)));
local
val refine_witness =
- Proof.refine (Method.Basic (K (NO_CASES o
- K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (resolve_tac [Drule.protectI])))))))));
+ Proof.refine (Method.Basic (fn ctxt => NO_CASES o
+ K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (resolve_tac ctxt [Drule.protectI]))))))));
fun gen_witness_proof proof after_qed wit_propss eq_props =
let
--- a/src/Pure/Isar/local_defs.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/local_defs.ML Tue Feb 10 14:48:26 2015 +0100
@@ -225,7 +225,7 @@
ALLGOALS
(CONVERSION (meta_rewrite_conv ctxt'') THEN'
rewrite_goal_tac ctxt'' [def] THEN'
- resolve_tac [Drule.reflexive_thm]))
+ resolve_tac ctxt'' [Drule.reflexive_thm]))
handle ERROR msg => cat_error msg "Failed to prove definitional specification";
in (((c, T), rhs), prove) end;
--- a/src/Pure/Isar/locale.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/locale.ML Tue Feb 10 14:48:26 2015 +0100
@@ -636,7 +636,7 @@
(* Tactics *)
fun gen_intro_locales_tac intros_tac eager ctxt =
- intros_tac
+ intros_tac ctxt
(get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));
val intro_locales_tac = gen_intro_locales_tac Method.intros_tac;
--- a/src/Pure/Isar/method.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/method.ML Tue Feb 10 14:48:26 2015 +0100
@@ -23,7 +23,7 @@
val unfold: thm list -> Proof.context -> method
val fold: thm list -> Proof.context -> method
val atomize: bool -> Proof.context -> method
- val this: method
+ val this: Proof.context -> method
val fact: thm list -> Proof.context -> method
val assm_tac: Proof.context -> int -> tactic
val all_assm_tac: Proof.context -> tactic
@@ -32,8 +32,8 @@
val trace: Proof.context -> thm list -> unit
val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic
val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic
- val intros_tac: thm list -> thm list -> tactic
- val try_intros_tac: thm list -> thm list -> tactic
+ val intros_tac: Proof.context -> thm list -> thm list -> tactic
+ val try_intros_tac: Proof.context -> thm list -> thm list -> tactic
val rule: Proof.context -> thm list -> method
val erule: Proof.context -> int -> thm list -> method
val drule: Proof.context -> int -> thm list -> method
@@ -99,7 +99,7 @@
local
fun cut_rule_tac rule =
- resolve_tac [Drule.forall_intr_vars rule COMP_INCR revcut_rl];
+ resolve0_tac [Drule.forall_intr_vars rule COMP_INCR revcut_rl];
in
@@ -120,7 +120,7 @@
fun cheating ctxt int = METHOD (fn _ => fn st =>
if int orelse Config.get ctxt quick_and_dirty then
- ALLGOALS Skip_Proof.cheat_tac st
+ ALLGOALS (Skip_Proof.cheat_tac ctxt) st
else error "Cheating requires quick_and_dirty mode!");
@@ -146,7 +146,7 @@
(* this -- resolve facts directly *)
-val this = METHOD (EVERY o map (HEADGOAL o resolve_tac o single));
+fun this ctxt = METHOD (EVERY o map (HEADGOAL o resolve_tac ctxt o single));
(* fact -- composition by facts from context *)
@@ -159,17 +159,17 @@
local
-fun cond_rtac cond rule = SUBGOAL (fn (prop, i) =>
+fun cond_rtac ctxt cond rule = SUBGOAL (fn (prop, i) =>
if cond (Logic.strip_assums_concl prop)
- then resolve_tac [rule] i else no_tac);
+ then resolve_tac ctxt [rule] i else no_tac);
in
fun assm_tac ctxt =
assume_tac ctxt APPEND'
Goal.assume_rule_tac ctxt APPEND'
- cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND'
- cond_rtac (can Logic.dest_term) Drule.termI;
+ cond_rtac ctxt (can Logic.dest_equals) Drule.reflexive_thm APPEND'
+ cond_rtac ctxt (can Logic.dest_term) Drule.termI;
fun all_assm_tac ctxt =
let
@@ -180,7 +180,7 @@
fun assumption ctxt = METHOD (HEADGOAL o
(fn [] => assm_tac ctxt
- | [fact] => solve_tac [fact]
+ | [fact] => solve_tac ctxt [fact]
| _ => K no_tac));
fun finish immed ctxt =
@@ -203,9 +203,9 @@
fun gen_rule_tac tac ctxt rules facts =
(fn i => fn st =>
- if null facts then tac rules i st
+ if null facts then tac ctxt rules i st
else
- Seq.maps (fn rule => (tac o single) rule i st)
+ Seq.maps (fn rule => (tac ctxt o single) rule i st)
(Drule.multi_resolves (SOME ctxt) facts rules))
THEN_ALL_NEW Goal.norm_hhf_tac ctxt;
@@ -238,9 +238,9 @@
(* intros_tac -- pervasive search spanned by intro rules *)
-fun gen_intros_tac goals intros facts =
+fun gen_intros_tac goals ctxt intros facts =
goals (insert_tac facts THEN'
- REPEAT_ALL_NEW (resolve_tac intros))
+ REPEAT_ALL_NEW (resolve_tac ctxt intros))
THEN Tactic.distinct_subgoals_tac;
val intros_tac = gen_intros_tac ALLGOALS;
@@ -313,7 +313,7 @@
fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r);
val succeed_text = Basic (K succeed);
val default_text = Source (Token.src ("default", Position.none) []);
-val this_text = Basic (K this);
+val this_text = Basic this;
val done_text = Basic (K (SIMPLE_METHOD all_tac));
fun sorry_text int = Basic (fn ctxt => cheating ctxt int);
@@ -603,7 +603,7 @@
setup @{binding erule} (xrule_meth erule) "apply rule in elimination manner (improper)" #>
setup @{binding drule} (xrule_meth drule) "apply rule in destruct manner (improper)" #>
setup @{binding frule} (xrule_meth frule) "apply rule in forward manner (improper)" #>
- setup @{binding this} (Scan.succeed (K this)) "apply current facts as rules" #>
+ setup @{binding this} (Scan.succeed this) "apply current facts as rules" #>
setup @{binding fact} (Attrib.thms >> fact) "composition by facts from context" #>
setup @{binding assumption} (Scan.succeed assumption)
"proof by assumption, preferring facts" #>
--- a/src/Pure/Isar/proof.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Isar/proof.ML Tue Feb 10 14:48:26 2015 +0100
@@ -440,12 +440,12 @@
Goal.norm_hhf_tac ctxt THEN'
SUBGOAL (fn (goal, i) =>
if can Logic.unprotect (Logic.strip_assums_concl goal) then
- eresolve_tac [Drule.protectI] i THEN finish_tac ctxt (n - 1) i
+ eresolve_tac ctxt [Drule.protectI] i THEN finish_tac ctxt (n - 1) i
else finish_tac ctxt (n - 1) (i + 1));
fun goal_tac ctxt rule =
Goal.norm_hhf_tac ctxt THEN'
- resolve_tac [rule] THEN'
+ resolve_tac ctxt [rule] THEN'
finish_tac ctxt (Thm.nprems_of rule);
fun FINDGOAL tac st =
@@ -883,9 +883,9 @@
in map (Logic.mk_term o Var) vars end;
fun refine_terms n =
- refine (Method.Basic (K (NO_CASES o
+ refine (Method.Basic (fn ctxt => NO_CASES o
K (HEADGOAL (PRECISE_CONJUNCTS n
- (HEADGOAL (CONJUNCTS (ALLGOALS (resolve_tac [Drule.termI])))))))))
+ (HEADGOAL (CONJUNCTS (ALLGOALS (resolve_tac ctxt [Drule.termI]))))))))
#> Seq.hd;
in
--- a/src/Pure/Tools/find_theorems.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Tools/find_theorems.ML Tue Feb 10 14:48:26 2015 +0100
@@ -188,9 +188,9 @@
fun limited_etac thm i =
Seq.take (Options.default_int @{system_option find_theorems_tactic_limit}) o
- eresolve_tac [thm] i;
+ eresolve_tac ctxt' [thm] i;
fun try_thm thm =
- if Thm.no_prems thm then resolve_tac [thm] 1 goal'
+ if Thm.no_prems thm then resolve_tac ctxt' [thm] 1 goal'
else
(limited_etac thm THEN_ALL_NEW (Goal.norm_hhf_tac ctxt' THEN' Method.assm_tac ctxt'))
1 goal';
--- a/src/Pure/Tools/rule_insts.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/Tools/rule_insts.ML Tue Feb 10 14:48:26 2015 +0100
@@ -338,13 +338,13 @@
(*warning: rule_tac etc. refer to dynamic subgoal context!*)
val _ = Theory.setup
- (Method.setup @{binding rule_tac} (method res_inst_tac (K resolve_tac))
+ (Method.setup @{binding rule_tac} (method res_inst_tac resolve_tac)
"apply rule (dynamic instantiation)" #>
- Method.setup @{binding erule_tac} (method eres_inst_tac (K eresolve_tac))
+ Method.setup @{binding erule_tac} (method eres_inst_tac eresolve_tac)
"apply rule in elimination manner (dynamic instantiation)" #>
- Method.setup @{binding drule_tac} (method dres_inst_tac (K dresolve_tac))
+ Method.setup @{binding drule_tac} (method dres_inst_tac dresolve_tac)
"apply rule in destruct manner (dynamic instantiation)" #>
- Method.setup @{binding frule_tac} (method forw_inst_tac (K forward_tac))
+ Method.setup @{binding frule_tac} (method forw_inst_tac forward_tac)
"apply rule in forward manner (dynamic instantiation)" #>
Method.setup @{binding cut_tac} (method cut_inst_tac (K cut_rules_tac))
"cut rule (dynamic instantiation)" #>
--- a/src/Pure/goal.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/goal.ML Tue Feb 10 14:48:26 2015 +0100
@@ -205,7 +205,7 @@
val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops);
fun tac' args st =
- if skip then ALLGOALS Skip_Proof.cheat_tac st before Skip_Proof.report ctxt
+ if skip then ALLGOALS (Skip_Proof.cheat_tac ctxt) st before Skip_Proof.report ctxt
else tac args st;
fun result () =
(case SINGLE (tac' {prems = prems, context = ctxt'}) (init stmt) of
@@ -255,7 +255,7 @@
fun prove_sorry ctxt xs asms prop tac =
if Config.get ctxt quick_and_dirty then
- prove ctxt xs asms prop (fn _ => ALLGOALS Skip_Proof.cheat_tac)
+ prove ctxt xs asms prop (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt))
else (if future_enabled 1 then prove_future_pri ~2 else prove) ctxt xs asms prop tac;
fun prove_sorry_global thy xs asms prop tac =
@@ -294,7 +294,7 @@
val adhoc_conjunction_tac = REPEAT_ALL_NEW
(SUBGOAL (fn (goal, i) =>
- if can Logic.dest_conjunction goal then resolve_tac [Conjunction.conjunctionI] i
+ if can Logic.dest_conjunction goal then resolve0_tac [Conjunction.conjunctionI] i
else no_tac));
val conjunction_tac = SUBGOAL (fn (goal, i) =>
@@ -318,7 +318,7 @@
(* hhf normal form *)
fun norm_hhf_tac ctxt =
- resolve_tac [Drule.asm_rl] (*cheap approximation -- thanks to builtin Logic.flatten_params*)
+ resolve_tac ctxt [Drule.asm_rl] (*cheap approximation -- thanks to builtin Logic.flatten_params*)
THEN' SUBGOAL (fn (t, i) =>
if Drule.is_norm_hhf t then all_tac
else rewrite_goal_tac ctxt Drule.norm_hhf_eqs i);
@@ -336,7 +336,7 @@
val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
val tacs = Rs |> map (fn R =>
- eresolve_tac [Raw_Simplifier.norm_hhf ctxt (Thm.trivial R)] THEN_ALL_NEW assume_tac ctxt);
+ eresolve_tac ctxt [Raw_Simplifier.norm_hhf ctxt (Thm.trivial R)] THEN_ALL_NEW assume_tac ctxt);
in fold_rev (curry op APPEND') tacs (K no_tac) i end);
end;
--- a/src/Pure/raw_simplifier.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/raw_simplifier.ML Tue Feb 10 14:48:26 2015 +0100
@@ -1309,7 +1309,7 @@
in bottomc (mode, Option.map (Drule.flexflex_unique (SOME ctxt)) oo prover, maxidx) ctxt ct end;
val simple_prover =
- SINGLE o (fn ctxt => ALLGOALS (resolve_tac (prems_of ctxt)));
+ SINGLE o (fn ctxt => ALLGOALS (resolve_tac ctxt (prems_of ctxt)));
fun rewrite _ _ [] = Thm.reflexive
| rewrite ctxt full thms =
--- a/src/Pure/simplifier.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/simplifier.ML Tue Feb 10 14:48:26 2015 +0100
@@ -384,7 +384,7 @@
val trivialities = Drule.reflexive_thm :: trivs;
fun unsafe_solver_tac ctxt =
- FIRST' [resolve_tac (trivialities @ Raw_Simplifier.prems_of ctxt), assume_tac ctxt];
+ FIRST' [resolve_tac ctxt (trivialities @ Raw_Simplifier.prems_of ctxt), assume_tac ctxt];
val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac;
(*no premature instantiation of variables during simplification*)
--- a/src/Pure/skip_proof.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/skip_proof.ML Tue Feb 10 14:48:26 2015 +0100
@@ -12,7 +12,7 @@
val report: Proof.context -> unit
val make_thm_cterm: cterm -> thm
val make_thm: theory -> term -> thm
- val cheat_tac: int -> tactic
+ val cheat_tac: Proof.context -> int -> tactic
end;
structure Skip_Proof: SKIP_PROOF =
@@ -37,7 +37,7 @@
(* cheat_tac *)
-fun cheat_tac i st =
- resolve_tac [make_thm (Thm.theory_of_thm st) (Var (("A", 0), propT))] i st;
+fun cheat_tac ctxt i st =
+ resolve_tac ctxt [make_thm (Thm.theory_of_thm st) (Var (("A", 0), propT))] i st;
end;
--- a/src/Pure/tactic.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Pure/tactic.ML Tue Feb 10 14:48:26 2015 +0100
@@ -12,18 +12,23 @@
val eq_assume_tac: int -> tactic
val compose_tac: Proof.context -> (bool * thm * int) -> int -> tactic
val make_elim: thm -> thm
- val biresolve_tac: (bool * thm) list -> int -> tactic
- val resolve_tac: thm list -> int -> tactic
- val eresolve_tac: thm list -> int -> tactic
- val forward_tac: thm list -> int -> tactic
- val dresolve_tac: thm list -> int -> tactic
+ val biresolve0_tac: (bool * thm) list -> int -> tactic
+ val biresolve_tac: Proof.context -> (bool * thm) list -> int -> tactic
+ val resolve0_tac: thm list -> int -> tactic
+ val resolve_tac: Proof.context -> thm list -> int -> tactic
+ val eresolve0_tac: thm list -> int -> tactic
+ val eresolve_tac: Proof.context -> thm list -> int -> tactic
+ val forward0_tac: thm list -> int -> tactic
+ val forward_tac: Proof.context -> thm list -> int -> tactic
+ val dresolve0_tac: thm list -> int -> tactic
+ val dresolve_tac: Proof.context -> thm list -> int -> tactic
val atac: int -> tactic
val rtac: thm -> int -> tactic
val dtac: thm -> int -> tactic
val etac: thm -> int -> tactic
val ftac: thm -> int -> tactic
val ares_tac: thm list -> int -> tactic
- val solve_tac: thm list -> int -> tactic
+ val solve_tac: Proof.context -> thm list -> int -> tactic
val bimatch_tac: Proof.context -> (bool * thm) list -> int -> tactic
val match_tac: Proof.context -> thm list -> int -> tactic
val ematch_tac: Proof.context -> thm list -> int -> tactic
@@ -50,7 +55,7 @@
val rotate_tac: int -> int -> tactic
val defer_tac: int -> tactic
val prefer_tac: int -> tactic
- val filter_prems_tac: (term -> bool) -> int -> tactic
+ val filter_prems_tac: Proof.context -> (term -> bool) -> int -> tactic
end;
signature TACTIC =
@@ -114,30 +119,35 @@
fun make_elim rl = zero_var_indexes (rl RS revcut_rl);
(*Attack subgoal i by resolution, using flags to indicate elimination rules*)
-fun biresolve_tac brules i = PRIMSEQ (Thm.biresolution NONE false brules i);
+fun biresolve0_tac brules i = PRIMSEQ (Thm.biresolution NONE false brules i);
+fun biresolve_tac ctxt brules i = PRIMSEQ (Thm.biresolution (SOME ctxt) false brules i);
(*Resolution: the simple case, works for introduction rules*)
-fun resolve_tac rules = biresolve_tac (map (pair false) rules);
+fun resolve0_tac rules = biresolve0_tac (map (pair false) rules);
+fun resolve_tac ctxt rules = biresolve_tac ctxt (map (pair false) rules);
(*Resolution with elimination rules only*)
-fun eresolve_tac rules = biresolve_tac (map (pair true) rules);
+fun eresolve0_tac rules = biresolve0_tac (map (pair true) rules);
+fun eresolve_tac ctxt rules = biresolve_tac ctxt (map (pair true) rules);
(*Forward reasoning using destruction rules.*)
-fun forward_tac rls = resolve_tac (map make_elim rls) THEN' atac;
+fun forward0_tac rls = resolve0_tac (map make_elim rls) THEN' atac;
+fun forward_tac ctxt rls = resolve_tac ctxt (map make_elim rls) THEN' atac;
(*Like forward_tac, but deletes the assumption after use.*)
-fun dresolve_tac rls = eresolve_tac (map make_elim rls);
+fun dresolve0_tac rls = eresolve0_tac (map make_elim rls);
+fun dresolve_tac ctxt rls = eresolve_tac ctxt (map make_elim rls);
(*Shorthand versions: for resolution with a single theorem*)
-fun rtac rl = resolve_tac [rl];
-fun dtac rl = dresolve_tac [rl];
-fun etac rl = eresolve_tac [rl];
-fun ftac rl = forward_tac [rl];
+fun rtac rl = resolve0_tac [rl];
+fun dtac rl = dresolve0_tac [rl];
+fun etac rl = eresolve0_tac [rl];
+fun ftac rl = forward0_tac [rl];
(*Use an assumption or some rules ... A popular combination!*)
-fun ares_tac rules = atac ORELSE' resolve_tac rules;
+fun ares_tac rules = atac ORELSE' resolve0_tac rules;
-fun solve_tac rules = resolve_tac rules THEN_ALL_NEW atac;
+fun solve_tac ctxt rules = resolve_tac ctxt rules THEN_ALL_NEW assume_tac ctxt;
(*Matching tactics -- as above, but forbid updating of state*)
fun bimatch_tac ctxt brules i = PRIMSEQ (Thm.biresolution (SOME ctxt) true brules i);
@@ -175,7 +185,7 @@
(*The conclusion of the rule gets assumed in subgoal i,
while subgoal i+1,... are the premises of the rule.*)
-fun cut_tac rule i = resolve_tac [cut_rl] i THEN resolve_tac [rule] (i + 1);
+fun cut_tac rule i = resolve0_tac [cut_rl] i THEN resolve0_tac [rule] (i + 1);
(*"Cut" a list of rules into the goal. Their premises will become new
subgoals.*)
@@ -300,12 +310,12 @@
fun prefer_tac i = PRIMITIVE (Thm.permute_prems (i - 1) 1 #> Thm.permute_prems 0 ~1);
(* remove premises that do not satisfy p; fails if all prems satisfy p *)
-fun filter_prems_tac p =
+fun filter_prems_tac ctxt p =
let fun Then NONE tac = SOME tac
| Then (SOME tac) tac' = SOME(tac THEN' tac');
fun thins H (tac,n) =
if p H then (tac,n+1)
- else (Then tac (rotate_tac n THEN' eresolve_tac [thin_rl]),0);
+ else (Then tac (rotate_tac n THEN' eresolve_tac ctxt [thin_rl]),0);
in SUBGOAL(fn (subg,n) =>
let val Hs = Logic.strip_assums_hyp subg
in case fst(fold thins Hs (NONE,0)) of
--- a/src/Sequents/LK.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Sequents/LK.thy Tue Feb 10 14:48:26 2015 +0100
@@ -180,7 +180,7 @@
apply (tactic {*
REPEAT (rtac @{thm cut} 1 THEN
DEPTH_SOLVE_1
- (resolve_tac [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
+ (resolve_tac @{context} [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
Cla.safe_tac @{context} 1) *})
done
@@ -193,7 +193,7 @@
apply (tactic {*
REPEAT (rtac @{thm cut} 1 THEN
DEPTH_SOLVE_1
- (resolve_tac [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
+ (resolve_tac @{context} [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
Cla.safe_tac @{context} 1) *})
done
--- a/src/Sequents/modal.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Sequents/modal.ML Tue Feb 10 14:48:26 2015 +0100
@@ -50,22 +50,22 @@
(*Like filt_resolve_tac, using could_resolve_seq
Much faster than resolve_tac when there are many rules.
Resolve subgoal i using the rules, unless more than maxr are compatible. *)
-fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) =>
+fun filseq_resolve_tac ctxt rules maxr = SUBGOAL(fn (prem,i) =>
let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules)
- in if length rls > maxr then no_tac else resolve_tac rls i
+ in if length rls > maxr then no_tac else resolve_tac ctxt rls i
end);
-fun fresolve_tac rls n = filseq_resolve_tac rls 999 n;
+fun fresolve_tac ctxt rls n = filseq_resolve_tac ctxt rls 999 n;
(* NB No back tracking possible with aside rules *)
val aside_net = Tactic.build_net Modal_Rule.aside_rls;
fun aside_tac ctxt n = DETERM (REPEAT (filt_resolve_from_net_tac ctxt 999 aside_net n));
-fun rule_tac ctxt rls n = fresolve_tac rls n THEN aside_tac ctxt n;
+fun rule_tac ctxt rls n = fresolve_tac ctxt rls n THEN aside_tac ctxt n;
-val fres_safe_tac = fresolve_tac Modal_Rule.safe_rls;
-fun fres_unsafe_tac ctxt = fresolve_tac Modal_Rule.unsafe_rls THEN' aside_tac ctxt;
-val fres_bound_tac = fresolve_tac Modal_Rule.bound_rls;
+fun fres_safe_tac ctxt = fresolve_tac ctxt Modal_Rule.safe_rls;
+fun fres_unsafe_tac ctxt = fresolve_tac ctxt Modal_Rule.unsafe_rls THEN' aside_tac ctxt;
+fun fres_bound_tac ctxt = fresolve_tac ctxt Modal_Rule.bound_rls;
fun UPTOGOAL n tf = let fun tac i = if i<n then all_tac
else tf(i) THEN tac(i-1)
@@ -75,16 +75,16 @@
fun solven_tac ctxt d n st = st |>
(if d < 0 then no_tac
else if nprems_of st = 0 then all_tac
- else (DETERM(fres_safe_tac n) THEN UPTOGOAL n (solven_tac ctxt d)) ORELSE
+ else (DETERM(fres_safe_tac ctxt n) THEN UPTOGOAL n (solven_tac ctxt d)) ORELSE
((fres_unsafe_tac ctxt n THEN UPTOGOAL n (solven_tac ctxt d)) APPEND
- (fres_bound_tac n THEN UPTOGOAL n (solven_tac ctxt (d - 1)))));
+ (fres_bound_tac ctxt n THEN UPTOGOAL n (solven_tac ctxt (d - 1)))));
fun solve_tac ctxt d =
rewrite_goals_tac ctxt Modal_Rule.rewrite_rls THEN solven_tac ctxt d 1;
fun step_tac ctxt n =
COND (has_fewer_prems 1) all_tac
- (DETERM(fres_safe_tac n) ORELSE
- (fres_unsafe_tac ctxt n APPEND fres_bound_tac n));
+ (DETERM(fres_safe_tac ctxt n) ORELSE
+ (fres_unsafe_tac ctxt n APPEND fres_bound_tac ctxt n));
end;
--- a/src/Sequents/prover.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Sequents/prover.ML Tue Feb 10 14:48:26 2015 +0100
@@ -146,13 +146,13 @@
(*Like filt_resolve_tac, using could_resolve_seq
Much faster than resolve_tac when there are many rules.
Resolve subgoal i using the rules, unless more than maxr are compatible. *)
-fun filseq_resolve_tac rules maxr = SUBGOAL(fn (prem,i) =>
+fun filseq_resolve_tac ctxt rules maxr = SUBGOAL(fn (prem,i) =>
let val rls = filter_thms could_resolve_seq (maxr+1, prem, rules)
in if length rls > maxr then no_tac
else (*((rtac derelict 1 THEN rtac impl 1
THEN (rtac identity 2 ORELSE rtac ll_mp 2)
THEN rtac context1 1)
- ORELSE *) resolve_tac rls i
+ ORELSE *) resolve_tac ctxt rls i
end);
@@ -170,23 +170,23 @@
The abstraction over state prevents needless divergence in recursion.
The 9999 should be a parameter, to delay treatment of flexible goals. *)
-fun RESOLVE_THEN rules =
+fun RESOLVE_THEN ctxt rules =
let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules;
fun tac nextac i state = state |>
- (filseq_resolve_tac rls0 9999 i
+ (filseq_resolve_tac ctxt rls0 9999 i
ORELSE
- (DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i))
+ (DETERM(filseq_resolve_tac ctxt rls1 9999 i) THEN TRY(nextac i))
ORELSE
- (DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1))
+ (DETERM(filseq_resolve_tac ctxt rls2 9999 i) THEN TRY(nextac(i+1))
THEN TRY(nextac i)))
in tac end;
(*repeated resolution applied to the designated goal*)
-fun reresolve_tac rules =
+fun reresolve_tac ctxt rules =
let
- val restac = RESOLVE_THEN rules; (*preprocessing done now*)
+ val restac = RESOLVE_THEN ctxt rules; (*preprocessing done now*)
fun gtac i = restac gtac i;
in gtac end;
@@ -194,8 +194,8 @@
fun repeat_goal_tac ctxt =
let
val (safes, unsafes) = get_rules ctxt;
- val restac = RESOLVE_THEN safes;
- val lastrestac = RESOLVE_THEN unsafes;
+ val restac = RESOLVE_THEN ctxt safes;
+ val lastrestac = RESOLVE_THEN ctxt unsafes;
fun gtac i =
restac gtac i ORELSE
(if Config.get ctxt trace then print_tac ctxt "" THEN lastrestac gtac i
@@ -204,11 +204,11 @@
(*Tries safe rules only*)
-fun safe_tac ctxt = reresolve_tac (#1 (get_rules ctxt));
+fun safe_tac ctxt = reresolve_tac ctxt (#1 (get_rules ctxt));
(*Tries a safe rule or else a unsafe rule. Single-step for tracing. *)
fun step_tac ctxt =
- safe_tac ctxt ORELSE' filseq_resolve_tac (#2 (get_rules ctxt)) 9999;
+ safe_tac ctxt ORELSE' filseq_resolve_tac ctxt (#2 (get_rules ctxt)) 9999;
(* Tactic for reducing a goal, using Predicate Calculus rules.
--- a/src/Sequents/simpdata.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Sequents/simpdata.ML Tue Feb 10 14:48:26 2015 +0100
@@ -43,7 +43,7 @@
(*Replace premises x=y, X<->Y by X==Y*)
fun mk_meta_prems ctxt =
rule_by_tactic ctxt
- (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
+ (REPEAT_FIRST (resolve_tac ctxt [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
(*Congruence rules for = or <-> (instead of ==)*)
fun mk_meta_cong ctxt rl =
@@ -58,7 +58,7 @@
@{thm iff_refl}, reflexive_thm];
fun unsafe_solver ctxt =
- FIRST' [resolve_tac (triv_rls @ Simplifier.prems_of ctxt), assume_tac ctxt];
+ FIRST' [resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt), assume_tac ctxt];
(*No premature instantiation of variables during simplification*)
fun safe_solver ctxt =
--- a/src/Tools/case_product.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Tools/case_product.ML Tue Feb 10 14:48:26 2015 +0100
@@ -70,9 +70,10 @@
val (p_cons1 :: p_cons2 :: premss) = unflat struc prems
val thm2' = thm2 OF p_cons2
in
- resolve_tac [thm1 OF p_cons1]
+ resolve_tac ctxt [thm1 OF p_cons1]
THEN' EVERY' (map (fn p =>
- resolve_tac [thm2'] THEN' EVERY' (map (Proof_Context.fact_tac ctxt o single) p)) premss)
+ resolve_tac ctxt [thm2'] THEN'
+ EVERY' (map (Proof_Context.fact_tac ctxt o single) p)) premss)
end
fun combine ctxt thm1 thm2 =
--- a/src/Tools/coherent.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Tools/coherent.ML Tue Feb 10 14:48:26 2015 +0100
@@ -215,7 +215,7 @@
(** external interface **)
fun coherent_tac ctxt rules = SUBPROOF (fn {prems, concl, params, context = ctxt', ...} =>
- resolve_tac [rulify_elim_conv ctxt' concl RS Drule.equal_elim_rule2] 1 THEN
+ resolve_tac ctxt' [rulify_elim_conv ctxt' concl RS Drule.equal_elim_rule2] 1 THEN
SUBPROOF (fn {prems = prems', concl, context = ctxt'', ...} =>
let
val xs =
@@ -227,7 +227,7 @@
valid ctxt'' (map (mk_rule ctxt'') (prems' @ prems @ rules)) (term_of concl)
(mk_dom xs) Net.empty 0 0 of
NONE => no_tac
- | SOME prf => resolve_tac [thm_of_cl_prf ctxt'' concl [] prf] 1)
+ | SOME prf => resolve_tac ctxt'' [thm_of_cl_prf ctxt'' concl [] prf] 1)
end) ctxt' 1) ctxt;
val _ = Theory.setup
--- a/src/Tools/eqsubst.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Tools/eqsubst.ML Tue Feb 10 14:48:26 2015 +0100
@@ -250,7 +250,7 @@
RW_Inst.rw ctxt m rule conclthm
|> unfix_frees cfvs
|> Conv.fconv_rule Drule.beta_eta_conversion
- |> (fn r => resolve_tac [r] i st);
+ |> (fn r => resolve_tac ctxt [r] i st);
(* substitute within the conclusion of goal i of gth, using a meta
equation rule. Note that we assume rule has var indicies zero'd *)
@@ -332,7 +332,8 @@
|> Conv.fconv_rule Drule.beta_eta_conversion; (* normal form *)
in
(* ~j because new asm starts at back, thus we subtract 1 *)
- Seq.map (Thm.rotate_rule (~ j) (Thm.nprems_of rule + i)) (dresolve_tac [preelimrule] i st2)
+ Seq.map (Thm.rotate_rule (~ j) (Thm.nprems_of rule + i))
+ (dresolve_tac ctxt [preelimrule] i st2)
end;
--- a/src/Tools/induct.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/Tools/induct.ML Tue Feb 10 14:48:26 2015 +0100
@@ -507,7 +507,7 @@
in
CASES (Rule_Cases.make_common (thy,
Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
- ((Method.insert_tac more_facts THEN' resolve_tac [rule'] THEN_ALL_NEW
+ ((Method.insert_tac more_facts THEN' resolve_tac ctxt [rule'] THEN_ALL_NEW
(if simp then TRY o trivial_tac ctxt else K all_tac)) i) st
end)
end;
@@ -680,7 +680,7 @@
(case goal_concl n [] goal of
SOME concl =>
(compose_tac ctxt (false, spec_rule (goal_prefix n goal) concl, 1) THEN'
- resolve_tac [asm_rl]) i
+ resolve_tac ctxt [asm_rl]) i
| NONE => all_tac)
end);
@@ -801,7 +801,7 @@
|> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
|> Seq.maps (fn rule' =>
CASES (rule_cases ctxt rule' cases)
- (resolve_tac [rule'] i THEN
+ (resolve_tac ctxt [rule'] i THEN
PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) st')))
end)
THEN_ALL_NEW_CASES
@@ -859,7 +859,7 @@
|> Seq.maps (fn rule' =>
CASES (Rule_Cases.make_common (thy,
Thm.prop_of (Rule_Cases.internalize_params rule')) cases)
- (Method.insert_tac more_facts i THEN resolve_tac [rule'] i) st))
+ (Method.insert_tac more_facts i THEN resolve_tac ctxt [rule'] i) st))
end);
end;
--- a/src/ZF/Datatype_ZF.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/Datatype_ZF.thy Tue Feb 10 14:48:26 2015 +0100
@@ -95,7 +95,7 @@
else ();
val goal = Logic.mk_equals (old, new)
val thm = Goal.prove ctxt [] [] goal
- (fn _ => resolve_tac @{thms iff_reflection} 1 THEN
+ (fn _ => resolve_tac ctxt @{thms iff_reflection} 1 THEN
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 ctxt goal);
--- a/src/ZF/OrdQuant.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/OrdQuant.thy Tue Feb 10 14:48:26 2015 +0100
@@ -370,14 +370,14 @@
fn _ => Quantifier1.rearrange_bex
(fn ctxt =>
unfold_tac ctxt @{thms rex_def} THEN
- Quantifier1.prove_one_point_ex_tac)
+ Quantifier1.prove_one_point_ex_tac ctxt)
*}
simproc_setup defined_rall ("\<forall>x[M]. P(x) \<longrightarrow> Q(x)") = {*
fn _ => Quantifier1.rearrange_ball
(fn ctxt =>
unfold_tac ctxt @{thms rall_def} THEN
- Quantifier1.prove_one_point_all_tac)
+ Quantifier1.prove_one_point_all_tac ctxt)
*}
end
--- a/src/ZF/Tools/datatype_package.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/Tools/datatype_package.ML Tue Feb 10 14:48:26 2015 +0100
@@ -286,9 +286,9 @@
(*Proves a single case equation. Could use simp_tac, but it's slower!*)
(fn {context = ctxt, ...} => EVERY
[rewrite_goals_tac ctxt [con_def],
- resolve_tac [case_trans] 1,
+ resolve_tac ctxt [case_trans] 1,
REPEAT
- (resolve_tac [@{thm refl}, split_trans,
+ (resolve_tac ctxt [@{thm refl}, split_trans,
Su.case_inl RS @{thm trans}, Su.case_inr RS @{thm trans}] 1)]);
val free_iffs = map Drule.export_without_context (con_defs RL [@{thm def_swap_iff}]);
@@ -327,7 +327,7 @@
(Ind_Syntax.traceIt "next recursor equation = " thy1 (mk_recursor_eqn arg))
(*Proves a single recursor equation.*)
(fn {context = ctxt, ...} => EVERY
- [resolve_tac [recursor_trans] 1,
+ [resolve_tac ctxt [recursor_trans] 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
--- a/src/ZF/Tools/inductive_package.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/Tools/inductive_package.ML Tue Feb 10 14:48:26 2015 +0100
@@ -187,8 +187,8 @@
val bnd_mono =
Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs))
- (fn _ => EVERY
- [resolve_tac [@{thm Collect_subset} RS @{thm bnd_monoI}] 1,
+ (fn {context = ctxt, ...} => EVERY
+ [resolve_tac ctxt [@{thm Collect_subset} RS @{thm bnd_monoI}] 1,
REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]);
val dom_subset = Drule.export_without_context (big_rec_def RS Fp.subs);
@@ -214,21 +214,21 @@
(*Type-checking is hardest aspect of proof;
disjIn selects the correct disjunct after unfolding*)
fun intro_tacsf disjIn ctxt =
- [DETERM (stac unfold 1),
- REPEAT (resolve_tac [@{thm Part_eqI}, @{thm CollectI}] 1),
+ [DETERM (stac ctxt unfold 1),
+ REPEAT (resolve_tac ctxt [@{thm Part_eqI}, @{thm CollectI}] 1),
(*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*)
- resolve_tac [disjIn] 2,
+ resolve_tac ctxt [disjIn] 2,
(*Not ares_tac, since refl must be tried before equality assumptions;
backtracking may occur if the premises have extra variables!*)
- DEPTH_SOLVE_1 (resolve_tac [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac ctxt 2),
+ DEPTH_SOLVE_1 (resolve_tac ctxt [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac ctxt 2),
(*Now solve the equations like Tcons(a,f) = Inl(?b4)*)
rewrite_goals_tac ctxt con_defs,
- REPEAT (resolve_tac @{thms refl} 2),
+ REPEAT (resolve_tac ctxt @{thms refl} 2),
(*Typechecking; this can fail*)
if !Ind_Syntax.trace then print_tac ctxt "The type-checking subgoal:"
else all_tac,
- REPEAT (FIRSTGOAL ( dresolve_tac rec_typechecks
- ORELSE' eresolve_tac (asm_rl :: @{thm PartE} :: @{thm SigmaE2} ::
+ REPEAT (FIRSTGOAL (dresolve_tac ctxt rec_typechecks
+ ORELSE' eresolve_tac ctxt (asm_rl :: @{thm PartE} :: @{thm SigmaE2} ::
type_elims)
ORELSE' hyp_subst_tac ctxt)),
if !Ind_Syntax.trace then print_tac ctxt "The subgoal after monos, type_elims:"
@@ -252,7 +252,7 @@
(*Breaks down logical connectives in the monotonic function*)
fun basic_elim_tac ctxt =
- REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ Su.free_SEs)
+ REPEAT (SOMEGOAL (eresolve_tac ctxt (Ind_Syntax.elim_rls @ Su.free_SEs)
ORELSE' bound_hyp_subst_tac ctxt))
THEN prune_params_tac ctxt
(*Mutual recursion: collapse references to Part(D,h)*)
@@ -329,25 +329,25 @@
(empty_simpset (Proof_Context.init_global thy)
|> Simplifier.set_mksimps (K (map mk_eq o ZF_atomize o gen_all)))
setSolver (mk_solver "minimal"
- (fn ctxt => resolve_tac (triv_rls @ Simplifier.prems_of ctxt)
+ (fn ctxt => resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt)
ORELSE' assume_tac ctxt
- ORELSE' eresolve_tac @{thms FalseE}));
+ ORELSE' eresolve_tac ctxt @{thms FalseE}));
val quant_induct =
- Goal.prove_global thy1 [] ind_prems
+ Goal.prove_global thy [] ind_prems
(FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred)))
(fn {context = ctxt, prems} => EVERY
[rewrite_goals_tac ctxt part_rec_defs,
- resolve_tac [@{thm impI} RS @{thm allI}] 1,
- DETERM (eresolve_tac [raw_induct] 1),
+ resolve_tac ctxt [@{thm impI} RS @{thm allI}] 1,
+ DETERM (eresolve_tac ctxt [raw_induct] 1),
(*Push Part inside Collect*)
full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1,
(*This CollectE and disjE separates out the introduction rules*)
- REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm disjE}])),
+ REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm disjE}])),
(*Now break down the individual cases. No disjE here in case
some premise involves disjunction.*)
- REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm exE}, @{thm conjE}]
- ORELSE' (bound_hyp_subst_tac ctxt1))),
+ REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm exE}, @{thm conjE}]
+ ORELSE' (bound_hyp_subst_tac ctxt))),
ind_tac (rev (map (rewrite_rule ctxt part_rec_defs) prems)) (length prems)]);
val dummy =
@@ -410,20 +410,21 @@
else ();
- val lemma_tac = FIRST' [eresolve_tac [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}],
- resolve_tac [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}],
- dresolve_tac [@{thm spec}, @{thm mp}, Pr.fsplitD]];
+ fun lemma_tac ctxt =
+ FIRST' [eresolve_tac ctxt [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}],
+ resolve_tac ctxt [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}],
+ dresolve_tac ctxt [@{thm spec}, @{thm mp}, Pr.fsplitD]];
val need_mutual = length rec_names > 1;
val lemma = (*makes the link between the two induction rules*)
if need_mutual then
(writeln " Proving the mutual induction rule...";
- Goal.prove_global thy1 [] []
+ Goal.prove_global thy [] []
(Logic.mk_implies (induct_concl, mutual_induct_concl))
(fn {context = ctxt, ...} => EVERY
[rewrite_goals_tac ctxt part_rec_defs,
- REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac 1)]))
+ REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac ctxt 1)]))
else (writeln " [ No mutual induction rule needed ]"; @{thm TrueI});
val dummy =
@@ -469,20 +470,20 @@
(mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1
THEN
(*unpackage and use "prem" in the corresponding place*)
- REPEAT (resolve_tac @{thms impI} 1) THEN
- resolve_tac [rewrite_rule ctxt all_defs prem] 1 THEN
+ REPEAT (resolve_tac ctxt @{thms impI} 1) THEN
+ resolve_tac ctxt [rewrite_rule ctxt all_defs prem] 1 THEN
(*prem must not be REPEATed below: could loop!*)
DEPTH_SOLVE (FIRSTGOAL (ares_tac [@{thm impI}] ORELSE'
- eresolve_tac (@{thm conjE} :: @{thm mp} :: cmonos))))
+ eresolve_tac ctxt (@{thm conjE} :: @{thm mp} :: cmonos))))
) i)
THEN mutual_ind_tac ctxt prems (i-1);
val mutual_induct_fsplit =
if need_mutual then
- Goal.prove_global thy1 [] (map (induct_prem (rec_tms~~preds)) intr_tms)
+ Goal.prove_global thy [] (map (induct_prem (rec_tms~~preds)) intr_tms)
mutual_induct_concl
(fn {context = ctxt, prems} => EVERY
- [resolve_tac [quant_induct RS lemma] 1,
+ [resolve_tac ctxt [quant_induct RS lemma] 1,
mutual_ind_tac ctxt (rev prems) (length prems)])
else @{thm TrueI};
@@ -493,8 +494,8 @@
The name "x.1" comes from the "RS spec" !*)
val inst =
case elem_frees of [_] => I
- | _ => Drule.instantiate_normalize ([], [(cterm_of thy1 (Var(("x",1), Ind_Syntax.iT)),
- cterm_of thy1 elem_tuple)]);
+ | _ => Drule.instantiate_normalize ([], [(cterm_of thy (Var(("x",1), Ind_Syntax.iT)),
+ cterm_of thy elem_tuple)]);
(*strip quantifier and the implication*)
val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));
--- a/src/ZF/Tools/primrec_package.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/Tools/primrec_package.ML Tue Feb 10 14:48:26 2015 +0100
@@ -177,7 +177,7 @@
eqn_terms |> map (fn t =>
Goal.prove_global thy1 [] [] (Ind_Syntax.traceIt "next primrec equation = " thy1 t)
(fn {context = ctxt, ...} =>
- EVERY [rewrite_goals_tac ctxt rewrites, resolve_tac @{thms refl} 1]));
+ EVERY [rewrite_goals_tac ctxt rewrites, resolve_tac ctxt @{thms refl} 1]));
val (eqn_thms', thy2) =
thy1
--- a/src/ZF/Tools/typechk.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/Tools/typechk.ML Tue Feb 10 14:48:26 2015 +0100
@@ -68,12 +68,12 @@
(* tactics *)
(*resolution using a net rather than rules*)
-fun net_res_tac maxr net =
+fun net_res_tac ctxt maxr net =
SUBGOAL
- (fn (prem,i) =>
+ (fn (prem, i) =>
let val rls = Net.unify_term net (Logic.strip_assums_concl prem)
in
- if length rls <= maxr then resolve_tac rls i else no_tac
+ if length rls <= maxr then resolve_tac ctxt rls i else no_tac
end);
fun is_rigid_elem (Const(@{const_name Trueprop},_) $ (Const(@{const_name mem},_) $ a $ _)) =
@@ -89,7 +89,7 @@
match_tac is too strict; would refuse to instantiate ?A*)
fun typecheck_step_tac ctxt =
let val TC {net, ...} = tcset_of ctxt
- in FIRSTGOAL (test_assume_tac ctxt ORELSE' net_res_tac 3 net) end;
+ in FIRSTGOAL (test_assume_tac ctxt ORELSE' net_res_tac ctxt 3 net) end;
fun typecheck_tac ctxt = REPEAT (typecheck_step_tac ctxt);
@@ -100,7 +100,7 @@
(*Instantiates variables in typing conditions.
drawback: does not simplify conjunctions*)
fun type_solver_tac ctxt hyps = SELECT_GOAL
- (DEPTH_SOLVE (eresolve_tac @{thms FalseE} 1
+ (DEPTH_SOLVE (eresolve_tac ctxt @{thms FalseE} 1
ORELSE resolve_from_net_tac ctxt basic_net 1
ORELSE (ares_tac hyps 1
APPEND typecheck_step_tac ctxt)));
--- a/src/ZF/UNITY/Constrains.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/UNITY/Constrains.thy Tue Feb 10 14:48:26 2015 +0100
@@ -472,7 +472,7 @@
(EVERY [REPEAT (Always_Int_tac ctxt 1),
REPEAT (etac @{thm Always_ConstrainsI} 1
ORELSE
- resolve_tac [@{thm StableI}, @{thm stableI},
+ resolve_tac ctxt [@{thm StableI}, @{thm stableI},
@{thm constrains_imp_Constrains}] 1),
rtac @{thm constrainsI} 1,
(* Three subgoals *)
--- a/src/ZF/arith_data.ML Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/arith_data.ML Tue Feb 10 14:48:26 2015 +0100
@@ -51,7 +51,7 @@
(*Apply the given rewrite (if present) just once*)
fun gen_trans_tac th2 NONE = all_tac
- | gen_trans_tac th2 (SOME th) = ALLGOALS (resolve_tac [th RS th2]);
+ | gen_trans_tac th2 (SOME th) = ALLGOALS (resolve0_tac [th RS th2]);
(*Use <-> or = depending on the type of t*)
fun mk_eq_iff(t,u) =
--- a/src/ZF/pair.thy Tue Feb 10 14:29:36 2015 +0100
+++ b/src/ZF/pair.thy Tue Feb 10 14:48:26 2015 +0100
@@ -22,14 +22,14 @@
fn _ => Quantifier1.rearrange_bex
(fn ctxt =>
unfold_tac ctxt @{thms Bex_def} THEN
- Quantifier1.prove_one_point_ex_tac)
+ Quantifier1.prove_one_point_ex_tac ctxt)
*}
simproc_setup defined_Ball ("\<forall>x\<in>A. P(x) \<longrightarrow> Q(x)") = {*
fn _ => Quantifier1.rearrange_ball
(fn ctxt =>
unfold_tac ctxt @{thms Ball_def} THEN
- Quantifier1.prove_one_point_all_tac)
+ Quantifier1.prove_one_point_all_tac ctxt)
*}