src/HOL/TLA/Memory/MemoryImplementation.thy
changeset 45133 2214ba5bdfff
parent 43596 78211f66cf8d
child 45605 a89b4bc311a5
equal deleted inserted replaced
45132:09de0d0de645 45133:2214ba5bdfff
   792 *)
   792 *)
   793 ML {*
   793 ML {*
   794 fun split_idle_tac ctxt =
   794 fun split_idle_tac ctxt =
   795   SELECT_GOAL
   795   SELECT_GOAL
   796    (TRY (rtac @{thm actionI} 1) THEN
   796    (TRY (rtac @{thm actionI} 1) THEN
   797     InductTacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" 1 THEN
   797     Induct_Tacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" 1 THEN
   798     rewrite_goals_tac @{thms action_rews} THEN
   798     rewrite_goals_tac @{thms action_rews} THEN
   799     forward_tac [temp_use @{thm Step1_4_7}] 1 THEN
   799     forward_tac [temp_use @{thm Step1_4_7}] 1 THEN
   800     asm_full_simp_tac (simpset_of ctxt) 1);
   800     asm_full_simp_tac (simpset_of ctxt) 1);
   801 *}
   801 *}
   802 
   802