diff -r a292751fb345 -r 43753eca324a src/HOL/IMP/Comp_Rev.thy --- a/src/HOL/IMP/Comp_Rev.thy Tue Nov 13 12:12:14 2012 +0100 +++ b/src/HOL/IMP/Comp_Rev.thy Wed Nov 14 14:11:47 2012 +0100 @@ -35,7 +35,7 @@ lemma exec_n_exec: "P \ c \^n c' \ P \ c \* c'" - by (induct n arbitrary: c) auto + by (induct n arbitrary: c) (auto simp del: exec1_def) lemma exec_0 [intro!]: "P \ c \^0 c" by simp @@ -45,7 +45,7 @@ lemma exec_exec_n: "P \ c \* c' \ \n. P \ c \^n c'" - by (induct rule: exec.induct) (auto intro: exec_Suc) + by (induct rule: exec.induct) (auto simp del: exec1_def intro: exec_Suc) lemma exec_eq_exec_n: "(P \ c \* c') = (\n. P \ c \^n c')" @@ -132,9 +132,9 @@ qed lemma succs_iexec1: - assumes "P!!i \i (i,s,stk) \ c'" "0 \ i" "i < isize P" + assumes "c' = iexec (P!!i) (i,s,stk)" "0 \ i" "i < isize P" shows "fst c' \ succs P 0" - using assms by (auto elim!: iexec1.cases simp: succs_def isuccs_def) + using assms by (auto simp: succs_def isuccs_def split: instr.split) lemma succs_shift: "(p - n \ succs P 0) = (p \ succs P n)" @@ -253,7 +253,7 @@ lemma exec1_split: "P @ c @ P' \ (isize P + i, s) \ (j,s') \ 0 \ i \ i < isize c \ c \ (i,s) \ (j - isize P, s')" - by (auto elim!: iexec1.cases) + by (auto split: instr.splits) lemma exec_n_split: assumes "P @ c @ P' \ (isize P + i, s) \^n (j, s')" @@ -294,7 +294,7 @@ assume "j0 \ {0 ..< isize c}" moreover from c j0 have "j0 \ succs c 0" - by (auto dest: succs_iexec1) + by (auto dest: succs_iexec1 simp del: iexec.simps) ultimately have "j0 \ exits c" by (simp add: exits_def) with c j0 rest @@ -331,7 +331,7 @@ have "n = isize P1 + (n - isize P1)" by simp then obtain n' where "n = isize P1 + n'" .. ultimately - show ?thesis using assms by clarsimp + show ?thesis using assms by (clarsimp simp del: iexec.simps) qed lemma exec_n_drop_left: @@ -346,13 +346,13 @@ obtain i' s'' stk'' where step: "P @ P' \ (i, s, stk) \ (i', s'', stk'')" and rest: "P @ P' \ (i', s'', stk'') \^k (n, s', stk')" - by auto + by (auto simp del: exec1_def) from step `isize P \ i` have "P' \ (i - isize P, s, stk) \ (i' - isize P, s'', stk'')" by (rule exec1_drop_left) moreover then have "i' - isize P \ succs P' 0" - by (fastforce dest!: succs_iexec1) + by (fastforce dest!: succs_iexec1 simp del: iexec.simps) with `exits P' \ {0..}` have "isize P \ i'" by (auto simp: exits_def) from rest this `exits P' \ {0..}`