src/HOL/TLA/Memory/MemoryImplementation.thy
changeset 27100 889613625e2b
parent 26342 0f65fa163304
child 27117 97e9dae57284
--- a/src/HOL/TLA/Memory/MemoryImplementation.thy	Mon Jun 09 17:31:25 2008 +0200
+++ b/src/HOL/TLA/Memory/MemoryImplementation.thy	Mon Jun 09 17:39:35 2008 +0200
@@ -764,7 +764,7 @@
 ML {*
 fun split_idle_tac ss simps i =
     EVERY [TRY (rtac @{thm actionI} i),
-           case_tac "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" i,
+           DatatypePackage.case_tac "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" i,
            rewrite_goals_tac @{thms action_rews},
            forward_tac [temp_use @{thm Step1_4_7}] i,
            asm_full_simp_tac (ss addsimps simps) i