diff -r 0c5cd369a643 -r 50b60f501b05 src/HOL/Bali/Evaln.thy --- a/src/HOL/Bali/Evaln.thy Tue Feb 10 14:29:36 2015 +0100 +++ b/src/HOL/Bali/Evaln.thy Tue Feb 10 14:48:26 2015 +0100 @@ -448,9 +448,9 @@ lemma evaln_nonstrict [rule_format (no_asm), elim]: "G\s \t\\n\ (w, s') \ \m. n\m \ G\s \t\\m\ (w, s')" apply (erule evaln.induct) -apply (tactic {* ALLGOALS (EVERY'[strip_tac, TRY o etac @{thm Suc_le_D_lemma}, +apply (tactic {* ALLGOALS (EVERY' [strip_tac @{context}, TRY o etac @{thm Suc_le_D_lemma}, REPEAT o smp_tac @{context} 1, - resolve_tac @{thms evaln.intros} THEN_ALL_NEW TRY o atac]) *}) + resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW TRY o atac]) *}) (* 3 subgoals *) apply (auto split del: split_if) done