merged
authorwenzelm
Sat, 18 Jul 2015 20:59:51 +0200
changeset 60755 cde2b5d084e6
parent 60742 4050b243fc60 (current diff)
parent 60754 02924903a6fd (diff)
child 60756 f122140b7195
merged
src/Pure/ML-Systems/ml_debugger_dummy.ML
--- a/src/CCL/CCL.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/CCL/CCL.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -268,7 +268,7 @@
   let
     fun mk_raw_dstnct_thm rls s =
       Goal.prove_global @{theory} [] [] (Syntax.read_prop_global @{theory} s)
-        (fn {context = ctxt, ...} => rtac @{thm notI} 1 THEN eresolve_tac ctxt rls 1)
+        (fn {context = ctxt, ...} => resolve_tac ctxt @{thms 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/Type.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/CCL/Type.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -398,7 +398,7 @@
 
 ML {*
 fun genIs_tac ctxt genXH gen_mono =
-  rtac (genXH RS @{thm iffD2}) THEN'
+  resolve_tac ctxt [genXH RS @{thm iffD2}] THEN'
   simp_tac ctxt THEN'
   TRY o fast_tac
     (ctxt addIs [genXH RS @{thm iffD2}, gen_mono RS @{thm coinduct3_mono_lemma} RS @{thm lfpI}])
@@ -436,7 +436,7 @@
 ML {*
 fun POgen_tac ctxt (rla, rlb) i =
   SELECT_GOAL (safe_tac ctxt) i THEN
-  rtac (rlb RS (rla RS @{thm ssubst_pair})) i THEN
+  resolve_tac ctxt [rlb RS (rla RS @{thm ssubst_pair})] i THEN
   (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}]) @
--- a/src/CCL/Wfd.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/CCL/Wfd.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -448,7 +448,7 @@
 in
 
 fun rcall_tac ctxt i =
-  let fun tac ps rl i = Rule_Insts.res_inst_tac ctxt ps [] rl i THEN atac i
+  let fun tac ps rl i = Rule_Insts.res_inst_tac ctxt ps [] rl i THEN assume_tac ctxt i
   in IHinst tac @{thms rcallTs} i end
   THEN eresolve_tac ctxt @{thms rcall_lemmas} i
 
@@ -468,7 +468,7 @@
 (*** Clean up Correctness Condictions ***)
 
 fun clean_ccs_tac ctxt =
-  let fun tac ps rl i = Rule_Insts.eres_inst_tac ctxt ps [] rl i THEN atac i in
+  let fun tac ps rl i = Rule_Insts.eres_inst_tac ctxt ps [] rl i THEN assume_tac ctxt i in
     TRY (REPEAT_FIRST (IHinst tac @{thms hyprcallTs} ORELSE'
       eresolve_tac ctxt ([asm_rl, @{thm SubtypeE}] @ @{thms rmIHs}) ORELSE'
       hyp_subst_tac ctxt))
@@ -525,7 +525,7 @@
   apply (rule 1 [THEN canonical])
   apply (tactic {*
     REPEAT (DEPTH_SOLVE_1 (resolve_tac @{context} (@{thms assms} @ @{thms eval_rls}) 1 ORELSE
-      etac @{thm substitute} 1)) *})
+      eresolve_tac @{context} @{thms substitute} 1)) *})
   done
 
 lemma fixV: "f(fix(f)) ---> c \<Longrightarrow> fix(f) ---> c"
--- a/src/CCL/ex/Stream.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/CCL/ex/Stream.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -82,7 +82,7 @@
   apply EQgen
    prefer 2
    apply blast
-  apply (tactic {* DEPTH_SOLVE (etac (XH_to_E @{thm ListsXH}) 1
+  apply (tactic {* DEPTH_SOLVE (eresolve_tac @{context} [XH_to_E @{thm ListsXH}] 1
     THEN EQgen_tac @{context} [] 1) *})
   done
 
--- a/src/CTT/CTT.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/CTT/CTT.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -421,25 +421,25 @@
     The (rtac EqE i) lets them apply to equality judgements. **)
 
 fun NE_tac ctxt sp i =
-  TRY (rtac @{thm EqE} i) THEN
+  TRY (resolve_tac ctxt @{thms EqE} i) THEN
   Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm NE} i
 
 fun SumE_tac ctxt sp i =
-  TRY (rtac @{thm EqE} i) THEN
+  TRY (resolve_tac ctxt @{thms EqE} i) THEN
   Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm SumE} i
 
 fun PlusE_tac ctxt sp i =
-  TRY (rtac @{thm EqE} i) THEN
+  TRY (resolve_tac ctxt @{thms EqE} i) THEN
   Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm PlusE} i
 
 (** Predicate logic reasoning, WITH THINNING!!  Procedures adapted from NJ. **)
 
 (*Finds f:Prod(A,B) and a:A in the assumptions, concludes there is z:B(a) *)
 fun add_mp_tac ctxt i =
-    rtac @{thm subst_prodE} i  THEN  assume_tac ctxt i  THEN  assume_tac ctxt i
+  resolve_tac ctxt @{thms subst_prodE} i  THEN  assume_tac ctxt i  THEN  assume_tac ctxt i
 
 (*Finds P-->Q and P in the assumptions, replaces implication by Q *)
-fun mp_tac ctxt i = etac @{thm subst_prodE} i  THEN  assume_tac ctxt i
+fun mp_tac ctxt i = eresolve_tac ctxt @{thms subst_prodE} i  THEN  assume_tac ctxt i
 
 (*"safe" when regarded as predicate calculus rules*)
 val safe_brls = sort (make_ord lessb)
--- a/src/Doc/Codegen/Evaluation.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Doc/Codegen/Evaluation.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -232,7 +232,7 @@
 ML_prf 
 (*>*) \<open>val thm =
   Reification.conv @{context} @{thms interp.simps} @{cterm "x < y \<and> x < z"}\<close> (*<*)
-by (tactic \<open>ALLGOALS (rtac thm)\<close>)
+by (tactic \<open>ALLGOALS (resolve_tac @{context} [thm])\<close>)
 (*>*) 
 
 text \<open>
--- a/src/Doc/Implementation/Isar.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Doc/Implementation/Isar.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -365,7 +365,7 @@
   assume a: A and b: B
 
   have "A \<and> B"
-    apply (tactic \<open>rtac @{thm conjI} 1\<close>)
+    apply (tactic \<open>resolve_tac @{context} @{thms conjI} 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
@@ -374,7 +374,7 @@
     using a and b
     ML_val \<open>@{Isar.goal}\<close>
     apply (tactic \<open>Method.insert_tac facts 1\<close>)
-    apply (tactic \<open>(rtac @{thm conjI} THEN_ALL_NEW atac) 1\<close>)
+    apply (tactic \<open>(resolve_tac @{context} @{thms conjI} THEN_ALL_NEW assume_tac @{context}) 1\<close>)
     done
 end
 
--- a/src/Doc/Implementation/Proof.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Doc/Implementation/Proof.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -492,7 +492,7 @@
   ML_prf %"ML"
    \<open>val ctxt0 = @{context};
     val (([(_, x)], [B]), ctxt1) = ctxt0
-      |> Obtain.result (fn _ => etac @{thm exE} 1) [@{thm ex}];\<close>
+      |> Obtain.result (fn _ => eresolve_tac ctxt0 @{thms exE} 1) [@{thm ex}];\<close>
   ML_prf %"ML"
    \<open>singleton (Proof_Context.export ctxt1 ctxt0) @{thm refl};\<close>
   ML_prf %"ML"
--- a/src/Doc/Tutorial/Protocol/Event.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Doc/Tutorial/Protocol/Event.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -261,7 +261,7 @@
 ML
 {*
 fun analz_mono_contra_tac ctxt =
-  rtac @{thm analz_impI} THEN' 
+  resolve_tac ctxt @{thms analz_impI} THEN' 
   REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
   THEN' mp_tac ctxt
 *}
@@ -331,7 +331,7 @@
 
 
 fun synth_analz_mono_contra_tac ctxt = 
-  rtac @{thm syan_impI} THEN'
+  resolve_tac ctxt @{thms syan_impI} THEN'
   REPEAT1 o 
     (dresolve_tac ctxt
      [@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
--- a/src/FOLP/FOLP.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/FOLP/FOLP.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -56,7 +56,7 @@
     and r2: "!!y. y:Q ==> g(y):R"
   shows "?p : R"
   apply (rule excluded_middle [THEN disjE])
-   apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE
+   apply (tactic {* DEPTH_SOLVE (assume_tac @{context} 1 ORELSE
        resolve_tac @{context} [@{thm r1}, @{thm r2}, @{thm major} RS @{thm mp}] 1) *})
   done
 
@@ -80,8 +80,9 @@
   apply (insert major)
   apply (unfold iff_def)
   apply (rule conjE)
-  apply (tactic {* DEPTH_SOLVE_1 (etac @{thm impCE} 1 ORELSE
-      eresolve_tac @{context} [@{thm notE}, @{thm impE}] 1 THEN atac 1 ORELSE atac 1 ORELSE
+  apply (tactic {* DEPTH_SOLVE_1 (eresolve_tac @{context} @{thms impCE} 1 ORELSE
+      eresolve_tac @{context} [@{thm notE}, @{thm impE}] 1 THEN assume_tac @{context} 1 ORELSE
+      assume_tac @{context} 1 ORELSE
       resolve_tac @{context} [@{thm r1}, @{thm r2}] 1) *})+
   done
 
@@ -107,7 +108,7 @@
   val mp = @{thm mp}
   val not_elim = @{thm notE}
   val swap = @{thm swap}
-  val hyp_subst_tacs = [hyp_subst_tac]
+  fun hyp_subst_tacs ctxt = [hyp_subst_tac ctxt]
 );
 open Cla;
 
--- a/src/FOLP/IFOLP.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/FOLP/IFOLP.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -504,19 +504,19 @@
 schematic_lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')"
   apply (rule iffI)
    apply (tactic {*
-     DEPTH_SOLVE (atac 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
+     DEPTH_SOLVE (assume_tac @{context} 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 @{context} [@{thm subst}, @{thm ssubst}] 1) *})
+     DEPTH_SOLVE (assume_tac @{context} 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 @{context} [@{thm subst}, @{thm ssubst}] 1) *})
+     DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
   done
 
 lemmas pred_congs = pred1_cong pred2_cong pred3_cong
@@ -543,7 +543,7 @@
   assumes major: "p:(P|Q)-->S"
     and minor: "!!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R"
   shows "?p:R"
-  apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE
+  apply (tactic {* DEPTH_SOLVE (assume_tac @{context} 1 ORELSE
       resolve_tac @{context} [@{thm disjI1}, @{thm disjI2}, @{thm impI},
         @{thm major} RS @{thm mp}, @{thm minor}] 1) *})
   done
--- a/src/FOLP/classical.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/FOLP/classical.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -22,7 +22,7 @@
   val not_elim: thm             (* [| ~P;  P |] ==> R *)
   val swap: thm                 (* ~P ==> (~Q ==> P) ==> Q *)
   val sizef : thm -> int        (* size function for BEST_FIRST *)
-  val hyp_subst_tacs: (int -> tactic) list
+  val hyp_subst_tacs: Proof.context -> (int -> tactic) list
   end;
 
 (*Higher precedence than := facilitates use of references*)
@@ -70,7 +70,7 @@
 val imp_elim = make_elim mp;
 
 (*Solve goal that assumes both P and ~P. *)
-fun contr_tac ctxt = etac not_elim THEN'  assume_tac ctxt;
+fun contr_tac ctxt = eresolve_tac ctxt [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 ctxt ([not_elim,imp_elim]) i  THEN  assume_tac ctxt i;
@@ -84,7 +84,7 @@
 (*Uses introduction rules in the normal way, or on negated assumptions,
   trying rules in order. *)
 fun swap_res_tac ctxt rls = 
-    let fun tacf rl = rtac rl ORELSE' etac (rl RSN (2,swap))
+    let fun tacf rl = resolve_tac ctxt [rl] ORELSE' eresolve_tac ctxt [rl RSN (2, swap)]
     in  assume_tac ctxt ORELSE' contr_tac ctxt ORELSE' FIRST' (map tacf rls)
     end;
 
@@ -152,7 +152,7 @@
   FIRST' [uniq_assume_tac ctxt,
           uniq_mp_tac ctxt,
           biresolve_tac ctxt safe0_brls,
-          FIRST' hyp_subst_tacs,
+          FIRST' (hyp_subst_tacs ctxt),
           biresolve_tac ctxt safep_brls] ;
 
 (*Repeatedly attack subgoals using safe inferences*)
--- a/src/FOLP/hypsubst.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/FOLP/hypsubst.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -19,8 +19,8 @@
 
 signature HYPSUBST =
   sig
-  val bound_hyp_subst_tac : int -> tactic
-  val hyp_subst_tac       : int -> tactic
+  val bound_hyp_subst_tac : Proof.context -> int -> tactic
+  val hyp_subst_tac       : Proof.context -> int -> tactic
     (*exported purely for debugging purposes*)
   val eq_var              : bool -> term -> int * thm
   val inspect_pair        : bool -> term * term -> thm
@@ -68,16 +68,16 @@
 
 (*Select a suitable equality assumption and substitute throughout the subgoal
   Replaces only Bound variables if bnd is true*)
-fun gen_hyp_subst_tac bnd = SUBGOAL(fn (Bi,i) =>
+fun gen_hyp_subst_tac bnd ctxt = SUBGOAL(fn (Bi,i) =>
       let val n = length(Logic.strip_assums_hyp Bi) - 1
           val (k,symopt) = eq_var bnd Bi
       in
          DETERM
-           (EVERY [REPEAT_DETERM_N k (etac rev_mp i),
-                   etac revcut_rl i,
-                   REPEAT_DETERM_N (n-k) (etac rev_mp i),
-                   etac (symopt RS subst) i,
-                   REPEAT_DETERM_N n (rtac imp_intr i)])
+           (EVERY [REPEAT_DETERM_N k (eresolve_tac ctxt [rev_mp] i),
+                   eresolve_tac ctxt [revcut_rl] i,
+                   REPEAT_DETERM_N (n-k) (eresolve_tac ctxt [rev_mp] i),
+                   eresolve_tac ctxt [symopt RS subst] i,
+                   REPEAT_DETERM_N n (resolve_tac ctxt [imp_intr] i)])
       end
       handle THM _ => no_tac | EQ_VAR => no_tac);
 
--- a/src/FOLP/intprover.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/FOLP/intprover.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -55,7 +55,7 @@
 fun safe_step_tac ctxt = FIRST' [uniq_assume_tac ctxt,
                             int_uniq_mp_tac ctxt,
                             biresolve_tac ctxt safe0_brls,
-                            hyp_subst_tac,
+                            hyp_subst_tac ctxt,
                             biresolve_tac ctxt safep_brls] ;
 
 (*Repeatedly attack subgoals using safe inferences*)
--- a/src/FOLP/simp.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/FOLP/simp.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -39,7 +39,7 @@
   val setauto   : simpset * (int -> tactic) -> simpset
   val ASM_SIMP_CASE_TAC : Proof.context -> simpset -> int -> tactic
   val ASM_SIMP_TAC      : Proof.context -> simpset -> int -> tactic
-  val CASE_TAC          : simpset -> int -> tactic
+  val CASE_TAC          : Proof.context -> simpset -> int -> tactic
   val SIMP_CASE2_TAC    : Proof.context -> simpset -> int -> tactic
   val SIMP_THM          : Proof.context -> simpset -> thm -> thm
   val SIMP_TAC          : Proof.context -> simpset -> int -> tactic
@@ -228,7 +228,7 @@
       in Variable.tradeT (K (map (add_norms o mk_trans) o maps mk_rew_rules)) ctxt [thm] end
   in normed end;
 
-fun NORM norm_lhs_tac = EVERY'[rtac red2 , norm_lhs_tac, refl_tac];
+fun NORM ctxt norm_lhs_tac = EVERY' [resolve_tac ctxt [red2], norm_lhs_tac, refl_tac];
 
 val trans_norms = map mk_trans normE_thms;
 
@@ -331,9 +331,9 @@
           | find_if(_) = false;
     in find_if(tm,0) end;
 
-fun IF1_TAC cong_tac i =
+fun IF1_TAC ctxt cong_tac i =
     let fun seq_try (ifth::ifths,ifc::ifcs) thm =
-                (COND (if_rewritable ifc i) (DETERM(rtac ifth i))
+                (COND (if_rewritable ifc i) (DETERM(resolve_tac ctxt [ifth] i))
                         (seq_try(ifths,ifcs))) thm
               | seq_try([],_) thm = no_tac thm
         and try_rew thm = (seq_try(case_rews,case_consts) ORELSE one_subt) thm
@@ -346,9 +346,9 @@
                 in (cong_tac THEN loop) thm end
     in COND (may_match(case_consts,i)) try_rew no_tac end;
 
-fun CASE_TAC (SS{cong_net,...}) i =
+fun CASE_TAC ctxt (SS{cong_net,...}) i =
 let val cong_tac = net_tac cong_net i
-in NORM (IF1_TAC cong_tac) i end;
+in NORM ctxt (IF1_TAC ctxt cong_tac) i end;
 
 (* Rewriting Automaton *)
 
@@ -441,7 +441,7 @@
                      thm', anet, ats, (ss,thm,anet,ats,seq',more)::cs)
             end
     | NONE => if more
-            then rew((lhs_net_tac anet i THEN atac i) thm,
+            then rew((lhs_net_tac anet i THEN assume_tac ctxt i) thm,
                      thm,ss,anet,ats,cs,false)
             else (ss,thm,anet,ats,cs);
 
@@ -457,7 +457,7 @@
               end;
 
 fun if_exp(thm,ss,anet,ats,cs) =
-        case Seq.pull (IF1_TAC (cong_tac i) i thm) of
+        case Seq.pull (IF1_TAC ctxt (cong_tac i) i thm) of
                 SOME(thm',_) => (SIMP_LHS::IF::ss,thm',anet,ats,cs)
               | NONE => (ss,thm,anet,ats,cs);
 
--- a/src/HOL/Auth/Event.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/Event.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -271,7 +271,7 @@
 ML
 {*
 fun analz_mono_contra_tac ctxt = 
-  rtac @{thm analz_impI} THEN' 
+  resolve_tac ctxt @{thms analz_impI} THEN' 
   REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
   THEN' (mp_tac ctxt)
 *}
@@ -287,7 +287,7 @@
 ML
 {*
 fun synth_analz_mono_contra_tac ctxt = 
-  rtac @{thm syan_impI} THEN'
+  resolve_tac ctxt @{thms syan_impI} THEN'
   REPEAT1 o 
     (dresolve_tac ctxt
      [@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
--- a/src/HOL/Auth/OtwayReesBella.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/OtwayReesBella.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -250,7 +250,7 @@
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
-          REPEAT_FIRST (rtac @{thm analz_image_freshCryptK_lemma}),
+          REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshCryptK_lemma}),
           ALLGOALS (asm_simp_tac
             (put_simpset OtwayReesBella.analz_image_freshK_ss ctxt))]))) *}
   "for proving useful rewrite rule"
--- a/src/HOL/Auth/Public.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/Public.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -434,7 +434,7 @@
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
-          REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
+          REPEAT_FIRST (resolve_tac ctxt @{thms 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	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/Shared.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -238,7 +238,7 @@
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
-          REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
+          REPEAT_FIRST (resolve_tac ctxt @{thms 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	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -753,8 +753,8 @@
 
 fun analz_prepare_tac ctxt =
          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 
+         dresolve_tac ctxt @{thms Gets_imp_knows_Spy_analz_Snd} 18 THEN 
+ (*SR9*) dresolve_tac ctxt @{thms Gets_imp_knows_Spy_analz_Snd} 19 THEN 
          REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
 
 end
@@ -818,7 +818,7 @@
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST
        (resolve_tac ctxt [allI, ballI, impI]),
-        REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
+        REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}),
         ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt
           addsimps [@{thm knows_Spy_Inputs_secureM_sr_Spy},
                     @{thm knows_Spy_Outpts_secureM_sr_Spy},
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -762,8 +762,8 @@
 
 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 
+         dresolve_tac ctxt @{thms Gets_imp_knows_Spy_analz_Snd} 18 THEN 
+ (*SR_U9*) dresolve_tac ctxt @{thms Gets_imp_knows_Spy_analz_Snd} 19 THEN 
          REPEAT_FIRST (eresolve_tac ctxt [asm_rl, conjE] ORELSE' hyp_subst_tac ctxt)
 
 end
@@ -827,7 +827,7 @@
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
-          REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
+          REPEAT_FIRST (resolve_tac ctxt @{thms analz_image_freshK_lemma}),
           ALLGOALS (asm_simp_tac (put_simpset Smartcard.analz_image_freshK_ss ctxt
               addsimps [@{thm knows_Spy_Inputs_secureM_srb_Spy},
                   @{thm knows_Spy_Outpts_secureM_srb_Spy},
--- a/src/HOL/Auth/Smartcard/Smartcard.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -404,7 +404,7 @@
     Scan.succeed (fn ctxt =>
      (SIMPLE_METHOD
       (EVERY [REPEAT_FIRST (resolve_tac ctxt [allI, ballI, impI]),
-          REPEAT_FIRST (rtac @{thm analz_image_freshK_lemma}),
+          REPEAT_FIRST (resolve_tac ctxt @{thms 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/AxCompl.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Bali/AxCompl.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -1370,7 +1370,7 @@
 apply -
 apply (induct_tac "n")
 apply  (tactic "ALLGOALS (clarsimp_tac @{context})")
-apply  (tactic {* dtac (Thm.permute_prems 0 1 @{thm card_seteq}) 1 *})
+apply  (tactic {* dresolve_tac @{context} [Thm.permute_prems 0 1 @{thm card_seteq}] 1 *})
 apply    simp
 apply   (erule finite_imageI)
 apply  (simp add: MGF_asm ax_derivs_asm)
--- a/src/HOL/Bali/AxExample.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Bali/AxExample.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -47,7 +47,7 @@
   | NONE => Seq.empty);
 
 fun ax_tac ctxt =
-  REPEAT o rtac allI THEN'
+  REPEAT o resolve_tac ctxt [allI] THEN'
   resolve_tac ctxt
     @{thms ax_Skip ax_StatRef ax_MethdN ax_Alloc ax_Alloc_Arr ax_SXAlloc_Normal ax_derivs.intros(8-)};
 *}
--- a/src/HOL/Bali/AxSem.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Bali/AxSem.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -692,8 +692,9 @@
 apply (erule ax_derivs.induct)
 (*42 subgoals*)
 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 {* ALLGOALS(REPEAT o (EVERY'[dresolve_tac @{context} @{thms subset_singletonD},
+         eresolve_tac @{context} [disjE],
+         fast_tac (@{context} addSIs @{thms ax_derivs.empty})]))*})
 apply       (tactic "TRYALL (hyp_subst_tac @{context})")
 apply       (simp, rule ax_derivs.empty)
 apply      (drule subset_insertD)
--- a/src/HOL/Bali/Evaln.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Bali/Evaln.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -448,9 +448,10 @@
 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 @{context}, TRY o etac @{thm Suc_le_D_lemma},
+apply (tactic {* ALLGOALS (EVERY' [strip_tac @{context},
+  TRY o eresolve_tac @{context} @{thms Suc_le_D_lemma},
   REPEAT o smp_tac @{context} 1, 
-  resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW TRY o atac]) *})
+  resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW TRY o assume_tac @{context}]) *})
 (* 3 subgoals *)
 apply (auto split del: split_if)
 done
--- a/src/HOL/Bali/TypeRel.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Bali/TypeRel.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -535,7 +535,7 @@
 apply      safe
 apply            (rule widen.int_obj)
 prefer          6 apply (drule implmt_is_class) apply simp
-apply (tactic "ALLGOALS (etac thin_rl)")
+apply (erule_tac [!] thin_rl)
 prefer         6 apply simp
 apply          (rule_tac [9] widen.arr_obj)
 apply         (rotate_tac [9] -1)
--- a/src/HOL/Bali/TypeSafe.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Bali/TypeSafe.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -3889,7 +3889,7 @@
 proof -
   note inj_term_simps [simp]
   from eval 
-  show "\<And> L accC T A. \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
+  have "\<And>L accC T A. \<lbrakk>\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
                        \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>t\<guillemotright>A\<rbrakk>  
         \<Longrightarrow> P L accC s0 t v s1" (is "PROP ?Hyp s0 t v s1")
   proof (induct)
--- a/src/HOL/Datatype_Examples/Stream_Processor.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Datatype_Examples/Stream_Processor.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -128,16 +128,27 @@
   "rel_sp\<^sub>\<mu> R1 R2 (map_sp\<^sub>\<mu> f1 f2 sp) (map_sp\<^sub>\<mu> g1 g2 sp') =
   rel_sp\<^sub>\<mu> (BNF_Def.vimage2p f1 g1 R1) (BNF_Def.vimage2p f2 g2 R2) sp sp'"
 by (tactic {*
-  let val ks = 1 upto 2;
+  let
+    val ks = 1 upto 2;
+    val ctxt = @{context};
   in
-    BNF_Tactics.unfold_thms_tac @{context}
+    BNF_Tactics.unfold_thms_tac ctxt
       @{thms sp\<^sub>\<mu>.rel_compp sp\<^sub>\<mu>.rel_conversep sp\<^sub>\<mu>.rel_Grp vimage2p_Grp} THEN
-    HEADGOAL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
-      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 @{context} @{thms relcomppE conversepE GrpE},
-      hyp_subst_tac @{context}, atac])
+    HEADGOAL (EVERY' [resolve_tac ctxt [iffI],
+      resolve_tac ctxt @{thms relcomppI},
+      resolve_tac ctxt @{thms GrpI},
+      resolve_tac ctxt [refl],
+      resolve_tac ctxt [CollectI],
+      BNF_Util.CONJ_WRAP' (K (resolve_tac ctxt @{thms subset_UNIV})) ks,
+      resolve_tac ctxt @{thms relcomppI}, assume_tac ctxt,
+      resolve_tac ctxt @{thms conversepI},
+      resolve_tac ctxt @{thms GrpI},
+      resolve_tac ctxt [refl],
+      resolve_tac ctxt [CollectI],
+      BNF_Util.CONJ_WRAP' (K (resolve_tac ctxt @{thms subset_UNIV})) ks,
+      REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE conversepE GrpE},
+      hyp_subst_tac ctxt,
+      assume_tac ctxt])
   end
 *})
 
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -4102,7 +4102,7 @@
   THEN' CSUBGOAL (fn (g, i) =>
     let
       val th = frpar_oracle (ctxt, alternative, T, ps, g);
-    in rtac (th RS iffD2) i end);
+    in resolve_tac ctxt [th RS iffD2] i end);
 
 fun method alternative =
   let
--- a/src/HOL/Decision_Procs/approximation.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Decision_Procs/approximation.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -12,7 +12,7 @@
 structure Approximation: APPROXIMATION =
 struct
 
-fun reorder_bounds_tac prems i =
+fun reorder_bounds_tac ctxt prems i =
   let
     fun variable_of_bound (Const (@{const_name Trueprop}, _) $
                            (Const (@{const_name Set.member}, _) $
@@ -35,8 +35,8 @@
                 |> Graph.strong_conn |> map the_single |> rev
                 |> map_filter (AList.lookup (op =) variable_bounds)
 
-    fun prepend_prem th tac
-      = tac THEN rtac (th RSN (2, @{thm mp})) i
+    fun prepend_prem th tac =
+      tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i
   in
     fold prepend_prem order all_tac
   end
@@ -49,9 +49,10 @@
   |> Thm.prop_of |> Logic.dest_equals |> snd;
 
 (* Should be in HOL.thy ? *)
-fun gen_eval_tac conv ctxt = CONVERSION
-  (Object_Logic.judgment_conv ctxt (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
-  THEN' rtac TrueI
+fun gen_eval_tac conv ctxt =
+  CONVERSION
+    (Object_Logic.judgment_conv ctxt (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))
+  THEN' resolve_tac ctxt [TrueI]
 
 val form_equations = @{thms interpret_form_equations};
 
@@ -78,10 +79,10 @@
                |> HOLogic.mk_list @{typ nat}
                |> Thm.cterm_of ctxt
      in
-       (rtac (Thm.instantiate ([], [((("n", 0), @{typ nat}), n),
+       (resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), @{typ nat}), n),
                                    ((("prec", 0), @{typ nat}), p),
                                    ((("ss", 0), @{typ "nat list"}), s)])
-            @{thm "approx_form"}) i
+            @{thm approx_form}] i
         THEN simp_tac (put_simpset (simpset_of @{context}) ctxt) i) st
      end
 
@@ -95,10 +96,10 @@
        val s = vs |> map lookup_splitting |> hd
             |> Thm.cterm_of ctxt
      in
-       rtac (Thm.instantiate ([], [((("s", 0), @{typ nat}), s),
+       resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), @{typ nat}), s),
                                    ((("t", 0), @{typ nat}), t),
                                    ((("prec", 0), @{typ nat}), p)])
-            @{thm "approx_tse_form"}) i st
+            @{thm approx_tse_form}] i st
      end
   end
 
@@ -190,8 +191,10 @@
   |> the |> Thm.prems_of |> hd
 
 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
+    REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE},
+      eresolve_tac ctxt @{thms meta_eqE},
+      resolve_tac ctxt @{thms impI}] 1)
+    THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems 1) ctxt 1
     THEN DETERM (TRY (filter_prems_tac ctxt (K false) 1)))
 
 fun reify_form ctxt term = apply_tactic ctxt term
@@ -249,10 +252,10 @@
       >> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t)));
 
 fun approximation_tac prec splitting taylor ctxt i =
-  REPEAT (FIRST' [etac @{thm intervalE},
-                  etac @{thm meta_eqE},
-                  rtac @{thm impI}] i)
-  THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) ctxt i
+  REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE},
+                  eresolve_tac ctxt @{thms meta_eqE},
+                  resolve_tac ctxt @{thms impI}] i)
+  THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt 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
--- a/src/HOL/Decision_Procs/cooper_tac.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Decision_Procs/cooper_tac.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -88,9 +88,9 @@
             val pth = linzqe_oracle (ctxt, Envir.eta_long [] t1)
           in
             ((pth RS iffD2) RS pre_thm,
-              assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
+              assm_tac (i + 1) THEN (if q then I else TRY) (resolve_tac ctxt [TrueI] i))
           end
       | _ => (pre_thm, assm_tac i))
-  in rtac (mp_step nh (spec_step np th)) i THEN tac end);
+  in resolve_tac ctxt [mp_step nh (spec_step np th)] i THEN tac end);
 
 end
--- a/src/HOL/Decision_Procs/ferrack_tac.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Decision_Procs/ferrack_tac.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -67,9 +67,9 @@
     let val pth = linr_oracle (ctxt, Envir.eta_long [] t1)
     in 
        ((pth RS iffD2) RS pre_thm,
-        assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
+        assm_tac (i + 1) THEN (if q then I else TRY) (resolve_tac ctxt [TrueI] i))
     end
       | _ => (pre_thm, assm_tac i)
-  in rtac ((mp_step nh o spec_step np) th) i THEN tac end);
+  in resolve_tac ctxt [(mp_step nh o spec_step np) th] i THEN tac end);
 
 end
--- a/src/HOL/Decision_Procs/mir_tac.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Decision_Procs/mir_tac.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -116,9 +116,9 @@
       val pth = mirfr_oracle (ctxt, Envir.eta_long [] t1)
     in 
        ((pth RS iffD2) RS pre_thm,
-        assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
+        assm_tac (i + 1) THEN (if q then I else TRY) (resolve_tac ctxt [TrueI] i))
     end
       | _ => (pre_thm, assm_tac i)
-  in rtac (((mp_step nh) o (spec_step np)) th) i THEN tac end);
+  in resolve_tac ctxt [((mp_step nh) o (spec_step np)) th] i THEN tac end);
 
 end
--- a/src/HOL/Filter.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Filter.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -250,7 +250,7 @@
           (Rule_Cases.internalize_params (raw_elim_thm RS Goal.init (Thm.cterm_of ctxt goal)))
       val cases = Rule_Cases.make_common ctxt cases_prop [(("elim", []), [])]
     in
-      CASES cases (rtac raw_elim_thm i)
+      CASES cases (resolve_tac ctxt [raw_elim_thm] i)
     end)
 *}
 
--- a/src/HOL/HOLCF/Domain.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Domain.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -142,12 +142,12 @@
 
 setup {*
   fold Sign.add_const_constraint
-  [ (@{const_name defl}, SOME @{typ "'a::domain itself \<Rightarrow> udom defl"})
-  , (@{const_name emb}, SOME @{typ "'a::domain \<rightarrow> udom"})
-  , (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::domain"})
-  , (@{const_name liftdefl}, SOME @{typ "'a::predomain itself \<Rightarrow> udom u defl"})
-  , (@{const_name liftemb}, SOME @{typ "'a::predomain u \<rightarrow> udom u"})
-  , (@{const_name liftprj}, SOME @{typ "udom u \<rightarrow> 'a::predomain u"}) ]
+   [(@{const_name defl}, SOME @{typ "'a::domain itself \<Rightarrow> udom defl"}),
+    (@{const_name emb}, SOME @{typ "'a::domain \<rightarrow> udom"}),
+    (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::domain"}),
+    (@{const_name liftdefl}, SOME @{typ "'a::predomain itself \<Rightarrow> udom u defl"}),
+    (@{const_name liftemb}, SOME @{typ "'a::predomain u \<rightarrow> udom u"}),
+    (@{const_name liftprj}, SOME @{typ "udom u \<rightarrow> 'a::predomain u"})]
 *}
 
 ML_file "Tools/domaindef.ML"
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoTraces.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/IOA/meta_theory/CompoTraces.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -477,7 +477,7 @@
 apply (rule_tac x = "schA" in spec)
 apply (rule_tac x = "schB" in spec)
 apply (rule_tac x = "tr" in spec)
-apply (tactic "thin_tac' 5 1")
+apply (tactic "thin_tac' @{context} 5 1")
 apply (rule nat_less_induct)
 apply (rule allI)+
 apply (rename_tac tr schB schA)
@@ -487,7 +487,7 @@
 apply (case_tac "Forall (%x. x:act B & x~:act A) tr")
 
 apply (rule seq_take_lemma [THEN iffD2, THEN spec])
-apply (tactic "thin_tac' 5 1")
+apply (tactic "thin_tac' @{context} 5 1")
 
 
 apply (case_tac "Finite tr")
@@ -697,7 +697,7 @@
 apply (rule_tac x = "schA" in spec)
 apply (rule_tac x = "schB" in spec)
 apply (rule_tac x = "tr" in spec)
-apply (tactic "thin_tac' 5 1")
+apply (tactic "thin_tac' @{context} 5 1")
 apply (rule nat_less_induct)
 apply (rule allI)+
 apply (rename_tac tr schB schA)
@@ -707,7 +707,7 @@
 apply (case_tac "Forall (%x. x:act A & x~:act B) tr")
 
 apply (rule seq_take_lemma [THEN iffD2, THEN spec])
-apply (tactic "thin_tac' 5 1")
+apply (tactic "thin_tac' @{context} 5 1")
 
 apply (case_tac "Finite tr")
 
--- a/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -1098,7 +1098,7 @@
   THEN simp_tac (ctxt addsimps rws) i;
 
 fun Seq_Finite_induct_tac ctxt i =
-  etac @{thm Seq_Finite_ind} i
+  eresolve_tac ctxt @{thms Seq_Finite_ind} i
   THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt i)));
 
 fun pair_tac ctxt s =
--- a/src/HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -48,9 +48,9 @@
 
 
 ML {*
-fun thin_tac' j =
+fun thin_tac' ctxt j =
   rotate_tac (j - 1) THEN'
-  etac thin_rl THEN'
+  eresolve_tac ctxt [thin_rl] THEN'
   rotate_tac (~ (j - 1))
 *}
 
@@ -180,7 +180,7 @@
 apply (intro strip)
 apply (rule mp)
 prefer 2 apply (assumption)
-apply (tactic "thin_tac' 1 1")
+apply (tactic "thin_tac' @{context} 1 1")
 apply (rule_tac x = "s" in spec)
 apply (rule nat_less_induct)
 apply (intro strip)
--- a/src/HOL/HOLCF/Lift.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Lift.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -46,7 +46,7 @@
 
 method_setup defined = {*
   Scan.succeed (fn ctxt => SIMPLE_METHOD'
-    (etac @{thm lift_definedE} THEN' asm_simp_tac ctxt))
+    (eresolve_tac ctxt @{thms lift_definedE} THEN' asm_simp_tac ctxt))
 *}
 
 lemma DefE: "Def x = \<bottom> \<Longrightarrow> R"
--- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -142,7 +142,8 @@
         val cs = ContProc.cont_thms lam
         val betas = map (fn c => mk_meta_eq (c RS @{thm beta_cfun})) cs
       in
-        prove thy (def_thm::betas) goal (K [rtac reflexive_thm 1])
+        prove thy (def_thm::betas) goal
+          (fn {context = ctxt, ...} => [resolve_tac ctxt [reflexive_thm] 1])
       end
 end
 
@@ -228,9 +229,9 @@
       val goal = mk_trp (foldr1 mk_disj (mk_undef y :: map one_con spec'))
       (* first rules replace "y = bottom \/ P" with "rep$y = bottom \/ P" *)
       fun tacs ctxt = [
-          rtac (iso_locale RS @{thm iso.casedist_rule}) 1,
+          resolve_tac ctxt [iso_locale RS @{thm iso.casedist_rule}] 1,
           rewrite_goals_tac ctxt [mk_meta_eq (iso_locale RS @{thm iso.iso_swap})],
-          rtac thm3 1]
+          resolve_tac ctxt [thm3] 1]
     in
       val nchotomy = prove thy con_betas goal (tacs o #context)
       val exhaust =
@@ -245,8 +246,8 @@
         val rules = @{thms compact_sinl compact_sinr compact_spair
                            compact_up compact_ONE}
         fun tacs ctxt =
-          [rtac (iso_locale RS @{thm iso.compact_abs}) 1,
-           REPEAT (resolve_tac ctxt rules 1 ORELSE atac 1)]
+          [resolve_tac ctxt [iso_locale RS @{thm iso.compact_abs}] 1,
+           REPEAT (resolve_tac ctxt rules 1 ORELSE assume_tac ctxt 1)]
         fun con_compact (con, args) =
           let
             val vs = vars_of args
@@ -709,7 +710,10 @@
     local
       fun dis_strict dis =
         let val goal = mk_trp (mk_strict dis)
-        in prove thy dis_defs goal (K [rtac (hd case_rews) 1]) end
+        in
+          prove thy dis_defs goal
+            (fn {context = ctxt, ...} => [resolve_tac ctxt [hd case_rews] 1])
+        end
     in
       val dis_stricts = map dis_strict dis_consts
     end
@@ -739,9 +743,9 @@
           val x = Free ("x", lhsT)
           val simps = dis_apps @ @{thms dist_eq_tr}
           fun tacs ctxt =
-            [rtac @{thm iffI} 1,
+            [resolve_tac ctxt @{thms iffI} 1,
              asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps dis_stricts) 2,
-             rtac exhaust 1, atac 1,
+             resolve_tac ctxt [exhaust] 1, assume_tac ctxt 1,
              ALLGOALS (asm_full_simp_tac (put_simpset simple_ss ctxt addsimps simps))]
           val goal = mk_trp (mk_eq (mk_undef (dis ` x), mk_undef x))
         in prove thy [] goal (tacs o #context) end
--- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -161,11 +161,11 @@
               val rules = prems @ con_rews @ @{thms simp_thms}
               val simplify = asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules)
               fun arg_tac (lazy, _) =
-                  rtac (if lazy then allI else case_UU_allI) 1
+                  resolve_tac ctxt [if lazy then allI else case_UU_allI] 1
               fun tacs ctxt =
                   rewrite_goals_tac ctxt @{thms atomize_all atomize_imp} ::
                   map arg_tac args @
-                  [REPEAT (rtac impI 1), ALLGOALS simplify]
+                  [REPEAT (resolve_tac ctxt [impI] 1), ALLGOALS simplify]
             in
               Goal.prove ctxt [] [] subgoal (EVERY o tacs o #context)
             end
@@ -181,7 +181,7 @@
             TRY (safe_tac (put_claset HOL_cs ctxt))]
           fun con_tac _ = 
             asm_simp_tac (put_simpset take_ss ctxt) 1 THEN
-            (resolve_tac ctxt prems' THEN_ALL_NEW etac spec) 1
+            (resolve_tac ctxt prems' THEN_ALL_NEW eresolve_tac ctxt [spec]) 1
           fun cases_tacs (cons, exhaust) =
             Rule_Insts.res_inst_tac ctxt [((("y", 0), Position.none), "x")] [] exhaust 1 ::
             asm_simp_tac (put_simpset take_ss ctxt addsimps prems) 1 ::
@@ -201,9 +201,9 @@
       fun tacf {prems, context = ctxt} =
         let
           fun finite_tac (take_induct, fin_ind) =
-              rtac take_induct 1 THEN
+              resolve_tac ctxt [take_induct] 1 THEN
               (if is_finite then all_tac else resolve_tac ctxt prems 1) THEN
-              (rtac fin_ind THEN_ALL_NEW solve_tac ctxt prems) 1
+              (resolve_tac ctxt [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
@@ -326,10 +326,10 @@
           val prems' = Project_Rule.projections ctxt prem'
           val dests = map (fn th => th RS spec RS spec RS mp) prems'
           fun one_tac (dest, rews) =
-              dtac dest 1 THEN safe_tac (put_claset HOL_cs ctxt) THEN
+              dresolve_tac ctxt [dest] 1 THEN safe_tac (put_claset HOL_cs ctxt) THEN
               ALLGOALS (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps rews))
         in
-          rtac @{thm nat.induct} 1 THEN
+          resolve_tac ctxt @{thms nat.induct} 1 THEN
           simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1 THEN
           safe_tac (put_claset HOL_cs ctxt) THEN
           EVERY (map one_tac (dests ~~ take_rews))
@@ -350,7 +350,7 @@
         let
           val rule = hd prems RS coind_lemma
         in
-          rtac take_lemma 1 THEN
+          resolve_tac ctxt [take_lemma] 1 THEN
           asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (rule :: prems)) 1
         end
     in
--- a/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -310,12 +310,12 @@
         Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} =>
          EVERY
           [rewrite_goals_tac ctxt map_apply_thms,
-           rtac (map_cont_thm RS @{thm cont_fix_ind}) 1,
+           resolve_tac ctxt [map_cont_thm RS @{thm cont_fix_ind}] 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 ctxt (deflation_rules @ prems) 1 ORELSE atac 1)])
+           REPEAT (eresolve_tac ctxt @{thms conjE} 1),
+           REPEAT (resolve_tac ctxt (deflation_rules @ prems) 1 ORELSE assume_tac ctxt 1)])
       end
     fun conjuncts [] _ = []
       | conjuncts (n::[]) thm = [(n, thm)]
@@ -623,14 +623,13 @@
         Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} =>
          EVERY
           [rewrite_goals_tac ctxt (defl_apply_thms @ map_apply_thms),
-           rtac (@{thm cont_parallel_fix_ind}
-             OF [defl_cont_thm, map_cont_thm]) 1,
+           resolve_tac ctxt [@{thm cont_parallel_fix_ind} OF [defl_cont_thm, map_cont_thm]] 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 ctxt (isodefl_rules @ prems) 1 ORELSE atac 1)])
+           REPEAT (eresolve_tac ctxt @{thms conjE} 1),
+           REPEAT (resolve_tac ctxt (isodefl_rules @ prems) 1 ORELSE assume_tac ctxt 1)])
       end
     val isodefl_binds = map (Binding.prefix_name "isodefl_") dbinds
     fun conjuncts [] _ = []
@@ -655,9 +654,9 @@
         val lhs = list_ccomb (map_const, map mk_ID (filter is_cpo Ts))
         val goal = mk_eqs (lhs, mk_ID lhsT)
         fun tac ctxt = EVERY
-          [rtac @{thm isodefl_DEFL_imp_ID} 1,
+          [resolve_tac ctxt @{thms isodefl_DEFL_imp_ID} 1,
            stac ctxt DEFL_thm 1,
-           rtac isodefl_thm 1,
+           resolve_tac ctxt [isodefl_thm] 1,
            REPEAT (resolve_tac ctxt @{thms isodefl_ID_DEFL isodefl_LIFTDEFL} 1)]
       in
         Goal.prove_global thy [] [] goal (tac o #context)
@@ -700,8 +699,8 @@
             EVERY
             [simp_tac (put_simpset HOL_basic_ss ctxt addsimps start_rules) 1,
              simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms fix_def2}) 1,
-             rtac @{thm lub_eq} 1,
-             rtac @{thm nat.induct} 1,
+             resolve_tac ctxt @{thms lub_eq} 1,
+             resolve_tac ctxt @{thms nat.induct} 1,
              simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules0) 1,
              asm_full_simp_tac (put_simpset beta_ss ctxt addsimps rules1) 1]
       in
@@ -713,12 +712,13 @@
       let
         val n = Free ("n", natT)
         val goal = mk_eqs (mk_lub (lambda n (take_const $ n)), mk_ID lhsT)
-        val tac =
+        fun tac ctxt =
             EVERY
-            [rtac @{thm trans} 1, rtac map_ID_thm 2,
+            [resolve_tac ctxt @{thms trans} 1,
+             resolve_tac ctxt [map_ID_thm] 2,
              cut_tac lub_take_lemma 1,
-             REPEAT (etac @{thm Pair_inject} 1), atac 1]
-        val lub_take_thm = Goal.prove_global thy [] [] goal (K tac)
+             REPEAT (eresolve_tac ctxt @{thms Pair_inject} 1), assume_tac ctxt 1]
+        val lub_take_thm = Goal.prove_global thy [] [] goal (tac o #context)
       in
         add_qualified_thm "lub_take" (dbind, lub_take_thm) thy
       end
--- a/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -320,12 +320,12 @@
       in
         Goal.prove_global thy [] [] goal (fn {context = ctxt, ...} =>
          EVERY
-          [rtac @{thm nat.induct} 1,
+          [resolve_tac ctxt @{thms nat.induct} 1,
            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
+           REPEAT (eresolve_tac ctxt @{thms conjE} 1
                    ORELSE resolve_tac ctxt deflation_rules 1
-                   ORELSE atac 1)])
+                   ORELSE assume_tac ctxt 1)])
       end
     fun conjuncts [] _ = []
       | conjuncts (n::[]) thm = [(n, thm)]
@@ -444,7 +444,7 @@
             take_Suc_thms @ decisive_abs_rep_thms
             @ @{thms decisive_ID decisive_ssum_map decisive_sprod_map}
         fun tac ctxt = EVERY [
-            rtac @{thm nat.induct} 1,
+            resolve_tac ctxt @{thms nat.induct} 1,
             simp_tac (put_simpset HOL_ss ctxt addsimps rules0) 1,
             asm_simp_tac (put_simpset HOL_ss ctxt addsimps rules1) 1]
       in Goal.prove_global thy [] [] goal (tac o #context) end
@@ -461,7 +461,7 @@
         fun tac ctxt =
             EVERY [
             rewrite_goals_tac ctxt finite_defs,
-            rtac @{thm lub_ID_finite} 1,
+            resolve_tac ctxt @{thms lub_ID_finite} 1,
             resolve_tac ctxt chain_take_thms 1,
             resolve_tac ctxt lub_take_thms 1,
             resolve_tac ctxt decisive_thms 1]
--- a/src/HOL/HOLCF/Tools/cont_proc.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/cont_proc.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -102,7 +102,7 @@
     fun new_cont_tac f' i =
       case all_cont_thms f' of
         [] => no_tac
-      | (c::_) => rtac c i
+      | (c::_) => resolve_tac ctxt [c] i
 
     fun cont_tac_of_term (Const (@{const_name cont}, _) $ f) =
       let
--- a/src/HOL/HOLCF/Tools/cpodef.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/cpodef.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -73,8 +73,8 @@
     val cpo_thms = map (Thm.transfer thy) [type_definition, below_def, admissible]
     val (full_tname, Ts) = dest_Type newT
     val lhs_sorts = map (snd o dest_TFree) Ts
-    val tac = rtac (@{thm typedef_cpo} OF cpo_thms) 1
-    val thy = Axclass.prove_arity (full_tname, lhs_sorts, @{sort cpo}) (K tac) thy
+    fun tac ctxt = resolve_tac ctxt [@{thm typedef_cpo} OF cpo_thms] 1
+    val thy = Axclass.prove_arity (full_tname, lhs_sorts, @{sort cpo}) tac thy
     (* transfer thms so that they will know about the new cpo instance *)
     val cpo_thms' = map (Thm.transfer thy) cpo_thms
     fun make thm = Drule.zero_var_indexes (thm OF cpo_thms')
@@ -112,8 +112,8 @@
     val pcpo_thms = map (Thm.transfer thy) [type_definition, below_def, bottom_mem]
     val (full_tname, Ts) = dest_Type newT
     val lhs_sorts = map (snd o dest_TFree) Ts
-    val tac = rtac (@{thm typedef_pcpo} OF pcpo_thms) 1
-    val thy = Axclass.prove_arity (full_tname, lhs_sorts, @{sort pcpo}) (K tac) thy
+    fun tac ctxt = resolve_tac ctxt [@{thm typedef_pcpo} OF pcpo_thms] 1
+    val thy = Axclass.prove_arity (full_tname, lhs_sorts, @{sort pcpo}) tac thy
     val pcpo_thms' = map (Thm.transfer thy) pcpo_thms
     fun make thm = Drule.zero_var_indexes (thm OF pcpo_thms')
     val Rep_strict = make @{thm typedef_Rep_strict}
@@ -182,7 +182,7 @@
     val below_def = singleton (Proof_Context.export lthy ctxt_thy) below_ldef
     val thy = lthy
       |> Class.prove_instantiation_exit
-          (K (rtac (@{thm typedef_po} OF [type_definition, below_def]) 1))
+          (fn ctxt => resolve_tac ctxt [@{thm typedef_po} OF [type_definition, below_def]] 1)
   in ((info, below_def), thy) end
 
 fun prepare_cpodef
@@ -205,7 +205,7 @@
     fun cpodef_result nonempty admissible thy =
       let
         val ((info as (_, {type_definition, ...}), below_def), thy) = thy
-          |> add_podef typ set opt_morphs (fn _ => rtac nonempty 1)
+          |> add_podef typ set opt_morphs (fn ctxt => resolve_tac ctxt [nonempty] 1)
         val (cpo_info, thy) = thy
           |> prove_cpo name newT morphs type_definition below_def admissible
       in
@@ -237,7 +237,7 @@
 
     fun pcpodef_result bottom_mem admissible thy =
       let
-        fun tac _ = rtac exI 1 THEN rtac bottom_mem 1
+        fun tac ctxt = resolve_tac ctxt [exI] 1 THEN resolve_tac ctxt [bottom_mem] 1
         val ((info as (_, {type_definition, ...}), below_def), thy) = thy
           |> add_podef typ set opt_morphs tac
         val (cpo_info, thy) = thy
--- a/src/HOL/HOLCF/Tools/domaindef.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/domaindef.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -108,8 +108,8 @@
     val set = @{term "defl_set :: udom defl => udom set"} $ defl
 
     (*pcpodef*)
-    fun tac1 _ = rtac @{thm defl_set_bottom} 1
-    fun tac2 _ = rtac @{thm adm_defl_set} 1
+    fun tac1 ctxt = resolve_tac ctxt @{thms defl_set_bottom} 1
+    fun tac2 ctxt = resolve_tac ctxt @{thms adm_defl_set} 1
     val ((info, cpo_info, pcpo_info), thy) = thy
       |> Cpodef.add_pcpodef typ set (SOME morphs) (tac1, tac2)
 
@@ -165,7 +165,7 @@
       liftemb_def, liftprj_def, liftdefl_def]
     val thy = lthy
       |> Class.prove_instantiation_instance
-          (K (rtac (@{thm typedef_domain_class} OF typedef_thms) 1))
+          (fn ctxt => resolve_tac ctxt [@{thm typedef_domain_class} OF typedef_thms] 1)
       |> Local_Theory.exit_global
 
     (*other theorems*)
--- a/src/HOL/HOLCF/Tools/fixrec.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/HOLCF/Tools/fixrec.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -301,7 +301,7 @@
         val unfold_thm = the (Symtab.lookup tab c)
         val rule = unfold_thm RS @{thm ssubst_lhs}
       in
-        CHANGED (rtac rule i THEN eta_tac i THEN asm_simp_tac ctxt i)
+        CHANGED (resolve_tac ctxt [rule] i THEN eta_tac i THEN asm_simp_tac ctxt i)
       end
   in
     SUBGOAL (fn ti => the_default no_tac (try tac ti))
@@ -311,7 +311,7 @@
 fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) =
   let
     val rule = unfold_thm RS @{thm ssubst_lhs}
-    val tac = rtac rule 1 THEN eta_tac 1 THEN asm_simp_tac ctxt 1
+    val tac = resolve_tac ctxt [rule] 1 THEN eta_tac 1 THEN asm_simp_tac ctxt 1
     fun prove_term t = Goal.prove ctxt [] [] t (K tac)
     fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t)
   in
--- a/src/HOL/Hoare/hoare_tac.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Hoare/hoare_tac.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -132,9 +132,9 @@
 fun set_to_pred_tac ctxt var_names = SUBGOAL (fn (_, i) =>
   before_set2pred_simp_tac ctxt i THEN_MAYBE
   EVERY [
-    rtac subsetI i,
-    rtac CollectI i,
-    dtac CollectD i,
+    resolve_tac ctxt [subsetI] i,
+    resolve_tac ctxt [CollectI] i,
+    dresolve_tac ctxt [CollectD] i,
     TRY (split_all_tac ctxt i) THEN_MAYBE
      (rename_tac var_names i THEN
       full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i)]);
@@ -148,7 +148,8 @@
 (*******************************************************************************)
 
 fun max_simp_tac ctxt var_names tac =
-  FIRST' [rtac subset_refl, set_to_pred_tac ctxt var_names THEN_MAYBE' tac];
+  FIRST' [resolve_tac ctxt [subset_refl],
+    set_to_pred_tac ctxt var_names THEN_MAYBE' tac];
 
 fun basic_simp_tac ctxt var_names tac =
   simp_tac
@@ -163,26 +164,28 @@
   let
     val var_names = map (fst o dest_Free) vars;
     fun wlp_tac i =
-      rtac @{thm SeqRule} i THEN rule_tac false (i + 1)
+      resolve_tac ctxt @{thms SeqRule} i THEN rule_tac false (i + 1)
     and rule_tac pre_cond i st = st |> (*abstraction over st prevents looping*)
       ((wlp_tac i THEN rule_tac pre_cond i)
         ORELSE
         (FIRST [
-          rtac @{thm SkipRule} i,
-          rtac @{thm AbortRule} i,
+          resolve_tac ctxt @{thms SkipRule} i,
+          resolve_tac ctxt @{thms AbortRule} i,
           EVERY [
-            rtac @{thm BasicRule} i,
-            rtac Mlem i,
+            resolve_tac ctxt @{thms BasicRule} i,
+            resolve_tac ctxt [Mlem] i,
             split_simp_tac ctxt i],
           EVERY [
-            rtac @{thm CondRule} i,
+            resolve_tac ctxt @{thms CondRule} i,
             rule_tac false (i + 2),
             rule_tac false (i + 1)],
           EVERY [
-            rtac @{thm WhileRule} i,
+            resolve_tac ctxt @{thms WhileRule} i,
             basic_simp_tac ctxt var_names tac (i + 2),
             rule_tac true (i + 1)]]
-         THEN (if pre_cond then basic_simp_tac ctxt var_names tac i else rtac subset_refl i)));
+         THEN (
+           if pre_cond then basic_simp_tac ctxt var_names tac i
+           else resolve_tac ctxt [subset_refl] i)));
   in rule_tac end;
 
 
--- a/src/HOL/Hoare_Parallel/Gar_Coll.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -775,11 +775,21 @@
 --\<open>20 subgoals left\<close>
 apply(tactic\<open>TRYALL (clarify_tac @{context})\<close>)
 apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
-apply(tactic \<open>TRYALL (etac disjE)\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
 apply simp_all
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac @{thm Graph3},force_tac @{context}, assume_tac @{context}])\<close>)
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{context}])\<close>)
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
+    resolve_tac @{context} [subset_trans],
+    eresolve_tac @{context} @{thms Graph3},
+    force_tac @{context},
+    assume_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
+    eresolve_tac @{context} [subset_trans],
+    resolve_tac @{context} @{thms Graph9},
+    force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI1],
+    eresolve_tac @{context} @{thms psubset_subset_trans},
+    resolve_tac @{context} @{thms Graph9},
+    force_tac @{context}])\<close>)
 done
 
 subsubsection \<open>Interference freedom Mutator-Collector\<close>
--- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -1178,40 +1178,62 @@
 --\<open>14 subgoals left\<close>
 apply(tactic \<open>TRYALL (clarify_tac @{context})\<close>)
 apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
-apply(tactic \<open>TRYALL (rtac conjI)\<close>)
-apply(tactic \<open>TRYALL (rtac impI)\<close>)
-apply(tactic \<open>TRYALL (etac disjE)\<close>)
-apply(tactic \<open>TRYALL (etac conjE)\<close>)
-apply(tactic \<open>TRYALL (etac disjE)\<close>)
-apply(tactic \<open>TRYALL (etac disjE)\<close>)
+apply(tactic \<open>TRYALL (resolve_tac @{context} [conjI])\<close>)
+apply(tactic \<open>TRYALL (resolve_tac @{context} [impI])\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [conjE])\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
 --\<open>72 subgoals left\<close>
 apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
 --\<open>35 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI1,rtac subset_trans,etac @{thm Graph3},force_tac @{context}, assume_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI1],
+    resolve_tac @{context} [subset_trans],
+    eresolve_tac @{context} @{thms Graph3},
+    force_tac @{context},
+    assume_tac @{context}])\<close>)
 --\<open>28 subgoals left\<close>
-apply(tactic \<open>TRYALL (etac conjE)\<close>)
-apply(tactic \<open>TRYALL (etac disjE)\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [conjE])\<close>)
+apply(tactic \<open>TRYALL (eresolve_tac @{context} [disjE])\<close>)
 --\<open>34 subgoals left\<close>
 apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
 apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update)
 apply(case_tac [!] "M x!(T (Muts x ! j))=Black")
 apply(simp_all add:Graph10)
 --\<open>47 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[REPEAT o (rtac disjI2),etac @{thm subset_psubset_trans}, etac @{thm Graph11},force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[REPEAT o resolve_tac @{context} [disjI2],
+    eresolve_tac @{context} @{thms subset_psubset_trans},
+    eresolve_tac @{context} @{thms Graph11},
+    force_tac @{context}])\<close>)
 --\<open>41 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans},
-    force_tac (@{context} addsimps
-      [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
+    resolve_tac @{context} [disjI1],
+    eresolve_tac @{context} @{thms le_trans},
+    force_tac (@{context} addsimps @{thms Queue_def less_Suc_eq_le le_length_filter_update})])\<close>)
 --\<open>35 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
+    resolve_tac @{context} [disjI1],
+    eresolve_tac @{context} @{thms psubset_subset_trans},
+    resolve_tac @{context} @{thms Graph9},
+    force_tac @{context}])\<close>)
 --\<open>31 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac @{thm subset_psubset_trans},etac @{thm Graph11},force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
+    resolve_tac @{context} [disjI1],
+    eresolve_tac @{context} @{thms subset_psubset_trans},
+    eresolve_tac @{context} @{thms Graph11},
+    force_tac @{context}])\<close>)
 --\<open>29 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[REPEAT o (rtac disjI2),etac @{thm subset_psubset_trans},etac @{thm subset_psubset_trans},etac @{thm Graph11},force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[REPEAT o resolve_tac @{context} [disjI2],
+    eresolve_tac @{context} @{thms subset_psubset_trans},
+    eresolve_tac @{context} @{thms subset_psubset_trans},
+    eresolve_tac @{context} @{thms Graph11},
+    force_tac @{context}])\<close>)
 --\<open>25 subgoals left\<close>
-apply(tactic \<open>TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans},
-    force_tac (@{context} addsimps
-      [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[resolve_tac @{context} [disjI2],
+    resolve_tac @{context} [disjI2],
+    resolve_tac @{context} [disjI1],
+    eresolve_tac @{context} @{thms le_trans},
+    force_tac (@{context} addsimps @{thms Queue_def less_Suc_eq_le le_length_filter_update})])\<close>)
 --\<open>10 subgoals left\<close>
 apply(rule disjI2,rule disjI2,rule conjI,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update, rule disjI1, rule less_imp_le, erule less_le_trans, force simp add:Queue_def less_Suc_eq_le le_length_filter_update)+
 done
--- a/src/HOL/Hoare_Parallel/OG_Examples.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Examples.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -434,7 +434,7 @@
 --\<open>112 subgoals left\<close>
 apply(simp_all (no_asm))
 --\<open>43 subgoals left\<close>
-apply(tactic \<open>ALLGOALS (conjI_Tac (K all_tac))\<close>)
+apply(tactic \<open>ALLGOALS (conjI_Tac @{context} (K all_tac))\<close>)
 --\<open>419 subgoals left\<close>
 apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
 --\<open>99 subgoals left\<close>
--- a/src/HOL/Hoare_Parallel/OG_Tactics.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -291,9 +291,9 @@
 @{text n} subgoals, one for each conjunct:\<close>
 
 ML \<open>
-fun conjI_Tac tac i st = st |>
-       ( (EVERY [rtac conjI i,
-          conjI_Tac tac (i+1),
+fun conjI_Tac ctxt tac i st = st |>
+       ( (EVERY [resolve_tac ctxt [conjI] i,
+          conjI_Tac ctxt tac (i+1),
           tac i]) ORELSE (tac i) )
 \<close>
 
@@ -326,119 +326,123 @@
 
 ML \<open>
 
-fun WlpTac ctxt i = (rtac (@{thm SeqRule}) i) THEN (HoareRuleTac ctxt false (i+1))
+fun WlpTac ctxt i = resolve_tac ctxt @{thms SeqRule} i THEN HoareRuleTac ctxt false (i + 1)
 and HoareRuleTac ctxt precond i st = st |>
     ( (WlpTac ctxt i THEN HoareRuleTac ctxt precond i)
       ORELSE
-      (FIRST[rtac (@{thm SkipRule}) i,
-             rtac (@{thm BasicRule}) i,
-             EVERY[rtac (@{thm ParallelConseqRule}) i,
+      (FIRST[resolve_tac ctxt @{thms SkipRule} i,
+             resolve_tac ctxt @{thms BasicRule} i,
+             EVERY[resolve_tac ctxt @{thms ParallelConseqRule} i,
                    ParallelConseq ctxt (i+2),
                    ParallelTac ctxt (i+1),
                    ParallelConseq ctxt i],
-             EVERY[rtac (@{thm CondRule}) i,
+             EVERY[resolve_tac ctxt @{thms CondRule} i,
                    HoareRuleTac ctxt false (i+2),
                    HoareRuleTac ctxt false (i+1)],
-             EVERY[rtac (@{thm WhileRule}) i,
+             EVERY[resolve_tac ctxt @{thms WhileRule} i,
                    HoareRuleTac ctxt true (i+1)],
              K all_tac i ]
-       THEN (if precond then (K all_tac i) else (rtac (@{thm subset_refl}) i))))
+       THEN (if precond then (K all_tac i) else resolve_tac ctxt @{thms subset_refl} i)))
 
-and AnnWlpTac ctxt i = (rtac (@{thm AnnSeq}) i) THEN (AnnHoareRuleTac ctxt (i+1))
+and AnnWlpTac ctxt i = resolve_tac ctxt @{thms AnnSeq} i THEN AnnHoareRuleTac ctxt (i + 1)
 and AnnHoareRuleTac ctxt i st = st |>
     ( (AnnWlpTac ctxt i THEN AnnHoareRuleTac ctxt i )
      ORELSE
-      (FIRST[(rtac (@{thm AnnskipRule}) i),
-             EVERY[rtac (@{thm AnnatomRule}) i,
+      (FIRST[(resolve_tac ctxt @{thms AnnskipRule} i),
+             EVERY[resolve_tac ctxt @{thms AnnatomRule} i,
                    HoareRuleTac ctxt true (i+1)],
-             (rtac (@{thm AnnwaitRule}) i),
-             rtac (@{thm AnnBasic}) i,
-             EVERY[rtac (@{thm AnnCond1}) i,
+             (resolve_tac ctxt @{thms AnnwaitRule} i),
+             resolve_tac ctxt @{thms AnnBasic} i,
+             EVERY[resolve_tac ctxt @{thms AnnCond1} i,
                    AnnHoareRuleTac ctxt (i+3),
                    AnnHoareRuleTac ctxt (i+1)],
-             EVERY[rtac (@{thm AnnCond2}) i,
+             EVERY[resolve_tac ctxt @{thms AnnCond2} i,
                    AnnHoareRuleTac ctxt (i+1)],
-             EVERY[rtac (@{thm AnnWhile}) i,
+             EVERY[resolve_tac ctxt @{thms AnnWhile} i,
                    AnnHoareRuleTac ctxt (i+2)],
-             EVERY[rtac (@{thm AnnAwait}) i,
+             EVERY[resolve_tac ctxt @{thms AnnAwait} i,
                    HoareRuleTac ctxt true (i+1)],
              K all_tac i]))
 
-and ParallelTac ctxt i = EVERY[rtac (@{thm ParallelRule}) i,
+and ParallelTac ctxt i = EVERY[resolve_tac ctxt @{thms ParallelRule} i,
                           interfree_Tac ctxt (i+1),
                            MapAnn_Tac ctxt i]
 
 and MapAnn_Tac ctxt i st = st |>
-    (FIRST[rtac (@{thm MapAnnEmpty}) i,
-           EVERY[rtac (@{thm MapAnnList}) i,
+    (FIRST[resolve_tac ctxt @{thms MapAnnEmpty} i,
+           EVERY[resolve_tac ctxt @{thms MapAnnList} i,
                  MapAnn_Tac ctxt (i+1),
                  AnnHoareRuleTac ctxt i],
-           EVERY[rtac (@{thm MapAnnMap}) i,
-                 rtac (@{thm allI}) i, rtac (@{thm impI}) i,
+           EVERY[resolve_tac ctxt @{thms MapAnnMap} i,
+                 resolve_tac ctxt @{thms allI} i,
+                 resolve_tac ctxt @{thms impI} i,
                  AnnHoareRuleTac ctxt i]])
 
 and interfree_swap_Tac ctxt i st = st |>
-    (FIRST[rtac (@{thm interfree_swap_Empty}) i,
-           EVERY[rtac (@{thm interfree_swap_List}) i,
+    (FIRST[resolve_tac ctxt @{thms interfree_swap_Empty} i,
+           EVERY[resolve_tac ctxt @{thms interfree_swap_List} i,
                  interfree_swap_Tac ctxt (i+2),
                  interfree_aux_Tac ctxt (i+1),
                  interfree_aux_Tac ctxt i ],
-           EVERY[rtac (@{thm interfree_swap_Map}) i,
-                 rtac (@{thm allI}) i,rtac (@{thm impI}) i,
-                 conjI_Tac (interfree_aux_Tac ctxt) i]])
+           EVERY[resolve_tac ctxt @{thms interfree_swap_Map} i,
+                 resolve_tac ctxt @{thms allI} i,
+                 resolve_tac ctxt @{thms impI} i,
+                 conjI_Tac ctxt (interfree_aux_Tac ctxt) i]])
 
 and interfree_Tac ctxt i st = st |>
-   (FIRST[rtac (@{thm interfree_Empty}) i,
-          EVERY[rtac (@{thm interfree_List}) i,
+   (FIRST[resolve_tac ctxt @{thms interfree_Empty} i,
+          EVERY[resolve_tac ctxt @{thms interfree_List} i,
                 interfree_Tac ctxt (i+1),
                 interfree_swap_Tac ctxt i],
-          EVERY[rtac (@{thm interfree_Map}) i,
-                rtac (@{thm allI}) i,rtac (@{thm allI}) i,rtac (@{thm impI}) i,
+          EVERY[resolve_tac ctxt @{thms interfree_Map} i,
+                resolve_tac ctxt @{thms allI} i,
+                resolve_tac ctxt @{thms allI} i,
+                resolve_tac ctxt @{thms impI} i,
                 interfree_aux_Tac ctxt i ]])
 
 and interfree_aux_Tac ctxt i = (before_interfree_simp_tac ctxt i ) THEN
-        (FIRST[rtac (@{thm interfree_aux_rule1}) i,
+        (FIRST[resolve_tac ctxt @{thms interfree_aux_rule1} i,
                dest_assertions_Tac ctxt i])
 
 and dest_assertions_Tac ctxt i st = st |>
-    (FIRST[EVERY[rtac (@{thm AnnBasic_assertions}) i,
+    (FIRST[EVERY[resolve_tac ctxt @{thms AnnBasic_assertions} i,
                  dest_atomics_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnSeq_assertions}) i,
+           EVERY[resolve_tac ctxt @{thms AnnSeq_assertions} i,
                  dest_assertions_Tac ctxt (i+1),
                  dest_assertions_Tac ctxt i],
-           EVERY[rtac (@{thm AnnCond1_assertions}) i,
+           EVERY[resolve_tac ctxt @{thms AnnCond1_assertions} i,
                  dest_assertions_Tac ctxt (i+2),
                  dest_assertions_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnCond2_assertions}) i,
+           EVERY[resolve_tac ctxt @{thms AnnCond2_assertions} i,
                  dest_assertions_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnWhile_assertions}) i,
+           EVERY[resolve_tac ctxt @{thms AnnWhile_assertions} i,
                  dest_assertions_Tac ctxt (i+2),
                  dest_atomics_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnAwait_assertions}) i,
+           EVERY[resolve_tac ctxt @{thms AnnAwait_assertions} i,
                  dest_atomics_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
            dest_atomics_Tac ctxt i])
 
 and dest_atomics_Tac ctxt i st = st |>
-    (FIRST[EVERY[rtac (@{thm AnnBasic_atomics}) i,
+    (FIRST[EVERY[resolve_tac ctxt @{thms AnnBasic_atomics} i,
                  HoareRuleTac ctxt true i],
-           EVERY[rtac (@{thm AnnSeq_atomics}) i,
+           EVERY[resolve_tac ctxt @{thms AnnSeq_atomics} i,
                  dest_atomics_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnCond1_atomics}) i,
+           EVERY[resolve_tac ctxt @{thms AnnCond1_atomics} i,
                  dest_atomics_Tac ctxt (i+1),
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnCond2_atomics}) i,
+           EVERY[resolve_tac ctxt @{thms AnnCond2_atomics} i,
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm AnnWhile_atomics}) i,
+           EVERY[resolve_tac ctxt @{thms AnnWhile_atomics} i,
                  dest_atomics_Tac ctxt i],
-           EVERY[rtac (@{thm Annatom_atomics}) i,
+           EVERY[resolve_tac ctxt @{thms Annatom_atomics} i,
                  HoareRuleTac ctxt true i],
-           EVERY[rtac (@{thm AnnAwait_atomics}) i,
+           EVERY[resolve_tac ctxt @{thms AnnAwait_atomics} i,
                  HoareRuleTac ctxt true i],
                  K all_tac i])
 \<close>
@@ -482,18 +486,18 @@
 text \<open>Tactics useful for dealing with the generated verification conditions:\<close>
 
 method_setup conjI_tac = \<open>
-  Scan.succeed (K (SIMPLE_METHOD' (conjI_Tac (K all_tac))))\<close>
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (conjI_Tac ctxt (K all_tac)))\<close>
   "verification condition generator for interference freedom tests"
 
 ML \<open>
-fun disjE_Tac tac i st = st |>
-       ( (EVERY [etac disjE i,
-          disjE_Tac tac (i+1),
+fun disjE_Tac ctxt tac i st = st |>
+       ( (EVERY [eresolve_tac ctxt [disjE] i,
+          disjE_Tac ctxt tac (i+1),
           tac i]) ORELSE (tac i) )
 \<close>
 
 method_setup disjE_tac = \<open>
-  Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac))))\<close>
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (disjE_Tac ctxt (K all_tac)))\<close>
   "verification condition generator for interference freedom tests"
 
 end
--- a/src/HOL/IMPP/Com.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/IMPP/Com.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -83,7 +83,10 @@
   "WT_bodies = (!(pn,b):set bodies. WT b)"
 
 
-ML {* val make_imp_tac = EVERY'[rtac mp, fn i => atac (i+1), etac thin_rl] *}
+ML {*
+  fun make_imp_tac ctxt =
+    EVERY' [resolve_tac ctxt [mp], fn i => assume_tac ctxt (i + 1), eresolve_tac ctxt [thin_rl]]
+*}
 
 lemma finite_dom_body: "finite (dom body)"
 apply (unfold body_def)
--- a/src/HOL/IMPP/Hoare.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/IMPP/Hoare.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -278,7 +278,7 @@
 
 lemma hoare_sound: "G||-ts ==> G||=ts"
 apply (erule hoare_derivs.induct)
-apply              (tactic {* TRYALL (eresolve_tac @{context} [@{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 assume_tac @{context}) *})
 apply            (unfold hoare_valids_def)
 apply            blast
 apply           blast
@@ -351,7 +351,8 @@
   rule_tac [3] P1 = "%Z' s. s= (Z[Loc loc::=fun Z])" in hoare_derivs.Local [THEN conseq1],
   erule_tac [3] conseq12)
 apply         (erule_tac [5] hoare_derivs.Comp, erule_tac [5] conseq12)
-apply         (tactic {* (rtac @{thm hoare_derivs.If} THEN_ALL_NEW etac @{thm conseq12}) 6 *})
+apply         (tactic {* (resolve_tac @{context} @{thms hoare_derivs.If} THEN_ALL_NEW
+                eresolve_tac @{context} @{thms conseq12}) 6 *})
 apply          (rule_tac [8] hoare_derivs.Loop [THEN conseq2], erule_tac [8] conseq12)
 apply           auto
 done
@@ -435,7 +436,7 @@
 apply (frule finite_subset)
 apply (rule finite_dom_body [THEN finite_imageI])
 apply (rotate_tac 2)
-apply (tactic "make_imp_tac 1")
+apply (tactic "make_imp_tac @{context} 1")
 apply (erule finite_induct)
 apply  (clarsimp intro!: hoare_derivs.empty)
 apply (clarsimp intro!: hoare_derivs.insert simp del: range_composition)
--- a/src/HOL/IMPP/Natural.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/IMPP/Natural.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -111,7 +111,9 @@
 
 lemma evaln_evalc: "<c,s> -n-> t ==> <c,s> -c-> t"
 apply (erule evaln.induct)
-apply (tactic {* ALLGOALS (resolve_tac @{context} @{thms evalc.intros} THEN_ALL_NEW atac) *})
+apply (tactic {*
+  ALLGOALS (resolve_tac @{context} @{thms evalc.intros} THEN_ALL_NEW assume_tac @{context})
+*})
 done
 
 lemma Suc_le_D_lemma: "[| Suc n <= m'; (!!m. n <= m ==> P (Suc m)) |] ==> P m'"
@@ -137,11 +139,12 @@
 
 lemma evalc_evaln: "<c,s> -c-> t ==> ? n. <c,s> -n-> t"
 apply (erule evalc.induct)
-apply (tactic {* ALLGOALS (REPEAT o etac exE) *})
-apply (tactic {* TRYALL (EVERY' [dtac @{thm evaln_max2}, assume_tac @{context},
+apply (tactic {* ALLGOALS (REPEAT o eresolve_tac @{context} [exE]) *})
+apply (tactic {* TRYALL (EVERY' [dresolve_tac @{context} @{thms evaln_max2}, assume_tac @{context},
   REPEAT o eresolve_tac @{context} [exE, conjE]]) *})
 apply (tactic
-  {* ALLGOALS (rtac exI THEN' resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW atac) *})
+  {* ALLGOALS (resolve_tac @{context} [exI] THEN'
+    resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW assume_tac @{context}) *})
 done
 
 lemma eval_eq: "<c,s> -c-> t = (? n. <c,s> -n-> t)"
--- a/src/HOL/IOA/Solve.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/IOA/Solve.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -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 @{context} [conjI, impI] 1 ORELSE etac conjE 1) THEN
+    REPEAT((resolve_tac @{context} [conjI, impI] 1 ORELSE eresolve_tac @{context} [conjE] 1) THEN
       asm_full_simp_tac(@{context} addsimps [@{thm comp1_reachable}, @{thm comp2_reachable}]) 1) *})
   done
 
--- a/src/HOL/Import/import_rule.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Import/import_rule.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -220,7 +220,7 @@
     val tnames = sort_strings (map fst tfrees)
     val ((_, typedef_info), thy') =
      Typedef.add_typedef_global false (Binding.name tycname, map (rpair dummyS) tnames, NoSyn) c
-       (SOME (Binding.name rep_name, Binding.name abs_name)) (fn _ => rtac th2 1) thy
+       (SOME (Binding.name rep_name, Binding.name abs_name)) (fn ctxt => resolve_tac ctxt [th2] 1) thy
     val aty = #abs_type (#1 typedef_info)
     val th = freezeT thy' (#type_definition (#2 typedef_info))
     val (th_s, _) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th))
--- a/src/HOL/Library/Countable.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/Countable.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -186,10 +186,10 @@
         val rules = @{thms finite_item.intros}
       in
         SOLVED' (fn i => EVERY
-          [rtac @{thm countable_datatype} i,
-           rtac typedef_thm i,
-           etac induct_thm' i,
-           REPEAT (resolve_tac ctxt rules i ORELSE atac i)]) 1
+          [resolve_tac ctxt @{thms countable_datatype} i,
+           resolve_tac ctxt [typedef_thm] i,
+           eresolve_tac ctxt [induct_thm'] i,
+           REPEAT (resolve_tac ctxt rules i ORELSE assume_tac ctxt i)]) 1
       end)
 \<close>
 
--- a/src/HOL/Library/Multiset.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/Multiset.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -1898,23 +1898,25 @@
             Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $
                   mk_mset T [x] $ mk_mset T xs
 
-    fun mset_member_tac m i =
+    fun mset_member_tac ctxt m i =
       if m <= 0 then
-        rtac @{thm multi_member_this} i ORELSE rtac @{thm multi_member_last} i
+        resolve_tac ctxt @{thms multi_member_this} i ORELSE
+        resolve_tac ctxt @{thms multi_member_last} i
       else
-        rtac @{thm multi_member_skip} i THEN mset_member_tac (m - 1) i
-
-    val mset_nonempty_tac =
-      rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single}
+        resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i
+
+    fun mset_nonempty_tac ctxt =
+      resolve_tac ctxt @{thms nonempty_plus} ORELSE'
+      resolve_tac ctxt @{thms nonempty_single}
 
     fun regroup_munion_conv ctxt =
       Function_Lib.regroup_conv ctxt @{const_abbrev Mempty} @{const_name plus}
         (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral}))
 
-    fun unfold_pwleq_tac i =
-      (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st))
-        ORELSE (rtac @{thm pw_leq_lstep} i)
-        ORELSE (rtac @{thm pw_leq_empty} i)
+    fun unfold_pwleq_tac ctxt i =
+      (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st))
+        ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i)
+        ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i)
 
     val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union},
                         @{thm Un_insert_left}, @{thm Un_empty_left}]
@@ -1925,7 +1927,7 @@
       mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac,
       mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps,
       smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
-      reduction_pair= @{thm ms_reduction_pair}
+      reduction_pair = @{thm ms_reduction_pair}
     })
   end
 \<close>
--- a/src/HOL/Library/Old_SMT/old_smt_solver.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/Old_SMT/old_smt_solver.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -355,14 +355,14 @@
   fun tag_rules thms = map_index (apsnd (pair NONE)) thms
   fun tag_prems thms = map (pair ~1 o pair NONE) thms
 
-  fun resolve (SOME thm) = rtac thm 1
-    | resolve NONE = no_tac
+  fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1
+    | resolve _ NONE = no_tac
 
   fun tac prove ctxt rules =
     CONVERSION (Old_SMT_Normalize.atomize_conv ctxt)
-    THEN' rtac @{thm ccontr}
-    THEN' SUBPROOF (fn {context, prems, ...} =>
-      resolve (prove context (tag_rules rules @ tag_prems prems))) ctxt
+    THEN' resolve_tac ctxt @{thms ccontr}
+    THEN' SUBPROOF (fn {context = ctxt', prems, ...} =>
+      resolve ctxt' (prove ctxt' (tag_rules rules @ tag_prems prems))) ctxt
 in
 
 val smt_tac = tac safe_solve
--- a/src/HOL/Library/Old_SMT/old_z3_proof_methods.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/Old_SMT/old_z3_proof_methods.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -33,7 +33,7 @@
 fun prove_ite ctxt =
   Old_Z3_Proof_Tools.by_tac ctxt (
     CONVERSION (Conv.arg_conv (Conv.arg1_conv pull_ites_conv))
-    THEN' rtac @{thm refl})
+    THEN' resolve_tac ctxt @{thms refl})
 
 
 
@@ -71,7 +71,7 @@
     val rule = disjE OF [Object_Logic.rulify ctxt' (Thm.assume lhs)]
   in
     Goal.init (mk_inj_prop ctxt' (Thm.dest_arg ct))
-    |> apply (rtac @{thm injI})
+    |> apply (resolve_tac ctxt' @{thms injI})
     |> apply (Tactic.solve_tac ctxt' [rule, rule RS @{thm sym}])
     |> Goal.norm_result ctxt' o Goal.finish ctxt'
     |> singleton (Variable.export ctxt' ctxt)
@@ -81,7 +81,7 @@
   Old_Z3_Proof_Tools.by_tac ctxt (
     CONVERSION (Conv.top_sweep_conv (K (Conv.rewr_conv def)) ctxt)
     THEN' REPEAT_ALL_NEW (match_tac ctxt @{thms allI})
-    THEN' rtac (@{thm inv_f_f} OF [prove_inj_prop ctxt def lhs]))
+    THEN' resolve_tac ctxt [@{thm inv_f_f} OF [prove_inj_prop ctxt def lhs]])
 
 
 fun expand thm ct =
@@ -142,7 +142,7 @@
 fun prove_injectivity ctxt =
   Old_Z3_Proof_Tools.by_tac ctxt (
     CONVERSION (Old_SMT_Utils.prop_conv (norm_conv ctxt))
-    THEN' CSUBGOAL (uncurry (rtac o prove_inj_eq ctxt)))
+    THEN' CSUBGOAL (uncurry (resolve_tac ctxt o single o prove_inj_eq ctxt)))
 
 end
 
--- a/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -413,7 +413,7 @@
     @{lemma "(!!x. ~P x == Q x) ==> ~(EX x. P x) == ALL x. Q x" by simp}])
 
   fun nnf_quant_tac ctxt thm (qs as (qs1, qs2)) i st = (
-    rtac thm ORELSE'
+    resolve_tac ctxt [thm] ORELSE'
     (match_tac ctxt qs1 THEN' nnf_quant_tac ctxt thm qs) ORELSE'
     (match_tac ctxt qs2 THEN' nnf_quant_tac ctxt thm qs)) i st
 
@@ -602,7 +602,7 @@
 in
 fun quant_inst ctxt = Thm o Old_Z3_Proof_Tools.by_tac ctxt (
   REPEAT_ALL_NEW (match_tac ctxt [rule])
-  THEN' rtac @{thm excluded_middle})
+  THEN' resolve_tac ctxt @{thms excluded_middle})
 end
 
 
@@ -640,10 +640,10 @@
   apfst Thm oo close vars (yield_singleton Assumption.add_assumes)
 
 fun discharge_sk_tac ctxt i st = (
-  rtac @{thm trans} i
+  resolve_tac ctxt @{thms trans} i
   THEN resolve_tac ctxt sk_rules i
-  THEN (rtac @{thm refl} ORELSE' discharge_sk_tac ctxt) (i+1)
-  THEN rtac @{thm refl} i) st
+  THEN (resolve_tac ctxt @{thms refl} ORELSE' discharge_sk_tac ctxt) (i+1)
+  THEN resolve_tac ctxt @{thms refl} i) st
 
 end
 
@@ -715,14 +715,14 @@
         THEN_ALL_NEW NAMED ctxt' "fast (logic)" (fast_tac ctxt'))),
     Old_Z3_Proof_Tools.by_abstraction 0 (false, true) ctxt [] (fn ctxt' =>
       Old_Z3_Proof_Tools.by_tac ctxt' (
-        (rtac @{thm iff_allI} ORELSE' K all_tac)
+        (resolve_tac ctxt' @{thms iff_allI} ORELSE' K all_tac)
         THEN' NAMED ctxt' "simp (theory)" (Simplifier.simp_tac (put_simpset simpset ctxt'))
         THEN_ALL_NEW (
           NAMED ctxt' "fast (theory)" (HOL_fast_tac ctxt')
           ORELSE' NAMED ctxt' "arith (theory)" (Arith_Data.arith_tac ctxt')))),
     Old_Z3_Proof_Tools.by_abstraction 0 (true, true) ctxt [] (fn ctxt' =>
       Old_Z3_Proof_Tools.by_tac ctxt' (
-        (rtac @{thm iff_allI} ORELSE' K all_tac)
+        (resolve_tac ctxt' @{thms iff_allI} ORELSE' K all_tac)
         THEN' NAMED ctxt' "simp (full)" (Simplifier.simp_tac (put_simpset simpset ctxt'))
         THEN_ALL_NEW (
           NAMED ctxt' "fast (full)" (HOL_fast_tac ctxt')
@@ -731,7 +731,7 @@
     Old_Z3_Proof_Tools.by_abstraction abstraction_depth (true, true) ctxt []
       (fn ctxt' =>
         Old_Z3_Proof_Tools.by_tac ctxt' (
-          (rtac @{thm iff_allI} ORELSE' K all_tac)
+          (resolve_tac ctxt' @{thms iff_allI} ORELSE' K all_tac)
           THEN' NAMED ctxt' "simp (deepen)" (Simplifier.simp_tac (put_simpset simpset ctxt'))
           THEN_ALL_NEW (
             NAMED ctxt' "fast (deepen)" (HOL_fast_tac ctxt')
--- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -1007,12 +1007,12 @@
       else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs))
   in () end
 
-fun core_sos_tac print_cert prover = SUBPROOF (fn {concl, context, ...} =>
+fun core_sos_tac print_cert prover = SUBPROOF (fn {concl, context = ctxt, ...} =>
   let
     val _ = check_sos known_sos_constants concl
-    val (ths, certificates) = real_sos prover context (Thm.dest_arg concl)
+    val (th, certificates) = real_sos prover ctxt (Thm.dest_arg concl)
     val _ = print_cert certificates
-  in rtac ths 1 end);
+  in resolve_tac ctxt [th] 1 end);
 
 fun default_SOME _ NONE v = SOME v
   | default_SOME _ (SOME v) _ = SOME v;
@@ -1050,7 +1050,7 @@
           instantiate' [] [SOME d, SOME (Thm.dest_arg P)]
             (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto}
              else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast})
-      in rtac th i THEN Simplifier.asm_full_simp_tac simp_ctxt i end));
+      in resolve_tac ctxt [th] i THEN Simplifier.asm_full_simp_tac simp_ctxt i end));
 
 fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
 
--- a/src/HOL/Library/bnf_lfp_countable.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/bnf_lfp_countable.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -29,11 +29,13 @@
 
 fun meta_spec_mp_tac ctxt 0 = K all_tac
   | meta_spec_mp_tac ctxt depth =
-    dtac ctxt meta_spec THEN' meta_spec_mp_tac ctxt (depth - 1) THEN' dtac ctxt meta_mp THEN' atac;
+    dtac ctxt meta_spec THEN' meta_spec_mp_tac ctxt (depth - 1) THEN'
+    dtac ctxt meta_mp THEN' assume_tac ctxt;
 
 fun use_induction_hypothesis_tac ctxt =
   DEEPEN (1, 64 (* large number *))
-    (fn depth => meta_spec_mp_tac ctxt depth THEN' etac ctxt allE THEN' etac ctxt impE THEN' atac THEN' atac) 0;
+    (fn depth => meta_spec_mp_tac ctxt depth THEN' etac ctxt allE THEN' etac ctxt impE THEN'
+      assume_tac ctxt THEN' assume_tac ctxt) 0;
 
 val same_ctr_simps = @{thms sum_encode_eq prod_encode_eq sum.inject prod.inject to_nat_split
   id_apply snd_conv simp_thms};
@@ -44,7 +46,7 @@
       (ss_only (injects @ recs @ map_congs' @ same_ctr_simps) ctxt) THEN_MAYBE'
     TRY o REPEAT_ALL_NEW (rtac ctxt conjI) THEN_ALL_NEW
     REPEAT_ALL_NEW (eresolve_tac ctxt (conjE :: inj_map_strongs')) THEN_ALL_NEW
-    (atac ORELSE' use_induction_hypothesis_tac ctxt));
+    (assume_tac ctxt ORELSE' use_induction_hypothesis_tac ctxt));
 
 fun distinct_ctrs_tac ctxt recs =
   HEADGOAL (asm_full_simp_tac (ss_only (recs @ distinct_ctrs_simps) ctxt));
--- a/src/HOL/Library/old_recdef.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Library/old_recdef.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -262,19 +262,20 @@
 struct
 
 (* make a casethm from an induction thm *)
-val cases_thm_of_induct_thm =
-     Seq.hd o (ALLGOALS (fn i => REPEAT (etac Drule.thin_rl i)));
+fun cases_thm_of_induct_thm ctxt =
+  Seq.hd o (ALLGOALS (fn i => REPEAT (eresolve_tac ctxt [Drule.thin_rl] i)));
 
 (* get the case_thm (my version) from a type *)
-fun case_thm_of_ty thy ty  =
+fun case_thm_of_ty ctxt ty  =
     let
+      val thy = Proof_Context.theory_of ctxt
       val ty_str = case ty of
                      Type(ty_str, _) => ty_str
                    | TFree(s,_)  => error ("Free type: " ^ s)
                    | TVar((s,_),_) => error ("Free variable: " ^ s)
       val {induct, ...} = BNF_LFP_Compat.the_info thy [BNF_LFP_Compat.Keep_Nesting] ty_str
     in
-      cases_thm_of_induct_thm induct
+      cases_thm_of_induct_thm ctxt induct
     end;
 
 
@@ -287,7 +288,7 @@
       val x = Free(vstr,ty);
       val abst = Abs(vstr, ty, Term.abstract_over (x, gt));
 
-      val case_thm = case_thm_of_ty thy ty;
+      val case_thm = case_thm_of_ty ctxt ty;
 
       val abs_ct = Thm.cterm_of ctxt abst;
       val free_ct = Thm.cterm_of ctxt x;
@@ -2640,7 +2641,8 @@
 (*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 ctxt [allI, impI, conjI] ORELSE' etac conjE)));
+  rule_by_tactic ctxt
+    (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' eresolve_tac ctxt [conjE])));
 
 (*Strip off the outer !P*)
 val spec'=
--- a/src/HOL/List.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/List.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -621,23 +621,25 @@
     | NONE => NONE)
   end
 
-fun tac _ [] = rtac set_singleton 1 ORELSE rtac inst_Collect_mem_eq 1
+fun tac ctxt [] =
+      resolve_tac ctxt [set_singleton] 1 ORELSE
+      resolve_tac ctxt [inst_Collect_mem_eq] 1
   | tac ctxt (If :: cont) =
       Splitter.split_tac ctxt @{thms split_if} 1
-      THEN rtac @{thm conjI} 1
-      THEN rtac @{thm impI} 1
+      THEN resolve_tac ctxt @{thms conjI} 1
+      THEN resolve_tac ctxt @{thms impI} 1
       THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
         CONVERSION (right_hand_set_comprehension_conv (K
           (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
            then_conv
            rewr_conv' @{lemma "(True \<and> P) = P" by simp})) ctxt') 1) ctxt 1
       THEN tac ctxt cont
-      THEN rtac @{thm impI} 1
+      THEN resolve_tac ctxt @{thms impI} 1
       THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
           CONVERSION (right_hand_set_comprehension_conv (K
             (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
              then_conv rewr_conv' @{lemma "(False \<and> P) = False" by simp})) ctxt') 1) ctxt 1
-      THEN rtac set_Nil_I 1
+      THEN resolve_tac ctxt [set_Nil_I] 1
   | tac ctxt (Case (T, i) :: cont) =
       let
         val SOME {injects, distincts, case_thms, split, ...} =
@@ -646,9 +648,9 @@
         (* do case distinction *)
         Splitter.split_tac ctxt [split] 1
         THEN EVERY (map_index (fn (i', _) =>
-          (if i' < length case_thms - 1 then rtac @{thm conjI} 1 else all_tac)
-          THEN REPEAT_DETERM (rtac @{thm allI} 1)
-          THEN rtac @{thm impI} 1
+          (if i' < length case_thms - 1 then resolve_tac ctxt @{thms conjI} 1 else all_tac)
+          THEN REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1)
+          THEN resolve_tac ctxt @{thms impI} 1
           THEN (if i' = i then
             (* continue recursively *)
             Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
@@ -683,7 +685,7 @@
                         Conv.repeat_conv
                           (Conv.bottom_conv
                             (K (rewr_conv' @{lemma "(\<exists>x. P) = P" by simp})) ctxt'')) ctxt'))) 1) ctxt 1
-            THEN rtac set_Nil_I 1)) case_thms)
+            THEN resolve_tac ctxt [set_Nil_I] 1)) case_thms)
       end
 
 in
--- a/src/HOL/Metis_Examples/Type_Encodings.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Metis_Examples/Type_Encodings.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -54,7 +54,7 @@
     fun tac [] st = all_tac st
       | tac (type_enc :: type_encs) st =
         st (* |> tap (fn _ => tracing (@{make_string} type_enc)) *)
-           |> ((if null type_encs then all_tac else rtac @{thm fork} 1)
+           |> ((if null type_encs then all_tac else resolve_tac ctxt @{thms fork} 1)
                THEN Metis_Tactic.metis_tac [type_enc]
                     ATP_Problem_Generate.combsN ctxt ths 1
                THEN COND (has_fewer_prems 2) all_tac no_tac
--- a/src/HOL/MicroJava/J/JTypeSafe.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/MicroJava/J/JTypeSafe.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -178,8 +178,9 @@
 
 ML{*
 fun forward_hyp_tac ctxt =
-  ALLGOALS (TRY o (EVERY' [dtac spec, mp_tac ctxt,
-    (mp_tac ctxt ORELSE' (dtac spec THEN' mp_tac ctxt)), REPEAT o (etac conjE)]))
+  ALLGOALS (TRY o (EVERY' [dresolve_tac ctxt [spec], mp_tac ctxt,
+    (mp_tac ctxt ORELSE' (dresolve_tac ctxt [spec] THEN' mp_tac ctxt)),
+    REPEAT o (eresolve_tac ctxt [conjE])]))
 *}
 
 
@@ -202,7 +203,7 @@
 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}])")
+apply(tactic "ALLGOALS (EVERY' [REPEAT o (eresolve_tac @{context} [conjE]), REPEAT o hyp_subst_tac @{context}])")
 
 -- "Level 7"
 -- "15 NewC"
@@ -241,7 +242,7 @@
 
 -- "for FAss"
 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*})
+       THEN_ALL_NEW (full_simp_tac @{context}), REPEAT o (eresolve_tac @{context} [conjE]), hyp_subst_tac @{context}] 3*})
 
 -- "for if"
 apply( tactic {* (Induct_Tacs.case_tac @{context} "the_Bool v" [] NONE THEN_ALL_NEW
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -127,8 +127,8 @@
               @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}])
   fun vector_arith_tac ctxt ths =
     simp_tac (put_simpset ss1 ctxt)
-    THEN' (fn i => rtac @{thm Cartesian_Euclidean_Space.setsum_cong_aux} i
-         ORELSE rtac @{thm setsum.neutral} i
+    THEN' (fn i => resolve_tac ctxt @{thms Cartesian_Euclidean_Space.setsum_cong_aux} i
+         ORELSE resolve_tac ctxt @{thms setsum.neutral} i
          ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i)
     (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
     THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths)
--- a/src/HOL/Multivariate_Analysis/normarith.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/normarith.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -401,6 +401,6 @@
  fun norm_arith_tac ctxt =
    clarify_tac (put_claset HOL_cs ctxt) THEN'
    Object_Logic.full_atomize_tac ctxt THEN'
-   CSUBGOAL ( fn (p,i) => rtac (norm_arith ctxt (Thm.dest_arg p )) i);
+   CSUBGOAL ( fn (p,i) => resolve_tac ctxt [norm_arith ctxt (Thm.dest_arg p )] i);
 
 end;
--- a/src/HOL/Nominal/nominal_atoms.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_atoms.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -117,10 +117,10 @@
               val simp1 = @{thm inj_on_def} :: injects;
               
               fun proof1 ctxt = EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1,
-                                          rtac @{thm ballI} 1,
-                                          rtac @{thm ballI} 1,
-                                          rtac @{thm impI} 1,
-                                          atac 1]
+                resolve_tac ctxt @{thms ballI} 1,
+                resolve_tac ctxt @{thms ballI} 1,
+                resolve_tac ctxt @{thms impI} 1,
+                assume_tac ctxt 1]
              
               val (inj_thm,thy2) = 
                    add_thms_string [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 (proof1 o #context)), [])] thy1
@@ -133,8 +133,8 @@
               val proof2 = fn {prems, context = ctxt} =>
                 Induct_Tacs.case_tac ctxt "y" [] NONE 1 THEN
                 asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1 THEN
-                rtac @{thm exI} 1 THEN
-                rtac @{thm refl} 1
+                resolve_tac ctxt @{thms exI} 1 THEN
+                resolve_tac ctxt @{thms refl} 1
 
               (* third statement *)
               val (inject_thm,thy3) =
@@ -149,9 +149,9 @@
               val simp3 = [@{thm UNIV_def}]
 
               fun proof3 ctxt = EVERY [cut_facts_tac inj_thm 1,
-                                          dtac @{thm range_inj_infinite} 1,
-                                          asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp2) 1,
-                                          simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp3) 1]
+                dresolve_tac ctxt @{thms range_inj_infinite} 1,
+                asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp2) 1,
+                simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp3) 1]
            
               val (inf_thm,thy4) =  
                     add_thms_string [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 (proof3 o #context)), [])] thy3
@@ -437,7 +437,10 @@
              fun proof ctxt =
               simp_tac (put_simpset HOL_basic_ss ctxt
                   addsimps maps (Global_Theory.get_thms thy') ["cp_def"]) 1
-                THEN EVERY [rtac allI 1, rtac allI 1, rtac allI 1, rtac cp1 1];
+                THEN EVERY [resolve_tac ctxt [allI] 1,
+                  resolve_tac ctxt [allI] 1,
+                  resolve_tac ctxt [allI] 1,
+                  resolve_tac ctxt [cp1] 1];
            in
              yield_singleton add_thms_string ((name,
                Goal.prove_global thy' [] [] statement (proof o #context)), []) thy'
@@ -504,10 +507,10 @@
            val at_inst  = Global_Theory.get_thm thy' ("at_" ^ ak_name' ^ "_inst");
 
            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];
+             resolve_tac ctxt [(at_inst RS at_pt_inst) RS pt1] 1,
+             resolve_tac ctxt [(at_inst RS at_pt_inst) RS pt2] 1,
+             resolve_tac ctxt [(at_inst RS at_pt_inst) RS pt3] 1,
+             assume_tac ctxt 1];
            fun proof2 ctxt =
              Class.intro_classes_tac ctxt [] THEN
              REPEAT (asm_simp_tac
@@ -537,7 +540,10 @@
 
           fun pt_proof thm ctxt =
               EVERY [Class.intro_classes_tac ctxt [],
-                     rtac (thm RS pt1) 1, rtac (thm RS pt2) 1, rtac (thm RS pt3) 1, atac 1];
+                resolve_tac ctxt [thm RS pt1] 1,
+                resolve_tac ctxt [thm RS pt2] 1,
+                resolve_tac ctxt [thm RS pt3] 1,
+                assume_tac ctxt 1];
 
           val pt_thm_fun   = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst));
           val pt_thm_set   = pt_inst RS pt_set_inst;
@@ -581,9 +587,10 @@
            fun proof ctxt =
                (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 ctxt [],
-                             rtac ((at_thm RS fs_at_inst) RS fs1) 1] end
+                  let val at_thm = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst") in
+                    EVERY [Class.intro_classes_tac ctxt [],
+                      resolve_tac ctxt [(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 =
@@ -605,7 +612,8 @@
         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 ctxt [], rtac (thm RS fs1) 1];
+          fun fs_proof thm ctxt =
+            EVERY [Class.intro_classes_tac ctxt [], resolve_tac ctxt [thm RS fs1] 1];
 
           val fs_thm_unit  = fs_unit_inst;
           val fs_thm_prod  = fs_inst RS (fs_inst RS fs_prod_inst);
@@ -652,7 +660,7 @@
                     val at_inst  = Global_Theory.get_thm thy'' ("at_"^ak_name''^"_inst");
                   in
                    EVERY [Class.intro_classes_tac ctxt [],
-                          rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1]
+                     resolve_tac ctxt [at_inst RS (pt_inst RS pt_perm_compose)] 1]
                   end)
                 else
                   (let
@@ -686,7 +694,8 @@
             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 ctxt [], rtac (thm RS cp1) 1];
+            fun cp_proof thm ctxt =
+              EVERY [Class.intro_classes_tac ctxt [], resolve_tac ctxt [thm RS cp1] 1];
           
             val cp_thm_unit = cp_unit_inst;
             val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst);
--- a/src/HOL/Nominal/nominal_datatype.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -586,7 +586,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 ctxt => rtac exI 1 THEN rtac CollectI 1 THEN
+            (fn ctxt =>
+              resolve_tac ctxt [exI] 1 THEN
+              resolve_tac ctxt [CollectI] 1 THEN
               QUIET_BREADTH_FIRST (has_fewer_prems 1)
               (resolve_tac ctxt rep_intrs 1)) thy |> (fn ((_, r), thy) =>
         let
@@ -663,8 +665,8 @@
                  (if atom1 = atom2 then []
                   else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
               cong_tac ctxt 1,
-              rtac refl 1,
-              rtac cp1' 1]) thy)
+              resolve_tac ctxt [refl] 1,
+              resolve_tac ctxt [cp1'] 1]) thy)
         (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
            tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
       end;
@@ -812,7 +814,7 @@
       in Goal.prove_global_future thy8 [] [] eqn (fn {context = ctxt, ...} => EVERY
         [resolve_tac ctxt inj_thms 1,
          rewrite_goals_tac ctxt rewrites,
-         rtac refl 3,
+         resolve_tac ctxt [refl] 3,
          resolve_tac ctxt rep_intrs 2,
          REPEAT (resolve_tac ctxt Rep_thms 1)])
       end;
@@ -1046,11 +1048,12 @@
         (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems),
          HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls)))
          (fn {context = ctxt, ...} => EVERY
-           [REPEAT (etac conjE 1),
+           [REPEAT (eresolve_tac ctxt [conjE] 1),
             REPEAT (EVERY
-              [TRY (rtac conjI 1),
+              [TRY (resolve_tac ctxt [conjI] 1),
                full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps Rep_inverse_thms) 1,
-               etac mp 1, resolve_tac ctxt Rep_thms 1])]);
+               eresolve_tac ctxt [mp] 1,
+               resolve_tac ctxt Rep_thms 1])]);
 
     val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule_lemma)));
     val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
@@ -1064,7 +1067,7 @@
     val dt_induct = Goal.prove_global_future thy8 []
       (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
       (fn {prems, context = ctxt} => EVERY
-        [rtac indrule_lemma' 1,
+        [resolve_tac ctxt [indrule_lemma'] 1,
          (Old_Datatype_Aux.ind_tac ctxt rep_induct [] THEN_ALL_NEW
            Object_Logic.atomize_prems_tac ctxt) 1,
          EVERY (map (fn (prem, r) => (EVERY
@@ -1262,9 +1265,9 @@
                fin_set_supp @ ths)) 1]);
         val (([(_, cx)], ths), ctxt') = Obtain.result
           (fn ctxt' => EVERY
-            [etac exE 1,
+            [eresolve_tac ctxt' [exE] 1,
              full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1,
-             REPEAT (etac conjE 1)])
+             REPEAT (eresolve_tac ctxt' [conjE] 1)])
           [ex] ctxt
       in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
 
@@ -1320,7 +1323,7 @@
              EVERY (Old_Datatype_Aux.ind_tac context1 dt_induct tnames 1 ::
                maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
                  map (fn ((cname, cargs), is) =>
-                   REPEAT (rtac allI 1) THEN
+                   REPEAT (resolve_tac context1 [allI] 1) THEN
                    SUBPROOF (fn {prems = iprems, params, concl,
                        context = context2, ...} =>
                      let
@@ -1354,14 +1357,15 @@
                               list_comb (cnstr, maps (fn ((bs, t), cs) =>
                                 cs @ [fold_rev (mk_perm []) (map perm_of_pair
                                   (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
-                           (fn _ => EVERY
+                           (fn {context = ctxt, ...} => EVERY
                               (simp_tac (put_simpset HOL_ss context3 addsimps flat inject_thms) 1 ::
-                               REPEAT (FIRSTGOAL (rtac conjI)) ::
+                               REPEAT (FIRSTGOAL (resolve_tac ctxt [conjI])) ::
                                maps (fn ((bs, t), cs) =>
                                  if null bs then []
-                                 else rtac sym 1 :: maps (fn (b, c) =>
-                                   [rtac trans 1, rtac sym 1,
-                                    rtac (fresh_fresh_inst thy9 b c) 1,
+                                 else resolve_tac ctxt [sym] 1 :: maps (fn (b, c) =>
+                                   [resolve_tac ctxt [trans] 1,
+                                    resolve_tac ctxt [sym] 1,
+                                    resolve_tac ctxt [fresh_fresh_inst thy9 b c] 1,
                                     simp_tac ind_ss1' 1,
                                     simp_tac ind_ss2 1,
                                     simp_tac ind_ss3' 1]) (bs ~~ cs))
@@ -1385,8 +1389,8 @@
         EVERY
           [cut_facts_tac [th] 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)]
+           REPEAT (eresolve_tac context [allE] 1),
+           REPEAT (TRY (resolve_tac context [conjI] 1) THEN asm_full_simp_tac ind_ss5 1)]
       end);
 
     val induct_aux' = Thm.instantiate ([],
@@ -1404,10 +1408,10 @@
       (map (augment_sort thy9 fs_cp_sort) ind_prems)
       (augment_sort thy9 fs_cp_sort ind_concl)
       (fn {prems, context = ctxt} => EVERY
-         [rtac induct_aux' 1,
+         [resolve_tac ctxt [induct_aux'] 1,
           REPEAT (resolve_tac ctxt fs_atoms 1),
           REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW
-            (etac @{thm meta_spec} ORELSE'
+            (eresolve_tac ctxt @{thms meta_spec} ORELSE'
               full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [fresh_def]))) 1)])
 
     val (_, thy10) = thy9 |>
@@ -1551,7 +1555,8 @@
             (Goal.prove_global_future thy11 [] []
               (augment_sort thy1 pt_cp_sort
                 (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
-              (fn {context = ctxt, ...} => rtac rec_induct 1 THEN REPEAT
+              (fn {context = ctxt, ...} =>
+                 resolve_tac ctxt [rec_induct] 1 THEN REPEAT
                  (simp_tac (put_simpset HOL_basic_ss ctxt
                     addsimps flat perm_simps'
                     addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
@@ -1561,9 +1566,10 @@
           Goal.prove_global_future thy11 [] []
             (augment_sort thy1 pt_cp_sort
               (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
-            (fn {context = ctxt, ...} => dtac (Thm.instantiate ([],
+            (fn {context = ctxt, ...} =>
+              dresolve_tac ctxt [Thm.instantiate ([],
                  [((("pi", 0), permT),
-                   Thm.global_cterm_of thy11 (Const (@{const_name rev}, permT --> permT) $ pi))]) th) 1 THEN
+                   Thm.global_cterm_of thy11 (Const (@{const_name rev}, permT --> permT) $ pi))]) th] 1 THEN
                NominalPermeq.perm_simp_tac (put_simpset HOL_ss ctxt) 1)) (ps ~~ ths)
       in (ths, ths') end) dt_atomTs);
 
@@ -1594,7 +1600,7 @@
                    end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
                      (1 upto length recTs))))))
             (fn {prems = fins, context = ctxt} =>
-              (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
+              (resolve_tac ctxt [rec_induct] THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
                (NominalPermeq.finite_guess_tac (put_simpset HOL_ss ctxt addsimps [fs_name]) 1))))
       end) dt_atomTs;
 
@@ -1640,25 +1646,25 @@
                    val _ $ (_ $ (_ $ S $ _)) $ _ = Thm.prop_of supports_fresh;
                    val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
                  in EVERY
-                   [rtac (Drule.cterm_instantiate
+                   [resolve_tac context [Drule.cterm_instantiate
                       [(Thm.global_cterm_of thy11 S,
                         Thm.global_cterm_of thy11 (Const (@{const_name Nominal.supp},
                           fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
-                      supports_fresh) 1,
+                      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 context [allI, impI] 1),
-                    REPEAT_DETERM (etac conjE 1),
-                    rtac unique 1,
-                    SUBPROOF (fn {prems = prems', params = [(_, a), (_, b)], ...} => EVERY
-                      [cut_facts_tac [rec_prem] 1,
-                       rtac (Thm.instantiate ([],
+                    REPEAT_DETERM (eresolve_tac context [conjE] 1),
+                    resolve_tac context [unique] 1,
+                    SUBPROOF (fn {context = ctxt, prems = prems', params = [(_, a), (_, b)], ...} =>
+                      EVERY [cut_facts_tac [rec_prem] 1,
+                       resolve_tac ctxt [Thm.instantiate ([],
                          [((("pi", 0), mk_permT aT),
                            Thm.global_cterm_of thy11
-                            (perm_of_pair (Thm.term_of a, Thm.term_of b)))]) eqvt_th) 1,
+                            (perm_of_pair (Thm.term_of a, Thm.term_of b)))]) eqvt_th] 1,
                        asm_simp_tac (put_simpset HOL_ss context addsimps
                          (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
-                    rtac rec_prem 1,
+                    resolve_tac context [rec_prem] 1,
                     simp_tac (put_simpset HOL_ss context addsimps (fs_name ::
                       supp_prod :: finite_Un :: finite_prems)) 1,
                     simp_tac (put_simpset HOL_ss context addsimps (Thm.symmetric fresh_def ::
@@ -1705,10 +1711,10 @@
              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
-            [etac exE 1,
-             full_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_prod :: fresh_atm)) 1,
-             REPEAT (etac conjE 1)])
+          (fn ctxt' => EVERY
+            [eresolve_tac ctxt' [exE] 1,
+             full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1,
+             REPEAT (eresolve_tac ctxt [conjE] 1)])
           [ex] ctxt
       in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
 
@@ -1740,13 +1746,13 @@
                         (rec_unique_frees'' ~~ rec_unique_frees' ~~
                            rec_sets ~~ rec_preds)))))
                (fn {context = ctxt, ...} =>
-                  rtac rec_induct 1 THEN
+                  resolve_tac ctxt [rec_induct] 1 THEN
                   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);
          in EVERY
-           ([rtac induct_aux_rec 1] @
+           ([resolve_tac context [induct_aux_rec] 1] @
             maps (fn ((_, finite_ths), finite_th) =>
               [cut_facts_tac (finite_th :: finite_ths) 1,
                asm_simp_tac (put_simpset HOL_ss context addsimps [supp_prod, finite_Un]) 1])
@@ -1754,10 +1760,10 @@
             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 context @{thms conjE ex1E} 1),
-               rtac @{thm ex1I} 1,
-               (resolve_tac context rec_intrs THEN_ALL_NEW atac) 1,
+               resolve_tac context @{thms ex1I} 1,
+               (resolve_tac context rec_intrs THEN_ALL_NEW assume_tac context) 1,
                rotate_tac ~1 1,
-               ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
+               ((DETERM o eresolve_tac context [elim]) THEN_ALL_NEW full_simp_tac
                   (put_simpset HOL_ss context addsimps flat distinct_thms)) 1] @
                (if null idxs then [] else [hyp_subst_tac context 1,
                 SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
@@ -1825,10 +1831,10 @@
                     val pi1_pi2_eq = Goal.prove context'' [] []
                       (HOLogic.mk_Trueprop (HOLogic.mk_eq
                         (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
-                      (fn _ => EVERY
+                      (fn {context = ctxt, ...} => EVERY
                          [cut_facts_tac constr_fresh_thms 1,
                           asm_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh) 1,
-                          rtac prem 1]);
+                          resolve_tac ctxt [prem] 1]);
 
                     (** pi1 o ts = pi2 o us **)
                     val _ = warning "step 4: pi1 o ts = pi2 o us";
@@ -1866,21 +1872,21 @@
                         val k = find_index (equal s) rec_set_names;
                         val pi = rpi1 @ pi2;
                         fun mk_pi z = fold_rev (mk_perm []) pi z;
-                        fun eqvt_tac p =
+                        fun eqvt_tac ctxt p =
                           let
                             val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
                             val l = find_index (equal T) dt_atomTs;
                             val th = nth (nth rec_equiv_thms' l) k;
                             val th' = Thm.instantiate ([],
                               [((("pi", 0), U), Thm.global_cterm_of thy11 p)]) th;
-                          in rtac th' 1 end;
+                          in resolve_tac ctxt [th'] 1 end;
                         val th' = Goal.prove context'' [] []
                           (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
-                          (fn _ => EVERY
-                             (map eqvt_tac pi @
+                          (fn {context = ctxt, ...} => EVERY
+                             (map (eqvt_tac ctxt) pi @
                               [simp_tac (put_simpset HOL_ss context'' addsimps (fresh_prems' @ freshs2' @
                                  perm_swap @ perm_fresh_fresh)) 1,
-                               rtac th 1]))
+                               resolve_tac ctxt [th] 1]))
                       in
                         Simplifier.simplify
                           (put_simpset HOL_basic_ss context'' addsimps rpi1_pi2_eqs) th'
@@ -1924,12 +1930,13 @@
                             in
                               Goal.prove context'' [] []
                                 (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
-                                (fn _ => EVERY
-                                   (rtac (nth (nth rec_fresh_thms l) k) 1 ::
-                                    map (fn th => rtac th 1)
+                                (fn {context = ctxt, ...} => EVERY
+                                   (resolve_tac ctxt [nth (nth rec_fresh_thms l) k] 1 ::
+                                    map (fn th => resolve_tac ctxt [th] 1)
                                       (snd (nth finite_thss l)) @
-                                    [rtac rec_prem 1, rtac ih 1,
-                                     REPEAT_DETERM (resolve_tac context fresh_prems 1)]))
+                                    [resolve_tac ctxt [rec_prem] 1,
+                                     resolve_tac ctxt [ih] 1,
+                                     REPEAT_DETERM (resolve_tac ctxt fresh_prems 1)]))
                             end) atoms
                         end) (rec_prems1 ~~ ihs);
 
@@ -1954,8 +1961,9 @@
                         (HOLogic.mk_Trueprop (fresh_const aT rT $
                             fold_rev (mk_perm []) (rpi2 @ pi1) a $
                             fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
-                        (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps fresh_bij) 1 THEN
-                           rtac th 1)
+                        (fn {context = ctxt, ...} =>
+                          simp_tac (put_simpset HOL_ss context'' addsimps fresh_bij) 1 THEN
+                          resolve_tac ctxt [th] 1)
                       in
                         Goal.prove context'' [] []
                           (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
@@ -2048,14 +2056,14 @@
         val prems' = flat finite_premss @ finite_ctxt_prems @
           rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
         fun solve ctxt rules prems = resolve_tac ctxt rules THEN_ALL_NEW
-          (resolve_tac ctxt prems THEN_ALL_NEW atac)
+          (resolve_tac ctxt prems THEN_ALL_NEW assume_tac ctxt)
       in
         Goal.prove_global_future thy12 []
           (map (augment_sort thy12 fs_cp_sort) prems')
           (augment_sort thy12 fs_cp_sort concl')
           (fn {context = ctxt, prems} => EVERY
             [rewrite_goals_tac ctxt reccomb_defs,
-             rtac @{thm the1_equality} 1,
+             resolve_tac ctxt @{thms the1_equality} 1,
              solve ctxt rec_unique_thms prems 1,
              resolve_tac ctxt rec_intrs 1,
              REPEAT (solve ctxt (prems @ rec_total_thms) prems 1)])
--- a/src/HOL/Nominal/nominal_fresh_fun.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -79,8 +79,8 @@
  in
    fn st =>
    (cut_inst_tac_term' ctxt [(x,s)] exists_fresh' 1 THEN
-   rtac fs_name_thm 1 THEN
-   etac exE 1) st
+   resolve_tac ctxt [fs_name_thm] 1 THEN
+   eresolve_tac ctxt [exE] 1) st
   handle List.Empty  => all_tac st (* if we collected no variables then we do nothing *)
  end) 1;
 
@@ -163,8 +163,8 @@
     (* The tactics which solve the subgoals generated
        by the conditionnal rewrite rule. *)
     val post_rewrite_tacs =
-          [rtac pt_name_inst,
-           rtac at_name_inst,
+          [resolve_tac ctxt [pt_name_inst],
+           resolve_tac ctxt [at_name_inst],
            TRY o SOLVED' (NominalPermeq.finite_guess_tac simp_ctxt),
            inst_fresh vars params THEN'
            (TRY o SOLVED' (NominalPermeq.fresh_guess_tac simp_ctxt)) THEN'
--- a/src/HOL/Nominal/nominal_induct.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_induct.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -121,7 +121,7 @@
               (finish_rule (Induct.internalize ctxt more_consumes rule)) i st'
             |> Seq.maps (fn rule' =>
               CASES (rule_cases ctxt rule' cases)
-                (rtac (rename_params_rule false [] rule') i THEN
+                (resolve_tac ctxt [rename_params_rule false [] rule'] i THEN
                   PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) st'))))
     THEN_ALL_NEW_CASES
       ((if simp then Induct.simplify_tac ctxt THEN' (TRY o Induct.trivial_tac ctxt)
--- a/src/HOL/Nominal/nominal_inductive.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -104,8 +104,8 @@
   | inst_conj_all _ _ _ _ _ = NONE;
 
 fun inst_conj_all_tac ctxt k = EVERY
-  [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]),
-   REPEAT_DETERM_N k (etac allE 1),
+  [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]),
+   REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1),
    simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1];
 
 fun map_term f t u = (case f t u of
@@ -129,9 +129,9 @@
     val prop = Thm.prop_of th;
     fun prove t =
       Goal.prove ctxt [] [] t (fn _ =>
-        EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
+        EVERY [cut_facts_tac [th] 1, eresolve_tac ctxt [rev_mp] 1,
           REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)),
-          REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
+          REPEAT_DETERM (resolve_tac ctxt [impI] 1 THEN (assume_tac ctxt 1 ORELSE tac))])
   in Option.map prove (map_term f prop (the_default prop opt)) end;
 
 val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion;
@@ -300,19 +300,20 @@
              resolve_tac ctxt fs_atoms 1]);
         val (([(_, cx)], ths), ctxt') = Obtain.result
           (fn ctxt' => EVERY
-            [etac exE 1,
+            [eresolve_tac ctxt' [exE] 1,
              full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1,
              full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1,
-             REPEAT (etac conjE 1)])
+             REPEAT (eresolve_tac ctxt [conjE] 1)])
           [ex] ctxt
       in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
 
     fun mk_ind_proof ctxt' thss =
       Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} =>
         let val th = Goal.prove ctxt [] [] concl (fn {context, ...} =>
-          rtac raw_induct 1 THEN
+          resolve_tac context [raw_induct] 1 THEN
           EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) =>
-            [REPEAT (rtac allI 1), simp_tac (put_simpset eqvt_ss context) 1,
+            [REPEAT (resolve_tac context [allI] 1),
+             simp_tac (put_simpset eqvt_ss context) 1,
              SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
                let
                  val (params', (pis, z)) =
@@ -354,7 +355,7 @@
                       if null (preds_of ps t) then (SOME th, mk_pi th)
                       else
                         (map_thm ctxt' (split_conj (K o I) names)
-                           (etac conjunct1 1) monos NONE th,
+                           (eresolve_tac ctxt' [conjunct1] 1) monos NONE th,
                          mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis''))
                            (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th))))
                       (gprems ~~ oprems)) |>> map_filter I;
@@ -370,7 +371,7 @@
                          (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs)
                             (fold_rev (NominalDatatype.mk_perm []) pis rhs)))
                        (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps
-                          (fresh_bij @ perm_bij)) 1 THEN rtac th' 1)
+                          (fresh_bij @ perm_bij)) 1 THEN resolve_tac ctxt' [th'] 1)
                    in Simplifier.simplify (put_simpset eqvt_ss ctxt'' addsimps fresh_atm) th'' end)
                      vc_compat_ths;
                  val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths';
@@ -382,7 +383,8 @@
                  val th = Goal.prove ctxt'' [] []
                    (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
                      map (fold (NominalDatatype.mk_perm []) pis') (tl ts))))
-                   (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, rtac ihyp' 1,
+                   (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1,
+                     resolve_tac ctxt'' [ihyp'] 1,
                      REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems)
                        (simp_tac swap_simps_simpset 1),
                      REPEAT_DETERM_N (length gprems)
@@ -398,9 +400,10 @@
                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
+          cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac ctxt [conjE] 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
+            eresolve_tac ctxt [impE] 1 THEN assume_tac ctxt 1 THEN
+            REPEAT (eresolve_tac ctxt @{thms allE_Nil} 1) THEN
             asm_full_simp_tac ctxt 1)
         end) |> singleton (Proof_Context.export ctxt' ctxt);
 
@@ -466,11 +469,11 @@
         prems') =
       (name, Goal.prove ctxt' [] (prem :: prems') concl
         (fn {prems = hyp :: hyps, context = ctxt1} =>
-        EVERY (rtac (hyp RS elim) 1 ::
+        EVERY (resolve_tac ctxt1 [hyp RS elim] 1 ::
           map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) =>
             SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} =>
               if null qs then
-                rtac (first_order_mrs case_hyps case_hyp) 1
+                resolve_tac ctxt2 [first_order_mrs case_hyps case_hyp] 1
               else
                 let
                   val params' = map (Thm.term_of o #2 o nth (rev params)) is;
@@ -518,8 +521,8 @@
                     (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj);
                   val th = Goal.prove ctxt3 [] [] (Thm.term_of concl)
                     (fn {context = ctxt4, ...} =>
-                       rtac (Thm.instantiate inst case_hyp) 1 THEN
-                       SUBPROOF (fn {prems = fresh_hyps, ...} =>
+                       resolve_tac ctxt4 [Thm.instantiate inst case_hyp] 1 THEN
+                       SUBPROOF (fn {context = ctxt5, prems = fresh_hyps, ...} =>
                          let
                            val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps;
                            val case_simpset = cases_eqvt_simpset addsimps freshs2' @
@@ -532,8 +535,8 @@
                            val case_hyps' = hyps1' @ hyps2'
                          in
                            simp_tac case_simpset 1 THEN
-                           REPEAT_DETERM (TRY (rtac conjI 1) THEN
-                             resolve_tac ctxt4 case_hyps' 1)
+                           REPEAT_DETERM (TRY (resolve_tac ctxt5 [conjI] 1) THEN
+                             resolve_tac ctxt5 case_hyps' 1)
                          end) ctxt4 1)
                   val final = Proof_Context.export ctxt3 ctxt2 [th]
                 in resolve_tac ctxt2 final 1 end) ctxt1 1)
@@ -629,13 +632,13 @@
         val res = SUBPROOF (fn {context = ctxt'', prems, params, ...} =>
           let
             val prems' = map (fn th => the_default th (map_thm ctxt''
-              (split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems;
+              (split_conj (K I) names) (eresolve_tac ctxt'' [conjunct2] 1) monos NONE th)) prems;
             val prems'' = map (fn th => Simplifier.simplify eqvt_simpset
               (mk_perm_bool (Thm.global_cterm_of thy pi) th)) prems';
             val intr' = Drule.cterm_instantiate (map (Thm.global_cterm_of thy) vs ~~
                map (Thm.global_cterm_of thy o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params)
                intr
-          in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac ctxt'' prems'')) 1
+          in (resolve_tac ctxt'' [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
@@ -655,7 +658,7 @@
               HOLogic.mk_imp (p, list_comb (h, ts1 @
                 map (NominalDatatype.mk_perm [] pi') ts2))
             end) ps)))
-          (fn _ => EVERY (rtac raw_induct 1 :: map (fn intr_vs =>
+          (fn _ => EVERY (resolve_tac ctxt' [raw_induct] 1 :: map (fn intr_vs =>
               full_simp_tac eqvt_simpset 1 THEN
               eqvt_tac pi' intr_vs) intrs')) |>
           singleton (Proof_Context.export ctxt' ctxt)))
--- a/src/HOL/Nominal/nominal_inductive2.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -109,8 +109,8 @@
   | inst_conj_all _ _ _ _ _ = NONE;
 
 fun inst_conj_all_tac ctxt k = EVERY
-  [TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]),
-   REPEAT_DETERM_N k (etac allE 1),
+  [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]),
+   REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1),
    simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1];
 
 fun map_term f t u = (case f t u of
@@ -134,9 +134,9 @@
     val prop = Thm.prop_of th;
     fun prove t =
       Goal.prove ctxt [] [] t (fn _ =>
-        EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
+        EVERY [cut_facts_tac [th] 1, eresolve_tac ctxt [rev_mp] 1,
           REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)),
-          REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))])
+          REPEAT_DETERM (resolve_tac ctxt [impI] 1 THEN (assume_tac ctxt 1 ORELSE tac))])
   in Option.map prove (map_term f prop (the_default prop opt)) end;
 
 fun abs_params params t =
@@ -321,11 +321,11 @@
           ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding});
         val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result
           (fn ctxt' => EVERY
-            [rtac avoid_th 1,
+            [resolve_tac ctxt' [avoid_th] 1,
              full_simp_tac (put_simpset HOL_ss ctxt' addsimps [@{thm fresh_star_prod_set}]) 1,
              full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1,
              rotate_tac 1 1,
-             REPEAT (etac conjE 1)])
+             REPEAT (eresolve_tac ctxt' [conjE] 1)])
           [] ctxt;
         val (Ts1, _ :: Ts2) = take_prefix (not o equal T) (map snd sets);
         val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2);
@@ -348,10 +348,10 @@
     fun mk_ind_proof ctxt' thss =
       Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} =>
         let val th = Goal.prove ctxt [] [] concl (fn {context, ...} =>
-          rtac raw_induct 1 THEN
+          resolve_tac ctxt [raw_induct] 1 THEN
           EVERY (maps (fn (((((_, sets, oprems, _),
               vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) =>
-            [REPEAT (rtac allI 1), simp_tac (put_simpset eqvt_ss context) 1,
+            [REPEAT (resolve_tac ctxt [allI] 1), simp_tac (put_simpset eqvt_ss context) 1,
              SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
                let
                  val (cparams', (pis, z)) =
@@ -366,7 +366,7 @@
                    if null (preds_of ps t) then SOME th
                    else
                      map_thm ctxt' (split_conj (K o I) names)
-                       (etac conjunct1 1) monos NONE th)
+                       (eresolve_tac ctxt' [conjunct1] 1) monos NONE th)
                    (gprems ~~ oprems);
                  val vc_compat_ths' = map2 (fn th => fn p =>
                    let
@@ -378,7 +378,7 @@
                        (HOLogic.mk_Trueprop (list_comb (h,
                           map (fold_rev (NominalDatatype.mk_perm []) pis) ts)))
                        (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt' addsimps
-                          (fresh_star_bij @ finite_ineq)) 1 THEN rtac th' 1)
+                          (fresh_star_bij @ finite_ineq)) 1 THEN resolve_tac ctxt' [th'] 1)
                    end) vc_compat_ths vc_compat_vs;
                  val (vc_compat_ths1, vc_compat_ths2) =
                    chop (length vc_compat_ths - length sets) vc_compat_ths';
@@ -416,8 +416,9 @@
                  val th = Goal.prove ctxt'' [] []
                    (HOLogic.mk_Trueprop (list_comb (P $ hd ts,
                      map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts))))
-                   (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, rtac ihyp' 1] @
-                     map (fn th => rtac th 1) fresh_ths3 @
+                   (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1,
+                     resolve_tac ctxt'' [ihyp'] 1] @
+                     map (fn th => resolve_tac ctxt'' [th] 1) fresh_ths3 @
                      [REPEAT_DETERM_N (length gprems)
                        (simp_tac (put_simpset HOL_basic_ss ctxt''
                           addsimps [inductive_forall_def']
@@ -431,9 +432,10 @@
                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
+          cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac ctxt' [conjE] 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
+            eresolve_tac ctxt' [impE] 1 THEN
+            assume_tac ctxt' 1 THEN REPEAT (eresolve_tac ctxt' @{thms allE_Nil} 1) THEN
             asm_full_simp_tac ctxt 1)
         end) |>
         fresh_postprocess ctxt' |>
--- a/src/HOL/Nominal/nominal_permeq.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_permeq.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -218,7 +218,7 @@
 fun perm_compose_tac ctxt i = 
   ("analysing permutation compositions on the lhs",
    fn st => EVERY
-     [rtac trans i,
+     [resolve_tac ctxt [trans] i,
       asm_full_simp_tac (empty_simpset ctxt addsimprocs [perm_compose_simproc]) i,
       asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [perm_aux_fold]) i] st);
 
@@ -230,14 +230,15 @@
 (*     pi o f = rhs                           *)  
 (* is transformed to                          *)
 (*     %x. pi o (f ((rev pi) o x)) = rhs      *)
-fun unfold_perm_fun_def_tac i =
+fun unfold_perm_fun_def_tac ctxt i =
     ("unfolding of permutations on functions", 
-      rtac (perm_fun_def RS meta_eq_to_obj_eq RS trans) i)
+      resolve_tac ctxt [perm_fun_def RS meta_eq_to_obj_eq RS trans] i)
 
 (* applies the ext-rule such that      *)
 (*                                     *)
 (*    f = g   goes to  /\x. f x = g x  *)
-fun ext_fun_tac i = ("extensionality expansion of functions", rtac @{thm ext} i);
+fun ext_fun_tac ctxt i =
+  ("extensionality expansion of functions", resolve_tac ctxt @{thms ext} i);
 
 
 (* perm_extend_simp_tac_i is perm_simp plus additional tactics        *)
@@ -248,13 +249,14 @@
   let fun perm_extend_simp_tac_aux tactical ctxt n = 
           if n=0 then K all_tac
           else DETERM o 
-               (FIRST'[fn i => tactical ctxt ("splitting conjunctions on the rhs", rtac conjI i),
-                       fn i => tactical ctxt (perm_simp asm_full_simp_tac ctxt i),
-                       fn i => tactical ctxt (perm_compose_tac ctxt i),
-                       fn i => tactical ctxt (apply_cong_tac ctxt i), 
-                       fn i => tactical ctxt (unfold_perm_fun_def_tac i),
-                       fn i => tactical ctxt (ext_fun_tac i)]
-                      THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ctxt (n-1))))
+               (FIRST'
+                 [fn i => tactical ctxt ("splitting conjunctions on the rhs", resolve_tac ctxt [conjI] i),
+                  fn i => tactical ctxt (perm_simp asm_full_simp_tac ctxt i),
+                  fn i => tactical ctxt (perm_compose_tac ctxt i),
+                  fn i => tactical ctxt (apply_cong_tac ctxt i), 
+                  fn i => tactical ctxt (unfold_perm_fun_def_tac ctxt i),
+                  fn i => tactical ctxt (ext_fun_tac ctxt i)]
+                THEN_ALL_NEW (TRY o (perm_extend_simp_tac_aux tactical ctxt (n-1))))
   in perm_extend_simp_tac_aux tactical ctxt 10 end;
 
 
@@ -266,10 +268,10 @@
      val simps        = [supports_def, Thm.symmetric fresh_def, fresh_prod]
   in
       EVERY [tactical ctxt ("unfolding of supports   ", simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps) i),
-             tactical ctxt ("stripping of foralls    ", REPEAT_DETERM (rtac allI i)),
-             tactical ctxt ("geting rid of the imps  ", rtac impI i),
-             tactical ctxt ("eliminating conjuncts   ", REPEAT_DETERM (etac  conjE i)),
-             tactical ctxt ("applying eqvt_simp      ", eqvt_simp_tac_gen_i asm_full_simp_tac tactical ctxt i )]
+             tactical ctxt ("stripping of foralls    ", REPEAT_DETERM (resolve_tac ctxt [allI] i)),
+             tactical ctxt ("geting rid of the imps  ", resolve_tac ctxt [impI] i),
+             tactical ctxt ("eliminating conjuncts   ", REPEAT_DETERM (eresolve_tac ctxt [conjE] i)),
+             tactical ctxt ("applying eqvt_simp      ", eqvt_simp_tac_gen_i asm_full_simp_tac tactical ctxt i)]
   end;
 
 
--- a/src/HOL/Nominal/nominal_thmdecls.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -58,12 +58,12 @@
     val mypifree = Thm.global_cterm_of thy (Const (@{const_name "rev"}, T --> T) $ pi')
     val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"
   in
-    EVERY1 [tactic ctxt ("iffI applied", rtac @{thm iffI}),
-            tactic ctxt ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
-            tactic ctxt ("solve with orig_thm", etac orig_thm),
+    EVERY1 [tactic ctxt ("iffI applied", resolve_tac ctxt @{thms iffI}),
+            tactic ctxt ("remove pi with perm_boolE", dresolve_tac ctxt @{thms perm_boolE}),
+            tactic ctxt ("solve with orig_thm", eresolve_tac ctxt [orig_thm]),
             tactic ctxt ("applies orig_thm instantiated with rev pi",
-                       dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)),
-            tactic ctxt ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
+                       dresolve_tac ctxt [Drule.cterm_instantiate [(mypi,mypifree)] orig_thm]),
+            tactic ctxt ("getting rid of the pi on the right", resolve_tac ctxt @{thms perm_boolI}),
             tactic ctxt ("getting rid of all remaining perms",
                        full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps perm_pi_simp))]
   end;
--- a/src/HOL/Prolog/prolog.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Prolog/prolog.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -74,7 +74,7 @@
 (*val hyp_resolve_tac = Subgoal.FOCUS_PREMS (fn {prems, ...} =>
                                   resolve_tac (maps atomizeD prems) 1);
   -- is nice, but cannot instantiate unknowns in the assumptions *)
-val hyp_resolve_tac = SUBGOAL (fn (subgoal, i) =>
+fun hyp_resolve_tac ctxt = SUBGOAL (fn (subgoal, i) =>
   let
         fun ap (Const(@{const_name All},_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a  ,t))
         |   ap (Const(@{const_name HOL.implies},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
@@ -91,9 +91,9 @@
         val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal);
         fun drot_tac k i = DETERM (rotate_tac k i);
         fun spec_tac 0 i = all_tac
-        |   spec_tac k i = EVERY' [dtac spec, drot_tac ~1, spec_tac (k-1)] i;
+        |   spec_tac k i = EVERY' [dresolve_tac ctxt [spec], drot_tac ~1, spec_tac (k-1)] i;
         fun dup_spec_tac k i = if k = 0 then all_tac else EVERY'
-                      [DETERM o (etac all_dupE), drot_tac ~2, spec_tac (k-1)] i;
+                      [DETERM o (eresolve_tac ctxt [all_dupE]), drot_tac ~2, spec_tac (k-1)] i;
         fun same_head _ (Const (x,_)) (Const (y,_)) = x = y
         |   same_head k (s$_)         (t$_)         = same_head k s t
         |   same_head k (Bound i)     (Bound j)     = i = j + k
@@ -101,10 +101,10 @@
         fun mapn f n []      = []
         |   mapn f n (x::xs) = f n x::mapn f (n+1) xs;
         fun pres_tac (k,arrow,t) n i = drot_tac n i THEN (
-                if same_head k t concl
-                then dup_spec_tac k i THEN
-                     (if arrow then etac mp i THEN drot_tac (~n) i else atac i)
-                else no_tac);
+          if same_head k t concl
+          then dup_spec_tac k i THEN
+               (if arrow then eresolve_tac ctxt [mp] i THEN drot_tac (~n) i else assume_tac ctxt i)
+          else no_tac);
         val ptacs = mapn (fn n => fn t =>
                           pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems;
   in Library.foldl (op APPEND) (no_tac, ptacs) end);
@@ -112,16 +112,16 @@
 fun ptac ctxt prog = let
   val proga = maps (atomizeD ctxt) prog         (* atomize the prog *)
   in    (REPEAT_DETERM1 o FIRST' [
-                rtac TrueI,                     (* "True" *)
-                rtac conjI,                     (* "[| P; Q |] ==> P & Q" *)
-                rtac allI,                      (* "(!!x. P x) ==> ! x. P x" *)
-                rtac exI,                       (* "P x ==> ? x. P x" *)
-                rtac impI THEN'                 (* "(P ==> Q) ==> P --> Q" *)
+                resolve_tac ctxt [TrueI],                     (* "True" *)
+                resolve_tac ctxt [conjI],                     (* "[| P; Q |] ==> P & Q" *)
+                resolve_tac ctxt [allI],                      (* "(!!x. P x) ==> ! x. P x" *)
+                resolve_tac ctxt [exI],                       (* "P x ==> ? x. P x" *)
+                resolve_tac ctxt [impI] THEN'                 (* "(P ==> Q) ==> P --> Q" *)
                   asm_full_simp_tac (put_simpset atomize_ss ctxt) THEN'    (* atomize the asms *)
-                  (REPEAT_DETERM o (etac conjE))        (* split the asms *)
+                  (REPEAT_DETERM o (eresolve_tac ctxt [conjE]))        (* split the asms *)
                 ])
         ORELSE' resolve_tac ctxt [disjI1,disjI2]     (* "P ==> P | Q","Q ==> P | Q"*)
-        ORELSE' ((resolve_tac ctxt proga APPEND' hyp_resolve_tac)
+        ORELSE' ((resolve_tac ctxt proga APPEND' hyp_resolve_tac ctxt)
                  THEN' (fn _ => check_HOHH_tac2))
 end;
 
--- a/src/HOL/ROOT	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/ROOT	Sat Jul 18 20:59:51 2015 +0200
@@ -467,6 +467,7 @@
     AxSound
     AxCompl
     Trans
+    TypeSafe
   document_files "root.tex"
 
 session "HOL-IOA" in IOA = HOL +
--- a/src/HOL/SET_Protocol/Cardholder_Registration.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/SET_Protocol/Cardholder_Registration.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -540,7 +540,7 @@
     (fn i =>
       EVERY [forward_tac ctxt @{thms Gets_certificate_valid} i,
              assume_tac ctxt i,
-             etac conjE i, REPEAT (hyp_subst_tac ctxt i)]))
+             eresolve_tac ctxt [conjE] i, REPEAT (hyp_subst_tac ctxt i)]))
 *}
 
 text{*The @{text "(no_asm)"} attribute is essential, since it retains
--- a/src/HOL/SET_Protocol/Event_SET.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/SET_Protocol/Event_SET.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -183,7 +183,7 @@
 ML
 {*
 fun analz_mono_contra_tac ctxt = 
-  rtac @{thm analz_impI} THEN' 
+  resolve_tac ctxt @{thms analz_impI} THEN' 
   REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
   THEN' mp_tac ctxt
 *}
--- a/src/HOL/SPARK/Tools/spark_vcs.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -206,22 +206,22 @@
     val UNIV_eq = Goal.prove lthy [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq
          (HOLogic.mk_UNIV T, HOLogic.mk_set T cs)))
-      (fn _ =>
-         rtac @{thm subset_antisym} 1 THEN
-         rtac @{thm subsetI} 1 THEN
-         Old_Datatype_Aux.exh_tac lthy (K (#exhaust (BNF_LFP_Compat.the_info
-           (Proof_Context.theory_of lthy) [BNF_LFP_Compat.Keep_Nesting] tyname'))) 1 THEN
-         ALLGOALS (asm_full_simp_tac lthy));
+      (fn {context = ctxt, ...} =>
+         resolve_tac ctxt @{thms subset_antisym} 1 THEN
+         resolve_tac ctxt @{thms subsetI} 1 THEN
+         Old_Datatype_Aux.exh_tac ctxt (K (#exhaust (BNF_LFP_Compat.the_info
+           (Proof_Context.theory_of ctxt) [BNF_LFP_Compat.Keep_Nesting] tyname'))) 1 THEN
+         ALLGOALS (asm_full_simp_tac ctxt));
 
     val finite_UNIV = Goal.prove lthy [] []
       (HOLogic.mk_Trueprop (Const (@{const_name finite},
          HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T))
-      (fn _ => simp_tac (lthy addsimps [UNIV_eq]) 1);
+      (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [UNIV_eq]) 1);
 
     val card_UNIV = Goal.prove lthy [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq
          (card, HOLogic.mk_number HOLogic.natT k)))
-      (fn _ => simp_tac (lthy addsimps [UNIV_eq]) 1);
+      (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [UNIV_eq]) 1);
 
     val range_pos = Goal.prove lthy [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq
@@ -232,20 +232,20 @@
             HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $
               HOLogic.mk_number HOLogic.intT 0 $
               (@{term int} $ card))))
-      (fn _ =>
-         simp_tac (lthy addsimps [card_UNIV]) 1 THEN
-         simp_tac (lthy addsimps [UNIV_eq, def1]) 1 THEN
-         rtac @{thm subset_antisym} 1 THEN
-         simp_tac lthy 1 THEN
-         rtac @{thm subsetI} 1 THEN
-         asm_full_simp_tac (lthy addsimps @{thms interval_expand}
+      (fn {context = ctxt, ...} =>
+         simp_tac (ctxt addsimps [card_UNIV]) 1 THEN
+         simp_tac (ctxt addsimps [UNIV_eq, def1]) 1 THEN
+         resolve_tac ctxt @{thms subset_antisym} 1 THEN
+         simp_tac ctxt 1 THEN
+         resolve_tac ctxt @{thms subsetI} 1 THEN
+         asm_full_simp_tac (ctxt addsimps @{thms interval_expand}
            delsimps @{thms atLeastLessThan_iff}) 1);
 
     val lthy' =
       Class.prove_instantiation_instance (fn ctxt =>
         Class.intro_classes_tac ctxt [] THEN
-        rtac finite_UNIV 1 THEN
-        rtac range_pos 1 THEN
+        resolve_tac ctxt [finite_UNIV] 1 THEN
+        resolve_tac ctxt [range_pos] 1 THEN
         simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def3]) 1 THEN
         simp_tac (put_simpset HOL_basic_ss ctxt addsimps [def2]) 1) lthy;
 
@@ -254,23 +254,24 @@
         val n = HOLogic.mk_number HOLogic.intT i;
         val th = Goal.prove lthy' [] []
           (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n)))
-          (fn _ => simp_tac (lthy' addsimps [def1]) 1);
+          (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [def1]) 1);
         val th' = Goal.prove lthy' [] []
           (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c)))
-          (fn _ =>
-             rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN
-             simp_tac (lthy' addsimps [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
+          (fn {context = ctxt, ...} =>
+             resolve_tac ctxt [@{thm inj_pos} RS @{thm injD}] 1 THEN
+             simp_tac (ctxt addsimps [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
       in (th, th') end) cs);
 
     val first_el = Goal.prove lthy' [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq
          (Const (@{const_name first_el}, T), hd cs)))
-      (fn _ => simp_tac (lthy' addsimps [@{thm first_el_def}, hd val_eqs]) 1);
+      (fn {context = ctxt, ...} => simp_tac (ctxt addsimps [@{thm first_el_def}, hd val_eqs]) 1);
 
     val last_el = Goal.prove lthy' [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq
          (Const (@{const_name last_el}, T), List.last cs)))
-      (fn _ => simp_tac (lthy' addsimps [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
+      (fn {context = ctxt, ...} =>
+        simp_tac (ctxt addsimps [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
   in
     lthy' |>
     Local_Theory.note
--- a/src/HOL/Statespace/distinct_tree_prover.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Statespace/distinct_tree_prover.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -345,7 +345,7 @@
       Const (@{const_name Trueprop}, _) $
           (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ x $ y)) =>
         (case get_fst_success (neq_x_y ctxt x y) names of
-          SOME neq => rtac neq i
+          SOME neq => resolve_tac ctxt [neq] i
         | NONE => no_tac)
     | _ => no_tac))
 
--- a/src/HOL/Statespace/state_fun.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Statespace/state_fun.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -248,7 +248,7 @@
                   val eq1 =
                     Goal.prove ctxt0 [] []
                       (Logic.list_all (vars, Logic.mk_equals (trm, trm')))
-                      (fn _ => rtac meta_ext 1 THEN simp_tac ctxt1 1);
+                      (fn _ => resolve_tac ctxt0 [meta_ext] 1 THEN simp_tac ctxt1 1);
                   val eq2 = Simplifier.asm_full_rewrite ctxt2 (Thm.dest_equals_rhs (Thm.cprop_of eq1));
                 in SOME (Thm.transitive eq1 eq2) end
             | _ => NONE)
--- a/src/HOL/Statespace/state_space.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Statespace/state_space.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -204,7 +204,7 @@
       (Const (@{const_name Not}, _) $
         (Const (@{const_name HOL.eq}, _) $ (x as Free _) $ (y as Free _))) =>
       (case neq_x_y ctxt x y of
-        SOME neq => rtac neq i
+        SOME neq => resolve_tac ctxt [neq] i
       | NONE => no_tac)
   | _ => no_tac));
 
@@ -222,7 +222,7 @@
       let
         val distinct_thm = Proof_Context.get_thm ctxt dist_thm_name;
         val rule = DistinctTreeProver.distinct_implProver ctxt distinct_thm goal;
-      in rtac rule i end);
+      in resolve_tac ctxt [rule] i end);
 
     fun tac ctxt =
       Locale.intro_locales_tac true ctxt [] THEN ALLGOALS (solve_tac ctxt);
--- a/src/HOL/TLA/Memory/MemoryImplementation.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/TLA/Memory/MemoryImplementation.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -227,7 +227,7 @@
   val config_fast_solver = Attrib.setup_config_bool @{binding fast_solver} (K false);
   val fast_solver = mk_solver "fast_solver" (fn ctxt =>
     if Config.get ctxt config_fast_solver
-    then assume_tac ctxt ORELSE' (etac notE)
+    then assume_tac ctxt ORELSE' (eresolve_tac ctxt [notE])
     else K no_tac);
 \<close>
 
@@ -797,7 +797,7 @@
 ML \<open>
 fun split_idle_tac ctxt =
   SELECT_GOAL
-   (TRY (rtac @{thm actionI} 1) THEN
+   (TRY (resolve_tac ctxt @{thms actionI} 1) THEN
     Induct_Tacs.case_tac ctxt "(s,t) \<Turnstile> unchanged (e p, c p, r p, m p, rmhist!p)" [] NONE 1 THEN
     rewrite_goals_tac ctxt @{thms action_rews} THEN
     forward_tac ctxt [temp_use ctxt @{thm Step1_4_7}] 1 THEN
--- a/src/HOL/TLA/TLA.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/TLA/TLA.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -284,23 +284,24 @@
 lemma box_thin: "\<lbrakk> sigma \<Turnstile> \<box>F; PROP W \<rbrakk> \<Longrightarrow> PROP W" .
 
 ML \<open>
-fun merge_box_tac i =
-   REPEAT_DETERM (EVERY [etac @{thm box_conjE} i, atac i, etac @{thm box_thin} i])
+fun merge_box_tac ctxt i =
+   REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE} i, assume_tac ctxt i,
+    eresolve_tac ctxt @{thms box_thin} i])
 
 fun merge_temp_box_tac ctxt i =
-  REPEAT_DETERM (EVERY [etac @{thm box_conjE_temp} i, atac i,
+  REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_temp} i, assume_tac ctxt i,
     Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "behavior")] [] @{thm box_thin} i])
 
 fun merge_stp_box_tac ctxt i =
-  REPEAT_DETERM (EVERY [etac @{thm box_conjE_stp} i, atac i,
+  REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_stp} i, assume_tac ctxt i,
     Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "state")] [] @{thm box_thin} i])
 
 fun merge_act_box_tac ctxt i =
-  REPEAT_DETERM (EVERY [etac @{thm box_conjE_act} i, atac i,
+  REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_act} i, assume_tac ctxt i,
     Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "state * state")] [] @{thm box_thin} i])
 \<close>
 
-method_setup merge_box = \<open>Scan.succeed (K (SIMPLE_METHOD' merge_box_tac))\<close>
+method_setup merge_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_box_tac)\<close>
 method_setup merge_temp_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_temp_box_tac)\<close>
 method_setup merge_stp_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_stp_box_tac)\<close>
 method_setup merge_act_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_act_box_tac)\<close>
@@ -577,9 +578,9 @@
   SELECT_GOAL
     (EVERY
      [auto_tac ctxt,
-      TRY (merge_box_tac 1),
-      rtac (temp_use ctxt @{thm INV1}) 1, (* fail if the goal is not a box *)
-      TRYALL (etac @{thm Stable})]);
+      TRY (merge_box_tac ctxt 1),
+      resolve_tac ctxt [temp_use ctxt @{thm INV1}] 1, (* fail if the goal is not a box *)
+      TRYALL (eresolve_tac ctxt @{thms Stable})]);
 
 (* auto_inv_tac applies inv_tac and then tries to attack the subgoals
    in simple cases it may be able to handle goals like \<turnstile> MyProg \<longrightarrow> \<box>Inv.
--- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -975,9 +975,9 @@
             raise SKELETON (*FIXME or classify it as a Caboose: TRY (HEADGOAL atac) *)
           else
             case hd skel of
-                Assumed => TRY (HEADGOAL atac) THEN rest (tl skel) memo
-              | Caboose => TRY (HEADGOAL atac)
-              | Unconjoin => rtac @{thm conjI} 1 THEN rest (tl skel) memo
+                Assumed => TRY (HEADGOAL (assume_tac ctxt)) THEN rest (tl skel) memo
+              | Caboose => TRY (HEADGOAL (assume_tac ctxt))
+              | Unconjoin => resolve_tac ctxt @{thms conjI} 1 THEN rest (tl skel) memo
               | Split (split_node, solved_node, antes) =>
                   let
                     val split_fmla = node_info (#meta pannot) #fmla split_node
@@ -993,13 +993,13 @@
                     val split_thm =
                       simulate_split ctxt split_fmla minor_prems_assumps conclusion
                   in
-                    rtac split_thm 1 THEN rest (tl skel) memo
+                    resolve_tac ctxt [split_thm] 1 THEN rest (tl skel) memo
                   end
               | Step s =>
                   let
-                    val (tac, memo') = tac_and_memo s memo
+                    val (th, memo') = tac_and_memo s memo
                   in
-                    rtac tac 1 THEN rest (tl skel) memo'
+                    resolve_tac ctxt [th] 1 THEN rest (tl skel) memo'
                   end
               | Definition n =>
                   let
@@ -1008,7 +1008,7 @@
                           NONE => error ("Did not find definition: " ^ n)
                         | SOME binding => Global_Theory.get_thm thy (Binding.name_of binding)
                   in
-                    rtac def_thm 1 THEN rest (tl skel) memo
+                    resolve_tac ctxt [def_thm] 1 THEN rest (tl skel) memo
                   end
               | Axiom n =>
                   let
@@ -1017,7 +1017,7 @@
                           NONE => error ("Did not find axiom: " ^ n)
                         | SOME binding => Global_Theory.get_thm thy (Binding.name_of binding)
                   in
-                    rtac ax_thm 1 THEN rest (tl skel) memo
+                    resolve_tac ctxt [ax_thm] 1 THEN rest (tl skel) memo
                   end
               | _ => raise SKELETON
       in tactic st end
@@ -1065,7 +1065,7 @@
              let
                fun thm ctxt'' = prover_tac ctxt'' prob_name (hd skel |> stock_to_string)
                val reconstructed_inference = thm ctxt'
-               fun rec_inf_tac st = HEADGOAL (rtac (thm ctxt')) st
+               fun rec_inf_tac st = HEADGOAL (resolve_tac ctxt' [thm ctxt']) st
              in (reconstructed_inference,
                  rec_inf_tac)
              end)
@@ -1088,17 +1088,17 @@
                 (hd skel,
                  Thm.prop_of @{thm asm_rl}
                  |> SOME,
-                 SOME (@{thm asm_rl}, TRY (HEADGOAL atac))) :: rest memo ctxt
+                 SOME (@{thm asm_rl}, TRY (HEADGOAL (assume_tac ctxt)))) :: rest memo ctxt
             | Caboose =>
                 [(Caboose,
                   Thm.prop_of @{thm asm_rl}
                   |> SOME,
-                  SOME (@{thm asm_rl}, TRY (HEADGOAL atac)))]
+                  SOME (@{thm asm_rl}, TRY (HEADGOAL (assume_tac ctxt))))]
             | Unconjoin =>
                 (hd skel,
                  Thm.prop_of @{thm conjI}
                  |> SOME,
-                 SOME (@{thm conjI}, rtac @{thm conjI} 1)) :: rest memo ctxt
+                 SOME (@{thm conjI}, resolve_tac ctxt @{thms conjI} 1)) :: rest memo ctxt
             | Split (split_node, solved_node, antes) =>
                 let
                   val split_fmla = node_info (#meta pannot) #fmla split_node
@@ -1117,7 +1117,7 @@
                   (hd skel,
                    Thm.prop_of split_thm
                    |> SOME,
-                   SOME (split_thm, rtac split_thm 1)) :: rest memo ctxt
+                   SOME (split_thm, resolve_tac ctxt [split_thm] 1)) :: rest memo ctxt
                 end
             | Step node =>
                 let
@@ -1180,7 +1180,7 @@
                   (hd skel,
                    Thm.prop_of (def_thm thy)
                    |> SOME,
-                   SOME (def_thm thy, HEADGOAL (rtac (def_thm thy)))) :: rest memo ctxt
+                   SOME (def_thm thy, HEADGOAL (resolve_tac ctxt [def_thm thy]))) :: rest memo ctxt
                 end
             | Axiom n =>
                 let
@@ -1192,7 +1192,7 @@
                   (hd skel,
                    Thm.prop_of ax_thm
                    |> SOME,
-                   SOME (ax_thm, rtac ax_thm 1)) :: rest memo ctxt
+                   SOME (ax_thm, resolve_tac ctxt [ax_thm] 1)) :: rest memo ctxt
                 end
       end
 
@@ -1216,7 +1216,7 @@
          TPTP_Proof.is_inference_called "solved_all_splits"
           (the last_inference_info)
         then (@{thm asm_rl}, all_tac)
-        else (solved_all_splits, TRY (rtac solved_all_splits 1))
+        else (solved_all_splits, TRY (resolve_tac ctxt [solved_all_splits] 1))
       end
 in
   (*Build a tactic from a skeleton. This is naive because it uses the naive skeleton.
@@ -1226,8 +1226,8 @@
     let
       val thy = Proof_Context.theory_of ctxt
     in
-      rtac @{thm ccontr} 1
-      THEN dtac neg_eq_false 1
+      resolve_tac ctxt @{thms ccontr} 1
+      THEN dresolve_tac ctxt [neg_eq_false] 1
       THEN (sas_if_needed_tac ctxt prob_name |> #2)
       THEN skel_to_naive_tactic ctxt prover_tac prob_name
        (make_skeleton ctxt
@@ -1241,9 +1241,9 @@
       val thy = Proof_Context.theory_of ctxt
     in
       (Synth_step "ccontr", Thm.prop_of @{thm ccontr} |> SOME,
-       SOME (@{thm ccontr}, rtac @{thm ccontr} 1)) ::
+       SOME (@{thm ccontr}, resolve_tac ctxt @{thms ccontr} 1)) ::
       (Synth_step "neg_eq_false", Thm.prop_of neg_eq_false |> SOME,
-       SOME (neg_eq_false, dtac neg_eq_false 1)) ::
+       SOME (neg_eq_false, dresolve_tac ctxt [neg_eq_false] 1)) ::
       (Synth_step "sas_if_needed_tac", Thm.prop_of @{thm asm_rl} (*FIXME *) |> SOME,
        SOME (sas_if_needed_tac ctxt prob_name)) ::
       skel_to_naive_tactic_dbg prover_tac ctxt prob_name
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -282,13 +282,13 @@
          val thm = Goal.prove ctxt [] []
            (Logic.mk_implies (hyp_clause, new_hyp))
            (fn _ =>
-              (REPEAT_DETERM (HEADGOAL (rtac @{thm allI})))
+              (REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI})))
               THEN (REPEAT_DETERM
                     (HEADGOAL
                      (nominal_inst_parametermatch_tac ctxt @{thm allE})))
-              THEN HEADGOAL atac)
+              THEN HEADGOAL (assume_tac ctxt))
       in
-        dtac thm i st
+        dresolve_tac ctxt [thm] i st
       end
     end
 *}
@@ -305,8 +305,8 @@
   (TPTP_Reconstruct.remove_polarity true t; true)
   handle TPTP_Reconstruct.UNPOLARISED _ => false
 
-val polarise_subgoal_hyps =
-  COND' (SOME #> TERMPRED is_polarised (fn _ => true)) (K no_tac) (dtac @{thm polarise})
+fun polarise_subgoal_hyps ctxt =
+  COND' (SOME #> TERMPRED is_polarised (fn _ => true)) (K no_tac) (dresolve_tac ctxt @{thms polarise})
 *}
 
 lemma simp_meta [rule_format]:
@@ -336,10 +336,10 @@
 
 lemma solved_all_splits: "False = True \<Longrightarrow> False" by simp
 ML {*
-val solved_all_splits_tac =
-  TRY (etac @{thm conjE} 1)
-  THEN rtac @{thm solved_all_splits} 1
-  THEN atac 1
+fun solved_all_splits_tac ctxt =
+  TRY (eresolve_tac ctxt @{thms conjE} 1)
+  THEN resolve_tac ctxt @{thms solved_all_splits} 1
+  THEN assume_tac ctxt 1
 *}
 
 lemma lots_of_logic_expansions_meta [rule_format]:
@@ -433,10 +433,10 @@
 
     fun instantiate_tac vars =
       instantiate_vars ctxt vars
-      THEN (HEADGOAL atac)
+      THEN (HEADGOAL (assume_tac ctxt))
   in
     HEADGOAL (canonicalise_qtfr_order ctxt)
-    THEN (REPEAT_DETERM (HEADGOAL (rtac @{thm allI})))
+    THEN (REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI})))
     THEN REPEAT_DETERM (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE}))
     (*now only the variable to instantiate should be left*)
     THEN FIRST (map instantiate_tac ordered_instances)
@@ -470,8 +470,8 @@
   let
     val default_tac =
       (TRY o CHANGED o (rewrite_goal_tac ctxt @{thms prop_normalise}))
-      THEN' rtac @{thm notI}
-      THEN' (REPEAT_DETERM o etac @{thm conjE})
+      THEN' resolve_tac ctxt @{thms notI}
+      THEN' (REPEAT_DETERM o eresolve_tac ctxt @{thms conjE})
       THEN' (TRY o (expander_animal ctxt))
   in
     default_tac ORELSE' resolve_tac ctxt @{thms flip}
@@ -646,8 +646,6 @@
 ML {*
 fun forall_neg_tac candidate_consts ctxt i = fn st =>
   let
-    val thy = Proof_Context.theory_of ctxt
-
     val gls =
       Thm.prop_of st
       |> Logic.strip_horn
@@ -749,7 +747,7 @@
           else
             raise (NO_SKOLEM_DEF (def_name, bnd_name, conclusion)))
   in
-    rtac (Drule.export_without_context thm) i st
+    resolve_tac ctxt [Drule.export_without_context thm] i st
   end
   handle SKOLEM_DEF _ => no_tac st
 *}
@@ -960,12 +958,12 @@
 ML {*
 fun new_skolem_tac ctxt consts_candidates =
   let
-    fun tec thm =
-      dtac thm
+    fun tac thm =
+      dresolve_tac ctxt [thm]
       THEN' instantiate_skols ctxt consts_candidates
   in
     if null consts_candidates then K no_tac
-    else FIRST' (map tec @{thms lift_exists})
+    else FIRST' (map tac @{thms lift_exists})
   end
 *}
 
@@ -1142,7 +1140,7 @@
             those free variables into logical variables.*)
           |> Thm.forall_intr_frees
           |> Drule.export_without_context
-      in dtac rule i st end
+      in dresolve_tac ctxt [rule] i st end
       handle NO_GOALS => no_tac st
 
     fun closure tac =
@@ -1151,10 +1149,10 @@
       SOLVE o (tac THEN' (batter_tac ctxt ORELSE' assume_tac ctxt))
     val search_tac =
       ASAP
-        (rtac @{thm disjI1} APPEND' rtac @{thm disjI2})
+        (resolve_tac ctxt @{thms disjI1} APPEND' resolve_tac ctxt @{thms disjI2})
         (FIRST' (map closure
                   [dresolve_tac ctxt @{thms dec_commut_eq},
-                   dtac @{thm dec_commut_disj},
+                   dresolve_tac ctxt @{thms dec_commut_disj},
                    elim_tac]))
   in
     (CHANGED o search_tac) i st
@@ -1306,14 +1304,15 @@
               v
   \<lbrakk>A = True; B = False\<rbrakk> \<Longrightarrow> R
 *)
-fun weak_conj_tac drule =
-  dtac drule THEN' (REPEAT_DETERM o etac @{thm conjE})
+fun weak_conj_tac ctxt drule =
+  dresolve_tac ctxt [drule] THEN'
+  (REPEAT_DETERM o eresolve_tac ctxt @{thms conjE})
 *}
 
 ML {*
-val uncurry_lit_neg_tac =
-  dtac @{lemma "(A \<longrightarrow> B \<longrightarrow> C) = False \<Longrightarrow> (A & B \<longrightarrow> C) = False" by auto}
-  #> REPEAT_DETERM
+fun uncurry_lit_neg_tac ctxt =
+  REPEAT_DETERM o
+    dresolve_tac ctxt [@{lemma "(A \<longrightarrow> B \<longrightarrow> C) = False \<Longrightarrow> (A & B \<longrightarrow> C) = False" by auto}]
 *}
 
 ML {*
@@ -1326,10 +1325,10 @@
             let
               val rule = mk_standard_cnf ctxt kind arity;
             in
-              (weak_conj_tac rule THEN' atac) i st
+              (weak_conj_tac ctxt rule THEN' assume_tac ctxt) i st
             end
   in
-    (uncurry_lit_neg_tac
+    (uncurry_lit_neg_tac ctxt
      THEN' TPTP_Reconstruct_Library.reassociate_conjs_tac ctxt
      THEN' core_tactic) i st
   end
@@ -1473,13 +1472,13 @@
 (*use as elim rule to remove premises*)
 lemma insa_prems: "\<lbrakk>Q; P\<rbrakk> \<Longrightarrow> P" by auto
 ML {*
-fun cleanup_skolem_defs feats =
+fun cleanup_skolem_defs ctxt feats =
   let
     (*remove hypotheses from skolem defs,
      after testing that they look like skolem defs*)
     val dehypothesise_skolem_defs =
       COND' (SOME #> TERMPRED (fn _ => true) conc_is_skolem_def)
-        (REPEAT_DETERM o etac @{thm insa_prems})
+        (REPEAT_DETERM o eresolve_tac ctxt @{thms insa_prems})
         (K no_tac)
   in
     if can_feature (CleanUp [RemoveHypothesesFromSkolemDefs]) feats then
@@ -1497,11 +1496,11 @@
 
 ML {*
 (*given a goal state, indicates the skolem constants committed-to in it (i.e. appearing in LHS of a skolem definition)*)
-val which_skolem_concs_used = fn st =>
+fun which_skolem_concs_used ctxt = fn st =>
   let
     val feats = [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]]
     val scrubup_tac =
-      cleanup_skolem_defs feats
+      cleanup_skolem_defs ctxt feats
       THEN remove_duplicates_tac feats
   in
     scrubup_tac st
@@ -1547,7 +1546,7 @@
 (*predicate over the type of the leading quantified variable*)
 
 ML {*
-val extcnf_forall_special_pos_tac =
+fun extcnf_forall_special_pos_tac ctxt =
   let
     val bool =
       ["True", "False"]
@@ -1555,16 +1554,16 @@
     val bool_to_bool =
       ["% _ . True", "% _ . False", "% x . x", "Not"]
 
-    val tecs =
+    val tacs =
       map (fn t_s =>  (* FIXME proper context!? *)
        Rule_Insts.eres_inst_tac @{context} [((("x", 0), Position.none), t_s)] [] @{thm allE}
-       THEN' atac)
+       THEN' assume_tac ctxt)
   in
-    (TRY o etac @{thm forall_pos_lift})
-    THEN' (atac
+    (TRY o eresolve_tac ctxt @{thms forall_pos_lift})
+    THEN' (assume_tac ctxt
            ORELSE' FIRST'
             (*FIXME could check the type of the leading quantified variable, instead of trying everything*)
-            (tecs (bool @ bool_to_bool)))
+            (tacs (bool @ bool_to_bool)))
   end
 *}
 
@@ -1573,9 +1572,9 @@
 
 lemma efq: "[|A = True; A = False|] ==> R" by auto
 ML {*
-val efq_tac =
-  (etac @{thm efq} THEN' atac)
-  ORELSE' atac
+fun efq_tac ctxt =
+  (eresolve_tac ctxt @{thms efq} THEN' assume_tac ctxt)
+  ORELSE' assume_tac ctxt
 *}
 
 ML {*
@@ -1586,13 +1585,13 @@
       consisting of a Skolem definition*)
     fun extcnf_combined_tac' ctxt i = fn st =>
       let
-        val skolem_consts_used_so_far = which_skolem_concs_used st
+        val skolem_consts_used_so_far = which_skolem_concs_used ctxt st
         val consts_diff' = subtract (op =) skolem_consts_used_so_far consts_diff
 
         fun feat_to_tac feat =
           case feat of
-              Close_Branch => trace_tac' ctxt "mark: closer" efq_tac
-            | ConjI => trace_tac' ctxt "mark: conjI" (rtac @{thm conjI})
+              Close_Branch => trace_tac' ctxt "mark: closer" (efq_tac ctxt)
+            | ConjI => trace_tac' ctxt "mark: conjI" (resolve_tac ctxt @{thms conjI})
             | King_Cong => trace_tac' ctxt "mark: expander_animal" (expander_animal ctxt)
             | Break_Hypotheses => trace_tac' ctxt "mark: break_hypotheses" (break_hypotheses_tac ctxt)
             | RemoveRedundantQuantifications => K all_tac
@@ -1603,28 +1602,28 @@
                  (REPEAT_DETERM o remove_redundant_quantification_in_lit)
 *)
 
-            | Assumption => atac
+            | Assumption => assume_tac ctxt
 (*FIXME both Existential_Free and Existential_Var run same code*)
             | Existential_Free => trace_tac' ctxt "mark: forall_neg" (exists_tac ctxt feats consts_diff')
             | Existential_Var => trace_tac' ctxt "mark: forall_neg" (exists_tac ctxt feats consts_diff')
             | Universal => trace_tac' ctxt "mark: forall_pos" (forall_tac ctxt feats)
-            | Not_pos => trace_tac' ctxt "mark: not_pos" (dtac @{thm leo2_rules(9)})
-            | Not_neg => trace_tac' ctxt "mark: not_neg" (dtac @{thm leo2_rules(10)})
-            | Or_pos => trace_tac' ctxt "mark: or_pos" (dtac @{thm leo2_rules(5)}) (*could add (6) for negated conjunction*)
-            | Or_neg => trace_tac' ctxt "mark: or_neg" (dtac @{thm leo2_rules(7)})
+            | Not_pos => trace_tac' ctxt "mark: not_pos" (dresolve_tac ctxt @{thms leo2_rules(9)})
+            | Not_neg => trace_tac' ctxt "mark: not_neg" (dresolve_tac ctxt @{thms leo2_rules(10)})
+            | Or_pos => trace_tac' ctxt "mark: or_pos" (dresolve_tac ctxt @{thms leo2_rules(5)}) (*could add (6) for negated conjunction*)
+            | Or_neg => trace_tac' ctxt "mark: or_neg" (dresolve_tac ctxt @{thms leo2_rules(7)})
             | Equal_pos => trace_tac' ctxt "mark: equal_pos" (dresolve_tac ctxt (@{thms eq_pos_bool} @ [@{thm leo2_rules(3)}, @{thm eq_pos_func}]))
             | Equal_neg => trace_tac' ctxt "mark: equal_neg" (dresolve_tac ctxt [@{thm eq_neg_bool}, @{thm leo2_rules(4)}])
             | Donkey_Cong => trace_tac' ctxt "mark: donkey_cong" (simper_animal ctxt THEN' ex_expander_tac ctxt)
 
-            | Extuni_Bool2 => trace_tac' ctxt "mark: extuni_bool2" (dtac @{thm extuni_bool2})
-            | Extuni_Bool1 => trace_tac' ctxt "mark: extuni_bool1" (dtac @{thm extuni_bool1})
-            | Extuni_Bind => trace_tac' ctxt "mark: extuni_triv" (etac @{thm extuni_triv})
-            | Extuni_Triv => trace_tac' ctxt "mark: extuni_triv" (etac @{thm extuni_triv})
+            | Extuni_Bool2 => trace_tac' ctxt "mark: extuni_bool2" (dresolve_tac ctxt @{thms extuni_bool2})
+            | Extuni_Bool1 => trace_tac' ctxt "mark: extuni_bool1" (dresolve_tac ctxt @{thms extuni_bool1})
+            | Extuni_Bind => trace_tac' ctxt "mark: extuni_triv" (eresolve_tac ctxt @{thms extuni_triv})
+            | Extuni_Triv => trace_tac' ctxt "mark: extuni_triv" (eresolve_tac ctxt @{thms extuni_triv})
             | Extuni_Dec => trace_tac' ctxt "mark: extuni_dec_tac" (extuni_dec_tac ctxt)
-            | Extuni_FlexRigid => trace_tac' ctxt "mark: extuni_flex_rigid" (atac ORELSE' asm_full_simp_tac ctxt)
-            | Extuni_Func => trace_tac' ctxt "mark: extuni_func" (dtac @{thm extuni_func})
+            | Extuni_FlexRigid => trace_tac' ctxt "mark: extuni_flex_rigid" (assume_tac ctxt ORELSE' asm_full_simp_tac ctxt)
+            | Extuni_Func => trace_tac' ctxt "mark: extuni_func" (dresolve_tac ctxt @{thms extuni_func})
             | Polarity_switch => trace_tac' ctxt "mark: polarity_switch" (eresolve_tac ctxt @{thms polarity_switch})
-            | Forall_special_pos => trace_tac' ctxt "mark: dorall_special_pos" extcnf_forall_special_pos_tac
+            | Forall_special_pos => trace_tac' ctxt "mark: dorall_special_pos" (extcnf_forall_special_pos_tac ctxt)
 
         val core_tac =
           get_loop_feats feats
@@ -1668,8 +1667,8 @@
    @{const_name Pure.imp}]
 
 fun strip_qtfrs_tac ctxt =
-  REPEAT_DETERM (HEADGOAL (rtac @{thm allI}))
-  THEN REPEAT_DETERM (HEADGOAL (etac @{thm exE}))
+  REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI}))
+  THEN REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt @{thms exE}))
   THEN HEADGOAL (canonicalise_qtfr_order ctxt)
   THEN
     ((REPEAT (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE})))
@@ -1846,7 +1845,7 @@
       then we should be left with skolem definitions:
         absorb them as axioms into the theory.*)
     val cleanup =
-      cleanup_skolem_defs feats
+      cleanup_skolem_defs ctxt feats
       THEN remove_duplicates_tac feats
       THEN (if can_feature AbsorbSkolemDefs feats then
               ALLGOALS (absorb_skolem_def ctxt prob_name_opt)
@@ -1868,9 +1867,9 @@
         can be simply eliminated -- they're redundant*)
       (*FIXME instead of just using allE, instantiate to a silly
          term, to remove opportunities for unification.*)
-      THEN (REPEAT_DETERM (etac @{thm allE} 1))
+      THEN (REPEAT_DETERM (eresolve_tac ctxt @{thms allE} 1))
 
-      THEN (REPEAT_DETERM (rtac @{thm allI} 1))
+      THEN (REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1))
 
       THEN (if have_loop_feats then
               REPEAT (CHANGED
@@ -1914,32 +1913,33 @@
       discharged.
      *)
     val kill_meta_eqs_tac =
-      dtac @{thm un_meta_polarise}
-      THEN' rtac @{thm meta_polarise}
-      THEN' (REPEAT_DETERM o (etac @{thm conjE}))
-      THEN' (REPEAT_DETERM o (rtac @{thm conjI} ORELSE' atac))
+      dresolve_tac ctxt @{thms un_meta_polarise}
+      THEN' resolve_tac ctxt @{thms meta_polarise}
+      THEN' (REPEAT_DETERM o (eresolve_tac ctxt @{thms conjE}))
+      THEN' (REPEAT_DETERM o (resolve_tac ctxt @{thms conjI} ORELSE' assume_tac ctxt))
 
     val continue_reducing_tac =
-      rtac @{thm meta_eq_to_obj_eq} 1
+      resolve_tac ctxt @{thms meta_eq_to_obj_eq} 1
       THEN (REPEAT_DETERM (ex_expander_tac ctxt 1))
-      THEN TRY (polarise_subgoal_hyps 1) (*no need to REPEAT_DETERM here, since there should only be one hypothesis*)
-      THEN TRY (dtac @{thm eq_reflection} 1)
+      THEN TRY (polarise_subgoal_hyps ctxt 1) (*no need to REPEAT_DETERM here, since there should only be one hypothesis*)
+      THEN TRY (dresolve_tac ctxt @{thms eq_reflection} 1)
       THEN (TRY ((CHANGED o rewrite_goal_tac ctxt
               (@{thm expand_iff} :: @{thms simp_meta})) 1))
-      THEN HEADGOAL (rtac @{thm reflexive}
-                     ORELSE' atac
+      THEN HEADGOAL (resolve_tac ctxt @{thms reflexive}
+                     ORELSE' assume_tac ctxt
                      ORELSE' kill_meta_eqs_tac)
 
     val tactic =
-      (rtac @{thm polarise} 1 THEN atac 1)
+      (resolve_tac ctxt @{thms polarise} 1 THEN assume_tac ctxt 1)
       ORELSE
-        (REPEAT_DETERM (etac @{thm conjE} 1 THEN etac @{thm drop_first_hypothesis} 1)
+        (REPEAT_DETERM (eresolve_tac ctxt @{thms conjE} 1 THEN
+          eresolve_tac ctxt @{thms drop_first_hypothesis} 1)
          THEN PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion)
          THEN (REPEAT_DETERM (ex_expander_tac ctxt 1))
          THEN (TRY ((CHANGED o rewrite_goal_tac ctxt @{thms simp_meta}) 1))
          THEN PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion)
          THEN
-           (HEADGOAL atac
+           (HEADGOAL (assume_tac ctxt)
            ORELSE
             (unfold_tac ctxt depends_on_defs
              THEN IF_UNSOLVED continue_reducing_tac)))
@@ -2117,12 +2117,12 @@
         *)
     | "copy" =>
          HEADGOAL
-          (atac
+          (assume_tac ctxt
            ORELSE'
-              (rtac @{thm polarise}
-               THEN' atac))
+              (resolve_tac ctxt @{thms polarise}
+               THEN' assume_tac ctxt))
     | "polarity_switch" => nonfull_extcnf_combined_tac [Polarity_switch]
-    | "solved_all_splits" => solved_all_splits_tac
+    | "solved_all_splits" => solved_all_splits_tac ctxt
     | "extcnf_not_pos" => nonfull_extcnf_combined_tac [Not_pos]
     | "extcnf_forall_pos" => nonfull_extcnf_combined_tac [Universal]
     | "negate_conjecture" => fail ctxt "Should not handle negate_conjecture here"
@@ -2138,7 +2138,7 @@
     | "extuni_bool2" => nonfull_extcnf_combined_tac [Extuni_Bool2]
     | "extuni_bool1" => nonfull_extcnf_combined_tac [Extuni_Bool1]
     | "extuni_dec" =>
-        HEADGOAL atac
+        HEADGOAL (assume_tac ctxt)
         ORELSE nonfull_extcnf_combined_tac [Extuni_Dec]
     | "extuni_bind" => nonfull_extcnf_combined_tac [Extuni_Bind]
     | "extuni_triv" => nonfull_extcnf_combined_tac [Extuni_Triv]
@@ -2161,7 +2161,7 @@
     | "split_preprocessing" =>
          (REPEAT (HEADGOAL (split_simp_tac ctxt)))
          THEN TRY (PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion))
-         THEN HEADGOAL atac
+         THEN HEADGOAL (assume_tac ctxt)
 
     (*FIXME some of these could eventually be handled specially*)
     | "fac_restr" => default_tac
--- a/src/HOL/Tools/BNF/bnf_comp.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_comp.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -456,7 +456,7 @@
     val wits = map (fn t => fold absfree (Term.add_frees t []) t)
       (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits);
 
-    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+    fun wit_tac ctxt = mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf);
 
     val (bnf', lthy') =
       bnf_def Smart_Inline (K Dont_Note) true qualify tacs wit_tac (SOME (Ds @ killedAs))
@@ -547,7 +547,7 @@
 
     val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
 
-    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+    fun wit_tac ctxt = mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf);
 
     val (bnf', lthy') =
       bnf_def Smart_Inline (K Dont_Note) true qualify tacs wit_tac (SOME Ds) Binding.empty
@@ -629,7 +629,7 @@
 
     val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
 
-    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+    fun wit_tac ctxt = mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf);
 
     val (bnf', lthy') =
       bnf_def Smart_Inline (K Dont_Note) true qualify tacs wit_tac (SOME Ds) Binding.empty
@@ -883,8 +883,9 @@
           (AbsA $ Term.list_comb (t, map Bound (0 upto length I - 1))))
       (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
 
-    fun wit_tac ctxt = ALLGOALS (dtac ctxt (type_definition RS @{thm type_copy_wit})) THEN
-      mk_simple_wit_tac (wit_thms_of_bnf bnf);
+    fun wit_tac ctxt =
+      ALLGOALS (dtac ctxt (type_definition RS @{thm type_copy_wit})) THEN
+      mk_simple_wit_tac ctxt (wit_thms_of_bnf bnf);
 
     val (bnf', lthy') =
       bnf_def Hardly_Inline (user_policy Dont_Note) true qualify tacs wit_tac (SOME all_deads)
--- a/src/HOL/Tools/BNF/bnf_comp_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_comp_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -30,7 +30,7 @@
 
   val mk_le_rel_OO_tac: Proof.context -> thm -> thm -> thm list -> tactic
   val mk_simple_rel_OO_Grp_tac: Proof.context -> thm -> thm -> tactic
-  val mk_simple_wit_tac: thm list -> tactic
+  val mk_simple_wit_tac: Proof.context -> thm list -> tactic
   val mk_simplified_set_tac: Proof.context -> thm -> tactic
   val bd_ordIso_natLeq_tac: tactic
 end;
@@ -98,7 +98,7 @@
             rtac ctxt trans_o_apply, rtac ctxt (@{thm collect_def} RS arg_cong_Union),
             rtac ctxt @{thm UnionI}, rtac ctxt @{thm UN_I}, REPEAT_DETERM_N i o rtac ctxt @{thm insertI2},
             rtac ctxt @{thm insertI1}, rtac ctxt (o_apply RS equalityD2 RS set_mp),
-            etac ctxt @{thm imageI}, atac])
+            etac ctxt @{thm imageI}, assume_tac ctxt])
           comp_set_alts))
       map_cong0s) 1)
   end;
@@ -158,8 +158,8 @@
     unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN
     (REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN
      REPEAT_DETERM (CHANGED ((
-       (rtac ctxt conjI THEN' (atac ORELSE' rtac ctxt subset_UNIV)) ORELSE'
-       atac ORELSE'
+       (rtac ctxt conjI THEN' (assume_tac ctxt ORELSE' rtac ctxt subset_UNIV)) ORELSE'
+       assume_tac ctxt ORELSE'
        (rtac ctxt subset_UNIV)) 1)) ORELSE rtac ctxt subset_UNIV 1));
 
 val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def UN_insert UN_empty Un_empty_right
@@ -170,13 +170,13 @@
   ALLGOALS (dtac ctxt @{thm in_Union_o_assoc}) THEN
   unfold_thms_tac ctxt [collect_set_map] THEN
   unfold_thms_tac ctxt comp_wit_thms THEN
-  REPEAT_DETERM ((atac ORELSE'
+  REPEAT_DETERM ((assume_tac ctxt ORELSE'
     REPEAT_DETERM o eresolve_tac ctxt @{thms UnionE UnE} THEN'
     etac ctxt imageE THEN' TRY o dresolve_tac ctxt Gwit_thms THEN'
     (etac ctxt FalseE ORELSE'
     hyp_subst_tac ctxt THEN'
     dresolve_tac ctxt Fwit_thms THEN'
-    (etac ctxt FalseE ORELSE' atac))) 1);
+    (etac ctxt FalseE ORELSE' assume_tac ctxt))) 1);
 
 
 (* Kill operation *)
@@ -190,9 +190,9 @@
   REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN
   REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE'
     rtac ctxt conjI THEN' rtac ctxt subset_UNIV) 1)) THEN
-  (rtac ctxt subset_UNIV ORELSE' atac) 1 THEN
+  (rtac ctxt subset_UNIV ORELSE' assume_tac ctxt) 1 THEN
   REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN
-  REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' atac) 1))) ORELSE
+  REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' assume_tac ctxt) 1))) ORELSE
   ((rtac ctxt @{thm UNIV_eq_I} THEN' rtac ctxt CollectI) 1 THEN
     REPEAT_DETERM (TRY (rtac ctxt conjI 1) THEN rtac ctxt subset_UNIV 1));
 
@@ -207,11 +207,11 @@
 fun lift_in_alt_tac ctxt =
   ((rtac ctxt @{thm Collect_cong} THEN' rtac ctxt iffI) 1 THEN
   REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN
-  REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' atac) 1)) THEN
+  REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE' assume_tac ctxt) 1)) THEN
   REPEAT_DETERM (CHANGED (etac ctxt conjE 1)) THEN
   REPEAT_DETERM (CHANGED ((etac ctxt conjI ORELSE'
     rtac ctxt conjI THEN' rtac ctxt @{thm empty_subsetI}) 1)) THEN
-  (rtac ctxt @{thm empty_subsetI} ORELSE' atac) 1) ORELSE
+  (rtac ctxt @{thm empty_subsetI} ORELSE' assume_tac ctxt) 1) ORELSE
   ((rtac ctxt sym THEN' rtac ctxt @{thm UNIV_eq_I} THEN' rtac ctxt CollectI) 1 THEN
     REPEAT_DETERM (TRY (rtac ctxt conjI 1) THEN rtac ctxt @{thm empty_subsetI} 1));
 
@@ -232,7 +232,8 @@
 fun mk_simple_rel_OO_Grp_tac ctxt rel_OO_Grp in_alt_thm =
   rtac ctxt (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' eresolve0_tac (@{thm emptyE} :: wit_thms));
+fun mk_simple_wit_tac ctxt wit_thms =
+  ALLGOALS (assume_tac ctxt 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]};
--- a/src/HOL/Tools/BNF/bnf_def_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -53,15 +53,15 @@
 fun mk_map_comp comp = @{thm comp_eq_dest_lhs} OF [mk_sym comp];
 fun mk_map_cong_tac ctxt cong0 =
   (hyp_subst_tac ctxt THEN' rtac ctxt cong0 THEN'
-   REPEAT_DETERM o (dtac ctxt meta_spec THEN' etac ctxt meta_mp THEN' atac)) 1;
+   REPEAT_DETERM o (dtac ctxt meta_spec THEN' etac ctxt meta_mp THEN' assume_tac ctxt)) 1;
 fun mk_set_map set_map0 = set_map0 RS @{thm comp_eq_dest};
 fun mk_in_mono_tac ctxt n = if n = 0 then rtac ctxt subset_UNIV 1
   else (rtac ctxt subsetI THEN'
   rtac ctxt CollectI) 1 THEN
   REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN
   REPEAT_DETERM_N (n - 1)
-    ((rtac ctxt conjI THEN' etac ctxt subset_trans THEN' atac) 1) THEN
-  (etac ctxt subset_trans THEN' atac) 1;
+    ((rtac ctxt conjI THEN' etac ctxt subset_trans THEN' assume_tac ctxt) 1) THEN
+  (etac ctxt subset_trans THEN' assume_tac ctxt) 1;
 
 fun mk_inj_map_tac ctxt n map_id map_comp map_cong0 map_cong =
   let
@@ -70,7 +70,7 @@
   in
     HEADGOAL (rtac ctxt @{thm injI} THEN' etac ctxt (map_cong' RS box_equals) THEN'
       REPEAT_DETERM_N 2 o (rtac ctxt (box_equals OF [map_cong0', map_comp RS sym, map_id]) THEN'
-        REPEAT_DETERM_N n o atac))
+        REPEAT_DETERM_N n o assume_tac ctxt))
   end;
 
 fun mk_inj_map_strong_tac ctxt rel_eq rel_maps rel_mono_strong =
@@ -104,10 +104,12 @@
         REPEAT_DETERM o
           eresolve_tac ctxt [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
         hyp_subst_tac ctxt, rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0,
-        REPEAT_DETERM_N n o EVERY' [rtac ctxt @{thm Collect_split_Grp_eqD}, etac ctxt @{thm set_mp}, atac],
+        REPEAT_DETERM_N n o
+          EVERY' [rtac ctxt @{thm Collect_split_Grp_eqD}, etac ctxt @{thm set_mp}, assume_tac ctxt],
         rtac ctxt CollectI,
         CONJ_WRAP' (fn thm => EVERY' [rtac ctxt (thm RS ord_eq_le_trans),
-          rtac ctxt @{thm image_subsetI}, rtac ctxt @{thm Collect_split_Grp_inD}, etac ctxt @{thm set_mp}, atac])
+          rtac ctxt @{thm image_subsetI}, rtac ctxt @{thm Collect_split_Grp_inD},
+          etac ctxt @{thm set_mp}, assume_tac ctxt])
         set_maps,
         rtac ctxt @{thm predicate2I}, REPEAT_DETERM o eresolve_tac ctxt [@{thm GrpE}, exE, conjE],
         hyp_subst_tac ctxt,
@@ -120,7 +122,7 @@
             rtac ctxt CollectI,
             CONJ_WRAP' (fn thm =>
               EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt @{thm image_subsetI},
-                rtac ctxt @{thm convol_mem_GrpI}, etac ctxt set_mp, atac])
+                rtac ctxt @{thm convol_mem_GrpI}, etac ctxt set_mp, assume_tac ctxt])
             set_maps])
           @{thms fst_convol snd_convol} [map_id, refl])] 1
   end;
@@ -146,11 +148,11 @@
       unfold_thms_tac ctxt [rel_compp, rel_conversep, rel_Grp, @{thm vimage2p_Grp}] THEN
       TRYALL (EVERY' [rtac ctxt iffI, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm GrpI},
         resolve_tac ctxt [map_id, refl], rtac ctxt CollectI,
-        CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, rtac ctxt @{thm relcomppI}, atac,
+        CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks, rtac ctxt @{thm relcomppI}, assume_tac ctxt,
         rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI}, resolve_tac ctxt [map_id, refl], rtac ctxt CollectI,
         CONJ_WRAP' (K (rtac ctxt @{thm subset_UNIV})) ks,
         REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE conversepE GrpE},
-        dtac ctxt (trans OF [sym, map_id]), hyp_subst_tac ctxt, atac])
+        dtac ctxt (trans OF [sym, map_id]), hyp_subst_tac ctxt, assume_tac ctxt])
     end;
 
 fun mk_rel_mono_tac ctxt rel_OO_Grps in_mono =
@@ -194,7 +196,7 @@
     val n = length set_maps;
     fun in_tac nthO_in = rtac ctxt CollectI THEN'
         CONJ_WRAP' (fn thm => EVERY' [rtac ctxt (thm RS ord_eq_le_trans),
-          rtac ctxt @{thm image_subsetI}, rtac ctxt nthO_in, etac ctxt set_mp, atac]) set_maps;
+          rtac ctxt @{thm image_subsetI}, rtac ctxt nthO_in, etac ctxt set_mp, assume_tac ctxt]) set_maps;
     val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
       else rtac ctxt (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
         rtac ctxt (@{thm arg_cong2[of _ _ _ _ "op OO"]} OF (replicate 2 (hd rel_OO_Grps RS sym)) RSN
@@ -213,8 +215,8 @@
         rtac ctxt @{thm GrpI}, rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0,
         REPEAT_DETERM_N n o EVERY' [rtac ctxt trans, rtac ctxt o_apply,
           rtac ctxt ballE, rtac ctxt subst,
-          rtac ctxt @{thm csquare_def}, rtac ctxt @{thm csquare_fstOp_sndOp}, atac, etac ctxt notE,
-          etac ctxt set_mp, atac],
+          rtac ctxt @{thm csquare_def}, rtac ctxt @{thm csquare_fstOp_sndOp}, assume_tac ctxt,
+          etac ctxt notE, etac ctxt set_mp, assume_tac ctxt],
         in_tac @{thm fstOp_in},
         rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm GrpI},
         rtac ctxt trans, rtac ctxt map_comp, rtac ctxt map_cong0,
@@ -226,14 +228,14 @@
   end;
 
 fun mk_rel_mono_strong0_tac ctxt in_rel set_maps =
-  if null set_maps then atac 1
+  if null set_maps then assume_tac ctxt 1
   else
     unfold_tac ctxt [in_rel] THEN
     REPEAT_DETERM (eresolve_tac ctxt [exE, CollectE, conjE] 1) THEN
     hyp_subst_tac ctxt 1 THEN
     EVERY' [rtac ctxt exI, rtac ctxt @{thm conjI[OF CollectI conjI[OF refl refl]]},
       CONJ_WRAP' (fn thm =>
-        (etac ctxt (@{thm Collect_split_mono_strong} OF [thm, thm]) THEN' atac))
+        (etac ctxt (@{thm Collect_split_mono_strong} OF [thm, thm]) THEN' assume_tac ctxt))
       set_maps] 1;
 
 fun mk_rel_transfer_tac ctxt in_rel rel_map rel_mono_strong =
@@ -241,7 +243,7 @@
     fun last_tac iffD =
       HEADGOAL (etac ctxt rel_mono_strong) THEN
       REPEAT_DETERM (HEADGOAL (etac ctxt (@{thm predicate2_transferD} RS iffD) THEN'
-        REPEAT_DETERM o atac));
+        REPEAT_DETERM o assume_tac ctxt));
   in
     REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI)) THEN
     (HEADGOAL (hyp_subst_tac ctxt THEN' rtac ctxt refl) ORELSE
@@ -262,7 +264,8 @@
   in
     REPEAT_DETERM_N n (HEADGOAL (rtac ctxt rel_funI)) THEN
     unfold_thms_tac ctxt @{thms rel_fun_iff_leq_vimage2p} THEN
-    HEADGOAL (EVERY' [rtac ctxt @{thm order_trans}, rtac ctxt rel_mono, REPEAT_DETERM_N n o atac,
+    HEADGOAL (EVERY' [rtac ctxt @{thm order_trans}, rtac ctxt rel_mono,
+      REPEAT_DETERM_N n o assume_tac ctxt,
       rtac ctxt @{thm predicate2I}, dtac ctxt (in_rel RS iffD1),
       REPEAT_DETERM o eresolve_tac ctxt [exE, CollectE, conjE], hyp_subst_tac ctxt,
       rtac ctxt @{thm vimage2pI}, rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt conjI, in_tac,
@@ -316,12 +319,12 @@
         else REPEAT_DETERM_N (live - 2) o (etac ctxt conjE THEN' rotate_tac ~1) THEN' etac ctxt conjE,
         rtac ctxt (Drule.rotate_prems 1 @{thm image_eqI}), rtac ctxt @{thm SigmaI}, rtac ctxt @{thm UNIV_I},
         CONJ_WRAP_GEN' (rtac ctxt @{thm SigmaI})
-          (K (etac ctxt @{thm If_the_inv_into_in_Func} THEN' atac)) set_maps,
+          (K (etac ctxt @{thm If_the_inv_into_in_Func} THEN' assume_tac ctxt)) set_maps,
         rtac ctxt sym,
         rtac ctxt (Drule.rotate_prems 1
            ((@{thm box_equals} OF [map_cong0 OF replicate live @{thm If_the_inv_into_f_f},
              map_comp RS sym, map_id]) RSN (2, trans))),
-        REPEAT_DETERM_N (2 * live) o atac,
+        REPEAT_DETERM_N (2 * live) o assume_tac ctxt,
         REPEAT_DETERM_N live o rtac ctxt (@{thm prod.case} RS trans),
         rtac ctxt refl,
         rtac ctxt @{thm surj_imp_ordLeq}, rtac ctxt subsetI, rtac ctxt (Drule.rotate_prems 1 @{thm image_eqI}),
@@ -335,8 +338,10 @@
   end;
 
 fun mk_trivial_wit_tac ctxt wit_defs set_maps =
-  unfold_thms_tac ctxt wit_defs THEN HEADGOAL (EVERY' (map (fn thm =>
-    dtac ctxt (thm RS equalityD1 RS set_mp) THEN' etac ctxt imageE THEN' atac) set_maps)) THEN ALLGOALS atac;
+  unfold_thms_tac ctxt wit_defs THEN
+  HEADGOAL (EVERY' (map (fn thm =>
+    dtac ctxt (thm RS equalityD1 RS set_mp) THEN' etac ctxt imageE THEN' assume_tac ctxt) set_maps)) THEN
+  ALLGOALS (assume_tac ctxt);
 
 fun mk_set_transfer_tac ctxt in_rel set_maps =
   Goal.conjunction_tac 1 THEN
@@ -345,7 +350,7 @@
     @{thms exE conjE CollectE}))) THEN
   HEADGOAL (hyp_subst_tac ctxt THEN' rtac ctxt (@{thm iffD2[OF arg_cong2]} OF [set_map, set_map]) THEN'
     rtac ctxt @{thm rel_setI}) THEN
-  REPEAT (HEADGOAL (etac ctxt imageE THEN' dtac ctxt @{thm set_mp} THEN' atac THEN'
+  REPEAT (HEADGOAL (etac ctxt imageE THEN' dtac ctxt @{thm set_mp} THEN' assume_tac ctxt THEN'
     REPEAT_DETERM o (eresolve_tac ctxt @{thms CollectE case_prodE}) THEN' hyp_subst_tac ctxt THEN'
     rtac ctxt bexI THEN' etac ctxt @{thm subst_Pair[OF _ refl]} THEN' etac ctxt imageI))) set_maps);
 
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -113,39 +113,40 @@
     unfold_thms_tac ctxt cases THEN
     ALLGOALS (fn k => (select_prem_tac ctxt n (dtac ctxt asm_rl) k) k) THEN
     ALLGOALS (REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt rel_funD THEN'
-      (atac THEN' etac ctxt thin_rl ORELSE' rtac ctxt refl)) THEN' atac)
+      (assume_tac ctxt THEN' etac ctxt thin_rl ORELSE' rtac ctxt refl)) THEN' assume_tac ctxt)
   end;
 
 fun mk_ctr_transfer_tac ctxt rel_intros rel_eqs =
   HEADGOAL Goal.conjunction_tac THEN
   ALLGOALS (REPEAT o (resolve_tac ctxt (rel_funI :: rel_intros) THEN'
-    TRY o (REPEAT_DETERM1 o (atac ORELSE'
+    TRY o (REPEAT_DETERM1 o (assume_tac ctxt ORELSE'
       K (unfold_thms_tac ctxt rel_eqs) THEN' hyp_subst_tac ctxt THEN' rtac ctxt refl))));
 
 fun mk_disc_transfer_tac ctxt rel_sel exhaust_disc distinct_disc =
   let
     fun last_disc_tac iffD =
-      HEADGOAL (rtac ctxt (rotate_prems ~1 exhaust_disc) THEN' atac THEN'
-      REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt (rotate_prems 1 iffD) THEN' atac THEN'
-        rotate_tac ~1 THEN' etac ctxt (rotate_prems 1 notE) THEN' eresolve0_tac distinct_disc));
+      HEADGOAL (rtac ctxt (rotate_prems ~1 exhaust_disc) THEN' assume_tac ctxt THEN'
+      REPEAT_DETERM o (rotate_tac ~1 THEN' dtac ctxt (rotate_prems 1 iffD) THEN'
+        assume_tac ctxt THEN' rotate_tac ~1 THEN'
+        etac ctxt (rotate_prems 1 notE) THEN' eresolve0_tac distinct_disc));
   in
     HEADGOAL Goal.conjunction_tac THEN
     REPEAT_DETERM (HEADGOAL (rtac ctxt rel_funI THEN' dtac ctxt (rel_sel RS iffD1) THEN'
-      REPEAT_DETERM o (etac ctxt conjE) THEN' (atac ORELSE' rtac ctxt iffI))) THEN
+      REPEAT_DETERM o (etac ctxt conjE) THEN' (assume_tac ctxt ORELSE' rtac ctxt iffI))) THEN
     TRY (last_disc_tac iffD2) THEN TRY (last_disc_tac iffD1)
   end;
 
 fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
   unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac ctxt sumEN') THEN
   HEADGOAL (EVERY' (maps (fn k => [select_prem_tac ctxt n (rotate_tac 1) k,
-    REPEAT_DETERM o dtac ctxt meta_spec, etac ctxt meta_mp, atac]) (1 upto n)));
+    REPEAT_DETERM o dtac ctxt meta_spec, etac ctxt meta_mp, assume_tac ctxt]) (1 upto n)));
 
 fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
   HEADGOAL (rtac ctxt iffI THEN'
     EVERY' (@{map 3} (fn cTs => fn cx => fn th =>
       dtac ctxt (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
       SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
-      atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
+      assume_tac ctxt) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
 
 fun mk_half_distinct_tac ctxt ctor_inject abs_inject ctr_defs =
   unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN
@@ -204,7 +205,7 @@
                    [mk_rel_funDN 2 case_sum_transfer_eq, mk_rel_funDN 2 case_sum_transfer])) THEN
                  HEADGOAL (select_prem_tac ctxt total_n (dtac ctxt asm_rl) (acc + n)) THEN
                  HEADGOAL (SELECT_GOAL (HEADGOAL
-                   (REPEAT_DETERM o (atac ORELSE' resolve_tac ctxt
+                   (REPEAT_DETERM o (assume_tac ctxt ORELSE' resolve_tac ctxt
                        [mk_rel_funDN 1 case_prod_transfer_eq,
                         mk_rel_funDN 1 case_prod_transfer,
                         rel_funI]) THEN_ALL_NEW
@@ -255,12 +256,12 @@
       HEADGOAL (Inl_Inr_Pair_tac THEN'
         rtac ctxt (mk_rel_funDN 3 @{thm If_transfer}) THEN'
         select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN'
-        dtac ctxt rel_funD THEN' atac THEN' atac);
+        dtac ctxt rel_funD THEN' assume_tac ctxt THEN' assume_tac ctxt);
 
     fun mk_unfold_Inl_Inr_Pair_tac total pos =
       HEADGOAL (Inl_Inr_Pair_tac THEN'
         select_prem_tac ctxt total (dtac ctxt asm_rl) pos THEN'
-        dtac ctxt rel_funD THEN' atac THEN' atac);
+        dtac ctxt rel_funD THEN' assume_tac ctxt THEN' assume_tac ctxt);
 
     fun mk_unfold_arg_tac qs gs =
       EVERY (map (mk_unfold_If_tac num_pgs o prem_no_of) qs) THEN
@@ -307,7 +308,7 @@
 fun solve_prem_prem_tac ctxt =
   REPEAT o (eresolve_tac ctxt @{thms bexE rev_bexI} ORELSE' rtac ctxt @{thm rev_bexI[OF UNIV_I]} ORELSE'
     hyp_subst_tac ctxt ORELSE' resolve_tac ctxt @{thms disjI1 disjI2}) THEN'
-  (rtac ctxt refl ORELSE' atac ORELSE' rtac ctxt @{thm singletonI});
+  (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' rtac ctxt @{thm singletonI});
 
 fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
     pre_set_defs =
@@ -323,7 +324,7 @@
       REPEAT_DETERM_N m o (dtac ctxt meta_spec THEN' rotate_tac ~1)]) THEN
     EVERY [REPEAT_DETERM_N r
         (HEADGOAL (rotate_tac ~1 THEN' dtac ctxt meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
-      if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac,
+      if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL (assume_tac ctxt),
       mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
         pre_set_defs]
   end;
@@ -349,10 +350,10 @@
   SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
   SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: fp_abs_inverse :: abs_inverse :: dtor_ctor ::
     sels @ sumprod_thms_rel @ @{thms o_apply vimage2p_def})) THEN'
-  (atac ORELSE' REPEAT o etac ctxt conjE THEN'
+  (assume_tac ctxt ORELSE' REPEAT o etac ctxt conjE THEN'
      full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
      REPEAT o etac ctxt conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
-     REPEAT o (resolve_tac ctxt [refl, conjI] ORELSE' atac));
+     REPEAT o (resolve_tac ctxt [refl, conjI] ORELSE' assume_tac ctxt));
 
 fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
   let
@@ -367,8 +368,8 @@
     dtor_ctor exhaust ctr_defs discss selss =
   let val ks = 1 upto n in
     EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, select_prem_tac ctxt nn (dtac ctxt meta_spec) kk,
-        dtac ctxt meta_spec, dtac ctxt meta_mp, atac, rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 0),
-        hyp_subst_tac ctxt] @
+        dtac ctxt meta_spec, dtac ctxt meta_mp, assume_tac ctxt, rtac ctxt exhaust,
+        K (co_induct_inst_as_projs_tac ctxt 0), hyp_subst_tac ctxt] @
       @{map 4} (fn k => fn ctr_def => fn discs => fn sels =>
         EVERY' ([rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @
           map2 (fn k' => fn discs' =>
@@ -413,7 +414,7 @@
     True_implies_equals conj_imp_eq_imp_imp} @
     map (fn thm => thm RS eqFalseI) (distincts @ rel_distincts) @
     map (fn thm => thm RS eqTrueI) rel_injects) THEN
-  TRYALL (atac ORELSE' etac ctxt FalseE ORELSE'
+  TRYALL (assume_tac ctxt ORELSE' etac ctxt FalseE ORELSE'
     (REPEAT_DETERM o dtac ctxt @{thm meta_spec} THEN'
      TRY o filter_prems_tac ctxt
        (forall (curry (op <>) (HOLogic.mk_Trueprop @{term False})) o Logic.strip_imp_prems) THEN'
@@ -427,7 +428,7 @@
       (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW
          (dtac ctxt (rotate_prems ~1 (cterm_instantiate_pos [NONE, NONE, NONE, NONE, SOME ct]
             @{thm arg_cong2} RS iffD1)) THEN'
-          atac THEN' atac THEN' hyp_subst_tac ctxt THEN' dtac ctxt assm THEN'
+          assume_tac ctxt THEN' assume_tac ctxt THEN' hyp_subst_tac ctxt THEN' dtac ctxt assm THEN'
           REPEAT_DETERM o etac ctxt conjE))) 1 THEN
       unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @ sels @ simp_thms') THEN
       unfold_thms_tac ctxt (dtor_ctor :: rel_pre_def :: abs_inverse :: ctor_inject ::
@@ -435,7 +436,7 @@
         @{thms id_bnf_def rel_sum_simps rel_prod_apply vimage2p_def Inl_Inr_False
           iffD2[OF eq_False Inr_not_Inl] sum.inject prod.inject}) THEN
       REPEAT_DETERM (HEADGOAL ((REPEAT_DETERM o etac ctxt conjE) THEN' (REPEAT_DETERM o rtac ctxt conjI) THEN'
-        (rtac ctxt refl ORELSE' atac))))
+        (rtac ctxt refl ORELSE' assume_tac ctxt))))
     cts assms exhausts discss selss ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs
       abs_inverses);
 
@@ -445,13 +446,14 @@
       fn ctor_inject => fn rel_pre_list_def => fn Abs_inverse =>
         HEADGOAL (rtac ctxt exhaust THEN_ALL_NEW (rtac ctxt exhaust THEN_ALL_NEW
           (rtac ctxt (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2)
-            THEN' atac THEN' atac THEN' TRY o resolve_tac ctxt assms))) THEN
+            THEN' assume_tac ctxt THEN' assume_tac ctxt 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
         unfold_thms_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_apply Inl_Inr_False
           Inr_Inl_False  sum.inject prod.inject}) THEN
-        TRYALL (rtac ctxt refl ORELSE' etac ctxt FalseE ORELSE' (REPEAT_DETERM o etac ctxt conjE) THEN' atac))
+        TRYALL (rtac ctxt refl ORELSE' etac ctxt FalseE ORELSE'
+          (REPEAT_DETERM o etac ctxt conjE) THEN' assume_tac ctxt))
     cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses);
 
 fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts rel_eqs =
@@ -467,7 +469,7 @@
   TRYALL Goal.conjunction_tac THEN
   unfold_thms_tac ctxt (map (Drule.abs_def o Local_Defs.meta_rewrite_rule ctxt) sel_defs) THEN
   ALLGOALS (rtac ctxt (mk_rel_funDN n case_transfer) THEN_ALL_NEW
-    REPEAT_DETERM o (atac ORELSE' rtac ctxt rel_funI));
+    REPEAT_DETERM o (assume_tac ctxt ORELSE' rtac ctxt rel_funI));
 
 fun mk_set_sel_tac ctxt ct exhaust discs sels sets =
   TRYALL Goal.conjunction_tac THEN
@@ -480,7 +482,7 @@
     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 ctxt @{thm singletonI} ORELSE' atac));
+      (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt));
 
 fun mk_set_cases_tac ctxt ct assms exhaust sets =
   HEADGOAL (rtac ctxt (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN
@@ -488,13 +490,14 @@
   REPEAT_DETERM (HEADGOAL
     (eresolve_tac ctxt @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE'
      hyp_subst_tac ctxt ORELSE'
-     SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o atac)))));
+     SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac ctxt assms THEN' REPEAT_DETERM o assume_tac ctxt)))));
 
 fun mk_set_intros_tac ctxt sets =
   TRYALL Goal.conjunction_tac THEN unfold_thms_tac ctxt sets THEN
   TRYALL (REPEAT o
     (resolve_tac ctxt @{thms UnI1 UnI2} ORELSE'
-     eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN' (rtac ctxt @{thm singletonI} ORELSE' atac));
+     eresolve_tac ctxt @{thms UN_I UN_I[rotated]}) THEN'
+     (rtac ctxt @{thm singletonI} ORELSE' assume_tac ctxt));
 
 fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors
     Abs_pre_inverses =
@@ -502,7 +505,7 @@
     val assms_tac =
       let val assms' = map (unfold_thms ctxt (@{thm id_bnf_def} :: ctor_defs)) assms in
         fold (curry (op ORELSE')) (map (fn thm =>
-            funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' atac) (rtac ctxt thm)) assms')
+            funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' assume_tac ctxt) (rtac ctxt thm)) assms')
           (etac ctxt FalseE)
       end;
     val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
@@ -512,7 +515,7 @@
     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
+        THEN' assume_tac ctxt 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
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -53,33 +53,34 @@
 
 val exhaust_inst_as_projs_tac = PRIMITIVE oo exhaust_inst_as_projs;
 
-fun distinct_in_prems_tac distincts =
-  eresolve0_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
+fun distinct_in_prems_tac ctxt distincts =
+  eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN'
+  assume_tac ctxt;
 
 fun mk_primcorec_nchotomy_tac ctxt exhaust_discs =
   HEADGOAL (Method.insert_tac exhaust_discs THEN' clean_blast_tac ctxt);
 
 fun mk_primcorec_exhaust_tac ctxt frees n nchotomy =
   let val ks = 1 upto n in
-    HEADGOAL (atac ORELSE'
+    HEADGOAL (assume_tac ctxt ORELSE'
       cut_tac nchotomy THEN'
       K (exhaust_inst_as_projs_tac ctxt frees) THEN'
       EVERY' (map (fn k =>
           (if k < n then etac ctxt disjE else K all_tac) THEN'
-          REPEAT o (dtac ctxt meta_mp THEN' atac ORELSE'
-            etac ctxt conjE THEN' dtac ctxt meta_mp THEN' atac ORELSE'
-            atac))
+          REPEAT o (dtac ctxt meta_mp THEN' assume_tac ctxt ORELSE'
+            etac ctxt conjE THEN' dtac ctxt meta_mp THEN' assume_tac ctxt ORELSE'
+            assume_tac ctxt))
         ks))
   end;
 
 fun mk_primcorec_assumption_tac ctxt discIs =
   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 ctxt refl ORELSE' atac ORELSE' etac ctxt conjE ORELSE'
+    SOLVE (HEADGOAL (REPEAT o (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' etac ctxt conjE ORELSE'
     eresolve_tac ctxt falseEs ORELSE'
     resolve_tac ctxt @{thms TrueI conjI disjI1 disjI2} ORELSE'
-    dresolve_tac ctxt discIs THEN' atac ORELSE'
-    etac ctxt notE THEN' atac ORELSE'
+    dresolve_tac ctxt discIs THEN' assume_tac ctxt ORELSE'
+    etac ctxt notE THEN' assume_tac ctxt ORELSE'
     etac ctxt disjE))));
 
 fun ss_fst_snd_conv ctxt =
@@ -120,15 +121,16 @@
     EVERY' (map (fn [] => etac ctxt FalseE
         | fun_discs' as [fun_disc'] =>
           if eq_list Thm.eq_thm (fun_discs', fun_discs) then
-            REPEAT_DETERM o etac ctxt conjI THEN' (atac ORELSE' rtac ctxt TrueI)
+            REPEAT_DETERM o etac ctxt conjI THEN' (assume_tac ctxt ORELSE' rtac ctxt TrueI)
           else
             rtac ctxt FalseE THEN'
-            (rotate_tac 1 THEN' dtac ctxt fun_disc' THEN' REPEAT o atac ORELSE'
+            (rotate_tac 1 THEN' dtac ctxt fun_disc' THEN' REPEAT o assume_tac ctxt ORELSE'
              cut_tac fun_disc') THEN'
-            dresolve_tac ctxt distinct_discs THEN' etac ctxt notE THEN' atac)
+            dresolve_tac ctxt distinct_discs THEN' etac ctxt notE THEN' assume_tac ctxt)
       fun_discss) THEN'
     (etac ctxt FalseE ORELSE'
-     resolve_tac ctxt (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
+     resolve_tac ctxt
+      (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' assume_tac ctxt));
 
 fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms mapsx map_ident0s map_comps fun_sel k
     m excludesss =
@@ -139,8 +141,8 @@
     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 ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
-    etac ctxt notE THEN' atac ORELSE'
+    eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' assume_tac ctxt ORELSE'
+    etac ctxt notE THEN' assume_tac ctxt 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]]} @
        mapsx @ map_ident0s @ map_comps))) ORELSE'
@@ -150,7 +152,7 @@
 
 fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs =
   HEADGOAL (rtac ctxt ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
-    (the_default (K all_tac) (Option.map (rtac ctxt) disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN
+    (the_default (K all_tac) (Option.map (rtac ctxt) disc_fun_opt)) THEN' REPEAT_DETERM_N m o assume_tac ctxt) THEN
   unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac ctxt refl);
 
 fun inst_split_eq ctxt split =
@@ -174,13 +176,13 @@
     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 ctxt refl ORELSE' atac ORELSE'
+      (rtac ctxt refl ORELSE' assume_tac ctxt 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 ctxt discIs) THEN' etac ctxt notE THEN' atac)))))
+       distinct_in_prems_tac ctxt distincts ORELSE'
+       (TRY o dresolve_tac ctxt discIs) THEN' etac ctxt notE THEN' assume_tac ctxt)))))
   end;
 
 fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'});
@@ -207,11 +209,11 @@
 fun mk_primcorec_code_tac ctxt distincts splits raw =
   HEADGOAL (rtac ctxt raw ORELSE' rtac ctxt (raw RS trans) THEN'
     SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o
-    (rtac ctxt refl ORELSE' atac ORELSE'
+    (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE'
      resolve_tac ctxt split_connectI ORELSE'
      Splitter.split_tac ctxt (split_if :: splits) ORELSE'
-     distinct_in_prems_tac distincts ORELSE'
-     rtac ctxt sym THEN' atac ORELSE'
-     etac ctxt notE THEN' atac));
+     distinct_in_prems_tac ctxt distincts ORELSE'
+     rtac ctxt sym THEN' assume_tac ctxt ORELSE'
+     etac ctxt notE THEN' assume_tac ctxt));
 
 end;
--- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -116,15 +116,15 @@
 fun mk_coalg_set_tac ctxt coalg_def =
   dtac ctxt (coalg_def RS iffD1) 1 THEN
   REPEAT_DETERM (etac ctxt conjE 1) THEN
-  EVERY' [dtac ctxt rev_bspec, atac] 1 THEN
-  REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN atac 1;
+  EVERY' [dtac ctxt rev_bspec, assume_tac ctxt] 1 THEN
+  REPEAT_DETERM (eresolve0_tac [CollectE, conjE] 1) THEN assume_tac ctxt 1;
 
 fun mk_mor_elim_tac ctxt mor_def =
   (dtac ctxt (mor_def RS iffD1) THEN'
   REPEAT o etac ctxt conjE THEN'
   TRY o rtac ctxt @{thm image_subsetI} THEN'
   etac ctxt bspec THEN'
-  atac) 1;
+  assume_tac ctxt) 1;
 
 fun mk_mor_incl_tac ctxt mor_def map_ids =
   (rtac ctxt (mor_def RS iffD2) THEN'
@@ -137,10 +137,11 @@
 fun mk_mor_comp_tac ctxt mor_def mor_images morEs map_comp_ids =
   let
     fun fbetw_tac image = EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS @{thm ssubst_mem}), etac ctxt image,
-      etac ctxt image, atac];
+      etac ctxt image, assume_tac ctxt];
     fun mor_tac ((mor_image, morE), map_comp_id) =
       EVERY' [rtac ctxt ballI, stac ctxt o_apply, rtac ctxt trans, rtac ctxt (map_comp_id RS sym), rtac ctxt trans,
-        etac ctxt (morE RS arg_cong), atac, etac ctxt morE, etac ctxt mor_image, atac];
+        etac ctxt (morE RS arg_cong), assume_tac ctxt, etac ctxt morE,
+        etac ctxt mor_image, assume_tac ctxt];
   in
     (rtac ctxt (mor_def RS iffD2) THEN' rtac ctxt conjI THEN'
     CONJ_WRAP' fbetw_tac mor_images THEN'
@@ -171,7 +172,7 @@
 
 fun mk_set_Jset_incl_Jset_tac ctxt n rec_Suc i =
   EVERY' (map (rtac ctxt) [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2,
-    rec_Suc, UnI2, mk_UnIN n i] @ [etac ctxt @{thm UN_I}, atac]) 1;
+    rec_Suc, UnI2, mk_UnIN n i] @ [etac ctxt @{thm UN_I}, assume_tac ctxt]) 1;
 
 fun mk_col_minimal_tac ctxt m cts rec_0s rec_Sucs =
   EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct),
@@ -185,13 +186,13 @@
         else (rtac ctxt @{thm Un_least} THEN' Goal.assume_rule_tac ctxt),
         CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least}))
           (K (EVERY' [rtac ctxt @{thm UN_least}, REPEAT_DETERM o eresolve_tac ctxt [allE, conjE],
-            rtac ctxt subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
+            rtac ctxt subset_trans, assume_tac ctxt, Goal.assume_rule_tac ctxt])) rec_0s])
       rec_Sucs] 1;
 
 fun mk_Jset_minimal_tac ctxt n col_minimal =
   (CONJ_WRAP' (K (EVERY' [rtac ctxt @{thm UN_least}, rtac ctxt rev_mp, rtac ctxt col_minimal,
     EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac ctxt impI,
-    REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], atac])) (1 upto n)) 1
+    REPEAT_DETERM o eresolve_tac ctxt [allE, conjE], assume_tac ctxt])) (1 upto n)) 1
 
 fun mk_mor_col_tac ctxt m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss =
   EVERY' [rtac ctxt (Drule.instantiate' [] cts nat_induct),
@@ -204,14 +205,16 @@
           if m = 0 then K all_tac
           else EVERY' [rtac ctxt Un_cong, rtac ctxt @{thm box_equals},
             rtac ctxt (nth passive_set_map0s (j - 1) RS sym),
-            rtac ctxt trans_fun_cong_image_id_id_apply, etac ctxt (morE RS arg_cong), atac],
+            rtac ctxt trans_fun_cong_image_id_id_apply, etac ctxt (morE RS arg_cong),
+            assume_tac ctxt],
           CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 Un_cong))
             (fn (i, (set_map0, coalg_set)) =>
               EVERY' [rtac ctxt sym, rtac ctxt trans, rtac ctxt (refl RSN (2, @{thm SUP_cong})),
-                etac ctxt (morE RS sym RS arg_cong RS trans), atac, rtac ctxt set_map0,
+                etac ctxt (morE RS sym RS arg_cong RS trans), assume_tac ctxt, rtac ctxt set_map0,
                 rtac ctxt (@{thm UN_simps(10)} RS trans), rtac ctxt (refl RS @{thm SUP_cong}),
-                ftac ctxt coalg_set, atac, dtac ctxt set_mp, atac, rtac ctxt mp, rtac ctxt (mk_conjunctN n i),
-                REPEAT_DETERM o etac ctxt allE, atac, atac])
+                ftac ctxt coalg_set, assume_tac ctxt, dtac ctxt set_mp, assume_tac ctxt,
+                rtac ctxt mp, rtac ctxt (mk_conjunctN n i),
+                REPEAT_DETERM o etac ctxt allE, assume_tac ctxt, assume_tac ctxt])
             (rev ((1 upto n) ~~ (active_set_map0s ~~ coalg_sets)))])
       (rec_Sucs ~~ (morEs ~~ (map (chop m) set_map0ss ~~ map (drop m) coalg_setss)))] 1;
 
@@ -222,12 +225,12 @@
 
     fun mk_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) =
       EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, dtac ctxt (mk_conjunctN n i),
-        etac ctxt allE, etac ctxt allE, etac ctxt impE, atac, etac ctxt bexE,
+        etac ctxt allE, etac ctxt allE, etac ctxt impE, assume_tac ctxt, etac ctxt bexE,
         REPEAT_DETERM o eresolve0_tac [CollectE, conjE],
         rtac ctxt (in_rel RS iffD2), rtac ctxt exI, rtac ctxt (Drule.rotate_prems 1 conjI),
         CONJ_WRAP' (fn thm => EVERY' [rtac ctxt trans, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt map_cong0,
           REPEAT_DETERM_N m o rtac ctxt thm, REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong),
-          atac])
+          assume_tac ctxt])
         @{thms fst_diag_id snd_diag_id},
         rtac ctxt CollectI,
         CONJ_WRAP' (fn (i, thm) =>
@@ -241,24 +244,24 @@
 
     fun mk_only_if_tac ((((i, map_comp0), map_cong0), set_map0s), in_rel) =
       EVERY' [dtac ctxt (mk_conjunctN n i), rtac ctxt allI, rtac ctxt allI, rtac ctxt impI,
-        etac ctxt allE, etac ctxt allE, etac ctxt impE, atac,
+        etac ctxt allE, etac ctxt allE, etac ctxt impE, assume_tac ctxt,
         dtac ctxt (in_rel RS @{thm iffD1}),
         REPEAT_DETERM o eresolve0_tac ([CollectE, conjE, exE] @
           @{thms CollectE Collect_split_in_rel_leE}),
         rtac ctxt bexI, rtac ctxt conjI, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0,
         REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong),
         REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong),
-        atac, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0,
+        assume_tac ctxt, rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0,
         REPEAT_DETERM_N m o rtac ctxt (@{thm id_comp} RS fun_cong),
         REPEAT_DETERM_N n o rtac ctxt (@{thm comp_id} RS fun_cong),
         rtac ctxt trans, rtac ctxt map_cong0,
-        REPEAT_DETERM_N m o EVERY' [rtac ctxt @{thm Collect_splitD}, etac ctxt set_mp, atac],
+        REPEAT_DETERM_N m o EVERY' [rtac ctxt @{thm Collect_splitD}, etac ctxt set_mp, assume_tac ctxt],
         REPEAT_DETERM_N n o rtac ctxt refl,
-        atac, rtac ctxt CollectI,
+        assume_tac ctxt, rtac ctxt CollectI,
         CONJ_WRAP' (fn (i, thm) =>
           if i <= m then rtac ctxt subset_UNIV
           else EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt trans, rtac ctxt thm,
-            rtac ctxt trans_fun_cong_image_id_id_apply, atac])
+            rtac ctxt trans_fun_cong_image_id_id_apply, assume_tac ctxt])
         (1 upto (m + n) ~~ set_map0s)];
   in
     EVERY' [rtac ctxt (bis_def RS trans),
@@ -283,7 +286,7 @@
 fun mk_bis_O_tac ctxt m bis_rel rel_congs le_rel_OOs =
   EVERY' [rtac ctxt (bis_rel RS iffD2), REPEAT_DETERM o dtac ctxt (bis_rel RS iffD1),
     REPEAT_DETERM o etac ctxt conjE, rtac ctxt conjI,
-    CONJ_WRAP' (K (EVERY' [etac ctxt @{thm relcomp_subset_Sigma}, atac])) rel_congs,
+    CONJ_WRAP' (K (EVERY' [etac ctxt @{thm relcomp_subset_Sigma}, assume_tac ctxt])) rel_congs,
     CONJ_WRAP' (fn (rel_cong, le_rel_OO) =>
       EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI,
         rtac ctxt (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
@@ -294,7 +297,7 @@
         REPEAT_DETERM o dtac ctxt prod_injectD,
         etac ctxt conjE, hyp_subst_tac ctxt,
         REPEAT_DETERM o etac ctxt allE, rtac ctxt @{thm relcomppI},
-        etac ctxt mp, atac, etac ctxt mp, atac]) (rel_congs ~~ le_rel_OOs)] 1;
+        etac ctxt mp, assume_tac ctxt, etac ctxt mp, assume_tac ctxt]) (rel_congs ~~ le_rel_OOs)] 1;
 
 fun mk_bis_Gr_tac ctxt bis_rel rel_Grps mor_images morEs coalg_ins =
   unfold_thms_tac ctxt (bis_rel :: @{thm eq_alt} :: @{thm in_rel_Gr} :: rel_Grps) THEN
@@ -313,14 +316,15 @@
     unfold_thms_tac ctxt [bis_def] THEN
     EVERY' [rtac ctxt conjI,
       CONJ_WRAP' (fn i =>
-        EVERY' [rtac ctxt @{thm UN_least}, dtac ctxt bspec, atac,
+        EVERY' [rtac ctxt @{thm UN_least}, dtac ctxt bspec, assume_tac ctxt,
           dtac ctxt conjunct1, etac ctxt (mk_conjunctN n i)]) ks,
       CONJ_WRAP' (fn (i, in_mono) =>
-        EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, etac ctxt @{thm UN_E}, dtac ctxt bspec, atac,
+        EVERY' [rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, etac ctxt @{thm UN_E},
+          dtac ctxt bspec, assume_tac ctxt,
           dtac ctxt conjunct2, dtac ctxt (mk_conjunctN n i), etac ctxt allE, etac ctxt allE, dtac ctxt mp,
-          atac, etac ctxt bexE, rtac ctxt bexI, atac, rtac ctxt in_mono,
+          assume_tac ctxt, etac ctxt bexE, rtac ctxt bexI, assume_tac ctxt, rtac ctxt in_mono,
           REPEAT_DETERM_N n o etac ctxt @{thm SUP_upper2[OF _ subset_refl]},
-          atac]) (ks ~~ in_monos)] 1
+          assume_tac ctxt]) (ks ~~ in_monos)] 1
   end;
 
 fun mk_sbis_lsbis_tac ctxt lsbis_defs bis_Union bis_cong =
@@ -334,15 +338,15 @@
 
 fun mk_incl_lsbis_tac ctxt n i lsbis_def =
   EVERY' [rtac ctxt @{thm xt1(3)}, rtac ctxt lsbis_def, rtac ctxt @{thm SUP_upper2}, rtac ctxt CollectI,
-    REPEAT_DETERM_N n o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, atac, rtac ctxt equalityD2,
-    rtac ctxt (mk_nth_conv n i)] 1;
+    REPEAT_DETERM_N n o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, assume_tac ctxt,
+    rtac ctxt equalityD2, rtac ctxt (mk_nth_conv n i)] 1;
 
 fun mk_equiv_lsbis_tac ctxt sbis_lsbis lsbis_incl incl_lsbis bis_Id_on bis_converse bis_O =
   EVERY' [rtac ctxt (@{thm equiv_def} RS iffD2),
 
     rtac ctxt conjI, rtac ctxt (@{thm refl_on_def} RS iffD2),
     rtac ctxt conjI, rtac ctxt lsbis_incl, rtac ctxt ballI, rtac ctxt set_mp,
-    rtac ctxt incl_lsbis, rtac ctxt bis_Id_on, atac, etac ctxt @{thm Id_onI},
+    rtac ctxt incl_lsbis, rtac ctxt bis_Id_on, assume_tac ctxt, etac ctxt @{thm Id_onI},
 
     rtac ctxt conjI, rtac ctxt (@{thm sym_def} RS iffD2),
     rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt set_mp,
@@ -351,7 +355,7 @@
     rtac ctxt (@{thm trans_def} RS iffD2),
     rtac ctxt allI, rtac ctxt allI, rtac ctxt allI, rtac ctxt impI, rtac ctxt impI, rtac ctxt set_mp,
     rtac ctxt incl_lsbis, rtac ctxt bis_O, rtac ctxt sbis_lsbis, rtac ctxt sbis_lsbis,
-    etac ctxt @{thm relcompI}, atac] 1;
+    etac ctxt @{thm relcompI}, assume_tac ctxt] 1;
 
 fun mk_coalgT_tac ctxt m defs strT_defs set_map0ss =
   let
@@ -369,20 +373,20 @@
             etac ctxt equalityD1, etac ctxt CollectD,
           rtac ctxt ballI,
             CONJ_WRAP' (fn i => EVERY' [rtac ctxt ballI, etac ctxt CollectE, dtac ctxt @{thm ShiftD},
-              dtac ctxt bspec, etac ctxt thin_rl, atac, dtac ctxt (mk_conjunctN n i),
+              dtac ctxt bspec, etac ctxt thin_rl, assume_tac ctxt, dtac ctxt (mk_conjunctN n i),
               dtac ctxt bspec, rtac ctxt CollectI, etac ctxt @{thm set_mp[OF equalityD1[OF Succ_Shift]]},
               REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac ctxt exI,
               rtac ctxt conjI, rtac ctxt (@{thm shift_def} RS fun_cong RS trans),
-              rtac ctxt (@{thm append_Cons} RS sym RS arg_cong RS trans), atac,
+              rtac ctxt (@{thm append_Cons} RS sym RS arg_cong RS trans), assume_tac ctxt,
               CONJ_WRAP' (K (EVERY' [etac ctxt trans, rtac ctxt @{thm Collect_cong},
                 rtac ctxt @{thm eqset_imp_iff}, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm Succ_Shift},
                 rtac ctxt (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks,
-          dtac ctxt bspec, atac, dtac ctxt (mk_conjunctN n i), dtac ctxt bspec,
-          etac ctxt @{thm set_mp[OF equalityD1]}, atac,
+          dtac ctxt bspec, assume_tac ctxt, dtac ctxt (mk_conjunctN n i), dtac ctxt bspec,
+          etac ctxt @{thm set_mp[OF equalityD1]}, assume_tac ctxt,
           REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], rtac ctxt exI,
           rtac ctxt conjI, rtac ctxt (@{thm shift_def} RS fun_cong RS trans),
           etac ctxt (@{thm append_Nil} RS sym RS arg_cong RS trans),
-          REPEAT_DETERM_N m o (rtac ctxt conjI THEN' atac),
+          REPEAT_DETERM_N m o (rtac ctxt conjI THEN' assume_tac ctxt),
           CONJ_WRAP' (K (EVERY' [etac ctxt trans, rtac ctxt @{thm Collect_cong},
             rtac ctxt @{thm eqset_imp_iff}, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm Succ_Shift},
             rtac ctxt (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)];
@@ -408,12 +412,13 @@
             (fn i =>
               EVERY' [REPEAT_DETERM o eresolve_tac ctxt [CollectE, exE, conjE], hyp_subst_tac ctxt,
                 rtac ctxt trans, rtac ctxt @{thm length_Cons}, rtac ctxt @{thm arg_cong[of _ _ Suc]},
-                REPEAT_DETERM o etac ctxt allE, dtac ctxt (mk_conjunctN n i), etac ctxt mp, atac]) ks])
+                REPEAT_DETERM o etac ctxt allE, dtac ctxt (mk_conjunctN n i),
+                etac ctxt mp, assume_tac ctxt]) ks])
       Lev_Sucs] 1
   end;
 
 fun mk_length_Lev'_tac ctxt length_Lev =
-  EVERY' [ftac ctxt length_Lev, etac ctxt ssubst, atac] 1;
+  EVERY' [ftac ctxt length_Lev, etac ctxt ssubst, assume_tac ctxt] 1;
 
 fun mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss =
   let
@@ -434,7 +439,7 @@
           CONJ_WRAP' (fn i' => EVERY' [dtac ctxt (mk_conjunctN n i'), etac ctxt exE, rtac ctxt exI,
             rtac ctxt (@{thm append_Cons} RS arg_cong RS trans), rtac ctxt (rv_Cons RS trans),
             if n = 1 then K all_tac else etac ctxt (sum_case_cong_weak RS trans),
-            rtac ctxt (mk_sum_caseN n i RS trans), atac])
+            rtac ctxt (mk_sum_caseN n i RS trans), assume_tac ctxt])
           ks])
         rv_Conss)
       ks)] 1
@@ -474,12 +479,12 @@
                     EVERY' [rtac ctxt impI, rtac ctxt (Lev_Suc RS equalityD2 RS set_mp),
                       rtac ctxt @{thm ssubst_mem[OF append_Cons]}, rtac ctxt (mk_UnIN n i),
                       rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl,
-                      rtac ctxt conjI, atac, dtac ctxt (sym RS trans RS sym),
+                      rtac ctxt conjI, assume_tac ctxt, dtac ctxt (sym RS trans RS sym),
                       rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS trans),
                       etac ctxt (from_to_sbd RS arg_cong), REPEAT_DETERM o etac ctxt allE,
-                      dtac ctxt (mk_conjunctN n i), dtac ctxt mp, atac,
-                      dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, atac,
-                      dtac ctxt (mk_conjunctN n i''), etac ctxt mp, atac])
+                      dtac ctxt (mk_conjunctN n i), dtac ctxt mp, assume_tac ctxt,
+                      dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, assume_tac ctxt,
+                      dtac ctxt (mk_conjunctN n i''), etac ctxt mp, assume_tac ctxt])
                   ks)
                 ks])
           (rev (ks ~~ from_to_sbds))])
@@ -528,14 +533,14 @@
                   (if k = i
                   then EVERY' [dtac ctxt (mk_InN_inject n i),
                     dtac ctxt (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
-                    atac, atac, hyp_subst_tac ctxt] THEN'
+                    assume_tac ctxt, assume_tac ctxt, hyp_subst_tac ctxt] THEN'
                     CONJ_WRAP' (fn i'' =>
                       EVERY' [rtac ctxt impI, dtac ctxt (sym RS trans),
                         rtac ctxt (rv_Cons RS trans), rtac ctxt (mk_sum_caseN n i RS arg_cong RS trans),
                         etac ctxt (from_to_sbd RS arg_cong),
                         REPEAT_DETERM o etac ctxt allE,
-                        dtac ctxt (mk_conjunctN n i), dtac ctxt mp, atac,
-                        dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, atac,
+                        dtac ctxt (mk_conjunctN n i), dtac ctxt mp, assume_tac ctxt,
+                        dtac ctxt (mk_conjunctN n i'), dtac ctxt mp, assume_tac ctxt,
                         dtac ctxt (mk_conjunctN n i''), etac ctxt mp, etac ctxt sym])
                     ks
                   else etac ctxt (mk_InN_not_InM i k RS notE)))
@@ -572,13 +577,13 @@
               CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
                 EVERY' [rtac ctxt (set_map0 RS trans), rtac ctxt equalityI, rtac ctxt @{thm image_subsetI},
                   rtac ctxt CollectI, rtac ctxt @{thm SuccI}, rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, etac ctxt set_Lev,
-                  if n = 1 then rtac ctxt refl else atac, atac, rtac ctxt subsetI,
+                  if n = 1 then rtac ctxt refl else assume_tac ctxt, assume_tac ctxt, rtac ctxt subsetI,
                   REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}],
                   REPEAT_DETERM_N 4 o etac ctxt thin_rl,
                   rtac ctxt set_image_Lev,
-                  atac, dtac ctxt length_Lev, hyp_subst_tac ctxt, dtac ctxt length_Lev',
+                  assume_tac ctxt, dtac ctxt length_Lev, hyp_subst_tac ctxt, dtac ctxt length_Lev',
                   etac ctxt @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
-                  if n = 1 then rtac ctxt refl else atac])
+                  if n = 1 then rtac ctxt refl else assume_tac ctxt])
               (drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))])
           (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_map0ss ~~
             (set_Levss ~~ set_image_Levss))))),
@@ -591,7 +596,7 @@
             EVERY' [rtac ctxt (set_map0 RS trans), rtac ctxt equalityI, rtac ctxt @{thm image_subsetI},
               rtac ctxt CollectI, rtac ctxt @{thm SuccI}, rtac ctxt @{thm UN_I}, rtac ctxt UNIV_I, rtac ctxt set_Lev,
               rtac ctxt (Lev_0 RS equalityD2 RS set_mp), rtac ctxt @{thm singletonI}, rtac ctxt rv_Nil,
-              atac, rtac ctxt subsetI,
+              assume_tac ctxt, rtac ctxt subsetI,
               REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm SuccE}, @{thm UN_E}],
               rtac ctxt set_image_Lev, rtac ctxt (Lev_0 RS equalityD2 RS set_mp),
               rtac ctxt @{thm singletonI}, dtac ctxt length_Lev',
@@ -621,12 +626,12 @@
                 dtac ctxt list_inject_iffD1, etac ctxt conjE,
                 if i' = i'' then EVERY' [dtac ctxt (mk_InN_inject n i'),
                   dtac ctxt (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
-                  atac, atac, hyp_subst_tac ctxt, etac ctxt @{thm UN_I[OF UNIV_I]}]
+                  assume_tac ctxt, assume_tac ctxt, hyp_subst_tac ctxt, etac ctxt @{thm UN_I[OF UNIV_I]}]
                 else etac ctxt (mk_InN_not_InM i' i'' RS notE)])
             (rev ks),
             rtac ctxt @{thm UN_least}, rtac ctxt subsetI, rtac ctxt CollectI, rtac ctxt @{thm UN_I[OF UNIV_I]},
             rtac ctxt (Lev_Suc RS equalityD2 RS set_mp), rtac ctxt (mk_UnIN n i'), rtac ctxt CollectI,
-            REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, etac ctxt conjI, atac,
+            REPEAT_DETERM o rtac ctxt exI, rtac ctxt conjI, rtac ctxt refl, etac ctxt conjI, assume_tac ctxt,
             rtac ctxt trans, rtac ctxt @{thm shift_def}, rtac ctxt iffD2, rtac ctxt @{thm fun_eq_iff}, rtac ctxt allI,
             CONVERSION (Conv.top_conv
               (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
@@ -653,7 +658,7 @@
     etac ctxt (sym RS arg_cong RS trans), rtac ctxt (map_comp_id RS trans),
     rtac ctxt (map_cong0 RS trans), REPEAT_DETERM_N m o rtac ctxt refl,
     EVERY' (map (fn equiv_LSBIS =>
-      EVERY' [rtac ctxt @{thm equiv_proj}, rtac ctxt equiv_LSBIS, etac ctxt set_mp, atac])
+      EVERY' [rtac ctxt @{thm equiv_proj}, rtac ctxt equiv_LSBIS, etac ctxt set_mp, assume_tac ctxt])
     equiv_LSBISs), rtac ctxt sym, rtac ctxt (o_apply RS trans),
     etac ctxt (sym RS arg_cong RS trans), rtac ctxt map_comp_id] 1;
 
@@ -673,11 +678,12 @@
 fun mk_mor_T_final_tac ctxt mor_def congruent_str_finals equiv_LSBISs =
   EVERY' [rtac ctxt (mor_def RS iffD2), rtac ctxt conjI,
     CONJ_WRAP' (fn equiv_LSBIS =>
-      EVERY' [rtac ctxt ballI, rtac ctxt iffD2, rtac ctxt @{thm proj_in_iff}, rtac ctxt equiv_LSBIS, atac])
+      EVERY' [rtac ctxt ballI, rtac ctxt iffD2, rtac ctxt @{thm proj_in_iff},
+        rtac ctxt equiv_LSBIS, assume_tac ctxt])
     equiv_LSBISs,
     CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) =>
       EVERY' [rtac ctxt ballI, rtac ctxt sym, rtac ctxt trans, rtac ctxt @{thm univ_commute}, rtac ctxt equiv_LSBIS,
-        rtac ctxt congruent_str_final, atac, rtac ctxt o_apply])
+        rtac ctxt congruent_str_final, assume_tac ctxt, rtac ctxt o_apply])
     (equiv_LSBISs ~~ congruent_str_finals)] 1;
 
 fun mk_mor_Rep_tac ctxt defs Reps Abs_inverses coalg_final_setss map_comp_ids map_cong0Ls =
@@ -716,7 +722,8 @@
   in
     EVERY' [rtac ctxt rev_mp, ftac ctxt (bis_def RS iffD1),
       REPEAT_DETERM o etac ctxt conjE, rtac ctxt bis_cong, rtac ctxt bis_O, rtac ctxt bis_converse,
-      rtac ctxt bis_Gr, rtac ctxt tcoalg, rtac ctxt mor_Rep, rtac ctxt bis_O, atac, rtac ctxt bis_Gr, rtac ctxt tcoalg,
+      rtac ctxt bis_Gr, rtac ctxt tcoalg, rtac ctxt mor_Rep, rtac ctxt bis_O, assume_tac ctxt,
+      rtac ctxt bis_Gr, rtac ctxt tcoalg,
       rtac ctxt mor_Rep, REPEAT_DETERM_N n o etac ctxt @{thm relImage_Gr},
       rtac ctxt impI, rtac ctxt rev_mp, rtac ctxt bis_cong, rtac ctxt bis_O, rtac ctxt bis_Gr, rtac ctxt coalgT,
       rtac ctxt mor_T_final, rtac ctxt bis_O, rtac ctxt sbis_lsbis, rtac ctxt bis_converse, rtac ctxt bis_Gr, rtac ctxt coalgT,
@@ -728,7 +735,7 @@
           rtac ctxt ord_eq_le_trans, rtac ctxt @{thm sym[OF relImage_relInvImage]},
           rtac ctxt @{thm xt1(3)}, rtac ctxt @{thm Sigma_cong},
           rtac ctxt @{thm proj_image}, rtac ctxt @{thm proj_image}, rtac ctxt lsbis_incl,
-          rtac ctxt subset_trans, rtac ctxt @{thm relImage_mono}, rtac ctxt incl_lsbis, atac,
+          rtac ctxt subset_trans, rtac ctxt @{thm relImage_mono}, rtac ctxt incl_lsbis, assume_tac ctxt,
           rtac ctxt @{thm relImage_proj}, rtac ctxt equiv_LSBIS, rtac ctxt @{thm relInvImage_Id_on},
           rtac ctxt Rep_inject])
       (Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1
@@ -787,7 +794,7 @@
       EVERY' (@{map 3} (fn i => fn set_Jset => fn set_Jset_Jsets =>
         EVERY' [rtac ctxt subsetI, rtac ctxt @{thm UnI2}, rtac ctxt (mk_UnIN n i), etac ctxt @{thm UN_I},
           etac ctxt UnE, etac ctxt set_Jset, REPEAT_DETERM_N (n - 1) o etac ctxt UnE,
-          EVERY' (map (fn thm => EVERY' [etac ctxt @{thm UN_E}, etac ctxt thm, atac]) set_Jset_Jsets)])
+          EVERY' (map (fn thm => EVERY' [etac ctxt @{thm UN_E}, etac ctxt thm, assume_tac ctxt]) set_Jset_Jsets)])
       (1 upto n) set_Jsets set_Jset_Jsetss)] 1;
 
 fun mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets =
@@ -840,7 +847,7 @@
              rtac ctxt CollectI, etac ctxt CollectE,
              REPEAT_DETERM o etac ctxt conjE,
              CONJ_WRAP' (fn set_Jset_Jset =>
-               EVERY' [rtac ctxt ballI, etac ctxt bspec, etac ctxt set_Jset_Jset, atac]) set_Jset_Jsets,
+               EVERY' [rtac ctxt ballI, etac ctxt bspec, etac ctxt set_Jset_Jset, assume_tac ctxt]) set_Jset_Jsets,
              rtac ctxt (conjI OF [refl, refl])])
            (drop m set_map0s ~~ set_Jset_Jsetss)])
         (map (fn th => th RS trans) map_comps ~~ map (fn th => th RS trans) dtor_maps ~~
@@ -912,7 +919,7 @@
 
 fun mk_wit_tac ctxt n dtor_ctors dtor_set wit coind_wits =
   ALLGOALS (TRY o (eresolve_tac ctxt coind_wits THEN' rtac ctxt refl)) THEN
-  REPEAT_DETERM (atac 1 ORELSE
+  REPEAT_DETERM (assume_tac ctxt 1 ORELSE
     EVERY' [dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt dtor_set,
     K (unfold_thms_tac ctxt dtor_ctors),
     REPEAT_DETERM_N n o etac ctxt UnE,
@@ -966,14 +973,14 @@
         CONJ_WRAP' (fn (in_Jrel, (set_map0, dtor_set_set_incls)) =>
           EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt set_map0, rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI,
             rtac ctxt @{thm case_prodI}, rtac ctxt (in_Jrel RS iffD2), rtac ctxt exI, rtac ctxt conjI, rtac ctxt CollectI,
-            CONJ_WRAP' (fn thm => etac ctxt (thm RS @{thm subset_trans}) THEN' atac) dtor_set_set_incls,
+            CONJ_WRAP' (fn thm => etac ctxt (thm RS @{thm subset_trans}) THEN' assume_tac ctxt) dtor_set_set_incls,
             rtac ctxt conjI, rtac ctxt refl, rtac ctxt refl])
         (in_Jrels ~~ (active_set_map0s ~~ dtor_set_set_inclss)),
         CONJ_WRAP' (fn conv =>
           EVERY' [rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0,
           REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]},
           REPEAT_DETERM_N n o EVERY' (map (rtac ctxt) [trans, o_apply, conv]),
-          rtac ctxt trans, rtac ctxt sym, rtac ctxt dtor_map, rtac ctxt (dtor_inject RS iffD2), atac])
+          rtac ctxt trans, rtac ctxt sym, rtac ctxt dtor_map, rtac ctxt (dtor_inject RS iffD2), assume_tac ctxt])
         @{thms fst_conv snd_conv}];
     val only_if_tac =
       EVERY' [dtac ctxt (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE],
@@ -981,7 +988,7 @@
         CONJ_WRAP' (fn (dtor_set, passive_set_map0) =>
           EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt dtor_set, rtac ctxt @{thm Un_least},
             rtac ctxt ord_eq_le_trans, rtac ctxt @{thm box_equals}, rtac ctxt passive_set_map0,
-            rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans_fun_cong_image_id_id_apply, atac,
+            rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans_fun_cong_image_id_id_apply, assume_tac ctxt,
             CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least}))
               (fn (active_set_map0, in_Jrel) => EVERY' [rtac ctxt ord_eq_le_trans,
                 rtac ctxt @{thm SUP_cong[OF _ refl]}, rtac ctxt @{thm box_equals[OF _ _ refl]},
@@ -993,20 +1000,20 @@
                 hyp_subst_tac ctxt,
                 dtac ctxt (in_Jrel RS iffD1),
                 dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE,
-                REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], atac])
+                REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], assume_tac ctxt])
             (rev (active_set_map0s ~~ in_Jrels))])
         (dtor_sets ~~ passive_set_map0s),
         rtac ctxt conjI,
         REPEAT_DETERM_N 2 o EVERY'[rtac ctxt (dtor_inject RS iffD1), rtac ctxt trans, rtac ctxt dtor_map,
           rtac ctxt @{thm box_equals}, rtac ctxt map_comp0, rtac ctxt (dtor_ctor RS sym RS arg_cong), rtac ctxt trans,
           rtac ctxt map_cong0, REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]},
-          EVERY' (map (fn in_Jrel => EVERY' [rtac ctxt trans, rtac ctxt o_apply, dtac ctxt set_rev_mp, atac,
+          EVERY' (map (fn in_Jrel => EVERY' [rtac ctxt trans, rtac ctxt o_apply, dtac ctxt set_rev_mp, assume_tac ctxt,
             dtac ctxt @{thm ssubst_mem[OF pair_collapse]},
             REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE ::
               @{thms case_prodE iffD1[OF prod.inject, elim_format]}),
             hyp_subst_tac ctxt, dtac ctxt (in_Jrel RS iffD1),
-            dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, atac]) in_Jrels),
-          atac]]
+            dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt]) in_Jrels),
+          assume_tac ctxt]]
   in
     EVERY' [rtac ctxt iffI, if_tac, only_if_tac] 1
   end;
@@ -1023,7 +1030,7 @@
           fn dtor_unfold => fn dtor_map => fn in_rel =>
         EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI, in_rel RS iffD2],
           REPEAT_DETERM o eresolve_tac ctxt [exE, conjE],
-          select_prem_tac ctxt (length ks) (dtac ctxt @{thm spec2}) i, dtac ctxt mp, atac,
+          select_prem_tac ctxt (length ks) (dtac ctxt @{thm spec2}) i, dtac ctxt mp, assume_tac ctxt,
           REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], hyp_subst_tac ctxt,
           rtac ctxt exI, rtac ctxt (Drule.rotate_prems 1 conjI), rtac ctxt conjI,
           rtac ctxt (map_comp0 RS trans), rtac ctxt (dtor_map RS trans RS sym),
@@ -1039,8 +1046,8 @@
           CONJ_WRAP' (fn set_map0 =>
             EVERY' [rtac ctxt (set_map0 RS ord_eq_le_trans),
               rtac ctxt @{thm image_subsetI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI},
-              FIRST' [rtac ctxt refl, EVERY'[rtac ctxt exI, rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, atac,
-                rtac ctxt (@{thm surjective_pairing} RS arg_cong)]]])
+              FIRST' [rtac ctxt refl, EVERY'[rtac ctxt exI, rtac ctxt conjI, etac ctxt Collect_splitD_set_mp,
+                assume_tac ctxt, rtac ctxt (@{thm surjective_pairing} RS arg_cong)]]])
           set_map0s])
       ks map_comp0s map_congs map_arg_congs set_map0ss dtor_unfolds dtor_maps in_rels)] 1
   end;
@@ -1053,7 +1060,7 @@
     rtac ctxt set_induct 1 THEN
     EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i =>
       EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt conjE,
-        select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, atac,
+        select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, assume_tac ctxt,
         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)),
@@ -1062,11 +1069,12 @@
     EVERY' (@{map 3} (fn unfold => fn set_map0s => fn i =>
       EVERY' (map (fn set_map0 =>
         EVERY' [REPEAT_DETERM o resolve_tac ctxt [allI, impI], etac ctxt conjE,
-        select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, atac,
+        select_prem_tac ctxt n (dtac ctxt @{thm spec2}) i, dtac ctxt mp, assume_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 ctxt imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp],
-        rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, atac, rtac ctxt (@{thm surjective_pairing} RS arg_cong)])
+        rtac ctxt conjI, etac ctxt Collect_splitD_set_mp, assume_tac ctxt,
+        rtac ctxt (@{thm surjective_pairing} RS arg_cong)])
       (drop m set_map0s)))
     unfolds set_map0ss ks) 1
   end;
@@ -1082,16 +1090,18 @@
         if null helper_inds then rtac ctxt UNIV_I
         else rtac ctxt CollectI THEN'
           CONJ_WRAP' (fn helper_ind =>
-            EVERY' [rtac ctxt (helper_ind RS rev_mp), REPEAT_DETERM_N n o atac,
+            EVERY' [rtac ctxt (helper_ind RS rev_mp), REPEAT_DETERM_N n o assume_tac ctxt,
               REPEAT_DETERM_N n o etac ctxt thin_rl, rtac ctxt impI,
               REPEAT_DETERM o resolve_tac ctxt [subsetI, CollectI, @{thm iffD2[OF split_beta]}],
-              dtac ctxt bspec, atac, REPEAT_DETERM o eresolve_tac ctxt [allE, mp, conjE],
+              dtac ctxt bspec, assume_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [allE, mp, conjE],
               etac ctxt (refl RSN (2, conjI))])
           helper_inds,
         rtac ctxt conjI,
-        rtac ctxt (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac ctxt thin_rl,
+        rtac ctxt (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o assume_tac ctxt,
+        REPEAT_DETERM_N n o etac ctxt thin_rl,
         rtac ctxt impI, etac ctxt mp, rtac ctxt exI, etac ctxt (refl RSN (2, conjI)),
-        rtac ctxt (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac ctxt thin_rl,
+        rtac ctxt (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o assume_tac ctxt,
+        REPEAT_DETERM_N n o etac ctxt thin_rl,
         rtac ctxt impI, etac ctxt mp, rtac ctxt exI, etac ctxt (refl RSN (2, conjI))])
     (in_Jrels ~~ (helper_indss ~~ (helper_coind1s ~~ helper_coind2s))) 1
   end;
--- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -92,7 +92,7 @@
 
 fun mk_alg_set_tac ctxt alg_def =
   EVERY' [dtac ctxt (alg_def RS iffD1), REPEAT_DETERM o etac ctxt conjE, etac ctxt bspec, rtac ctxt CollectI,
-   REPEAT_DETERM o (rtac ctxt (subset_UNIV RS conjI) ORELSE' etac ctxt conjI), atac] 1;
+   REPEAT_DETERM o (rtac ctxt (subset_UNIV RS conjI) ORELSE' etac ctxt conjI), assume_tac ctxt] 1;
 
 fun mk_alg_not_empty_tac ctxt alg_set alg_sets wits =
   (EVERY' [rtac ctxt notI, hyp_subst_tac ctxt, forward_tac ctxt [alg_set]] THEN'
@@ -100,7 +100,7 @@
     [EVERY' [rtac ctxt @{thm subset_emptyI}, eresolve_tac ctxt wits],
     EVERY' [rtac ctxt subsetI, rtac ctxt FalseE, eresolve_tac ctxt wits],
     EVERY' [rtac ctxt subsetI, dresolve_tac ctxt wits, hyp_subst_tac ctxt,
-      FIRST' (map (fn thm => rtac ctxt thm THEN' atac) alg_sets)]] THEN'
+      FIRST' (map (fn thm => rtac ctxt thm THEN' assume_tac ctxt) alg_sets)]] THEN'
   etac ctxt @{thm emptyE}) 1;
 
 fun mk_mor_elim_tac ctxt mor_def =
@@ -108,7 +108,7 @@
   REPEAT o etac ctxt conjE THEN'
   TRY o rtac ctxt @{thm image_subsetI} THEN'
   etac ctxt bspec THEN'
-  atac) 1;
+  assume_tac ctxt) 1;
 
 fun mk_mor_incl_tac ctxt mor_def map_ids =
   (rtac ctxt (mor_def RS iffD2) THEN'
@@ -121,15 +121,16 @@
 fun mk_mor_comp_tac ctxt mor_def set_maps map_comp_ids =
   let
     val fbetw_tac =
-      EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS @{thm ssubst_mem}), etac ctxt bspec, etac ctxt bspec, atac];
+      EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS @{thm ssubst_mem}),
+        etac ctxt bspec, etac ctxt bspec, assume_tac ctxt];
     fun mor_tac (set_map, map_comp_id) =
       EVERY' [rtac ctxt ballI, rtac ctxt (o_apply RS trans), rtac ctxt trans,
-        rtac ctxt trans, dtac ctxt rev_bspec, atac, etac ctxt arg_cong,
+        rtac ctxt trans, dtac ctxt rev_bspec, assume_tac ctxt, etac ctxt arg_cong,
          REPEAT o eresolve0_tac [CollectE, conjE], etac ctxt bspec, rtac ctxt CollectI] THEN'
       CONJ_WRAP' (fn thm =>
         FIRST' [rtac ctxt subset_UNIV,
           (EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt thm, rtac ctxt @{thm image_subsetI},
-            etac ctxt bspec, etac ctxt set_mp, atac])]) set_map THEN'
+            etac ctxt bspec, etac ctxt set_mp, assume_tac ctxt])]) set_map THEN'
       rtac ctxt (map_comp_id RS arg_cong);
   in
     (dtac ctxt (mor_def RS iffD1) THEN' dtac ctxt (mor_def RS iffD1) THEN' rtac ctxt (mor_def RS iffD2) THEN'
@@ -195,11 +196,12 @@
     CONJ_WRAP' (fn i =>
       EVERY' [etac ctxt bspec, REPEAT_DETERM_N i o rtac ctxt @{thm insertI2}, rtac ctxt @{thm insertI1}])
       (0 upto n - 1),
-    atac] 1;
+    assume_tac ctxt] 1;
 
 fun mk_min_algs_tac ctxt worel in_congs =
   let
-    val minG_tac = EVERY' [rtac ctxt @{thm SUP_cong}, rtac ctxt refl, dtac ctxt bspec, atac, etac ctxt arg_cong];
+    val minG_tac = EVERY' [rtac ctxt @{thm SUP_cong}, rtac ctxt refl, dtac ctxt bspec,
+      assume_tac ctxt, etac ctxt arg_cong];
     fun minH_tac thm =
       EVERY' [rtac ctxt Un_cong, minG_tac, rtac ctxt @{thm image_cong}, rtac ctxt thm,
         REPEAT_DETERM_N (length in_congs) o minG_tac, rtac ctxt refl];
@@ -212,8 +214,9 @@
 
 fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [rtac ctxt relChainD, rtac ctxt allI, rtac ctxt allI, rtac ctxt impI,
   rtac ctxt @{thm case_split}, rtac ctxt @{thm xt1(3)}, rtac ctxt min_algs, etac ctxt @{thm FieldI2}, rtac ctxt subsetI,
-  rtac ctxt UnI1, rtac ctxt @{thm UN_I}, etac ctxt @{thm underS_I}, atac, atac, rtac ctxt equalityD1,
-  dtac ctxt @{thm notnotD}, hyp_subst_tac ctxt, rtac ctxt refl] 1;
+  rtac ctxt UnI1, rtac ctxt @{thm UN_I}, etac ctxt @{thm underS_I}, assume_tac ctxt,
+  assume_tac ctxt, rtac ctxt equalityD1, dtac ctxt @{thm notnotD},
+  hyp_subst_tac ctxt, rtac ctxt refl] 1;
 
 fun mk_min_algs_card_of_tac ctxt cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
   suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite =
@@ -236,8 +239,10 @@
 
     val minG_tac = EVERY' [rtac ctxt @{thm UNION_Cinfinite_bound}, rtac ctxt @{thm ordLess_imp_ordLeq},
       rtac ctxt @{thm ordLess_transitive}, rtac ctxt @{thm card_of_underS}, rtac ctxt suc_Card_order,
-      atac, rtac ctxt suc_Asuc, rtac ctxt ballI, etac ctxt allE, dtac ctxt mp, etac ctxt @{thm underS_E},
-      dtac ctxt mp, etac ctxt @{thm underS_Field}, REPEAT o etac ctxt conjE, atac, rtac ctxt Asuc_Cinfinite]
+      assume_tac ctxt, rtac ctxt suc_Asuc, rtac ctxt ballI, etac ctxt allE,
+      dtac ctxt mp, etac ctxt @{thm underS_E},
+      dtac ctxt mp, etac ctxt @{thm underS_Field},
+      REPEAT o etac ctxt conjE, assume_tac ctxt, rtac ctxt Asuc_Cinfinite]
 
     fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac ctxt @{thm ordIso_ordLeq_trans},
       rtac ctxt @{thm card_of_ordIso_subst}, etac ctxt min_alg, rtac ctxt @{thm Un_Cinfinite_bound},
@@ -265,8 +270,9 @@
     val induct = worel RS
       Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
 
-    val minG_tac = EVERY' [rtac ctxt @{thm UN_least}, etac ctxt allE, dtac ctxt mp, etac ctxt @{thm underS_E},
-      dtac ctxt mp, etac ctxt @{thm underS_Field}, REPEAT_DETERM o etac ctxt conjE, atac];
+    val minG_tac =
+      EVERY' [rtac ctxt @{thm UN_least}, etac ctxt allE, dtac ctxt mp, etac ctxt @{thm underS_E},
+        dtac ctxt mp, etac ctxt @{thm underS_Field}, REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt];
 
     fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ctxt ord_eq_le_trans, etac ctxt min_alg,
       rtac ctxt @{thm Un_least}, minG_tac, rtac ctxt @{thm image_subsetI},
@@ -288,15 +294,17 @@
     fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
       EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve0_tac [CollectE, conjE],
         EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac ctxt bexE,
-        rtac ctxt bd_limit, REPEAT_DETERM_N (n - 1) o etac ctxt conjI, atac,
+        rtac ctxt bd_limit, REPEAT_DETERM_N (n - 1) o etac ctxt conjI, assume_tac ctxt,
         rtac ctxt (min_alg_def RS @{thm set_mp[OF equalityD2]}),
-        rtac ctxt @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac ctxt thin_rl, atac, rtac ctxt set_mp,
-        rtac ctxt equalityD2, rtac ctxt min_alg, atac, rtac ctxt UnI2, rtac ctxt @{thm image_eqI}, rtac ctxt refl,
+        rtac ctxt @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac ctxt thin_rl,
+        assume_tac ctxt, rtac ctxt set_mp,
+        rtac ctxt equalityD2, rtac ctxt min_alg, assume_tac ctxt, rtac ctxt UnI2,
+        rtac ctxt @{thm image_eqI}, rtac ctxt refl,
         rtac ctxt CollectI, REPEAT_DETERM_N m o dtac ctxt asm_rl, REPEAT_DETERM_N n o etac ctxt thin_rl,
         REPEAT_DETERM o etac ctxt conjE,
-        CONJ_WRAP' (K (FIRST' [atac,
+        CONJ_WRAP' (K (FIRST' [assume_tac ctxt,
           EVERY' [etac ctxt subset_trans, rtac ctxt subsetI, rtac ctxt @{thm UN_I},
-            etac ctxt @{thm underS_I}, atac, atac]]))
+            etac ctxt @{thm underS_I}, assume_tac ctxt, assume_tac ctxt]]))
           set_bds];
   in
     (rtac ctxt (alg_def RS iffD2) THEN'
@@ -307,24 +315,27 @@
   EVERY' [rtac ctxt @{thm ordIso_ordLeq_trans}, rtac ctxt (min_alg_def RS @{thm card_of_ordIso_subst}),
     rtac ctxt @{thm UNION_Cinfinite_bound}, rtac ctxt @{thm ordIso_ordLeq_trans},
     rtac ctxt @{thm card_of_Field_ordIso}, rtac ctxt suc_Card_order, rtac ctxt @{thm ordLess_imp_ordLeq},
-    rtac ctxt suc_Asuc, rtac ctxt ballI, dtac ctxt rev_mp, rtac ctxt card_of,  REPEAT_DETERM o etac ctxt conjE, atac,
-    rtac ctxt Asuc_Cinfinite] 1;
+    rtac ctxt suc_Asuc, rtac ctxt ballI, dtac ctxt rev_mp, rtac ctxt card_of,
+    REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt, rtac ctxt Asuc_Cinfinite] 1;
 
 fun mk_least_min_alg_tac ctxt min_alg_def least =
-  EVERY' [rtac ctxt (min_alg_def RS ord_eq_le_trans), rtac ctxt @{thm UN_least}, dtac ctxt least, dtac ctxt mp, atac,
-    REPEAT_DETERM o etac ctxt conjE, atac] 1;
+  EVERY' [rtac ctxt (min_alg_def RS ord_eq_le_trans), rtac ctxt @{thm UN_least},
+    dtac ctxt least, dtac ctxt mp, assume_tac ctxt,
+    REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt] 1;
 
 fun mk_alg_select_tac ctxt Abs_inverse =
   EVERY' [rtac ctxt 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;
+  unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN assume_tac ctxt 1;
 
 fun mk_mor_select_tac ctxt mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select alg_sets
     set_maps str_init_defs =
   let
     val n = length alg_sets;
     val fbetw_tac =
-      CONJ_WRAP' (K (EVERY' [rtac ctxt ballI, etac ctxt rev_bspec, etac ctxt CollectE, atac])) alg_sets;
+      CONJ_WRAP'
+        (K (EVERY' [rtac ctxt ballI, etac ctxt rev_bspec,
+          etac ctxt CollectE, assume_tac ctxt])) alg_sets;
     val mor_tac =
       CONJ_WRAP' (fn thm => EVERY' [rtac ctxt ballI, rtac ctxt thm]) str_init_defs;
     fun alg_epi_tac ((alg_set, str_init_def), set_map) =
@@ -332,7 +343,7 @@
         rtac ctxt ballI, ftac ctxt (alg_select RS bspec), rtac ctxt (str_init_def RS @{thm ssubst_mem}),
         etac ctxt alg_set, REPEAT_DETERM o EVERY' [rtac ctxt ord_eq_le_trans, resolve0_tac set_map,
           rtac ctxt subset_trans, etac ctxt @{thm image_mono}, rtac ctxt @{thm image_Collect_subsetI}, etac ctxt bspec,
-          atac]];
+          assume_tac ctxt]];
   in
     EVERY' [rtac ctxt mor_cong, REPEAT_DETERM_N n o (rtac ctxt sym THEN' rtac ctxt @{thm comp_id}),
       rtac ctxt (Thm.permute_prems 0 1 mor_comp), etac ctxt (Thm.permute_prems 0 1 mor_comp),
@@ -345,10 +356,11 @@
     val n = length card_of_min_algs;
   in
     EVERY' [Method.insert_tac (map (fn thm => thm RS @{thm ex_bij_betw}) card_of_min_algs),
-      REPEAT_DETERM o etac ctxt exE, rtac ctxt rev_mp, rtac ctxt copy, REPEAT_DETERM_N n o atac,
+      REPEAT_DETERM o etac ctxt exE, rtac ctxt rev_mp, rtac ctxt copy,
+      REPEAT_DETERM_N n o assume_tac ctxt,
       rtac ctxt impI, REPEAT_DETERM o eresolve0_tac [exE, conjE], REPEAT_DETERM_N n o rtac ctxt exI,
       rtac ctxt mor_comp, rtac ctxt mor_Rep, rtac ctxt mor_select, rtac ctxt CollectI, REPEAT_DETERM o rtac ctxt exI,
-      rtac ctxt conjI, rtac ctxt refl, atac,
+      rtac ctxt conjI, rtac ctxt refl, assume_tac ctxt,
       SELECT_GOAL (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)),
       etac ctxt mor_comp, rtac ctxt mor_incl, REPEAT_DETERM_N n o rtac ctxt subset_UNIV] 1
   end;
@@ -361,11 +373,11 @@
 
     fun mor_tac morE in_mono = EVERY' [etac ctxt morE, rtac ctxt set_mp, rtac ctxt in_mono,
       REPEAT_DETERM_N n o rtac ctxt @{thm Collect_restrict}, rtac ctxt CollectI,
-      REPEAT_DETERM_N (m + n) o (TRY o rtac ctxt conjI THEN' atac)];
+      REPEAT_DETERM_N (m + n) o (TRY o rtac ctxt conjI THEN' assume_tac ctxt)];
     fun cong_tac ct map_cong0 = EVERY'
       [rtac ctxt (map_cong0 RS cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong),
       REPEAT_DETERM_N m o rtac ctxt refl,
-      REPEAT_DETERM_N n o (etac ctxt @{thm prop_restrict} THEN' atac)];
+      REPEAT_DETERM_N n o (etac ctxt @{thm prop_restrict} THEN' assume_tac ctxt)];
 
     fun mk_alg_tac (ct, (alg_set, (in_mono, (morE, map_cong0)))) =
       EVERY' [rtac ctxt ballI, rtac ctxt CollectI,
@@ -391,9 +403,10 @@
       REPEAT_DETERM o eresolve0_tac [CollectE, conjE], rtac ctxt conjI, rtac ctxt (alg_min_alg RS alg_set),
       REPEAT_DETERM_N n o (etac ctxt subset_trans THEN' rtac ctxt @{thm Collect_restrict}),
       rtac ctxt mp, etac ctxt bspec, rtac ctxt CollectI,
-      REPEAT_DETERM_N m o (rtac ctxt conjI THEN' atac),
+      REPEAT_DETERM_N m o (rtac ctxt conjI THEN' assume_tac ctxt),
       CONJ_WRAP' (K (etac ctxt subset_trans THEN' rtac ctxt @{thm Collect_restrict})) alg_sets,
-      CONJ_WRAP' (K (rtac ctxt ballI THEN' etac ctxt @{thm prop_restrict} THEN' atac)) alg_sets];
+      CONJ_WRAP' (K (rtac ctxt ballI THEN' etac ctxt @{thm prop_restrict} THEN' assume_tac ctxt))
+        alg_sets];
 
     fun mk_induct_tac least_min_alg =
       rtac ctxt ballI THEN' etac ctxt @{thm prop_restrict} THEN' rtac ctxt least_min_alg THEN'
@@ -422,7 +435,8 @@
       EVERY' [rtac ctxt ballI, REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE],
         rtac ctxt (thm RS (cterm_instantiate_pos [NONE, NONE, SOME ct] arg_cong) RS sym),
         EVERY' (map (fn Abs_inverse =>
-          EVERY' [rtac ctxt (o_apply RS trans RS ballI), etac ctxt (set_mp RS Abs_inverse), atac])
+          EVERY' [rtac ctxt (o_apply RS trans RS ballI), etac ctxt (set_mp RS Abs_inverse),
+            assume_tac ctxt])
         Abs_inverses)])
     (cts ~~ map2 mk_trans map_comp_ids map_congLs)] 1;;
 
@@ -436,8 +450,8 @@
     fun mk_unique type_def =
       EVERY' [rtac ctxt @{thm surj_fun_eq}, rtac ctxt (type_def RS @{thm type_definition.Abs_image}),
         rtac ctxt ballI, resolve0_tac init_unique_mors,
-        EVERY' (map (fn thm => atac ORELSE' rtac ctxt thm) Reps),
-        rtac ctxt mor_comp, rtac ctxt mor_Abs, atac,
+        EVERY' (map (fn thm => assume_tac ctxt ORELSE' rtac ctxt thm) Reps),
+        rtac ctxt mor_comp, rtac ctxt mor_Abs, assume_tac ctxt,
         rtac ctxt mor_comp, rtac ctxt mor_Abs, rtac ctxt mor_fold];
   in
     CONJ_WRAP' mk_unique type_defs 1
@@ -469,13 +483,14 @@
     fun mk_IH_tac Rep_inv Abs_inv set_map =
       DETERM o EVERY' [dtac ctxt meta_mp, rtac ctxt (Rep_inv RS arg_cong RS iffD1), etac ctxt bspec,
         dtac ctxt set_rev_mp, rtac ctxt equalityD1, rtac ctxt set_map, etac ctxt imageE,
-        hyp_subst_tac ctxt, rtac ctxt (Abs_inv RS @{thm ssubst_mem}), etac ctxt set_mp, atac, atac];
+        hyp_subst_tac ctxt, rtac ctxt (Abs_inv RS @{thm ssubst_mem}), etac ctxt set_mp,
+        assume_tac ctxt, assume_tac ctxt];
 
     fun mk_closed_tac (k, (morE, set_maps)) =
       EVERY' [select_prem_tac ctxt n (dtac ctxt asm_rl) k, rtac ctxt ballI, rtac ctxt impI,
-        rtac ctxt (mor_Abs RS morE RS arg_cong RS iffD2), atac,
+        rtac ctxt (mor_Abs RS morE RS arg_cong RS iffD2), assume_tac ctxt,
         REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], dtac ctxt @{thm meta_spec},
-        EVERY' (@{map 3} mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), atac];
+        EVERY' (@{map 3} mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), assume_tac ctxt];
 
     fun mk_induct_tac (Rep, Rep_inv) =
       EVERY' [rtac ctxt (Rep_inv RS arg_cong RS iffD1), etac ctxt (Rep RSN (2, bspec))];
@@ -495,13 +510,14 @@
         select_prem_tac ctxt n (dtac ctxt @{thm meta_spec2}) i,
         REPEAT_DETERM_N n o
           EVERY' [dtac ctxt meta_mp THEN_ALL_NEW Goal.norm_hhf_tac ctxt,
-            REPEAT_DETERM o dtac ctxt @{thm meta_spec}, etac ctxt (spec RS meta_mp), atac],
-        atac];
+            REPEAT_DETERM o dtac ctxt @{thm meta_spec}, etac ctxt (spec RS meta_mp),
+            assume_tac ctxt],
+        assume_tac ctxt];
   in
     EVERY' [rtac ctxt rev_mp, rtac ctxt (Drule.instantiate' cTs cts ctor_induct),
       EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac ctxt impI,
       REPEAT_DETERM o eresolve_tac ctxt [conjE, allE],
-      CONJ_WRAP' (K atac) ks] 1
+      CONJ_WRAP' (K (assume_tac ctxt)) ks] 1
   end;
 
 fun mk_map_tac ctxt m n foldx map_comp_id map_cong0 =
@@ -515,7 +531,7 @@
 fun mk_ctor_map_unique_tac ctxt fold_unique sym_map_comps =
   rtac ctxt fold_unique 1 THEN
   unfold_thms_tac ctxt (sym_map_comps @ @{thms comp_assoc id_comp comp_id}) THEN
-  ALLGOALS atac;
+  ALLGOALS (assume_tac ctxt);
 
 fun mk_set_tac ctxt foldx = EVERY' [rtac ctxt @{thm ext}, rtac ctxt trans, rtac ctxt o_apply,
   rtac ctxt trans, rtac ctxt foldx, rtac ctxt sym, rtac ctxt o_apply] 1;
@@ -592,8 +608,8 @@
     EVERY' [rtac ctxt impI, etac ctxt (nchotomy RS @{thm nchotomy_relcomppE}),
       REPEAT_DETERM_N 2 o dtac ctxt (Irel RS iffD1), rtac ctxt (Irel RS iffD2),
       rtac ctxt rel_mono, rtac ctxt (le_rel_OO RS @{thm predicate2D}),
-      rtac ctxt @{thm relcomppI}, atac, atac,
-      REPEAT_DETERM_N m o EVERY' [rtac ctxt ballI, rtac ctxt ballI, rtac ctxt impI, atac],
+      rtac ctxt @{thm relcomppI}, assume_tac ctxt, assume_tac ctxt,
+      REPEAT_DETERM_N m o EVERY' [rtac ctxt ballI, rtac ctxt ballI, rtac ctxt impI, assume_tac ctxt],
       REPEAT_DETERM_N (length le_rel_OOs) o
         EVERY' [rtac ctxt ballI, rtac ctxt ballI, Goal.assume_rule_tac ctxt]])
   ctor_nchotomys ctor_Irels rel_mono_strong0s le_rel_OOs) 1;
@@ -627,7 +643,7 @@
   CONJ_WRAP_GEN' (rtac ctxt @{thm card_order_csum}) (rtac ctxt) bd_card_orders 1;
 
 fun mk_wit_tac ctxt n ctor_set wit =
-  REPEAT_DETERM (atac 1 ORELSE
+  REPEAT_DETERM (assume_tac ctxt 1 ORELSE
     EVERY' [dtac ctxt set_rev_mp, rtac ctxt equalityD1, resolve_tac ctxt ctor_set,
     REPEAT_DETERM o
       (TRY o REPEAT_DETERM o etac ctxt UnE THEN' TRY o etac ctxt @{thm UN_E} THEN'
@@ -676,7 +692,7 @@
         CONJ_WRAP' (fn (ctor_set, passive_set_map0) =>
           EVERY' [rtac ctxt ord_eq_le_trans, rtac ctxt ctor_set, rtac ctxt @{thm Un_least},
             rtac ctxt ord_eq_le_trans, rtac ctxt @{thm box_equals[OF _ refl]},
-            rtac ctxt passive_set_map0, rtac ctxt trans_fun_cong_image_id_id_apply, atac,
+            rtac ctxt passive_set_map0, rtac ctxt trans_fun_cong_image_id_id_apply, assume_tac ctxt,
             CONJ_WRAP_GEN' (rtac ctxt (Thm.permute_prems 0 1 @{thm Un_least}))
               (fn (active_set_map0, in_Irel) => EVERY' [rtac ctxt ord_eq_le_trans,
                 rtac ctxt @{thm SUP_cong[OF _ refl]}, rtac ctxt active_set_map0, rtac ctxt @{thm UN_least},
@@ -686,21 +702,23 @@
                   @{thms case_prodE iffD1[OF prod.inject, elim_format]}),
                 hyp_subst_tac ctxt,
                 dtac ctxt (in_Irel RS iffD1), dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE,
-                REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], atac])
+                REPEAT_DETERM o eresolve_tac ctxt [CollectE, conjE], assume_tac ctxt])
             (rev (active_set_map0s ~~ in_Irels))])
         (ctor_sets ~~ passive_set_map0s),
         rtac ctxt conjI,
         REPEAT_DETERM_N 2 o EVERY' [rtac ctxt trans, rtac ctxt ctor_map, rtac ctxt (ctor_inject RS iffD2),
           rtac ctxt trans, rtac ctxt map_comp0, rtac ctxt trans, rtac ctxt map_cong0,
           REPEAT_DETERM_N m o rtac ctxt @{thm fun_cong[OF comp_id]},
-          EVERY' (map (fn in_Irel => EVERY' [rtac ctxt trans, rtac ctxt o_apply, dtac ctxt set_rev_mp, atac,
+          EVERY' (map (fn in_Irel => EVERY' [rtac ctxt trans, rtac ctxt o_apply,
+            dtac ctxt set_rev_mp, assume_tac ctxt,
             dtac ctxt @{thm ssubst_mem[OF pair_collapse]},
             REPEAT_DETERM o eresolve_tac ctxt (CollectE :: conjE ::
               @{thms case_prodE iffD1[OF prod.inject, elim_format]}),
             hyp_subst_tac ctxt,
-            dtac ctxt (in_Irel RS iffD1), dtac ctxt @{thm someI_ex}, REPEAT_DETERM o etac ctxt conjE, atac])
+            dtac ctxt (in_Irel RS iffD1), dtac ctxt @{thm someI_ex},
+            REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt])
           in_Irels),
-          atac]]
+          assume_tac ctxt]]
   in
     EVERY' [rtac ctxt iffI, if_tac, only_if_tac] 1
   end;
@@ -753,7 +771,7 @@
         rtac ctxt (funpow (m + n + 1) (fn thm => thm RS rel_funD) map_transfer),
         REPEAT_DETERM_N m o rtac ctxt @{thm id_transfer},
         REPEAT_DETERM_N n o rtac ctxt @{thm vimage2p_rel_fun},
-        atac])
+        assume_tac ctxt])
       map_transfers)])
   end;
 
--- a/src/HOL/Tools/BNF/bnf_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/BNF/bnf_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -100,7 +100,8 @@
 
 fun mk_map_cong0L_tac ctxt passive map_cong0 map_id =
   (rtac ctxt trans THEN' rtac ctxt map_cong0 THEN' EVERY' (replicate passive (rtac ctxt refl))) 1 THEN
-  REPEAT_DETERM (EVERY' [rtac ctxt trans, etac ctxt bspec, atac, rtac ctxt sym, rtac ctxt @{thm id_apply}] 1) THEN
+  REPEAT_DETERM (EVERY' [rtac ctxt trans, etac ctxt bspec, assume_tac ctxt,
+      rtac ctxt sym, rtac ctxt @{thm id_apply}] 1) THEN
   rtac ctxt map_id 1;
 
 end;
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -54,17 +54,21 @@
 
 fun mk_nchotomy_tac ctxt n exhaust =
   HEADGOAL (rtac ctxt allI THEN' rtac ctxt exhaust THEN'
-    EVERY' (maps (fn k => [rtac ctxt (mk_disjIN n k), REPEAT_DETERM o rtac ctxt exI, atac]) (1 upto n)));
+    EVERY' (maps (fn k => [rtac ctxt (mk_disjIN n k), REPEAT_DETERM o rtac ctxt exI, assume_tac ctxt])
+      (1 upto n)));
 
 fun mk_unique_disc_def_tac ctxt m uexhaust =
-  HEADGOAL (EVERY' [rtac ctxt iffI, rtac ctxt uexhaust, REPEAT_DETERM_N m o rtac ctxt exI, atac, rtac ctxt refl]);
+  HEADGOAL (EVERY'
+    [rtac ctxt iffI, rtac ctxt uexhaust, REPEAT_DETERM_N m o rtac ctxt exI,
+      assume_tac ctxt, rtac ctxt refl]);
 
 fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust =
   HEADGOAL (EVERY' ([rtac ctxt (other_disc_def RS @{thm arg_cong[of _ _ Not]} RS trans),
     rtac ctxt @{thm iffI_np}, REPEAT_DETERM o etac ctxt exE,
     hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac ctxt allI,
     rtac ctxt distinct, rtac ctxt uexhaust] @
-    (([etac ctxt notE, REPEAT_DETERM o rtac ctxt exI, atac], [REPEAT_DETERM o rtac ctxt exI, atac])
+    (([etac ctxt notE, REPEAT_DETERM o rtac ctxt exI, assume_tac ctxt],
+      [REPEAT_DETERM o rtac ctxt exI, assume_tac ctxt])
      |> k = 1 ? swap |> op @)));
 
 fun mk_half_distinct_disc_tac ctxt m discD disc' =
@@ -77,7 +81,7 @@
 fun mk_exhaust_disc_or_sel_tac ctxt n exhaust destIs =
   HEADGOAL (rtac ctxt exhaust THEN'
     EVERY' (map2 (fn k => fn destI => dtac ctxt destI THEN'
-      select_prem_tac ctxt n (etac ctxt meta_mp) k THEN' atac) (1 upto n) destIs));
+      select_prem_tac ctxt n (etac ctxt meta_mp) k THEN' assume_tac ctxt) (1 upto n) destIs));
 
 val mk_exhaust_disc_tac = mk_exhaust_disc_or_sel_tac;
 
@@ -88,7 +92,7 @@
 fun mk_collapse_tac ctxt m discD sels =
   HEADGOAL (dtac ctxt discD THEN'
     (if m = 0 then
-       atac
+       assume_tac ctxt
      else
        REPEAT_DETERM_N m o etac ctxt exE THEN' hyp_subst_tac ctxt THEN'
        SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac ctxt refl));
@@ -105,26 +109,29 @@
     distinct_discsss' =
   if ms = [0] then
     HEADGOAL (rtac ctxt (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN'
-      TRY o EVERY' [rtac ctxt uexhaust_disc, atac, rtac ctxt vexhaust_disc, atac])
+      TRY o
+      EVERY' [rtac ctxt uexhaust_disc, assume_tac ctxt, rtac ctxt vexhaust_disc, assume_tac ctxt])
   else
     let val ks = 1 upto n in
       HEADGOAL (rtac ctxt uexhaust_disc THEN'
         EVERY' (@{map 5} (fn k => fn m => fn distinct_discss => fn distinct_discss' => fn uuncollapse =>
-          EVERY' [rtac ctxt (uuncollapse RS trans) THEN' TRY o atac, rtac ctxt sym, rtac ctxt vexhaust_disc,
+          EVERY' [rtac ctxt (uuncollapse RS trans) THEN'
+            TRY o assume_tac ctxt, rtac ctxt sym, rtac ctxt vexhaust_disc,
             EVERY' (@{map 4} (fn k' => fn distinct_discs => fn distinct_discs' => fn vuncollapse =>
               EVERY'
                 (if k' = k then
-                   [rtac ctxt (vuncollapse RS trans), TRY o atac] @
+                   [rtac ctxt (vuncollapse RS trans), TRY o assume_tac ctxt] @
                    (if m = 0 then
                       [rtac ctxt refl]
                     else
-                      [if n = 1 then K all_tac else EVERY' [dtac ctxt meta_mp, atac, dtac ctxt meta_mp, atac],
+                      [if n = 1 then K all_tac
+                       else EVERY' [dtac ctxt meta_mp, assume_tac ctxt, dtac ctxt meta_mp, assume_tac ctxt],
                        REPEAT_DETERM_N (Int.max (0, m - 1)) o etac ctxt conjE,
                        asm_simp_tac (ss_only [] ctxt)])
                  else
                    [dtac ctxt (the_single (if k = n then distinct_discs else distinct_discs')),
                     etac ctxt (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
-                    atac, atac]))
+                    assume_tac ctxt, assume_tac ctxt]))
               ks distinct_discss distinct_discss' uncollapses)])
           ks ms distinct_discsss distinct_discsss' uncollapses))
     end;
@@ -132,7 +139,7 @@
 fun mk_case_same_ctr_tac ctxt injects =
   REPEAT_DETERM o etac ctxt exE THEN' etac ctxt conjE THEN'
     (case injects of
-      [] => atac
+      [] => assume_tac ctxt
     | [inject] => dtac ctxt (inject RS iffD1) THEN' REPEAT_DETERM o etac ctxt conjE THEN'
         hyp_subst_tac ctxt THEN' rtac ctxt refl);
 
@@ -176,8 +183,10 @@
          flat (nth distinctsss (k - 1))) ctxt)) k) THEN
     ALLGOALS (etac ctxt thin_rl THEN' rtac ctxt iffI THEN'
       REPEAT_DETERM o rtac ctxt allI THEN' rtac ctxt impI THEN' REPEAT_DETERM o etac ctxt conjE THEN'
-      hyp_subst_tac ctxt THEN' atac THEN' REPEAT_DETERM o etac ctxt allE THEN' etac ctxt impE THEN'
-      REPEAT_DETERM o (rtac ctxt conjI THEN' rtac ctxt refl) THEN' rtac ctxt refl THEN' atac)
+      hyp_subst_tac ctxt THEN' assume_tac ctxt THEN'
+      REPEAT_DETERM o etac ctxt allE THEN' etac ctxt impE THEN'
+      REPEAT_DETERM o (rtac ctxt conjI THEN' rtac ctxt refl) THEN'
+      rtac ctxt refl THEN' assume_tac ctxt)
   end;
 
 val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -275,8 +275,8 @@
   let val (butlast, last) = split_last xs;
   in WRAP' (fn thm => conj_tac THEN' gen_tac thm) (K (K all_tac)) butlast (gen_tac last) end;
 
-fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (rtac conjI 1) gen_tac;
-fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (rtac conjI) gen_tac;
+fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (resolve0_tac [conjI] 1) gen_tac;
+fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (resolve0_tac [conjI]) gen_tac;
 
 fun rtac ctxt thm = resolve_tac ctxt [thm];
 fun etac ctxt thm = eresolve_tac ctxt [thm];
--- a/src/HOL/Tools/Function/function_core.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Function/function_core.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -692,7 +692,7 @@
       subset_induct_rule
       |> Thm.forall_intr (Thm.cterm_of ctxt D)
       |> Thm.forall_elim (Thm.cterm_of ctxt acc_R)
-      |> atac 1 |> Seq.hd
+      |> assume_tac ctxt 1 |> Seq.hd
       |> (curry op COMP) (acc_downward
         |> (instantiate' [SOME (Thm.ctyp_of ctxt domT)] (map (SOME o Thm.cterm_of ctxt) [R, x, z]))
         |> Thm.forall_intr (Thm.cterm_of ctxt z)
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -104,17 +104,17 @@
 
 (** Proof Reconstruction **)
 
-fun prove_row (c :: cs) =
-  (case Lazy.force c of
-     Less thm =>
-       rtac @{thm "mlex_less"} 1
-       THEN PRIMITIVE (Thm.elim_implies thm)
-   | LessEq (thm, _) =>
-       rtac @{thm "mlex_leq"} 1
-       THEN PRIMITIVE (Thm.elim_implies thm)
-       THEN prove_row cs
-   | _ => raise General.Fail "lexicographic_order")
- | prove_row [] = no_tac;
+fun prove_row_tac ctxt (c :: cs) =
+      (case Lazy.force c of
+        Less thm =>
+          resolve_tac ctxt @{thms mlex_less} 1
+          THEN PRIMITIVE (Thm.elim_implies thm)
+      | LessEq (thm, _) =>
+          resolve_tac ctxt @{thms mlex_leq} 1
+          THEN PRIMITIVE (Thm.elim_implies thm)
+          THEN prove_row_tac ctxt cs
+      | _ => raise General.Fail "lexicographic_order")
+  | prove_row_tac _ [] = no_tac;
 
 
 (** Error reporting **)
@@ -202,9 +202,9 @@
         in (* 4: proof reconstruction *)
           st |>
           (PRIMITIVE (cterm_instantiate [apply2 (Thm.cterm_of ctxt) (rel, relation)])
-            THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
-            THEN (rtac @{thm "wf_empty"} 1)
-            THEN EVERY (map prove_row clean_table))
+            THEN (REPEAT (resolve_tac ctxt @{thms wf_mlex} 1))
+            THEN (resolve_tac ctxt @{thms wf_empty} 1)
+            THEN EVERY (map (prove_row_tac ctxt) clean_table))
         end
   end) 1 st;
 
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -17,9 +17,9 @@
      msetT : typ -> typ,
      mk_mset : typ -> term list -> term,
      mset_regroup_conv : Proof.context -> int list -> conv,
-     mset_member_tac : int -> int -> tactic,
-     mset_nonempty_tac : int -> tactic,
-     mset_pwleq_tac : int -> tactic,
+     mset_member_tac : Proof.context -> int -> int -> tactic,
+     mset_nonempty_tac : Proof.context -> int -> tactic,
+     mset_pwleq_tac : Proof.context -> int -> tactic,
      set_of_simps : thm list,
      smsI' : thm,
      wmsI2'' : thm,
@@ -49,9 +49,9 @@
    msetT : typ -> typ,
    mk_mset : typ -> term list -> term,
    mset_regroup_conv : Proof.context -> int list -> conv,
-   mset_member_tac : int -> int -> tactic,
-   mset_nonempty_tac : int -> tactic,
-   mset_pwleq_tac : int -> tactic,
+   mset_member_tac : Proof.context -> int -> int -> tactic,
+   mset_nonempty_tac : Proof.context -> int -> tactic,
+   mset_pwleq_tac : Proof.context -> int -> tactic,
    set_of_simps : thm list,
    smsI' : thm,
    wmsI2'' : thm,
@@ -116,13 +116,13 @@
 
 (* General reduction pair application *)
 fun rem_inv_img ctxt =
-  rtac @{thm subsetI} 1
-  THEN etac @{thm CollectE} 1
-  THEN REPEAT (etac @{thm exE} 1)
+  resolve_tac ctxt @{thms subsetI} 1
+  THEN eresolve_tac ctxt @{thms CollectE} 1
+  THEN REPEAT (eresolve_tac ctxt @{thms exE} 1)
   THEN Local_Defs.unfold_tac ctxt @{thms inv_image_def}
-  THEN rtac @{thm CollectI} 1
-  THEN etac @{thm conjE} 1
-  THEN etac @{thm ssubst} 1
+  THEN resolve_tac ctxt @{thms CollectI} 1
+  THEN eresolve_tac ctxt @{thms conjE} 1
+  THEN eresolve_tac ctxt @{thms ssubst} 1
   THEN Local_Defs.unfold_tac ctxt @{thms split_conv triv_forall_equality sum.case}
 
 
@@ -130,15 +130,15 @@
 
 val setT = HOLogic.mk_setT
 
-fun set_member_tac m i =
-  if m = 0 then rtac @{thm insertI1} i
-  else rtac @{thm insertI2} i THEN set_member_tac (m - 1) i
+fun set_member_tac ctxt m i =
+  if m = 0 then resolve_tac ctxt @{thms insertI1} i
+  else resolve_tac ctxt @{thms insertI2} i THEN set_member_tac ctxt (m - 1) i
 
-val set_nonempty_tac = rtac @{thm insert_not_empty}
+fun set_nonempty_tac ctxt = resolve_tac ctxt @{thms insert_not_empty}
 
-fun set_finite_tac i =
-  rtac @{thm finite.emptyI} i
-  ORELSE (rtac @{thm finite.insertI} i THEN (fn st => set_finite_tac i st))
+fun set_finite_tac ctxt i =
+  resolve_tac ctxt @{thms finite.emptyI} i
+  ORELSE (resolve_tac ctxt @{thms finite.insertI} i THEN (fn st => set_finite_tac ctxt i st))
 
 
 (* Reconstruction *)
@@ -183,7 +183,7 @@
               then if b < a then @{thm pair_lessI2} else @{thm pair_lessI1}
               else if b <= a then @{thm pair_leqI2} else @{thm pair_leqI1}
           in
-            rtac rule 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
+            resolve_tac ctxt [rule] 1 THEN PRIMITIVE (Thm.elim_implies stored_thm)
             THEN (if tag_flag then Arith_Data.arith_tac ctxt 1 else all_tac)
           end
 
@@ -192,16 +192,16 @@
                 val (empty, step) = ord_intros_max strict
               in
                 if length lq = 0
-                then rtac empty 1 THEN set_finite_tac 1
-                     THEN (if strict then set_nonempty_tac 1 else all_tac)
+                then resolve_tac ctxt [empty] 1 THEN set_finite_tac ctxt 1
+                     THEN (if strict then set_nonempty_tac ctxt 1 else all_tac)
                 else
                   let
                     val (j, b) :: rest = lq
                     val (i, a) = the (covering g strict j)
-                    fun choose xs = set_member_tac (find_index (curry op = (i, a)) xs) 1
+                    fun choose xs = set_member_tac ctxt (find_index (curry op = (i, a)) xs) 1
                     val solve_tac = choose lp THEN less_proof strict (j, b) (i, a)
                   in
-                    rtac step 1 THEN solve_tac THEN steps_tac MAX strict rest lp
+                    resolve_tac ctxt [step] 1 THEN solve_tac THEN steps_tac MAX strict rest lp
                   end
               end
           | steps_tac MIN strict lq lp =
@@ -209,16 +209,16 @@
                 val (empty, step) = ord_intros_min strict
               in
                 if length lp = 0
-                then rtac empty 1
-                     THEN (if strict then set_nonempty_tac 1 else all_tac)
+                then resolve_tac ctxt [empty] 1
+                     THEN (if strict then set_nonempty_tac ctxt 1 else all_tac)
                 else
                   let
                     val (i, a) :: rest = lp
                     val (j, b) = the (covering g strict i)
-                    fun choose xs = set_member_tac (find_index (curry op = (j, b)) xs) 1
+                    fun choose xs = set_member_tac ctxt (find_index (curry op = (j, b)) xs) 1
                     val solve_tac = choose lq THEN less_proof strict (j, b) (i, a)
                   in
-                    rtac step 1 THEN solve_tac THEN steps_tac MIN strict lq rest
+                    resolve_tac ctxt [step] 1 THEN solve_tac THEN steps_tac MIN strict lq rest
                   end
               end
           | steps_tac MS strict lq lp =
@@ -243,10 +243,10 @@
                 end
               in
                 CONVERSION goal_rewrite 1
-                THEN (if strict then rtac smsI' 1
-                      else if qs = lq then rtac wmsI2'' 1
-                      else rtac wmsI1 1)
-                THEN mset_pwleq_tac 1
+                THEN (if strict then resolve_tac ctxt [smsI'] 1
+                      else if qs = lq then resolve_tac ctxt [wmsI2''] 1
+                      else resolve_tac ctxt [wmsI1] 1)
+                THEN mset_pwleq_tac ctxt 1
                 THEN EVERY (map2 (less_proof false) qs ps)
                 THEN (if strict orelse qs <> lq
                       then Local_Defs.unfold_tac ctxt set_of_simps
@@ -275,9 +275,9 @@
     in
       PROFILE "Proof Reconstruction"
         (CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv ctxt sl))) 1
-         THEN (rtac @{thm reduction_pair_lemma} 1)
-         THEN (rtac @{thm rp_inv_image_rp} 1)
-         THEN (rtac (order_rpair ms_rp label) 1)
+         THEN (resolve_tac ctxt @{thms reduction_pair_lemma} 1)
+         THEN (resolve_tac ctxt @{thms rp_inv_image_rp} 1)
+         THEN (resolve_tac ctxt [order_rpair ms_rp label] 1)
          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}
@@ -300,7 +300,7 @@
       NONE => no_tac
     | SOME cert =>
         SELECT_GOAL (reconstruct_tac ctxt D cs gp cert) i
-        THEN TRY (rtac @{thm wf_empty} i))
+        THEN TRY (resolve_tac ctxt @{thms wf_empty} i))
   end)
 
 fun gen_decomp_scnp_tac orders autom_tac ctxt =
@@ -315,7 +315,7 @@
 fun gen_sizechange_tac orders autom_tac ctxt =
   TRY (Function_Common.termination_rule_tac ctxt 1)
   THEN TRY (Termination.wf_union_tac ctxt)
-  THEN (rtac @{thm wf_empty} 1 ORELSE gen_decomp_scnp_tac orders autom_tac ctxt 1)
+  THEN (resolve_tac ctxt @{thms wf_empty} 1 ORELSE gen_decomp_scnp_tac orders autom_tac ctxt 1)
 
 fun sizechange_tac ctxt autom_tac =
   gen_sizechange_tac [MAX, MS, MIN] autom_tac ctxt
--- a/src/HOL/Tools/Function/termination.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Function/termination.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -289,13 +289,13 @@
         |> foldr1 (HOLogic.mk_binop @{const_name Lattices.sup})
 
     fun solve_membership_tac i =
-      (EVERY' (replicate (i - 2) (rtac @{thm UnI2}))  (* pick the right component of the union *)
-      THEN' (fn j => TRY (rtac @{thm UnI1} j))
-      THEN' (rtac @{thm CollectI})                    (* unfold comprehension *)
-      THEN' (fn i => REPEAT (rtac @{thm exI} i))      (* Turn existentials into schematic Vars *)
-      THEN' ((rtac @{thm refl})                       (* unification instantiates all Vars *)
-        ORELSE' ((rtac @{thm conjI})
-          THEN' (rtac @{thm refl})
+      (EVERY' (replicate (i - 2) (resolve_tac ctxt @{thms UnI2}))  (* pick the right component of the union *)
+      THEN' (fn j => TRY (resolve_tac ctxt @{thms UnI1} j))
+      THEN' (resolve_tac ctxt @{thms CollectI})                    (* unfold comprehension *)
+      THEN' (fn i => REPEAT (resolve_tac ctxt @{thms exI} i))      (* Turn existentials into schematic Vars *)
+      THEN' ((resolve_tac ctxt @{thms refl})                       (* unification instantiates all Vars *)
+        ORELSE' ((resolve_tac ctxt @{thms conjI})
+          THEN' (resolve_tac ctxt @{thms refl})
           THEN' (blast_tac ctxt)))    (* Solve rest of context... not very elegant *)
       ) i
   in
@@ -318,10 +318,10 @@
      then Term_Graph.add_edge (c2, c1) else I)
      cs cs
 
-fun ucomp_empty_tac T =
-  REPEAT_ALL_NEW (rtac @{thm union_comp_emptyR}
-    ORELSE' rtac @{thm union_comp_emptyL}
-    ORELSE' SUBGOAL (fn (_ $ (_ $ (_ $ c1 $ c2) $ _), i) => rtac (T c1 c2) i))
+fun ucomp_empty_tac ctxt T =
+  REPEAT_ALL_NEW (resolve_tac ctxt @{thms union_comp_emptyR}
+    ORELSE' resolve_tac ctxt @{thms union_comp_emptyL}
+    ORELSE' SUBGOAL (fn (_ $ (_ $ (_ $ c1 $ c2) $ _), i) => resolve_tac ctxt [T c1 c2] i))
 
 fun regroup_calls_tac ctxt cs = CALLS (fn (cs', i) =>
  let
@@ -332,11 +332,13 @@
  end)
 
 
-fun solve_trivial_tac D = CALLS (fn ([c], i) =>
-  (case get_chain D c c of
-     SOME (SOME thm) => rtac @{thm wf_no_loop} i
-                        THEN rtac thm i
-   | _ => no_tac)
+fun solve_trivial_tac ctxt D =
+  CALLS (fn ([c], i) =>
+    (case get_chain D c c of
+      SOME (SOME thm) =>
+        resolve_tac ctxt @{thms wf_no_loop} i THEN
+        resolve_tac ctxt [thm] i
+    | _ => no_tac)
   | _ => no_tac)
 
 fun decompose_tac ctxt D = CALLS (fn (cs, i) =>
@@ -344,17 +346,17 @@
     val G = mk_dgraph D cs
     val sccs = Term_Graph.strong_conn G
 
-    fun split [SCC] i = TRY (solve_trivial_tac D i)
+    fun split [SCC] i = TRY (solve_trivial_tac ctxt D i)
       | split (SCC::rest) i =
         regroup_calls_tac ctxt SCC i
-        THEN rtac @{thm wf_union_compatible} i
-        THEN rtac @{thm less_by_empty} (i + 2)
-        THEN ucomp_empty_tac (the o the oo get_chain D) (i + 2)
+        THEN resolve_tac ctxt @{thms wf_union_compatible} i
+        THEN resolve_tac ctxt @{thms less_by_empty} (i + 2)
+        THEN ucomp_empty_tac ctxt (the o the oo get_chain D) (i + 2)
         THEN split rest (i + 1)
-        THEN TRY (solve_trivial_tac D i)
+        THEN TRY (solve_trivial_tac ctxt D i)
   in
     if length sccs > 1 then split sccs i
-    else solve_trivial_tac D i
+    else solve_trivial_tac ctxt D i
   end)
 
 end
--- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -413,7 +413,7 @@
          case_tac exhaust_rule ctxt THEN_ALL_NEW (
         EVERY' [hyp_subst_tac ctxt, (* does not kill wits because = was rewritten to eq_onp top *)
         Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules), 
-        REPEAT_DETERM o etac ctxt conjE, atac])) i
+        REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt])) i
     val pred_simps = Transfer.lookup_pred_data lthy (Tname rty) |> the |> Transfer.pred_simps
     val sel_tac = lift_sel_tac (#exhaust ctr_sugar) (#case_thms ctr_sugar @ pred_simps)
     val sel_names = map (fn (k, xs) => map (fn k' => Binding.qualified true
@@ -560,7 +560,7 @@
 
         fun non_empty_typedef_tac non_empty_pred ctxt i =
           (Method.insert_tac [non_empty_pred] THEN' 
-            SELECT_GOAL (Local_Defs.unfold_tac ctxt [mem_Collect_eq]) THEN' atac) i
+            SELECT_GOAL (Local_Defs.unfold_tac ctxt [mem_Collect_eq]) THEN' assume_tac ctxt) i
         val uTname = unique_Tname (rty, qty)
         val Tdef_set = HOLogic.mk_Collect ("x", rty, pred $ Free("x", rty));
         val ((_, tcode_dt), lthy) = conceal_naming_result (typedef (Binding.concealed uTname, TFrees, NoSyn)
--- a/src/HOL/Tools/Old_Datatype/old_datatype.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -188,7 +188,9 @@
         (fn (((name, mx), tvs), c) =>
           Typedef.add_typedef_global false (name, tvs, mx)
             (Collect $ Const (c, UnivT')) NONE
-            (fn ctxt => rtac exI 1 THEN rtac CollectI 1 THEN
+            (fn ctxt =>
+              resolve_tac ctxt [exI] 1 THEN
+              resolve_tac ctxt [CollectI] 1 THEN
               QUIET_BREADTH_FIRST (has_fewer_prems 1)
               (resolve_tac ctxt rep_intrs 1)))
         (types_syntax ~~ tyvars ~~ take (length newTs) rep_set_names)
@@ -355,7 +357,8 @@
         val rewrites = def_thms @ map mk_meta_eq rec_rewrites;
         val char_thms' =
           map (fn eqn => Goal.prove_sorry_global thy' [] [] eqn
-            (fn {context = ctxt, ...} => EVERY [rewrite_goals_tac ctxt rewrites, rtac refl 1])) eqns;
+            (fn {context = ctxt, ...} =>
+              EVERY [rewrite_goals_tac ctxt rewrites, resolve_tac ctxt [refl] 1])) eqns;
 
       in (thy', char_thms' @ char_thms) end;
 
@@ -382,8 +385,11 @@
                    (map (pair "x") Ts, S $ Old_Datatype_Aux.app_bnds f i)),
                  HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (Term.abs o pair "x") Ts
                    (r $ (a $ Old_Datatype_Aux.app_bnds f i)), f))))
-              (fn {context = ctxt, ...} => EVERY [REPEAT_DETERM_N i (rtac @{thm ext} 1),
-                 REPEAT (etac allE 1), rtac thm 1, assume_tac ctxt 1])
+              (fn {context = ctxt, ...} =>
+                EVERY [REPEAT_DETERM_N i (resolve_tac ctxt @{thms ext} 1),
+                 REPEAT (eresolve_tac ctxt [allE] 1),
+                 resolve_tac ctxt [thm] 1,
+                 assume_tac ctxt 1])
           end
       in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
 
@@ -421,7 +427,7 @@
               [(Old_Datatype_Aux.ind_tac ctxt induct [] THEN_ALL_NEW
                   Object_Logic.atomize_prems_tac ctxt) 1,
                REPEAT (EVERY
-                 [rtac allI 1, rtac impI 1,
+                 [resolve_tac ctxt [allI] 1, resolve_tac ctxt [impI] 1,
                   Old_Datatype_Aux.exh_tac ctxt (exh_thm_of dt_info) 1,
                   REPEAT (EVERY
                     [hyp_subst_tac ctxt 1,
@@ -431,9 +437,9 @@
                      ORELSE (EVERY
                        [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 (cong_tac ctxt 1), resolve_tac ctxt [refl] 1,
                         REPEAT (assume_tac ctxt 1 ORELSE (EVERY
-                          [REPEAT (rtac @{thm ext} 1),
+                          [REPEAT (resolve_tac ctxt @{thms ext} 1),
                            REPEAT (eresolve_tac ctxt (mp :: allE ::
                              map make_elim (Suml_inject :: Sumr_inject ::
                                Lim_inject :: inj_thms') @ fun_congs) 1),
@@ -450,7 +456,7 @@
                   Object_Logic.atomize_prems_tac ctxt) 1,
                 rewrite_goals_tac ctxt rewrites,
                 REPEAT ((resolve_tac ctxt rep_intrs THEN_ALL_NEW
-                  ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
+                  ((REPEAT o eresolve_tac ctxt [allE]) THEN' ares_tac elem_thms)) 1)]);
 
       in (inj_thms'' @ inj_thms, elem_thms @ Old_Datatype_Aux.split_conj_thm elem_thm) end;
 
@@ -485,7 +491,7 @@
           (Goal.prove_sorry_global thy5 [] [] iso_t (fn {context = ctxt, ...} => EVERY
              [(Old_Datatype_Aux.ind_tac ctxt rep_induct [] THEN_ALL_NEW
                  Object_Logic.atomize_prems_tac ctxt) 1,
-              REPEAT (rtac TrueI 1),
+              REPEAT (resolve_tac ctxt [TrueI] 1),
               rewrite_goals_tac ctxt (mk_meta_eq @{thm choice_eq} ::
                 Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
               rewrite_goals_tac ctxt (map Thm.symmetric range_eqs),
@@ -493,7 +499,7 @@
                 [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 ctxt [sym RS range_eqI] 1,
                  resolve_tac ctxt iso_char_thms 1])])));
 
     val Abs_inverse_thms' =
@@ -518,7 +524,7 @@
         (fn {context = ctxt, ...} => EVERY
           [resolve_tac ctxt inj_thms 1,
            rewrite_goals_tac ctxt rewrites,
-           rtac refl 3,
+           resolve_tac ctxt [refl] 3,
            resolve_tac ctxt rep_intrs 2,
            REPEAT (resolve_tac ctxt iso_elem_thms 1)])
       end;
@@ -560,12 +566,14 @@
       in
         Goal.prove_sorry_global thy5 [] [] t
           (fn {context = ctxt, ...} => EVERY
-            [rtac iffI 1,
-             REPEAT (etac conjE 2), hyp_subst_tac ctxt 2, rtac refl 2,
-             dresolve_tac ctxt rep_congs 1, dtac @{thm box_equals} 1,
+            [resolve_tac ctxt [iffI] 1,
+             REPEAT (eresolve_tac ctxt [conjE] 2), hyp_subst_tac ctxt 2,
+             resolve_tac ctxt [refl] 2,
+             dresolve_tac ctxt rep_congs 1,
+             dresolve_tac ctxt @{thms 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 (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (resolve_tac ctxt @{thms ext} 1),
                REPEAT (eresolve_tac ctxt (make_elim fun_cong :: inj_thms) 1),
                assume_tac ctxt 1]))])
       end;
@@ -618,10 +626,10 @@
            HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls)))
         (fn {context = ctxt, ...} =>
           EVERY
-           [REPEAT (etac conjE 1),
+           [REPEAT (eresolve_tac ctxt [conjE] 1),
             REPEAT (EVERY
-              [TRY (rtac conjI 1), resolve_tac ctxt Rep_inverse_thms 1,
-               etac mp 1, resolve_tac ctxt iso_elem_thms 1])]);
+              [TRY (resolve_tac ctxt [conjI] 1), resolve_tac ctxt Rep_inverse_thms 1,
+               eresolve_tac ctxt [mp] 1, resolve_tac ctxt iso_elem_thms 1])]);
 
     val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule_lemma)));
     val frees =
@@ -637,14 +645,14 @@
       (Logic.strip_imp_concl dt_induct_prop)
       (fn {context = ctxt, prems, ...} =>
         EVERY
-          [rtac indrule_lemma' 1,
+          [resolve_tac ctxt [indrule_lemma'] 1,
            (Old_Datatype_Aux.ind_tac ctxt rep_induct [] THEN_ALL_NEW
               Object_Logic.atomize_prems_tac ctxt) 1,
            EVERY (map (fn (prem, r) => (EVERY
              [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)]))
+              DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE eresolve_tac ctxt [allE] 1)]))
                   (prems ~~ (constr_defs @ map mk_meta_eq iso_char_thms)))]);
 
     val ([(_, [dt_induct'])], thy7) =
--- a/src/HOL/Tools/Predicate_Compile/core_data.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Predicate_Compile/core_data.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -249,7 +249,7 @@
         val thesis =
           Thm.instantiate ([], inst_of_matches (Thm.prems_of case_th' ~~ map Thm.prop_of prems2))
             case_th' OF prems2
-      in rtac thesis 1 end
+      in resolve_tac ctxt2 [thesis] 1 end
   in
     Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule
       (fn {context = ctxt1, ...} =>
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -1106,7 +1106,7 @@
           val tac =
             Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps
               (@{thm if_False} :: @{thm Predicate.not_pred_eq} :: simprules)) 1
-            THEN rtac @{thm Predicate.singleI} 1
+            THEN resolve_tac ctxt @{thms Predicate.singleI} 1
         in SOME (Goal.prove ctxt (argnames @ hoarg_names') []
             neg_introtrm (fn _ => tac))
         end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -166,8 +166,9 @@
             val tac = Simplifier.simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [th]) 1
               THEN Inductive.select_disj_tac ctxt' (length disjuncts) (index + 1) 1
               THEN (EVERY (map (fn y =>
-                rtac (Drule.cterm_instantiate [(x, Thm.cterm_of ctxt' (Free y))] @{thm exI}) 1) ps))
-              THEN REPEAT_DETERM (rtac @{thm conjI} 1 THEN assume_tac ctxt' 1)
+                resolve_tac ctxt'
+                  [Drule.cterm_instantiate [(x, Thm.cterm_of ctxt' (Free y))] @{thm exI}] 1) ps))
+              THEN REPEAT_DETERM (resolve_tac ctxt' @{thms conjI} 1 THEN assume_tac ctxt' 1)
               THEN TRY (assume_tac ctxt' 1)
           in
             Goal.prove ctxt' (map fst ps) [] introrule (fn _ => tac)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -38,7 +38,7 @@
   simpset_of (put_simpset HOL_basic_ss @{context}
     addsimps @{thms simp_thms Pair_eq}
     setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
-    setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI})))
+    setSolver (mk_solver "True_solver" (fn ctxt => resolve_tac ctxt @{thms TrueI})))
 
 
 (* auxillary functions *)
@@ -74,13 +74,13 @@
           end) ctxt 1
       | Abs _ => raise Fail "prove_param: No valid parameter term")
   in
-    REPEAT_DETERM (rtac @{thm ext} 1)
+    REPEAT_DETERM (resolve_tac ctxt @{thms ext} 1)
     THEN trace_tac ctxt options "prove_param"
     THEN f_tac
     THEN trace_tac ctxt options "after prove_param"
     THEN (REPEAT_DETERM (assume_tac ctxt 1))
     THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
-    THEN REPEAT_DETERM (rtac @{thm refl} 1)
+    THEN REPEAT_DETERM (resolve_tac ctxt @{thms refl} 1)
   end
 
 fun prove_expr options ctxt nargs (premposition : int) (t, deriv) =
@@ -93,7 +93,7 @@
         val ho_args = ho_args_of mode args
       in
         trace_tac ctxt options "before intro rule:"
-        THEN rtac introrule 1
+        THEN resolve_tac ctxt [introrule] 1
         THEN trace_tac ctxt options "after intro rule"
         (* for the right assumption in first position *)
         THEN rotate_tac premposition 1
@@ -118,7 +118,7 @@
                 [rew_eq RS @{thm sym}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}])
               param_prem
           in
-            rtac param_prem' 1
+            resolve_tac ctxt' [param_prem'] 1
           end) ctxt 1
       THEN trace_tac ctxt options "after prove parameter call")
 
@@ -144,9 +144,9 @@
     asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt addsimps simprules) 1
     THEN trace_tac ctxt options "after prove_match:"
     THEN (DETERM (TRY
-           (rtac eval_if_P 1
+           (resolve_tac ctxt [eval_if_P] 1
            THEN (SUBPROOF (fn {prems, context = ctxt', ...} =>
-             (REPEAT_DETERM (rtac @{thm conjI} 1
+             (REPEAT_DETERM (resolve_tac ctxt' @{thms conjI} 1
              THEN (SOLVED (asm_simp_tac (put_simpset HOL_basic_ss' ctxt') 1))))
              THEN trace_tac ctxt' options "if condition to be solved:"
              THEN asm_simp_tac (put_simpset HOL_basic_ss' ctxt') 1
@@ -157,7 +157,7 @@
                   rewrite_goal_tac ctxt'
                     (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
                 end
-             THEN REPEAT_DETERM (rtac @{thm refl} 1))
+             THEN REPEAT_DETERM (resolve_tac ctxt' @{thms refl} 1))
              THEN trace_tac ctxt' options "after if simp; in SUBPROOF") ctxt 1))))
     THEN trace_tac ctxt options "after if simplification"
   end;
@@ -199,13 +199,13 @@
             in
               rewrite_goal_tac ctxt' (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
             end) ctxt 1
-        THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
+        THEN (resolve_tac ctxt [if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}] 1)
     | prove_prems out_ts ((p, deriv) :: ps) =
         let
           val premposition = (find_index (equal p) clauses) + nargs
           val mode = head_mode_of deriv
           val rest_tac =
-            rtac @{thm bindI} 1
+            resolve_tac ctxt @{thms bindI} 1
             THEN (case p of Prem t =>
               let
                 val (_, us) = strip_comb t
@@ -238,15 +238,15 @@
                 THEN (if (is_some name) then
                     trace_tac ctxt options "before applying not introduction rule"
                     THEN Subgoal.FOCUS_PREMS (fn {prems, ...} =>
-                      rtac (the neg_intro_rule) 1
-                      THEN rtac (nth prems premposition) 1) ctxt 1
+                      resolve_tac ctxt [the neg_intro_rule] 1
+                      THEN resolve_tac ctxt [nth prems premposition] 1) ctxt 1
                     THEN trace_tac ctxt options "after applying not introduction rule"
                     THEN (EVERY (map2 (prove_param options ctxt nargs) params param_derivations))
                     THEN (REPEAT_DETERM (assume_tac ctxt 1))
                   else
-                    rtac @{thm not_predI'} 1
+                    resolve_tac ctxt @{thms not_predI'} 1
                     (* test: *)
-                    THEN dtac @{thm sym} 1
+                    THEN dresolve_tac ctxt @{thms sym} 1
                     THEN asm_full_simp_tac
                       (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1)
                     THEN simp_tac
@@ -254,7 +254,7 @@
                 THEN rec_tac
               end
             | Sidecond t =>
-             rtac @{thm if_predI} 1
+             resolve_tac ctxt @{thms if_predI} 1
              THEN trace_tac ctxt options "before sidecond:"
              THEN prove_sidecond ctxt t
              THEN trace_tac ctxt options "after sidecond:"
@@ -265,14 +265,14 @@
     val prems_tac = prove_prems in_ts moded_ps
   in
     trace_tac ctxt options "Proving clause..."
-    THEN rtac @{thm bindI} 1
-    THEN rtac @{thm singleI} 1
+    THEN resolve_tac ctxt @{thms bindI} 1
+    THEN resolve_tac ctxt @{thms singleI} 1
     THEN prems_tac
   end
 
-fun select_sup 1 1 = []
-  | select_sup _ 1 = [rtac @{thm supI1}]
-  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
+fun select_sup _ 1 1 = []
+  | select_sup ctxt _ 1 = [resolve_tac ctxt @{thms supI1}]
+  | select_sup ctxt n i = resolve_tac ctxt @{thms supI2} :: select_sup ctxt (n - 1) (i - 1);
 
 fun prove_one_direction options ctxt clauses preds pred mode moded_clauses =
   let
@@ -282,11 +282,11 @@
   in
     REPEAT_DETERM (CHANGED (rewrite_goals_tac ctxt @{thms split_paired_all}))
     THEN trace_tac ctxt options "before applying elim rule"
-    THEN etac (predfun_elim_of ctxt pred mode) 1
-    THEN etac pred_case_rule 1
+    THEN eresolve_tac ctxt [predfun_elim_of ctxt pred mode] 1
+    THEN eresolve_tac ctxt [pred_case_rule] 1
     THEN trace_tac ctxt options "after applying elim rule"
     THEN (EVERY (map
-           (fn i => EVERY' (select_sup (length moded_clauses) i) i)
+           (fn i => EVERY' (select_sup ctxt (length moded_clauses) i) i)
              (1 upto (length moded_clauses))))
     THEN (EVERY (map2 (prove_clause options ctxt nargs mode) clauses moded_clauses))
     THEN trace_tac ctxt options "proved one direction"
@@ -313,15 +313,15 @@
             (* THEN (Simplifier.asm_full_simp_tac (put_simpset HOL_basic_ss ctxt) 1)
               THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
               THEN (REPEAT_DETERM_N (num_of_constrs - 1)
-                (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
+                (eresolve_tac ctxt @{thms botE} 1 ORELSE eresolve_tac ctxt @{thms botE} 2)))
             THEN assert_tac ctxt @{here} (fn st => Thm.nprems_of st <= 2)
             THEN (EVERY (map split_term_tac ts))
           end
       else all_tac
   in
     split_term_tac (HOLogic.mk_tuple out_ts)
-    THEN (DETERM (TRY ((Splitter.split_asm_tac ctxt [@{thm "split_if_asm"}] 1)
-    THEN (etac @{thm botE} 2))))
+    THEN (DETERM (TRY ((Splitter.split_asm_tac ctxt @{thms split_if_asm} 1)
+    THEN (eresolve_tac ctxt @{thms botE} 2))))
   end
 
 (* VERY LARGE SIMILIRATIY to function prove_param
@@ -357,15 +357,15 @@
           val param_derivations = param_derivations_of deriv
           val ho_args = ho_args_of mode args
         in
-          etac @{thm bindE} 1
+          eresolve_tac ctxt @{thms bindE} 1
           THEN (REPEAT_DETERM (CHANGED (rewrite_goals_tac ctxt @{thms split_paired_all})))
           THEN trace_tac ctxt options "prove_expr2-before"
-          THEN etac (predfun_elim_of ctxt name mode) 1
+          THEN eresolve_tac ctxt [predfun_elim_of ctxt name mode] 1
           THEN trace_tac ctxt options "prove_expr2"
           THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
           THEN trace_tac ctxt options "finished prove_expr2"
         end
-      | _ => etac @{thm bindE} 1)
+      | _ => eresolve_tac ctxt @{thms bindE} 1)
 
 fun prove_sidecond2 options ctxt t =
   let
@@ -399,15 +399,15 @@
       trace_tac ctxt options "before prove_match2 - last call:"
       THEN prove_match2 options ctxt out_ts
       THEN trace_tac ctxt options "after prove_match2 - last call:"
-      THEN (etac @{thm singleE} 1)
-      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+      THEN (eresolve_tac ctxt @{thms singleE} 1)
+      THEN (REPEAT_DETERM (eresolve_tac ctxt @{thms Pair_inject} 1))
       THEN (asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1)
       THEN TRY (
-        (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+        (REPEAT_DETERM (eresolve_tac ctxt @{thms Pair_inject} 1))
         THEN (asm_full_simp_tac (put_simpset HOL_basic_ss' ctxt) 1)
 
         THEN SOLVED (trace_tac ctxt options "state before applying intro rule:"
-        THEN (rtac pred_intro_rule
+        THEN (resolve_tac ctxt [pred_intro_rule]
         (* How to handle equality correctly? *)
         THEN_ALL_NEW (K (trace_tac ctxt options "state before assumption matching")
         THEN' (assume_tac ctxt ORELSE'
@@ -438,20 +438,20 @@
                 val ho_args = ho_args_of mode args
               in
                 trace_tac ctxt options "before neg prem 2"
-                THEN etac @{thm bindE} 1
+                THEN eresolve_tac ctxt @{thms bindE} 1
                 THEN (if is_some name then
                     full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps
                       [predfun_definition_of ctxt (the name) mode]) 1
-                    THEN etac @{thm not_predE} 1
+                    THEN eresolve_tac ctxt @{thms not_predE} 1
                     THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm not_False_eq_True}]) 1
                     THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
                   else
-                    etac @{thm not_predE'} 1)
+                    eresolve_tac ctxt @{thms not_predE'} 1)
                 THEN rec_tac
               end
           | Sidecond t =>
-              etac @{thm bindE} 1
-              THEN etac @{thm if_predE} 1
+              eresolve_tac ctxt @{thms bindE} 1
+              THEN eresolve_tac ctxt @{thms if_predE} 1
               THEN prove_sidecond2 options ctxt t
               THEN prove_prems2 [] ps)
       in
@@ -463,9 +463,9 @@
     val prems_tac = prove_prems2 in_ts ps
   in
     trace_tac ctxt options "starting prove_clause2"
-    THEN etac @{thm bindE} 1
-    THEN (etac @{thm singleE'} 1)
-    THEN (TRY (etac @{thm Pair_inject} 1))
+    THEN eresolve_tac ctxt @{thms bindE} 1
+    THEN (eresolve_tac ctxt @{thms singleE'} 1)
+    THEN (TRY (eresolve_tac ctxt @{thms Pair_inject} 1))
     THEN trace_tac ctxt options "after singleE':"
     THEN prems_tac
   end;
@@ -473,15 +473,15 @@
 fun prove_other_direction options ctxt pred mode moded_clauses =
   let
     fun prove_clause clause i =
-      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
+      (if i < length moded_clauses then eresolve_tac ctxt @{thms supE} 1 else all_tac)
       THEN (prove_clause2 options ctxt pred mode clause i)
   in
-    (DETERM (TRY (rtac @{thm unit.induct} 1)))
+    (DETERM (TRY (resolve_tac ctxt @{thms unit.induct} 1)))
      THEN (REPEAT_DETERM (CHANGED (rewrite_goals_tac ctxt @{thms split_paired_all})))
-     THEN (rtac (predfun_intro_of ctxt pred mode) 1)
-     THEN (REPEAT_DETERM (rtac @{thm refl} 2))
+     THEN (resolve_tac ctxt [predfun_intro_of ctxt pred mode] 1)
+     THEN (REPEAT_DETERM (resolve_tac ctxt @{thms refl} 2))
      THEN (
-       if null moded_clauses then etac @{thm botE} 1
+       if null moded_clauses then eresolve_tac ctxt @{thms botE} 1
        else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
   end
 
@@ -496,7 +496,7 @@
     Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
       (if not (skip_proof options) then
         (fn _ =>
-        rtac @{thm pred_iffI} 1
+        resolve_tac ctxt @{thms pred_iffI} 1
         THEN trace_tac ctxt options "after pred_iffI"
         THEN prove_one_direction options ctxt clauses preds pred mode moded_clauses
         THEN trace_tac ctxt options "proved one direction"
--- a/src/HOL/Tools/Qelim/cooper.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -741,7 +741,7 @@
         (case try (fn () =>
             Goal.prove_internal ctxt [] g (K (blast_tac (put_claset HOL_cs ctxt) 1))) () of
           NONE => no_tac
-        | SOME r => rtac r i)
+        | SOME r => resolve_tac ctxt [r] i)
     end)
 end;
 
@@ -860,11 +860,11 @@
        else Conv.arg_conv (conv ctxt) p
      val p' = Thm.rhs_of cpth
      val th = Thm.implies_intr p' (Thm.equal_elim (Thm.symmetric cpth) (Thm.assume p'))
-   in rtac th i end
+   in resolve_tac ctxt [th] i end
    handle COOPER _ => no_tac);
 
-fun finish_tac q = SUBGOAL (fn (_, i) =>
-  (if q then I else TRY) (rtac TrueI i));
+fun finish_tac ctxt q = SUBGOAL (fn (_, i) =>
+  (if q then I else TRY) (resolve_tac ctxt [TrueI] i));
 
 fun tac elim add_ths del_ths = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} =>
   let
@@ -885,7 +885,7 @@
     THEN_ALL_NEW CONVERSION Thm.eta_long_conversion
     THEN_ALL_NEW nat_to_int_tac ctxt
     THEN_ALL_NEW core_tac ctxt
-    THEN_ALL_NEW finish_tac elim
+    THEN_ALL_NEW finish_tac ctxt elim
   end 1);
 
 
--- a/src/HOL/Tools/Quickcheck/random_generators.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Quickcheck/random_generators.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -108,7 +108,7 @@
           [(cT_random_aux, Thm.cterm_of lthy' t_random_aux),
            (cT_rhs, Thm.cterm_of lthy' t_rhs)]);
     fun tac ctxt =
-      ALLGOALS (rtac rule)
+      ALLGOALS (resolve_tac ctxt [rule])
       THEN ALLGOALS (simp_tac (put_simpset rew_ss ctxt))
       THEN (ALLGOALS (Proof_Context.fact_tac ctxt eqs2));
     val simp = Goal.prove_sorry lthy' [v] [] eq (tac o #context);
--- a/src/HOL/Tools/Quotient/quotient_def.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Quotient/quotient_def.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -9,15 +9,12 @@
   val add_quotient_def:
     ((binding * mixfix) * Attrib.binding) * (term * term) -> thm ->
     local_theory -> Quotient_Info.quotconsts * local_theory
-
   val quotient_def:
     (binding * typ option * mixfix) option * (Attrib.binding * (term * term)) ->
     local_theory -> Proof.state
-
   val quotient_def_cmd:
     (binding * string option * mixfix) option * (Attrib.binding * (string * string)) ->
     local_theory -> Proof.state
-
 end;
 
 structure Quotient_Def: QUOTIENT_DEF =
@@ -189,7 +186,9 @@
           case thm_list of
             [] => the maybe_proven_rsp_thm
           | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm
-            (fn _ => rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt [thm] 1)
+            (fn _ =>
+              resolve_tac ctxt [readable_rsp_thm_eq] 1 THEN
+              Proof_Context.fact_tac ctxt [thm] 1)
       in
         snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy)
       end
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -61,7 +61,7 @@
 
 fun quotient_tac ctxt =
   (REPEAT_ALL_NEW (FIRST'
-    [rtac @{thm identity_quotient3},
+    [resolve_tac ctxt @{thms identity_quotient3},
      resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems quot_thm}))]))
 
 val quotient_solver = mk_solver "Quotient goal solver" quotient_tac
@@ -250,7 +250,7 @@
               val inst_thm = Drule.instantiate' ty_inst
                 ([NONE, NONE, NONE] @ t_inst) @{thm apply_rspQ3}
             in
-              (rtac inst_thm THEN' SOLVED' (quotient_tac ctxt)) 1
+              (resolve_tac ctxt [inst_thm] THEN' SOLVED' (quotient_tac ctxt)) 1
             end
       | _ => no_tac
     end)
@@ -266,7 +266,7 @@
       in
         case try (Drule.instantiate' [SOME (Thm.ctyp_of ctxt ty)]
             [SOME (Thm.cterm_of ctxt R)]) @{thm equals_rsp} of
-          SOME thm => rtac thm THEN' quotient_tac ctxt
+          SOME thm => resolve_tac ctxt [thm] THEN' quotient_tac ctxt
         | NONE => K no_tac
       end
   | _ => K no_tac
@@ -282,7 +282,7 @@
           case try (map (SOME o Thm.cterm_of ctxt)) [rel, abs, rep] of
             SOME t_inst =>
               (case try (Drule.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
-                SOME inst_thm => (rtac inst_thm THEN' quotient_tac ctxt) i
+                SOME inst_thm => (resolve_tac ctxt [inst_thm] THEN' quotient_tac ctxt) i
               | NONE => no_tac)
           | NONE => no_tac
         end
@@ -315,48 +315,48 @@
   (case bare_concl goal of
       (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *)
     (Const (@{const_name rel_fun}, _) $ _ $ _) $ (Abs _) $ (Abs _)
-        => rtac @{thm rel_funI} THEN' quot_true_tac ctxt unlam
+        => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam
 
       (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
   | (Const (@{const_name HOL.eq},_) $
       (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
       (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
-        => rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
+        => resolve_tac ctxt @{thms ball_rsp} THEN' dresolve_tac ctxt @{thms QT_all}
 
       (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *)
   | (Const (@{const_name rel_fun}, _) $ _ $ _) $
       (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
       (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
-        => rtac @{thm rel_funI} THEN' quot_true_tac ctxt unlam
+        => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam
 
       (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
   | Const (@{const_name HOL.eq},_) $
       (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
       (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
-        => rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
+        => resolve_tac ctxt @{thms bex_rsp} THEN' dresolve_tac ctxt @{thms QT_ex}
 
       (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *)
   | (Const (@{const_name rel_fun}, _) $ _ $ _) $
       (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
       (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
-        => rtac @{thm rel_funI} THEN' quot_true_tac ctxt unlam
+        => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam
 
   | (Const (@{const_name rel_fun}, _) $ _ $ _) $
       (Const(@{const_name Bex1_rel},_) $ _) $ (Const(@{const_name Bex1_rel},_) $ _)
-        => rtac @{thm bex1_rel_rsp} THEN' quotient_tac ctxt
+        => resolve_tac ctxt @{thms bex1_rel_rsp} THEN' quotient_tac ctxt
 
   | (_ $
       (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
       (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
-        => rtac @{thm babs_rsp} THEN' quotient_tac ctxt
+        => resolve_tac ctxt @{thms babs_rsp} THEN' quotient_tac ctxt
 
   | Const (@{const_name HOL.eq},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
-     (rtac @{thm refl} ORELSE'
+     (resolve_tac ctxt @{thms refl} ORELSE'
       (equals_rsp_tac R ctxt THEN' RANGE [
          quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
 
       (* reflexivity of operators arising from Cong_tac *)
-  | Const (@{const_name HOL.eq},_) $ _ $ _ => rtac @{thm refl}
+  | Const (@{const_name HOL.eq},_) $ _ $ _ => resolve_tac ctxt @{thms refl}
 
      (* respectfulness of constants; in particular of a simple relation *)
   | _ $ (Const _) $ (Const _)  (* rel_fun, list_rel, etc but not equality *)
@@ -366,7 +366,7 @@
       (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
       (* observe map_fun *)
   | _ $ _ $ _
-      => (rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
+      => (resolve_tac ctxt @{thms quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
          ORELSE' rep_abs_rsp_tac ctxt
 
   | _ => K no_tac) i)
@@ -388,7 +388,7 @@
     assume_tac ctxt,
 
     (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *)
-    rtac @{thm ext} THEN' quot_true_tac ctxt unlam,
+    resolve_tac ctxt @{thms ext} THEN' quot_true_tac ctxt unlam,
 
     (* reflexivity of the basic relations *)
     (* R ... ... *)
@@ -522,7 +522,7 @@
       map_fun_tac ctxt,
       lambda_prs_tac ctxt,
       simp_tac simpset,
-      TRY o rtac refl]
+      TRY o resolve_tac ctxt [refl]]
   end
 
 
@@ -531,8 +531,8 @@
 fun inst_spec ctrm =
   Drule.instantiate' [SOME (Thm.ctyp_of_cterm ctrm)] [NONE, SOME ctrm] @{thm spec}
 
-fun inst_spec_tac ctrms =
-  EVERY' (map (dtac o inst_spec) ctrms)
+fun inst_spec_tac ctxt ctrms =
+  EVERY' (map (dresolve_tac ctxt o single o inst_spec) ctrms)
 
 fun all_list xs trm =
   fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm
@@ -548,9 +548,9 @@
       val concl' = apply_under_Trueprop (all_list vrs) concl
       val goal = Logic.mk_implies (concl', concl)
       val rule = Goal.prove ctxt [] [] goal
-        (K (EVERY1 [inst_spec_tac (rev cvrs), assume_tac ctxt]))
+        (K (EVERY1 [inst_spec_tac ctxt (rev cvrs), assume_tac ctxt]))
     in
-      rtac rule i
+      resolve_tac ctxt [rule] i
     end)
 
 
@@ -632,7 +632,7 @@
         val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal
         val rule = procedure_inst ctxt rtrm  goal
       in
-        rtac rule i
+        resolve_tac ctxt [rule] i
       end)
   end
 
@@ -681,7 +681,7 @@
         val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal
         val rule = partiality_procedure_inst ctxt rtrm  goal
       in
-        rtac rule i
+        resolve_tac ctxt [rule] i
       end)
   end
 
@@ -722,7 +722,7 @@
 
         val rule = procedure_inst ctxt (Thm.prop_of rthm') goal
       in
-        (rtac rule THEN' rtac rthm') i
+        (resolve_tac ctxt [rule] THEN' resolve_tac ctxt [rthm']) i
       end)
   end
 
--- a/src/HOL/Tools/Quotient/quotient_type.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Quotient/quotient_type.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -42,8 +42,8 @@
 (* makes the new type definitions and proves non-emptyness *)
 fun typedef_make (vs, qty_name, mx, rel, rty) equiv_thm lthy =
   let
-    fun typedef_tac _ =
-      EVERY1 (map rtac [@{thm part_equivp_typedef}, equiv_thm])
+    fun typedef_tac ctxt =
+      EVERY1 (map (resolve_tac ctxt o single) [@{thm part_equivp_typedef}, equiv_thm])
   in
     Typedef.add_typedef false (qty_name, map (rpair dummyS) vs, mx)
       (typedef_term rel rty lthy) NONE typedef_tac lthy
@@ -58,12 +58,12 @@
     val abs_inv = #Abs_inverse typedef_info
     val rep_inj = #Rep_inject typedef_info
   in
-    (rtac @{thm quot_type.intro} THEN' RANGE [
-      rtac equiv_thm,
-      rtac rep_thm,
-      rtac rep_inv,
-      rtac abs_inv THEN' rtac mem_def2 THEN' assume_tac ctxt,
-      rtac rep_inj]) 1
+    (resolve_tac ctxt @{thms quot_type.intro} THEN' RANGE [
+      resolve_tac ctxt [equiv_thm],
+      resolve_tac ctxt [rep_thm],
+      resolve_tac ctxt [rep_inv],
+      resolve_tac ctxt [abs_inv] THEN' resolve_tac ctxt [mem_def2] THEN' assume_tac ctxt,
+      resolve_tac ctxt [rep_inj]]) 1
   end
 
 (* proves the quot_type theorem for the new type *)
--- a/src/HOL/Tools/SMT/conj_disj_perm.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/SMT/conj_disj_perm.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -6,7 +6,7 @@
 
 signature CONJ_DISJ_PERM =
 sig
-  val conj_disj_perm_tac: int -> tactic
+  val conj_disj_perm_tac: Proof.context -> int -> tactic
 end
 
 structure Conj_Disj_Perm: CONJ_DISJ_PERM =
@@ -118,10 +118,10 @@
   | (True, Disj) => contrapos (prove_contradiction_eq false) cp
   | _ => raise CTERM ("prove_conj_disj_perm", [ct]))
 
-val conj_disj_perm_tac = CSUBGOAL (fn (ct, i) => 
+fun conj_disj_perm_tac ctxt = CSUBGOAL (fn (ct, i) => 
   (case Thm.term_of ct of
     @{const HOL.Trueprop} $ (@{const HOL.eq(bool)} $ _ $ _) =>
-      rtac (prove_conj_disj_perm ct (Thm.dest_binop (Thm.dest_arg ct))) i
+      resolve_tac ctxt [prove_conj_disj_perm ct (Thm.dest_binop (Thm.dest_arg ct))] i
   | _ => no_tac))
 
 end
--- a/src/HOL/Tools/SMT/smt_solver.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/SMT/smt_solver.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -289,13 +289,14 @@
           "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")
     | SMT_Failure.SMT fail => error (str_of ctxt fail)
 
-  fun resolve (SOME thm) = rtac thm 1
-    | resolve NONE = no_tac
+  fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1
+    | resolve _ NONE = no_tac
 
   fun tac prove ctxt rules =
     CONVERSION (SMT_Normalize.atomize_conv ctxt)
-    THEN' rtac @{thm ccontr}
-    THEN' SUBPROOF (fn {context = ctxt', prems, ...} => resolve (prove ctxt' (rules @ prems))) ctxt
+    THEN' resolve_tac ctxt @{thms ccontr}
+    THEN' SUBPROOF (fn {context = ctxt', prems, ...} =>
+      resolve ctxt' (prove ctxt' (rules @ prems))) ctxt
 in
 
 val smt_tac = tac safe_solve
--- a/src/HOL/Tools/SMT/z3_replay.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/SMT/z3_replay.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -148,10 +148,10 @@
 in
 
 fun discharge_sk_tac ctxt i st =
-  (rtac @{thm trans} i
+  (resolve_tac ctxt @{thms trans} i
    THEN resolve_tac ctxt sk_rules i
-   THEN (rtac @{thm refl} ORELSE' discharge_sk_tac ctxt) (i+1)
-   THEN rtac @{thm refl} i) st
+   THEN (resolve_tac ctxt @{thms refl} ORELSE' discharge_sk_tac ctxt) (i+1)
+   THEN resolve_tac ctxt @{thms refl} i) st
 
 end
 
--- a/src/HOL/Tools/SMT/z3_replay_methods.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/SMT/z3_replay_methods.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -443,9 +443,9 @@
 fun lift_ite_rewrite ctxt t =
   prove ctxt t (fn ctxt => 
     CONVERSION (HOLogic.Trueprop_conv (Conv.binop_conv (if_context_conv ctxt)))
-    THEN' rtac @{thm refl})
+    THEN' resolve_tac ctxt @{thms refl})
 
-fun prove_conj_disj_perm ctxt t = prove ctxt t (fn _ => Conj_Disj_Perm.conj_disj_perm_tac)
+fun prove_conj_disj_perm ctxt t = prove ctxt t Conj_Disj_Perm.conj_disj_perm_tac
 
 fun rewrite ctxt _ = try_provers ctxt Z3_Proof.Rewrite [
   ("rules", apply_rule ctxt),
@@ -492,8 +492,8 @@
 val quant_inst_rule = @{lemma "~P x | Q ==> ~(ALL x. P x) | Q" by fast}
 
 fun quant_inst ctxt _ t = prove ctxt t (fn _ =>
-  REPEAT_ALL_NEW (rtac quant_inst_rule)
-  THEN' rtac @{thm excluded_middle})
+  REPEAT_ALL_NEW (resolve_tac ctxt [quant_inst_rule])
+  THEN' resolve_tac ctxt @{thms excluded_middle})
 
 
 (* propositional lemma *)
--- a/src/HOL/Tools/Transfer/transfer.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/Transfer/transfer.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -627,7 +627,7 @@
 
 fun eq_rules_tac ctxt eq_rules =
   TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_rules)
-  THEN_ALL_NEW rtac @{thm is_equality_eq}
+  THEN_ALL_NEW resolve_tac ctxt @{thms is_equality_eq}
 
 fun eq_tac ctxt = eq_rules_tac ctxt (get_relator_eq_raw ctxt)
 
@@ -646,8 +646,8 @@
     val end_tac = if equiv then K all_tac else K no_tac
     val err_msg = "Transfer failed to convert goal to an object-logic formula"
     fun main_tac (t, i) =
-      rtac start_rule i THEN
-      (rtac (transfer_rule_of_term ctxt equiv (HOLogic.dest_Trueprop t))
+      resolve_tac ctxt [start_rule] i THEN
+      (resolve_tac ctxt [transfer_rule_of_term ctxt equiv (HOLogic.dest_Trueprop t)]
         THEN_ALL_NEW
           (SOLVED'
             (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW
@@ -673,12 +673,13 @@
   in
     EVERY
       [CONVERSION prep_conv i,
-       rtac @{thm transfer_prover_start} i,
-       ((rtac rule1 ORELSE' (CONVERSION expand_eq_in_rel THEN' rtac rule1))
+       resolve_tac ctxt @{thms transfer_prover_start} i,
+       ((resolve_tac ctxt [rule1] ORELSE'
+         (CONVERSION expand_eq_in_rel THEN' resolve_tac ctxt [rule1]))
         THEN_ALL_NEW
          (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW
            (DETERM o eq_rules_tac ctxt eq_rules))) (i + 1),
-       rtac @{thm refl} i]
+       resolve_tac ctxt @{thms refl} i]
   end)
 
 (** Transfer attribute **)
@@ -702,7 +703,7 @@
     val rule = transfer_rule_of_lhs ctxt' t
     val tac =
       resolve_tac ctxt' [thm2 RS start_rule', thm2 RS start_rule] 1 THEN
-      (rtac rule
+      (resolve_tac ctxt' [rule]
         THEN_ALL_NEW
           (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules)
             THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1
@@ -737,8 +738,8 @@
     val t = HOLogic.dest_Trueprop (Thm.concl_of thm2)
     val rule = transfer_rule_of_term ctxt' true t
     val tac =
-      rtac (thm2 RS start_rule) 1 THEN
-      (rtac rule
+      resolve_tac ctxt' [thm2 RS start_rule] 1 THEN
+      (resolve_tac ctxt' [rule]
         THEN_ALL_NEW
           (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules)
             THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1
--- a/src/HOL/Tools/datatype_realizer.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/datatype_realizer.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -113,11 +113,11 @@
         (fn prems =>
            EVERY [
             rewrite_goals_tac ctxt (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]),
-            rtac (cterm_instantiate inst induct) 1,
+            resolve_tac ctxt [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 ctxt prems THEN_ALL_NEW (fn i =>
-              REPEAT (etac allE i) THEN assume_tac ctxt i)) 1)])
+              REPEAT (eresolve_tac ctxt [allE] i) THEN assume_tac ctxt i)) 1)])
       |> Drule.export_without_context;
 
     val ind_name = Thm.derivation_name induct;
@@ -190,7 +190,7 @@
           (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y']))))
         (fn prems =>
            EVERY [
-            rtac (cterm_instantiate [apply2 (Thm.cterm_of ctxt) (y, y')] exhaust) 1,
+            resolve_tac ctxt [cterm_instantiate [apply2 (Thm.cterm_of ctxt) (y, y')] exhaust] 1,
             ALLGOALS (EVERY'
               [asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites),
                resolve_tac ctxt prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])])
--- a/src/HOL/Tools/groebner.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/groebner.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -931,11 +931,11 @@
   Object_Logic.full_atomize_tac ctxt
   THEN' presimplify ctxt add_ths del_ths
   THEN' CSUBGOAL (fn (p, i) =>
-    rtac (let val form = Object_Logic.dest_judgment ctxt p
+    resolve_tac ctxt [let val form = Object_Logic.dest_judgment ctxt p
           in case get_ring_ideal_convs ctxt form of
            NONE => Thm.reflexive form
           | SOME thy => #ring_conv thy ctxt form
-          end) i
+          end] i
       handle TERM _ => no_tac
         | CTERM _ => no_tac
         | THM _ => no_tac);
@@ -944,8 +944,9 @@
  fun lhs t = case Thm.term_of t of
   Const(@{const_name HOL.eq},_)$_$_ => Thm.dest_arg1 t
  | _=> raise CTERM ("ideal_tac - lhs",[t])
- fun exitac NONE = no_tac
-   | exitac (SOME y) = rtac (instantiate' [SOME (Thm.ctyp_of_cterm y)] [NONE,SOME y] exI) 1
+ fun exitac _ NONE = no_tac
+   | exitac ctxt (SOME y) =
+      resolve_tac ctxt [instantiate' [SOME (Thm.ctyp_of_cterm y)] [NONE,SOME y] exI] 1
 
  val claset = claset_of @{context}
 in
@@ -964,7 +965,7 @@
      val ps = map_filter (try (lhs o Thm.dest_arg)) asms
      val cfs = (map swap o #multi_ideal thy evs ps)
                    (map Thm.dest_arg1 (conjuncts bod))
-     val ws = map (exitac o AList.lookup op aconvc cfs) evs
+     val ws = map (exitac ctxt o AList.lookup op aconvc cfs) evs
     in EVERY (rev ws) THEN Method.insert_tac prems 1
         THEN ring_tac add_ths del_ths ctxt 1
    end
--- a/src/HOL/Tools/inductive_realizer.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/inductive_realizer.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -393,11 +393,12 @@
         val thm = Goal.prove_global thy []
           (map attach_typeS prems) (attach_typeS concl)
           (fn {context = ctxt, prems} => EVERY
-          [rtac (#raw_induct ind_info) 1,
+          [resolve_tac ctxt [#raw_induct ind_info] 1,
            rewrite_goals_tac ctxt rews,
            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)]);
+              DEPTH_SOLVE_1 o
+              FIRST' [assume_tac ctxt, eresolve_tac ctxt [allE], eresolve_tac ctxt [impE]]]) 1)]);
         val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"
           (Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy;
         val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp)))
@@ -456,11 +457,12 @@
           (Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz')
           (fn {context = ctxt, prems, ...} => EVERY
             [cut_tac (hd prems) 1,
-             etac elimR 1,
+             eresolve_tac ctxt [elimR] 1,
              ALLGOALS (asm_simp_tac (put_simpset HOL_basic_ss ctxt)),
              rewrite_goals_tac ctxt rews,
              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)]);
+               DEPTH_SOLVE_1 o
+               FIRST' [assume_tac ctxt, eresolve_tac ctxt [allE], eresolve_tac ctxt [impE]])) 1)]);
         val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"
           (name_of_thm elim :: vs @ Ps @ ["correctness"])), thm) thy
       in
--- a/src/HOL/Tools/record.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/record.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -112,7 +112,7 @@
   in
     thy
     |> Typedef.add_typedef_global false (raw_tyco, vs, NoSyn)
-        (HOLogic.mk_UNIV repT) NONE (fn _ => rtac UNIV_witness 1)
+        (HOLogic.mk_UNIV repT) NONE (fn ctxt' => resolve_tac ctxt' [UNIV_witness] 1)
     |-> (fn (tyco, info) => get_typedef_info tyco vs info)
   end;
 
@@ -184,7 +184,7 @@
         (case Symtab.lookup isthms (#1 is) of
           SOME isthm => isthm
         | NONE => err "no thm found for constant" (Const is));
-    in rtac isthm i end);
+    in resolve_tac ctxt [isthm] i end);
 
 end;
 
@@ -1380,7 +1380,7 @@
         val thm = cterm_instantiate [(crec, cfree)] induct_thm;
       in
         simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) i THEN
-        rtac thm i THEN
+        resolve_tac ctxt [thm] i THEN
         simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_rulify}) i
       end;
 
@@ -1591,9 +1591,9 @@
           (fn {context = ctxt, ...} =>
             simp_tac (put_simpset HOL_basic_ss ctxt addsimps [ext_def]) 1 THEN
             REPEAT_DETERM
-              (rtac @{thm refl_conj_eq} 1 ORELSE
+              (resolve_tac ctxt @{thms refl_conj_eq} 1 ORELSE
                 Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1 ORELSE
-                rtac refl 1))));
+                resolve_tac ctxt [refl] 1))));
 
 
     (*We need a surjection property r = (| f = f r, g = g r ... |)
@@ -1610,7 +1610,8 @@
         val tactic1 =
           simp_tac (put_simpset HOL_basic_ss defs_ctxt addsimps [ext_def]) 1 THEN
           REPEAT_ALL_NEW (Iso_Tuple_Support.iso_tuple_intros_tac defs_ctxt) 1;
-        val tactic2 = REPEAT (rtac surject_assistI 1 THEN rtac refl 1);
+        val tactic2 =
+          REPEAT (resolve_tac defs_ctxt [surject_assistI] 1 THEN resolve_tac defs_ctxt [refl] 1);
         val [halfway] = Seq.list_of (tactic1 start);
         val [surject] = Seq.list_of (tactic2 (Drule.export_without_context halfway));
       in
@@ -1621,10 +1622,11 @@
       Goal.prove_sorry_global defs_thy [] [] split_meta_prop
         (fn {context = ctxt, ...} =>
           EVERY1
-           [rtac @{thm equal_intr_rule}, Goal.norm_hhf_tac ctxt,
-            etac @{thm meta_allE}, assume_tac ctxt,
-            rtac (@{thm prop_subst} OF [surject]),
-            REPEAT o etac @{thm meta_allE}, assume_tac ctxt]));
+           [resolve_tac ctxt @{thms equal_intr_rule},
+            Goal.norm_hhf_tac ctxt,
+            eresolve_tac ctxt @{thms meta_allE}, assume_tac ctxt,
+            resolve_tac ctxt [@{thm prop_subst} OF [surject]],
+            REPEAT o eresolve_tac ctxt @{thms meta_allE}, assume_tac ctxt]));
 
     val induct = timeit_msg ctxt "record extension induct proof:" (fn () =>
       let val (assm, concl) = induct_prop in
@@ -1748,7 +1750,7 @@
       fun tac ctxt eq_def =
         Class.intro_classes_tac ctxt []
         THEN rewrite_goals_tac ctxt [Simpdata.mk_eq eq_def]
-        THEN ALLGOALS (rtac @{thm refl});
+        THEN ALLGOALS (resolve_tac ctxt @{thms refl});
       fun mk_eq thy eq_def =
         rewrite_rule (Proof_Context.init_global thy)
           [Axclass.unoverload thy (Thm.symmetric (Simpdata.mk_eq eq_def))] inject;
@@ -1946,7 +1948,8 @@
           val cterm_vrs = Thm.cterm_of ext_ctxt r_rec0_Vars;
           val insts = [("v", cterm_rec), ("v'", cterm_vrs)];
           val init_thm = named_cterm_instantiate insts updacc_eq_triv;
-          val terminal = rtac updacc_eq_idI 1 THEN rtac refl 1;
+          val terminal =
+            resolve_tac ext_ctxt [updacc_eq_idI] 1 THEN resolve_tac ext_ctxt [refl] 1;
           val tactic =
             simp_tac (put_simpset HOL_basic_ss ext_ctxt addsimps ext_defs) 1 THEN
             REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac ext_ctxt 1 ORELSE terminal);
@@ -2131,11 +2134,11 @@
       Goal.prove_sorry_global defs_thy [] [] surjective_prop
         (fn {context = ctxt, ...} =>
           EVERY
-           [rtac surject_assist_idE 1,
+           [resolve_tac ctxt [surject_assist_idE] 1,
             simp_tac (put_simpset HOL_basic_ss ctxt addsimps ext_defs) 1,
             REPEAT
               (Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1 ORELSE
-                (rtac surject_assistI 1 THEN
+                (resolve_tac ctxt [surject_assistI] 1 THEN
                   simp_tac (put_simpset (get_sel_upd_defs defs_thy) ctxt
                     addsimps (sel_defs @ @{thms o_assoc id_apply id_o o_id})) 1))]));
 
@@ -2143,7 +2146,7 @@
       Goal.prove_sorry_global defs_thy [] [#1 cases_scheme_prop] (#2 cases_scheme_prop)
         (fn {context = ctxt, prems, ...} =>
           resolve_tac ctxt prems 1 THEN
-          rtac surjective 1));
+          resolve_tac ctxt [surjective] 1));
 
     val cases = timeit_msg ctxt "record cases proof:" (fn () =>
       Goal.prove_sorry_global defs_thy [] [] cases_prop
@@ -2156,26 +2159,26 @@
       Goal.prove_sorry_global defs_thy [] [] split_meta_prop
         (fn {context = ctxt', ...} =>
           EVERY1
-           [rtac @{thm equal_intr_rule}, Goal.norm_hhf_tac ctxt',
-            etac @{thm meta_allE}, assume_tac ctxt',
-            rtac (@{thm prop_subst} OF [surjective]),
-            REPEAT o etac @{thm meta_allE}, assume_tac ctxt']));
+           [resolve_tac ctxt' @{thms equal_intr_rule}, Goal.norm_hhf_tac ctxt',
+            eresolve_tac ctxt' @{thms meta_allE}, assume_tac ctxt',
+            resolve_tac ctxt' [@{thm prop_subst} OF [surjective]],
+            REPEAT o eresolve_tac ctxt' @{thms meta_allE}, assume_tac ctxt']));
 
     val split_object = timeit_msg ctxt "record split_object proof:" (fn () =>
       Goal.prove_sorry_global defs_thy [] [] split_object_prop
         (fn {context = ctxt, ...} =>
-          rtac @{lemma "Trueprop A == Trueprop B ==> A = B" by (rule iffI) unfold} 1 THEN
+          resolve_tac ctxt [@{lemma "Trueprop A \<equiv> Trueprop B \<Longrightarrow> A = B" by (rule iffI) unfold}] 1 THEN
           rewrite_goals_tac ctxt @{thms atomize_all [symmetric]} THEN
-          rtac split_meta 1));
+          resolve_tac ctxt [split_meta] 1));
 
     val split_ex = timeit_msg ctxt "record split_ex proof:" (fn () =>
       Goal.prove_sorry_global defs_thy [] [] split_ex_prop
         (fn {context = ctxt, ...} =>
           simp_tac
             (put_simpset HOL_basic_ss ctxt addsimps
-              (@{lemma "EX x. P x == ~ (ALL x. ~ P x)" by simp} ::
+              (@{lemma "\<exists>x. P x \<equiv> \<not> (\<forall>x. \<not> P x)" by simp} ::
                 @{thms not_not Not_eq_iff})) 1 THEN
-          rtac split_object 1));
+          resolve_tac ctxt [split_object] 1));
 
     val equality = timeit_msg ctxt "record equality proof:" (fn () =>
       Goal.prove_sorry_global defs_thy [] [] equality_prop
--- a/src/HOL/Tools/reification.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Tools/reification.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -32,7 +32,7 @@
     val thm = conv ct;
   in
     if Thm.is_reflexive thm then no_tac
-    else ALLGOALS (rtac (pure_subst OF [thm]))
+    else ALLGOALS (resolve_tac ctxt [pure_subst OF [thm]])
   end) ctxt;
 
 (* Make a congruence rule out of a defining equation for the interpretation
@@ -82,7 +82,7 @@
       (Goal.prove ctxt'' [] (map mk_def env)
         (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
         (fn {context, prems, ...} =>
-          Local_Defs.unfold_tac context (map tryext prems) THEN rtac th' 1)) RS sym;
+          Local_Defs.unfold_tac context (map tryext prems) THEN resolve_tac ctxt'' [th'] 1)) RS sym;
     val (cong' :: vars') =
       Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o Thm.cterm_of ctxt'') vs);
     val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars';
--- a/src/HOL/UNITY/Comp/Alloc.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/UNITY/Comp/Alloc.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -712,7 +712,7 @@
 fun rename_client_map_tac ctxt =
   EVERY [
     simp_tac (ctxt addsimps [@{thm rename_guarantees_eq_rename_inv}]) 1,
-    rtac @{thm guarantees_PLam_I} 1,
+    resolve_tac ctxt @{thms guarantees_PLam_I} 1,
     assume_tac ctxt 2,
          (*preserves: routine reasoning*)
     asm_simp_tac (ctxt addsimps [@{thm lift_preserves_sub}]) 2,
@@ -899,9 +899,9 @@
 text{*safety (1), step 4 (final result!) *}
 theorem System_safety: "System : system_safety"
   apply (unfold system_safety_def)
-  apply (tactic {* rtac (Always_Int_rule [@{thm System_sum_bounded},
+  apply (tactic {* resolve_tac @{context} [Always_Int_rule [@{thm System_sum_bounded},
     @{thm Always_tokens_giv_le_allocGiv}, @{thm Always_tokens_allocRel_le_rel}] RS
-    @{thm Always_weaken}) 1 *})
+    @{thm Always_weaken}] 1 *})
   apply auto
   apply (rule setsum_fun_mono [THEN order_trans])
   apply (drule_tac [2] order_trans)
@@ -942,8 +942,8 @@
 lemma System_Bounded_allocAsk: "System : Always {s. ALL i<Nclients.
                           ALL elt : set ((sub i o allocAsk) s). elt \<le> NbT}"
   apply (auto simp add: Collect_all_imp_eq)
-  apply (tactic {* rtac (Always_Int_rule [@{thm Always_allocAsk_le_ask},
-    @{thm System_Bounded_ask}] RS @{thm Always_weaken}) 1 *})
+  apply (tactic {* resolve_tac @{context} [Always_Int_rule [@{thm Always_allocAsk_le_ask},
+    @{thm System_Bounded_ask}] RS @{thm Always_weaken}] 1 *})
   apply (auto dest: set_mono)
   done
 
--- a/src/HOL/UNITY/Simple/NSP_Bad.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/UNITY/Simple/NSP_Bad.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -104,11 +104,11 @@
 fun ns_constrains_tac ctxt i =
   SELECT_GOAL
     (EVERY
-     [REPEAT (etac @{thm Always_ConstrainsI} 1),
+     [REPEAT (eresolve_tac ctxt @{thms Always_ConstrainsI} 1),
       REPEAT (resolve_tac ctxt [@{thm StableI}, @{thm stableI}, @{thm constrains_imp_Constrains}] 1),
-      rtac @{thm ns_constrainsI} 1,
+      resolve_tac ctxt @{thms ns_constrainsI} 1,
       full_simp_tac ctxt 1,
-      REPEAT (FIRSTGOAL (etac disjE)),
+      REPEAT (FIRSTGOAL (eresolve_tac ctxt [disjE])),
       ALLGOALS (clarify_tac (ctxt delrules [impI, @{thm impCE}])),
       REPEAT (FIRSTGOAL (analz_mono_contra_tac ctxt)),
       ALLGOALS (asm_simp_tac ctxt)]) i;
@@ -116,10 +116,10 @@
 (*Tactic for proving secrecy theorems*)
 fun ns_induct_tac ctxt =
   (SELECT_GOAL o EVERY)
-     [rtac @{thm AlwaysI} 1,
+     [resolve_tac ctxt @{thms AlwaysI} 1,
       force_tac ctxt 1,
       (*"reachable" gets in here*)
-      rtac (@{thm Always_reachable} RS @{thm Always_ConstrainsI} RS @{thm StableI}) 1,
+      resolve_tac ctxt [@{thm Always_reachable} RS @{thm Always_ConstrainsI} RS @{thm StableI}] 1,
       ns_constrains_tac ctxt 1];
 *}
 
--- a/src/HOL/UNITY/UNITY_tactics.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/UNITY/UNITY_tactics.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -7,7 +7,9 @@
 
 (*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
 fun Always_Int_tac ctxt =
-  dtac @{thm Always_Int_I} THEN' assume_tac ctxt THEN' etac @{thm Always_thin}
+  dresolve_tac ctxt @{thms Always_Int_I} THEN'
+  assume_tac ctxt THEN'
+  eresolve_tac ctxt @{thms Always_thin}
 
 (*Combines a list of invariance THEOREMS into one.*)
 val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I})
@@ -19,16 +21,16 @@
     (EVERY
      [REPEAT (Always_Int_tac ctxt 1),
       (*reduce the fancy safety properties to "constrains"*)
-      REPEAT (etac @{thm Always_ConstrainsI} 1
+      REPEAT (eresolve_tac ctxt @{thms Always_ConstrainsI} 1
               ORELSE
               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
         addsimps [@{thm mk_total_program_def}, @{thm totalize_constrains_iff}]) 1,
-      rtac @{thm constrainsI} 1,
+      resolve_tac ctxt @{thms constrainsI} 1,
       full_simp_tac ctxt 1,
-      REPEAT (FIRSTGOAL (etac disjE)),
+      REPEAT (FIRSTGOAL (eresolve_tac ctxt [disjE])),
       ALLGOALS (clarify_tac ctxt),
       ALLGOALS (asm_full_simp_tac ctxt)]) i;
 
@@ -37,7 +39,7 @@
   SELECT_GOAL
     (EVERY
      [REPEAT (Always_Int_tac ctxt 1),
-      etac @{thm Always_LeadsTo_Basis} 1
+      eresolve_tac ctxt @{thms Always_LeadsTo_Basis} 1
           ORELSE   (*subgoal may involve LeadsTo, leadsTo or ensures*)
           REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
                             @{thm EnsuresI}, @{thm ensuresI}] 1),
--- a/src/HOL/Word/Word.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/Word/Word.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -1562,8 +1562,8 @@
           |> fold Simplifier.add_cong @{thms power_False_cong})),
       rewrite_goals_tac ctxt @{thms word_size}, 
       ALLGOALS  (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN
-                         REPEAT (etac conjE n) THEN
-                         REPEAT (dtac @{thm word_of_int_inverse} n 
+                         REPEAT (eresolve_tac ctxt [conjE] n) THEN
+                         REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n 
                                  THEN assume_tac ctxt n 
                                  THEN assume_tac ctxt n)),
       TRYALL arith_tac' ]
@@ -2063,9 +2063,9 @@
           |> 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 ctxt [allI, impI] n) THEN      
-                         REPEAT (etac conjE n) THEN
-                         REPEAT (dtac @{thm of_nat_inverse} n THEN atac n)),
+      ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN
+                         REPEAT (eresolve_tac ctxt [conjE] n) THEN
+                         REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)),
       TRYALL arith_tac' ] 
   end
 
--- a/src/HOL/ex/Cartouche_Examples.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/ex/Cartouche_Examples.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -33,7 +33,7 @@
     [OF \<open>x = y\<close>]}.\<close>
 
   have "x = y"
-    by (tactic \<open>rtac @{thm \<open>x = y\<close>} 1\<close>)  -- \<open>more cartouches involving ML\<close>
+    by (tactic \<open>resolve_tac @{context} @{thms \<open>x = y\<close>} 1\<close>)  -- \<open>more cartouches involving ML\<close>
 end
 
 
@@ -230,10 +230,10 @@
 \<close>
 
 lemma "A \<and> B \<longrightarrow> B \<and> A"
-  apply (ml_tactic \<open>rtac @{thm impI} 1\<close>)
-  apply (ml_tactic \<open>etac @{thm conjE} 1\<close>)
-  apply (ml_tactic \<open>rtac @{thm conjI} 1\<close>)
-  apply (ml_tactic \<open>ALLGOALS atac\<close>)
+  apply (ml_tactic \<open>resolve_tac @{context} @{thms impI} 1\<close>)
+  apply (ml_tactic \<open>eresolve_tac @{context} @{thms conjE} 1\<close>)
+  apply (ml_tactic \<open>resolve_tac @{context} @{thms conjI} 1\<close>)
+  apply (ml_tactic \<open>ALLGOALS (assume_tac @{context})\<close>)
   done
 
 lemma "A \<and> B \<longrightarrow> B \<and> A" by (ml_tactic \<open>blast_tac ctxt 1\<close>)
@@ -251,14 +251,17 @@
 \<close>
 
 lemma "A \<and> B \<longrightarrow> B \<and> A"
-  apply \<open>rtac @{thm impI} 1\<close>
-  apply \<open>etac @{thm conjE} 1\<close>
-  apply \<open>rtac @{thm conjI} 1\<close>
-  apply \<open>ALLGOALS atac\<close>
+  apply \<open>resolve_tac @{context} @{thms impI} 1\<close>
+  apply \<open>eresolve_tac @{context} @{thms conjE} 1\<close>
+  apply \<open>resolve_tac @{context} @{thms conjI} 1\<close>
+  apply \<open>ALLGOALS (assume_tac @{context})\<close>
   done
 
 lemma "A \<and> B \<longrightarrow> B \<and> A"
-  by (\<open>rtac @{thm impI} 1\<close>, \<open>etac @{thm conjE} 1\<close>, \<open>rtac @{thm conjI} 1\<close>, \<open>atac 1\<close>+)
+  by (\<open>resolve_tac @{context} @{thms impI} 1\<close>,
+    \<open>eresolve_tac @{context} @{thms conjE} 1\<close>,
+    \<open>resolve_tac @{context} @{thms conjI} 1\<close>,
+    \<open>assume_tac @{context} 1\<close>+)
 
 
 subsection \<open>ML syntax\<close>
--- a/src/HOL/ex/Iff_Oracle.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/ex/Iff_Oracle.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -58,7 +58,7 @@
 method_setup iff =
   \<open>Scan.lift Parse.nat >> (fn n => fn ctxt =>
     SIMPLE_METHOD
-      (HEADGOAL (rtac (iff_oracle (Proof_Context.theory_of ctxt, n)))
+      (HEADGOAL (resolve_tac ctxt [iff_oracle (Proof_Context.theory_of ctxt, n)])
         handle Fail _ => no_tac))\<close>
 
 
--- a/src/HOL/ex/Meson_Test.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/HOL/ex/Meson_Test.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -29,7 +29,7 @@
     val nnf25 = Meson.make_nnf ctxt prem25;
     val xsko25 = Meson.skolemize ctxt nnf25;
   *}
-  apply (tactic {* cut_tac xsko25 1 THEN REPEAT (etac exE 1) *})
+  apply (tactic {* cut_tac xsko25 1 THEN REPEAT (eresolve_tac @{context} [exE] 1) *})
   ML_val {*
     val ctxt = @{context};
     val [_, sko25] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
@@ -39,7 +39,7 @@
 
     val ctxt' = fold Thm.declare_hyps (maps (#hyps o Thm.crep_thm) (go25 :: horns25)) ctxt;
     Goal.prove ctxt' [] [] @{prop False} (fn _ =>
-      rtac go25 1 THEN
+      resolve_tac ctxt' [go25] 1 THEN
       Meson.depth_prolog_tac ctxt' horns25);
   *}
   oops
@@ -53,7 +53,7 @@
     val nnf26 = Meson.make_nnf ctxt prem26;
     val xsko26 = Meson.skolemize ctxt nnf26;
   *}
-  apply (tactic {* cut_tac xsko26 1 THEN REPEAT (etac exE 1) *})
+  apply (tactic {* cut_tac xsko26 1 THEN REPEAT (eresolve_tac @{context} [exE] 1) *})
   ML_val {*
     val ctxt = @{context};
     val [_, sko26] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
@@ -65,7 +65,7 @@
 
     val ctxt' = fold Thm.declare_hyps (maps (#hyps o Thm.crep_thm) (go26 :: horns26)) ctxt;
     Goal.prove ctxt' [] [] @{prop False} (fn _ =>
-      rtac go26 1 THEN
+      resolve_tac ctxt' [go26] 1 THEN
       Meson.depth_prolog_tac ctxt' horns26);  (*7 ms*)
     (*Proof is of length 107!!*)
   *}
@@ -80,7 +80,7 @@
     val nnf43 = Meson.make_nnf ctxt prem43;
     val xsko43 = Meson.skolemize ctxt nnf43;
   *}
-  apply (tactic {* cut_tac xsko43 1 THEN REPEAT (etac exE 1) *})
+  apply (tactic {* cut_tac xsko43 1 THEN REPEAT (eresolve_tac @{context} [exE] 1) *})
   ML_val {*
     val ctxt = @{context};
     val [_, sko43] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal})));
@@ -92,7 +92,7 @@
 
     val ctxt' = fold Thm.declare_hyps (maps (#hyps o Thm.crep_thm) (go43 :: horns43)) ctxt;
     Goal.prove ctxt' [] [] @{prop False} (fn _ =>
-      rtac go43 1 THEN
+      resolve_tac ctxt' [go43] 1 THEN
       Meson.best_prolog_tac ctxt' Meson.size_of_subgoals horns43);   (*7ms*)
     *}
   oops
--- a/src/Provers/Arith/assoc_fold.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Provers/Arith/assoc_fold.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -44,8 +44,10 @@
     val (lits, others) = sift_terms plus (lhs, ([],[]))
     val _ = length lits < 2 andalso raise Assoc_fail (*we can't reduce the number of terms*)
     val rhs = plus $ mk_sum plus lits $ mk_sum plus others
-    val th = Goal.prove ctxt [] [] (Logic.mk_equals (lhs, rhs))
-      (fn _ => rtac Data.eq_reflection 1 THEN simp_tac (put_simpset Data.assoc_ss ctxt) 1)
+    val th =
+      Goal.prove ctxt [] [] (Logic.mk_equals (lhs, rhs)) (fn _ =>
+        resolve_tac ctxt [Data.eq_reflection] 1 THEN
+        simp_tac (put_simpset Data.assoc_ss ctxt) 1)
   in SOME th end handle Assoc_fail => NONE;
 
 end;
--- a/src/Provers/Arith/cancel_numeral_factor.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Provers/Arith/cancel_numeral_factor.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -72,7 +72,7 @@
   in
     Option.map (export o Data.simplify_meta_eq ctxt)
       (Data.prove_conv
-         [Data.trans_tac ctxt reshape, rtac Data.cancel 1,
+         [Data.trans_tac ctxt reshape, resolve_tac ctxt [Data.cancel] 1,
           Data.numeral_simp_tac ctxt] ctxt prems (t', rhs))
   end
   (* FIXME avoid handling of generic exceptions *)
--- a/src/Pure/Isar/attrib.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/Isar/attrib.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -527,6 +527,7 @@
   register_config Name_Space.names_unique_raw #>
   register_config ML_Options.source_trace_raw #>
   register_config ML_Options.exception_trace_raw #>
+  register_config ML_Options.debugger_raw #>
   register_config ML_Options.print_depth_raw #>
   register_config Proof_Context.show_abbrevs_raw #>
   register_config Goal_Display.goals_limit_raw #>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/ml_compiler_parameters.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -0,0 +1,17 @@
+(*  Title:      Pure/ML/ml_compiler_parameters.ML
+    Author:     Makarius
+
+Additional ML compiler parameters for Poly/ML.
+*)
+
+signature ML_COMPILER_PARAMETERS =
+sig
+  val debug: bool -> PolyML.Compiler.compilerParameters list
+end;
+
+structure ML_Compiler_Parameters: ML_COMPILER_PARAMETERS =
+struct
+
+fun debug _ = [];
+
+end;
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/ml_compiler_parameters_polyml-5.5.3.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -0,0 +1,12 @@
+(*  Title:      Pure/ML/ml_compiler_parameters_polyml-5.5.3.ML
+    Author:     Makarius
+
+Additional ML compiler parameters for Poly/ML 5.5.3, or later.
+*)
+
+structure ML_Compiler_Parameters: ML_COMPILER_PARAMETERS =
+struct
+
+fun debug b = [PolyML.Compiler.CPDebug b];
+
+end;
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/ml_debugger.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -0,0 +1,48 @@
+(*  Title:      Pure/ML/ml_debugger.ML
+    Author:     Makarius
+
+ML debugger interface.
+*)
+
+signature ML_DEBUGGER =
+sig
+  type location
+  val on_entry: (string * location -> unit) option -> unit
+  val on_exit: (string * location -> unit) option -> unit
+  val on_exit_exception: (string * location -> exn -> unit) option -> unit
+  val on_breakpoint: (location * bool Unsynchronized.ref -> unit) option -> unit
+  type state
+  val state: Thread.thread -> state list
+  val debug_function: state -> string
+  val debug_function_arg: state -> ML_Name_Space.valueVal
+  val debug_function_result: state -> ML_Name_Space.valueVal
+  val debug_location: state -> location
+  val debug_name_space: state -> ML_Name_Space.T
+end;
+
+structure ML_Debugger: ML_DEBUGGER =
+struct
+
+(* hooks *)
+
+type location = unit;
+fun on_entry _ = ();
+fun on_exit _ = ();
+fun on_exit_exception _ = ();
+fun on_breakpoint _ = ();
+
+
+(* debugger *)
+
+fun fail () = raise Fail "No debugger support on this ML platform";
+
+type state = unit;
+
+fun state _ = [];
+fun debug_function () = fail ();
+fun debug_function_arg () = fail ();
+fun debug_function_result () = fail ();
+fun debug_location () = fail ();
+fun debug_name_space () = fail ();
+
+end;
--- a/src/Pure/ML-Systems/ml_debugger_dummy.ML	Fri Jul 17 16:13:03 2015 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,51 +0,0 @@
-(*  Title:      Pure/ML/ml_debugger_dummy.ML
-    Author:     Makarius
-
-ML debugger interface -- dummy version.
-*)
-
-signature ML_DEBUGGER =
-sig
-  type location
-  val on_entry: (string * location -> unit) option -> unit
-  val on_exit: (string * location -> unit) option -> unit
-  val on_exit_exception: (string * location -> exn -> unit) option -> unit
-  val on_break_point: (location * bool Unsynchronized.ref -> unit) option -> unit
-  type state
-  val state: Thread.thread -> state list
-  val debug_function: state -> string
-  val debug_function_arg: state -> ML_Name_Space.valueVal
-  val debug_function_result: state -> ML_Name_Space.valueVal
-  val debug_location: state -> location
-  val debug_name_space: state -> ML_Name_Space.T
-end;
-
-structure ML_Debugger: ML_DEBUGGER =
-struct
-
-(* hooks *)
-
-type location = unit;
-fun on_entry _ = ();
-fun on_exit _ = ();
-fun on_exit_exception _ = ();
-fun on_break_point _ = ();
-
-
-(* debugger *)
-
-fun fail () = raise Fail "No debugger support on this ML platform";
-
-type state = unit;
-
-fun state _ = [];
-fun debug_function () = fail ();
-fun debug_function_arg () = fail ();
-fun debug_function_result () = fail ();
-fun debug_location () = fail ();
-fun debug_name_space () = fail ();
-
-end;
-
-
-fun ml_debugger_parameters (_: bool) = [];
--- a/src/Pure/ML-Systems/ml_debugger_polyml-5.5.3.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML-Systems/ml_debugger_polyml-5.5.3.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -14,7 +14,7 @@
 val on_entry = PolyML.DebuggerInterface.setOnEntry;
 val on_exit = PolyML.DebuggerInterface.setOnExit;
 val on_exit_exception = PolyML.DebuggerInterface.setOnExitException;
-val on_break_point = PolyML.DebuggerInterface.setOnBreakPoint;
+val on_breakpoint = PolyML.DebuggerInterface.setOnBreakPoint;
 
 
 (* debugger operations *)
@@ -29,6 +29,3 @@
 val debug_name_space = PolyML.DebuggerInterface.debugNameSpace;
 
 end;
-
-fun ml_debugger_parameters false = []
-  | ml_debugger_parameters true = [PolyML.Compiler.CPDebug true];
--- a/src/Pure/ML-Systems/ml_parse_tree.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML-Systems/ml_parse_tree.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -7,13 +7,13 @@
 signature ML_PARSE_TREE =
 sig
   val completions: PolyML.ptProperties -> string list option
-  val break_point: PolyML.ptProperties -> bool Unsynchronized.ref option
+  val breakpoint: PolyML.ptProperties -> bool Unsynchronized.ref option
 end;
 
 structure ML_Parse_Tree: ML_PARSE_TREE =
 struct
 
 fun completions _ = NONE;
-fun break_point _ = NONE;
+fun breakpoint _ = NONE;
 
 end;
\ No newline at end of file
--- a/src/Pure/ML-Systems/ml_parse_tree_polyml-5.5.3.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML-Systems/ml_parse_tree_polyml-5.5.3.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -10,7 +10,7 @@
 fun completions (PolyML.PTcompletions x) = SOME x
   | completions _ = NONE;
 
-fun break_point (PolyML.PTbreakPoint x) = SOME x
-  | break_point _ = NONE;
+fun breakpoint (PolyML.PTbreakPoint x) = SOME x
+  | breakpoint _ = NONE;
 
 end;
\ No newline at end of file
--- a/src/Pure/ML-Systems/polyml.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML-Systems/polyml.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -133,12 +133,17 @@
 fun ml_prompts p1 p2 = (PolyML.Compiler.prompt1 := p1; PolyML.Compiler.prompt2 := p2);
 
 use "ML-Systems/ml_parse_tree.ML";
-if ML_System.name = "polyml-5.5.3" then use "ML-Systems/ml_parse_tree_polyml-5.5.3.ML" else ();
+if ML_System.name = "polyml-5.5.3"
+then use "ML-Systems/ml_parse_tree_polyml-5.5.3.ML" else ();
+
+use "ML-Systems/ml_compiler_parameters.ML";
+if ML_System.name = "polyml-5.5.3"
+then use "ML-Systems/ml_compiler_parameters_polyml-5.5.3.ML" else ();
 
 
 (* ML debugger *)
 
-use "ML-Systems/ml_debugger_dummy.ML";
+use "ML-Systems/ml_debugger.ML";
 if ML_System.name = "polyml-5.5.3" then use "ML-Systems/ml_debugger_polyml-5.5.3.ML" else ();
 
 
--- a/src/Pure/ML-Systems/smlnj.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML-Systems/smlnj.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -162,7 +162,7 @@
 
 
 use "ML-Systems/unsynchronized.ML";
-use "ML-Systems/ml_debugger_dummy.ML";
+use "ML-Systems/ml_debugger.ML";
 
 
 (* ML system operations *)
--- a/src/Pure/ML/ml_compiler_polyml.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML/ml_compiler_polyml.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -20,6 +20,9 @@
       | NONE => true);
     fun is_reported pos = is_visible andalso Position.is_reported pos;
 
+
+    (* syntax reports *)
+
     fun reported_types loc types =
       let val pos = Exn_Properties.position_of loc in
         is_reported pos ?
@@ -53,12 +56,12 @@
         else I
       end;
 
-    fun reported loc (PolyML.PTtype types) = reported_types loc types
+    fun reported _ (PolyML.PTnextSibling tree) = reported_tree (tree ())
+      | reported _ (PolyML.PTfirstChild tree) = reported_tree (tree ())
+      | reported loc (PolyML.PTtype types) = reported_types loc types
       | reported loc (PolyML.PTdeclaredAt decl) = reported_entity Markup.ML_defN loc decl
       | reported loc (PolyML.PTopenedAt decl) = reported_entity Markup.ML_openN loc decl
       | reported loc (PolyML.PTstructureAt decl) = reported_entity Markup.ML_structureN loc decl
-      | reported _ (PolyML.PTnextSibling tree) = reported_tree (tree ())
-      | reported _ (PolyML.PTfirstChild tree) = reported_tree (tree ())
       | reported loc pt =
           (case ML_Parse_Tree.completions pt of
             SOME names => reported_completions loc names
@@ -71,14 +74,38 @@
       persistent_reports
       |> map (fn (pos, markup, text) => Position.reported_text pos (markup ()) (text ()))
       |> Output.report;
-  in
-    if not (null persistent_reports) andalso redirect andalso Multithreading.enabled ()
-    then
-      Execution.print
-        {name = "ML_Compiler.report", pos = Position.thread_data (), pri = Task_Queue.urgent_pri}
-        output
-    else output ()
-  end;
+    val _ =
+      if not (null persistent_reports) andalso redirect andalso Multithreading.enabled ()
+      then
+        Execution.print
+          {name = "ML_Compiler.report", pos = Position.thread_data (), pri = Task_Queue.urgent_pri}
+          output
+      else output ();
+
+
+    (* breakpoints *)
+
+    fun breakpoints _ (PolyML.PTnextSibling tree) = breakpoints_tree (tree ())
+      | breakpoints _ (PolyML.PTfirstChild tree) = breakpoints_tree (tree ())
+      | breakpoints loc pt =
+          (case ML_Parse_Tree.breakpoint pt of
+            SOME b =>
+              let val pos = Exn_Properties.position_of loc in
+                if is_reported pos then
+                  let val id = serial ();
+                  in cons ((pos, Markup.ML_breakpoint id), (id, (b, pos))) end
+                else I
+              end
+          | NONE => I)
+    and breakpoints_tree (loc, props) = fold (breakpoints loc) props;
+
+    val all_breakpoints = rev (breakpoints_tree parse_tree []);
+    val _ = Position.reports (map #1 all_breakpoints);
+    val _ =
+      if is_some (Context.thread_data ()) then
+        Context.>> (fold (ML_Env.add_breakpoint o #2) all_breakpoints)
+      else ();
+  in () end;
 
 
 (* eval ML source tokens *)
@@ -86,7 +113,7 @@
 fun eval (flags: flags) pos toks =
   let
     val _ = Secure.secure_mltext ();
-    val space = ML_Env.name_space {SML = #SML flags, exchange = #exchange flags}
+    val space = ML_Env.name_space {SML = #SML flags, exchange = #exchange flags};
     val opt_context = Context.thread_data ();
 
 
@@ -114,7 +141,7 @@
       | [] => Position.none);
 
 
-    (* output channels *)
+    (* output *)
 
     val writeln_buffer = Unsynchronized.ref Buffer.empty;
     fun write s = Unsynchronized.change writeln_buffer (Buffer.add s);
@@ -197,7 +224,7 @@
       PolyML.Compiler.CPPrintDepth ML_Options.get_print_depth,
       PolyML.Compiler.CPCompilerResultFun result_fun,
       PolyML.Compiler.CPPrintInAlphabeticalOrder false] @
-     ml_debugger_parameters (ML_Options.debugger_enabled opt_context);
+     ML_Compiler_Parameters.debug (ML_Options.debugger_enabled opt_context);
     val _ =
       (while not (List.null (! input_buffer)) do
         PolyML.compiler (get, parameters) ())
--- a/src/Pure/ML/ml_env.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ML/ml_env.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -9,6 +9,9 @@
 sig
   val inherit: Context.generic -> Context.generic -> Context.generic
   val forget_structure: string -> Context.generic -> Context.generic
+  val add_breakpoint: serial * (bool Unsynchronized.ref * Position.T) ->
+    Context.generic -> Context.generic
+  val get_breakpoint: Context.generic -> serial -> (bool Unsynchronized.ref * Position.T) option
   val name_space: {SML: bool, exchange: bool} -> ML_Name_Space.T
   val local_context: use_context
   val local_name_space: ML_Name_Space.T
@@ -38,10 +41,14 @@
    Symtab.merge (K true) (signature1, signature2),
    Symtab.merge (K true) (functor1, functor2));
 
-type data = {bootstrap: bool, tables: tables, sml_tables: tables};
+type data =
+ {bootstrap: bool,
+  tables: tables,
+  sml_tables: tables,
+  breakpoints: (bool Unsynchronized.ref * Position.T) Inttab.table};
 
-fun make_data (bootstrap, tables, sml_tables) : data =
-  {bootstrap = bootstrap, tables = tables, sml_tables = sml_tables};
+fun make_data (bootstrap, tables, sml_tables, breakpoints) : data =
+  {bootstrap = bootstrap, tables = tables, sml_tables = sml_tables, breakpoints = breakpoints};
 
 structure Env = Generic_Data
 (
@@ -54,23 +61,32 @@
        Symtab.make ML_Name_Space.initial_fixity,
        Symtab.make ML_Name_Space.initial_structure,
        Symtab.make ML_Name_Space.initial_signature,
-       Symtab.make ML_Name_Space.initial_functor));
-  fun extend (data : T) = make_data (false, #tables data, #sml_tables data);
+       Symtab.make ML_Name_Space.initial_functor),
+      Inttab.empty);
+  fun extend (data : T) = make_data (false, #tables data, #sml_tables data, #breakpoints data);
   fun merge (data : T * T) =
     make_data (false,
       merge_tables (apply2 #tables data),
-      merge_tables (apply2 #sml_tables data));
+      merge_tables (apply2 #sml_tables data),
+      Inttab.merge (K true) (apply2 #breakpoints data));
 );
 
 val inherit = Env.put o Env.get;
 
 fun forget_structure name =
-  Env.map (fn {bootstrap, tables, sml_tables} =>
+  Env.map (fn {bootstrap, tables, sml_tables, breakpoints} =>
     let
       val _ = if bootstrap then ML_Name_Space.forget_global_structure name else ();
       val tables' =
         (#1 tables, #2 tables, #3 tables, Symtab.delete_safe name (#4 tables), #5 tables, #6 tables);
-    in make_data (bootstrap, tables', sml_tables) end);
+    in make_data (bootstrap, tables', sml_tables, breakpoints) end);
+
+fun add_breakpoint breakpoint =
+  Env.map (fn {bootstrap, tables, sml_tables, breakpoints} =>
+    let val breakpoints' = Inttab.update_new breakpoint breakpoints;
+    in make_data (bootstrap, tables, sml_tables, breakpoints') end);
+
+val get_breakpoint = Inttab.lookup o #breakpoints o Env.get;
 
 
 (* name space *)
@@ -98,15 +114,15 @@
 
     fun enter ap1 sel2 entry =
       if SML <> exchange then
-        Context.>> (Env.map (fn {bootstrap, tables, sml_tables} =>
+        Context.>> (Env.map (fn {bootstrap, tables, sml_tables, breakpoints} =>
           let val sml_tables' = ap1 (Symtab.update entry) sml_tables
-          in make_data (bootstrap, tables, sml_tables') end))
+          in make_data (bootstrap, tables, sml_tables', breakpoints) end))
       else if is_some (Context.thread_data ()) then
-        Context.>> (Env.map (fn {bootstrap, tables, sml_tables} =>
+        Context.>> (Env.map (fn {bootstrap, tables, sml_tables, breakpoints} =>
           let
             val _ = if bootstrap then sel2 ML_Name_Space.global entry else ();
             val tables' = ap1 (Symtab.update entry) tables;
-          in make_data (bootstrap, tables', sml_tables) end))
+          in make_data (bootstrap, tables', sml_tables, breakpoints) end))
       else sel2 ML_Name_Space.global entry;
   in
    {lookupVal    = lookup #1 #lookupVal,
@@ -145,4 +161,3 @@
   else error ("Unknown ML functor: " ^ quote name);
 
 end;
-
--- a/src/Pure/PIDE/markup.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/PIDE/markup.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -19,6 +19,8 @@
   val nameN: string
   val name: string -> T -> T
   val kindN: string
+  val serialN: string
+  val serial_properties: int -> Properties.T
   val instanceN: string
   val languageN: string
   val symbolsN: string
@@ -107,6 +109,7 @@
   val ML_openN: string
   val ML_structureN: string
   val ML_typingN: string val ML_typing: T
+  val ML_breakpointN: string val ML_breakpoint: int -> T
   val antiquotedN: string val antiquoted: T
   val antiquoteN: string val antiquote: T
   val ML_antiquotationN: string
@@ -150,8 +153,6 @@
   val runningN: string val running: T
   val finishedN: string val finished: T
   val failedN: string val failed: T
-  val serialN: string
-  val serial_properties: int -> Properties.T
   val exec_idN: string
   val initN: string
   val statusN: string
@@ -268,6 +269,9 @@
 
 val kindN = "kind";
 
+val serialN = "serial";
+fun serial_properties i = [(serialN, print_int i)];
+
 val instanceN = "instance";
 
 
@@ -420,7 +424,7 @@
 val (typingN, typing) = markup_elem "typing";
 
 
-(* ML syntax *)
+(* ML *)
 
 val (ML_keyword1N, ML_keyword1) = markup_elem "ML_keyword1";
 val (ML_keyword2N, ML_keyword2) = markup_elem "ML_keyword2";
@@ -440,6 +444,8 @@
 val ML_structureN = "ML_structure";
 val (ML_typingN, ML_typing) = markup_elem "ML_typing";
 
+val (ML_breakpointN, ML_breakpoint) = markup_int "ML_breakpoint" serialN;
+
 
 (* antiquotations *)
 
@@ -539,9 +545,6 @@
 
 (* messages *)
 
-val serialN = "serial";
-fun serial_properties i = [(serialN, print_int i)];
-
 val exec_idN = "exec_id";
 
 val initN = "init";
--- a/src/Pure/PIDE/markup.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/PIDE/markup.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -42,6 +42,9 @@
   val KIND = "kind"
   val Kind = new Properties.String(KIND)
 
+  val SERIAL = "serial"
+  val Serial = new Properties.Long(SERIAL)
+
   val INSTANCE = "instance"
   val Instance = new Properties.String(INSTANCE)
 
@@ -69,6 +72,15 @@
       if (markup.name == name) Prop.unapply(markup.properties) else None
   }
 
+  class Markup_Long(val name: String, prop: String)
+  {
+    private val Prop = new Properties.Long(prop)
+
+    def apply(i: Long): Markup = Markup(name, Prop(i))
+    def unapply(markup: Markup): Option[Long] =
+      if (markup.name == name) Prop.unapply(markup.properties) else None
+  }
+
 
   /* formal entities */
 
@@ -231,7 +243,7 @@
   val TEXT_FOLD = "text_fold"
 
 
-  /* ML syntax */
+  /* ML */
 
   val ML_KEYWORD1 = "ML_keyword1"
   val ML_KEYWORD2 = "ML_keyword2"
@@ -251,6 +263,9 @@
   val ML_STRUCTURE = "ML_structure"
   val ML_TYPING = "ML_typing"
 
+  val ML_BREAKPOINT = "ML_breakpoint"
+  val ML_Breakpoint = new Markup_Long(ML_BREAKPOINT, SERIAL)
+
 
   /* outer syntax */
 
@@ -341,9 +356,6 @@
 
   /* messages */
 
-  val SERIAL = "serial"
-  val Serial = new Properties.Long(SERIAL)
-
   val INIT = "init"
   val STATUS = "status"
   val REPORT = "report"
--- a/src/Pure/PIDE/session.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/PIDE/session.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -202,9 +202,10 @@
   val phase_changed = new Session.Outlet[Session.Phase](dispatcher)
   val syslog_messages = new Session.Outlet[Prover.Output](dispatcher)
   val raw_output_messages = new Session.Outlet[Prover.Output](dispatcher)
-  val all_messages = new Session.Outlet[Prover.Message](dispatcher)  // potential bottle-neck
   val trace_events = new Session.Outlet[Simplifier_Trace.Event.type](dispatcher)
+  val debugger_events = new Session.Outlet[Debugger.Event.type](dispatcher)
 
+  val all_messages = new Session.Outlet[Prover.Message](dispatcher)  // potential bottle-neck!
 
 
   /** main protocol manager **/
--- a/src/Pure/Pure.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/Pure.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -107,6 +107,7 @@
 ML_file "Tools/find_theorems.ML"
 ML_file "Tools/find_consts.ML"
 ML_file "Tools/simplifier_trace.ML"
+ML_file "Tools/debugger.ML"
 ML_file "Tools/named_theorems.ML"
 
 
--- a/src/Pure/ROOT	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/ROOT	Sat Jul 18 20:59:51 2015 +0200
@@ -6,7 +6,9 @@
     "General/exn.ML"
     "ML-Systems/compiler_polyml.ML"
     "ML-Systems/exn_trace_polyml-5.5.1.ML"
-    "ML-Systems/ml_debugger_dummy.ML"
+    "ML-Systems/ml_compiler_parameters.ML"
+    "ML-Systems/ml_compiler_parameters_polyml-5.5.3.ML"
+    "ML-Systems/ml_debugger.ML"
     "ML-Systems/ml_debugger_polyml-5.5.3.ML"
     "ML-Systems/ml_name_space.ML"
     "ML-Systems/ml_parse_tree.ML"
@@ -37,7 +39,9 @@
     "General/exn.ML"
     "ML-Systems/compiler_polyml.ML"
     "ML-Systems/exn_trace_polyml-5.5.1.ML"
-    "ML-Systems/ml_debugger_dummy.ML"
+    "ML-Systems/ml_compiler_parameters.ML"
+    "ML-Systems/ml_compiler_parameters_polyml-5.5.3.ML"
+    "ML-Systems/ml_debugger.ML"
     "ML-Systems/ml_debugger_polyml-5.5.3.ML"
     "ML-Systems/ml_name_space.ML"
     "ML-Systems/ml_parse_tree.ML"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Tools/debugger.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -0,0 +1,16 @@
+(*  Title:      Pure/Tools/debugger.ML
+    Author:     Makarius
+
+Interactive debugger for Isabelle/ML.
+*)
+
+signature DEBUGGER =
+sig
+end;
+
+structure Debugger: DEBUGGER =
+struct
+
+val _ = Session.protocol_handler "isabelle.Debugger$Handler";
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Tools/debugger.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -0,0 +1,44 @@
+/*  Title:      Pure/Tools/debugger.scala
+    Author:     Makarius
+
+Interactive debugger for Isabelle/ML.
+*/
+
+package isabelle
+
+
+object Debugger
+{
+  /* GUI interaction */
+
+  case object Event
+
+
+  /* manager thread */
+
+  private lazy val manager: Consumer_Thread[Any] =
+    Consumer_Thread.fork[Any]("Debugger.manager", daemon = true)(
+      consume = (arg: Any) =>
+      {
+        // FIXME
+        true
+      },
+      finish = () =>
+      {
+        // FIXME
+      }
+    )
+
+
+  /* protocol handler */
+
+  class Handler extends Session.Protocol_Handler
+  {
+    override def stop(prover: Prover)
+    {
+      manager.shutdown()
+    }
+
+    val functions = Map.empty[String, (Prover, Prover.Protocol_Output) => Boolean]  // FIXME
+  }
+}
--- a/src/Pure/Tools/simplifier_trace.ML	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/Tools/simplifier_trace.ML	Sat Jul 18 20:59:51 2015 +0200
@@ -395,17 +395,17 @@
 
 val _ =
   Isabelle_Process.protocol_command "Simplifier_Trace.reply"
-    (fn [s, r] =>
+    (fn [serial_string, reply] =>
       let
-        val serial = Markup.parse_int s
-        fun lookup_delete tab =
-          (Inttab.lookup tab serial, Inttab.delete_safe serial tab)
-        fun apply_result (SOME promise) = Future.fulfill promise r
-          | apply_result NONE = () (* FIXME handle protocol failure, just like in active.ML? *)
+        val serial = Markup.parse_int serial_string
+        val result =
+          Synchronized.change_result futures
+            (fn tab => (Inttab.lookup tab serial, Inttab.delete_safe serial tab))
       in
-        (Synchronized.change_result futures lookup_delete |> apply_result)
-          handle exn => if Exn.is_interrupt exn then () (*sic!*) else reraise exn
-      end)
+        (case result of
+          SOME promise => Future.fulfill promise reply
+        | NONE => ()) (* FIXME handle protocol failure, just like in active.ML (!?) *)
+      end handle exn => if Exn.is_interrupt exn then () (*sic!*) else reraise exn)
 
 
 
--- a/src/Pure/build-jars	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Pure/build-jars	Sat Jul 18 20:59:51 2015 +0200
@@ -94,6 +94,7 @@
   Tools/build_doc.scala
   Tools/check_keywords.scala
   Tools/check_source.scala
+  Tools/debugger.scala
   Tools/doc.scala
   Tools/main.scala
   Tools/ml_statistics.scala
--- a/src/Sequents/LK.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Sequents/LK.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -178,7 +178,7 @@
   apply (lem p1)
   apply safe
    apply (tactic {*
-     REPEAT (rtac @{thm cut} 1 THEN
+     REPEAT (resolve_tac @{context} @{thms cut} 1 THEN
        DEPTH_SOLVE_1
          (resolve_tac @{context} [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
            Cla.safe_tac @{context} 1) *})
@@ -191,7 +191,7 @@
   apply (lem p1)
   apply safe
    apply (tactic {*
-     REPEAT (rtac @{thm cut} 1 THEN
+     REPEAT (resolve_tac @{context} @{thms cut} 1 THEN
        DEPTH_SOLVE_1
          (resolve_tac @{context} [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
            Cla.safe_tac @{context} 1) *})
--- a/src/Sequents/LK0.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Sequents/LK0.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -154,12 +154,12 @@
 (*Cut and thin, replacing the right-side formula*)
 fun cutR_tac ctxt s i =
   Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN
-  rtac @{thm thinR} i
+  resolve_tac ctxt @{thms thinR} i
 
 (*Cut and thin, replacing the left-side formula*)
 fun cutL_tac ctxt s i =
   Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN
-  rtac @{thm thinL} (i + 1)
+  resolve_tac ctxt @{thms thinL} (i + 1)
 *}
 
 
@@ -244,11 +244,11 @@
   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack LK_dup_pack ctxt))) *}
 
 method_setup lem = {*
-  Attrib.thm >> (fn th => fn _ =>
+  Attrib.thm >> (fn th => fn ctxt =>
     SIMPLE_METHOD' (fn i =>
-      rtac (@{thm thinR} RS @{thm cut}) i THEN
-      REPEAT (rtac @{thm thinL} i) THEN
-      rtac th i))
+      resolve_tac ctxt [@{thm thinR} RS @{thm cut}] i THEN
+      REPEAT (resolve_tac ctxt @{thms thinL} i) THEN
+      resolve_tac ctxt [th] i))
 *}
 
 
--- a/src/Tools/jEdit/lib/Tools/jedit	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/lib/Tools/jedit	Sat Jul 18 20:59:51 2015 +0200
@@ -12,6 +12,7 @@
   "src/bibtex_jedit.scala"
   "src/completion_popup.scala"
   "src/context_menu.scala"
+  "src/debugger_dockable.scala"
   "src/dockable.scala"
   "src/document_model.scala"
   "src/document_view.scala"
--- a/src/Tools/jEdit/src/Isabelle.props	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/Isabelle.props	Sat Jul 18 20:59:51 2015 +0200
@@ -30,6 +30,7 @@
 #menu actions and dockables
 plugin.isabelle.jedit.Plugin.menu.label=Isabelle
 plugin.isabelle.jedit.Plugin.menu= \
+  isabelle-debugger \
   isabelle-documentation \
   isabelle-monitor \
   isabelle-output \
@@ -42,6 +43,8 @@
   isabelle-syslog \
   isabelle-theories \
   isabelle-timing
+isabelle-debugger.label=Debugger panel
+isabelle-debugger.title=Debugger
 isabelle-documentation.label=Documentation panel
 isabelle-documentation.title=Documentation
 isabelle-graphview.label=Graphview panel
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/debugger_dockable.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -0,0 +1,105 @@
+/*  Title:      Tools/jEdit/src/debugger_dockable.scala
+    Author:     Makarius
+
+Dockable window for Isabelle/ML debugger.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.awt.BorderLayout
+import java.awt.event.{ComponentEvent, ComponentAdapter}
+
+import org.gjt.sp.jedit.View
+
+
+class Debugger_Dockable(view: View, position: String) extends Dockable(view, position)
+{
+  GUI_Thread.require {}
+
+
+  /* component state -- owned by GUI thread */
+
+  private var current_snapshot = Document.Snapshot.init
+  private var current_output: List[XML.Tree] = Nil
+
+
+  /* pretty text area */
+
+  val pretty_text_area = new Pretty_Text_Area(view)
+  set_content(pretty_text_area)
+
+  override def detach_operation = pretty_text_area.detach_operation
+
+
+  private val zoom = new Font_Info.Zoom_Box { def changed = handle_resize() }
+
+  private def handle_resize()
+  {
+    GUI_Thread.require {}
+
+    pretty_text_area.resize(
+      Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100))
+  }
+
+  private def handle_update()
+  {
+    GUI_Thread.require {}
+
+    val new_snapshot = PIDE.editor.current_node_snapshot(view).getOrElse(current_snapshot)
+    val new_output = List(XML.Text("FIXME"))
+
+    if (new_output != current_output)
+      pretty_text_area.update(new_snapshot, Command.Results.empty, Pretty.separate(new_output))
+
+    current_snapshot = new_snapshot
+    current_output = new_output
+  }
+
+
+  /* main */
+
+  private val main =
+    Session.Consumer[Any](getClass.getName) {
+      case _: Session.Global_Options =>
+        GUI_Thread.later { handle_resize() }
+
+      case Debugger.Event =>
+        GUI_Thread.later { handle_update() }
+    }
+
+  override def init()
+  {
+    PIDE.session.global_options += main
+    PIDE.session.debugger_events += main
+    handle_update()
+  }
+
+  override def exit()
+  {
+    PIDE.session.global_options -= main
+    PIDE.session.debugger_events -= main
+    delay_resize.revoke()
+  }
+
+
+  /* resize */
+
+  private val delay_resize =
+    GUI_Thread.delay_first(PIDE.options.seconds("editor_update_delay")) { handle_resize() }
+
+  addComponentListener(new ComponentAdapter {
+    override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
+    override def componentShown(e: ComponentEvent) { delay_resize.invoke() }
+  })
+
+
+  /* controls */
+
+  private val controls =
+    new Wrap_Panel(Wrap_Panel.Alignment.Right)(
+      pretty_text_area.search_label, pretty_text_area.search_field, zoom)
+  add(controls.peer, BorderLayout.NORTH)
+}
--- a/src/Tools/jEdit/src/dockables.xml	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/dockables.xml	Sat Jul 18 20:59:51 2015 +0200
@@ -2,6 +2,9 @@
 <!DOCTYPE DOCKABLES SYSTEM "dockables.dtd">
 
 <DOCKABLES>
+	<DOCKABLE NAME="isabelle-debugger" MOVABLE="TRUE">
+		new isabelle.jedit.Debugger_Dockable(view, position);
+	</DOCKABLE>
 	<DOCKABLE NAME="isabelle-documentation" MOVABLE="TRUE">
 		new isabelle.jedit.Documentation_Dockable(view, position);
 	</DOCKABLE>
--- a/src/Tools/jEdit/src/info_dockable.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/info_dockable.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -86,6 +86,7 @@
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
+    override def componentShown(e: ComponentEvent) { delay_resize.invoke() }
   })
 
   private val controls = new Wrap_Panel(Wrap_Panel.Alignment.Right)(
--- a/src/Tools/jEdit/src/isabelle.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/isabelle.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -93,6 +93,12 @@
 
   private def wm(view: View): DockableWindowManager = view.getDockableWindowManager
 
+  def debugger_dockable(view: View): Option[Debugger_Dockable] =
+    wm(view).getDockableWindow("isabelle-debugger") match {
+      case dockable: Debugger_Dockable => Some(dockable)
+      case _ => None
+    }
+
   def documentation_dockable(view: View): Option[Documentation_Dockable] =
     wm(view).getDockableWindow("isabelle-documentation") match {
       case dockable: Documentation_Dockable => Some(dockable)
--- a/src/Tools/jEdit/src/jEdit.props	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/jEdit.props	Sat Jul 18 20:59:51 2015 +0200
@@ -185,6 +185,7 @@
 home.shortcut=
 insert-newline-indent.shortcut=
 insert-newline.shortcut=ENTER
+isabelle-debugger.dock-position=floating
 isabelle-documentation.dock-position=right
 isabelle-output.dock-position=bottom
 isabelle-output.height=174
--- a/src/Tools/jEdit/src/output_dockable.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/output_dockable.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -122,6 +122,7 @@
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
+    override def componentShown(e: ComponentEvent) { delay_resize.invoke() }
   })
 
 
--- a/src/Tools/jEdit/src/query_dockable.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/query_dockable.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -319,6 +319,7 @@
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
+    override def componentShown(e: ComponentEvent) { delay_resize.invoke() }
   })
 
 
--- a/src/Tools/jEdit/src/simplifier_trace_dockable.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/simplifier_trace_dockable.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -22,6 +22,7 @@
 {
   GUI_Thread.require {}
 
+
   /* component state -- owned by GUI thread */
 
   private var current_snapshot = Document.State.init.snapshot()
@@ -36,9 +37,6 @@
 
   private def update_contents()
   {
-
-    GUI_Thread.require {}
-
     val snapshot = current_snapshot
     val context = Simplifier_Trace.handle_results(PIDE.session, current_id, current_results)
 
@@ -125,8 +123,6 @@
 
   override def init()
   {
-    GUI_Thread.require {}
-
     PIDE.session.global_options += main
     PIDE.session.commands_changed += main
     PIDE.session.caret_focus += main
@@ -136,8 +132,6 @@
 
   override def exit()
   {
-    GUI_Thread.require {}
-
     PIDE.session.global_options -= main
     PIDE.session.commands_changed -= main
     PIDE.session.caret_focus -= main
--- a/src/Tools/jEdit/src/sledgehammer_dockable.scala	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/Tools/jEdit/src/sledgehammer_dockable.scala	Sat Jul 18 20:59:51 2015 +0200
@@ -66,6 +66,7 @@
 
   addComponentListener(new ComponentAdapter {
     override def componentResized(e: ComponentEvent) { delay_resize.invoke() }
+    override def componentShown(e: ComponentEvent) { delay_resize.invoke() }
   })
 
 
--- a/src/ZF/UNITY/Constrains.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/ZF/UNITY/Constrains.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -460,7 +460,9 @@
 {*
 (*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
 fun Always_Int_tac ctxt =
-  dtac @{thm Always_Int_I} THEN' assume_tac ctxt THEN' etac @{thm Always_thin};
+  dresolve_tac ctxt @{thms Always_Int_I} THEN'
+  assume_tac ctxt THEN'
+  eresolve_tac ctxt @{thms Always_thin};
 
 (*Combines a list of invariance THEOREMS into one.*)
 val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I});
@@ -470,26 +472,28 @@
 fun constrains_tac ctxt =
    SELECT_GOAL
       (EVERY [REPEAT (Always_Int_tac ctxt 1),
-              REPEAT (etac @{thm Always_ConstrainsI} 1
+              REPEAT (eresolve_tac ctxt @{thms Always_ConstrainsI} 1
                       ORELSE
                       resolve_tac ctxt [@{thm StableI}, @{thm stableI},
                                    @{thm constrains_imp_Constrains}] 1),
-              rtac @{thm constrainsI} 1,
+              resolve_tac ctxt @{thms constrainsI} 1,
               (* Three subgoals *)
               rewrite_goal_tac ctxt [@{thm st_set_def}] 3,
               REPEAT (force_tac ctxt 2),
               full_simp_tac (ctxt addsimps (Named_Theorems.get ctxt @{named_theorems program})) 1,
               ALLGOALS (clarify_tac ctxt),
-              REPEAT (FIRSTGOAL (etac @{thm disjE})),
+              REPEAT (FIRSTGOAL (eresolve_tac ctxt @{thms disjE})),
               ALLGOALS (clarify_tac ctxt),
-              REPEAT (FIRSTGOAL (etac @{thm disjE})),
+              REPEAT (FIRSTGOAL (eresolve_tac ctxt @{thms disjE})),
               ALLGOALS (clarify_tac ctxt),
               ALLGOALS (asm_full_simp_tac ctxt),
               ALLGOALS (clarify_tac ctxt)]);
 
 (*For proving invariants*)
 fun always_tac ctxt i =
-    rtac @{thm AlwaysI} i THEN force_tac ctxt i THEN constrains_tac ctxt i;
+  resolve_tac ctxt @{thms AlwaysI} i THEN
+  force_tac ctxt i
+  THEN constrains_tac ctxt i;
 *}
 
 method_setup safety = {*
--- a/src/ZF/UNITY/SubstAx.thy	Fri Jul 17 16:13:03 2015 +0200
+++ b/src/ZF/UNITY/SubstAx.thy	Sat Jul 18 20:59:51 2015 +0200
@@ -353,7 +353,7 @@
 fun ensures_tac ctxt sact =
   SELECT_GOAL
     (EVERY [REPEAT (Always_Int_tac ctxt 1),
-            etac @{thm Always_LeadsTo_Basis} 1
+            eresolve_tac ctxt @{thms Always_LeadsTo_Basis} 1
                 ORELSE   (*subgoal may involve LeadsTo, leadsTo or ensures*)
                 REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
                                   @{thm EnsuresI}, @{thm ensuresI}] 1),
@@ -364,10 +364,11 @@
                (*simplify the command's domain*)
             simp_tac (ctxt addsimps [@{thm domain_def}]) 3,
             (* proving the domain part *)
-           clarify_tac ctxt 3, dtac @{thm swap} 3, force_tac ctxt 4,
-           rtac @{thm ReplaceI} 3, force_tac ctxt 3, force_tac ctxt 4,
-           asm_full_simp_tac ctxt 3, rtac @{thm conjI} 3, simp_tac ctxt 4,
-           REPEAT (rtac @{thm state_update_type} 3),
+           clarify_tac ctxt 3,
+           dresolve_tac ctxt @{thms swap} 3, force_tac ctxt 4,
+           resolve_tac ctxt @{thms ReplaceI} 3, force_tac ctxt 3, force_tac ctxt 4,
+           asm_full_simp_tac ctxt 3, resolve_tac ctxt @{thms conjI} 3, simp_tac ctxt 4,
+           REPEAT (resolve_tac ctxt @{thms state_update_type} 3),
            constrains_tac ctxt 1,
            ALLGOALS (clarify_tac ctxt),
            ALLGOALS (asm_full_simp_tac (ctxt addsimps [@{thm st_set_def}])),