src/HOL/TLA/Memory/MemoryImplementation.thy
changeset 59802 684cfaa12e47
parent 59498 50b60f501b05
child 59826 442b09c0f898
--- a/src/HOL/TLA/Memory/MemoryImplementation.thy	Tue Mar 24 20:07:27 2015 +0100
+++ b/src/HOL/TLA/Memory/MemoryImplementation.thy	Tue Mar 24 21:54:25 2015 +0100
@@ -798,7 +798,7 @@
 fun split_idle_tac ctxt =
   SELECT_GOAL
    (TRY (rtac @{thm actionI} 1) THEN
-    Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" 1 THEN
+    Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" [] 1 THEN
     rewrite_goals_tac ctxt @{thms action_rews} THEN
     forward_tac ctxt [temp_use ctxt @{thm Step1_4_7}] 1 THEN
     asm_full_simp_tac ctxt 1);