author | paulson <lp15@cam.ac.uk> |
Fri, 10 Apr 2020 22:50:59 +0100 | |
changeset 71743 | 0239bee6bffd |
parent 69597 | ff784d5a5bfb |
child 80612 | e65eed943bee |
permissions | -rw-r--r-- |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
1 |
(* Author: Gerwin Klein *) |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
2 |
|
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
68484
diff
changeset
|
3 |
section \<open>Compiler Correctness, Reverse Direction\<close> |
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
68484
diff
changeset
|
4 |
|
52400 | 5 |
theory Compiler2 |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
6 |
imports Compiler |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
7 |
begin |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
8 |
|
67406 | 9 |
text \<open> |
56224 | 10 |
The preservation of the source code semantics is already shown in the |
69505 | 11 |
parent theory \<open>Compiler\<close>. This here shows the second direction. |
67406 | 12 |
\<close> |
56224 | 13 |
|
67406 | 14 |
subsection \<open>Definitions\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
15 |
|
69597 | 16 |
text \<open>Execution in \<^term>\<open>n\<close> steps for simpler induction\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
17 |
primrec |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
18 |
exec_n :: "instr list \<Rightarrow> config \<Rightarrow> nat \<Rightarrow> config \<Rightarrow> bool" |
44000
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
19 |
("_/ \<turnstile> (_ \<rightarrow>^_/ _)" [65,0,1000,55] 55) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
20 |
where |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
21 |
"P \<turnstile> c \<rightarrow>^0 c' = (c'=c)" | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
22 |
"P \<turnstile> c \<rightarrow>^(Suc n) c'' = (\<exists>c'. (P \<turnstile> c \<rightarrow> c') \<and> P \<turnstile> c' \<rightarrow>^n c'')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
23 |
|
69597 | 24 |
text \<open>The possible successor PCs of an instruction at position \<^term>\<open>n\<close>\<close> |
67406 | 25 |
text_raw\<open>\snip{isuccsdef}{0}{1}{%\<close> |
53911 | 26 |
definition isuccs :: "instr \<Rightarrow> int \<Rightarrow> int set" where |
27 |
"isuccs i n = (case i of |
|
28 |
JMP j \<Rightarrow> {n + 1 + j} | |
|
29 |
JMPLESS j \<Rightarrow> {n + 1 + j, n + 1} | |
|
30 |
JMPGE j \<Rightarrow> {n + 1 + j, n + 1} | |
|
31 |
_ \<Rightarrow> {n +1})" |
|
67406 | 32 |
text_raw\<open>}%endsnip\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
33 |
|
67406 | 34 |
text \<open>The possible successors PCs of an instruction list\<close> |
53911 | 35 |
definition succs :: "instr list \<Rightarrow> int \<Rightarrow> int set" where |
36 |
"succs P n = {s. \<exists>i::int. 0 \<le> i \<and> i < size P \<and> s \<in> isuccs (P!!i) (n+i)}" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
37 |
|
67406 | 38 |
text \<open>Possible exit PCs of a program\<close> |
53911 | 39 |
definition exits :: "instr list \<Rightarrow> int set" where |
40 |
"exits P = succs P 0 - {0..< size P}" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
41 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
42 |
|
69597 | 43 |
subsection \<open>Basic properties of \<^term>\<open>exec_n\<close>\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
44 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
45 |
lemma exec_n_exec: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
46 |
"P \<turnstile> c \<rightarrow>^n c' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c'" |
61147 | 47 |
by (induct n arbitrary: c) (auto intro: star.step) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
48 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
49 |
lemma exec_0 [intro!]: "P \<turnstile> c \<rightarrow>^0 c" by simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
50 |
|
45218 | 51 |
lemma exec_Suc: |
44000
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
52 |
"\<lbrakk> P \<turnstile> c \<rightarrow> c'; P \<turnstile> c' \<rightarrow>^n c'' \<rbrakk> \<Longrightarrow> P \<turnstile> c \<rightarrow>^(Suc n) c''" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
53 |
by (fastforce simp del: split_paired_Ex) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
54 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
55 |
lemma exec_exec_n: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
56 |
"P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> \<exists>n. P \<turnstile> c \<rightarrow>^n c'" |
52915 | 57 |
by (induct rule: star.induct) (auto intro: exec_Suc) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
58 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
59 |
lemma exec_eq_exec_n: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
60 |
"(P \<turnstile> c \<rightarrow>* c') = (\<exists>n. P \<turnstile> c \<rightarrow>^n c')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
61 |
by (blast intro: exec_exec_n exec_n_exec) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
62 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
63 |
lemma exec_n_Nil [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
64 |
"[] \<turnstile> c \<rightarrow>^k c' = (c' = c \<and> k = 0)" |
52915 | 65 |
by (induct k) (auto simp: exec1_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
66 |
|
44004 | 67 |
lemma exec1_exec_n [intro!]: |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
68 |
"P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c \<rightarrow>^1 c'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
69 |
by (cases c') simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
70 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
71 |
|
67406 | 72 |
subsection \<open>Concrete symbolic execution steps\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
73 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
74 |
lemma exec_n_step: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
75 |
"n \<noteq> n' \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
76 |
P \<turnstile> (n,stk,s) \<rightarrow>^k (n',stk',s') = |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
77 |
(\<exists>c. P \<turnstile> (n,stk,s) \<rightarrow> c \<and> P \<turnstile> c \<rightarrow>^(k - 1) (n',stk',s') \<and> 0 < k)" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
78 |
by (cases k) auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
79 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
80 |
lemma exec1_end: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
81 |
"size P <= fst c \<Longrightarrow> \<not> P \<turnstile> c \<rightarrow> c'" |
52915 | 82 |
by (auto simp: exec1_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
83 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
84 |
lemma exec_n_end: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
85 |
"size P <= (n::int) \<Longrightarrow> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
86 |
P \<turnstile> (n,s,stk) \<rightarrow>^k (n',s',stk') = (n' = n \<and> stk'=stk \<and> s'=s \<and> k =0)" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
87 |
by (cases k) (auto simp: exec1_end) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
88 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
89 |
lemmas exec_n_simps = exec_n_step exec_n_end |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
90 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
91 |
|
69597 | 92 |
subsection \<open>Basic properties of \<^term>\<open>succs\<close>\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
93 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
94 |
lemma succs_simps [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
95 |
"succs [ADD] n = {n + 1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
96 |
"succs [LOADI v] n = {n + 1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
97 |
"succs [LOAD x] n = {n + 1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
98 |
"succs [STORE x] n = {n + 1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
99 |
"succs [JMP i] n = {n + 1 + i}" |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45218
diff
changeset
|
100 |
"succs [JMPGE i] n = {n + 1 + i, n + 1}" |
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45218
diff
changeset
|
101 |
"succs [JMPLESS i] n = {n + 1 + i, n + 1}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
102 |
by (auto simp: succs_def isuccs_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
103 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
104 |
lemma succs_empty [iff]: "succs [] n = {}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
105 |
by (simp add: succs_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
106 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
107 |
lemma succs_Cons: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
108 |
"succs (x#xs) n = isuccs x n \<union> succs xs (1+n)" (is "_ = ?x \<union> ?xs") |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
109 |
proof |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
110 |
let ?isuccs = "\<lambda>p P n i::int. 0 \<le> i \<and> i < size P \<and> p \<in> isuccs (P!!i) (n+i)" |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
111 |
have "p \<in> ?x \<union> ?xs" if assm: "p \<in> succs (x#xs) n" for p |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
112 |
proof - |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
113 |
from assm obtain i::int where isuccs: "?isuccs p (x#xs) n i" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
114 |
unfolding succs_def by auto |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
115 |
show ?thesis |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
116 |
proof cases |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
117 |
assume "i = 0" with isuccs show ?thesis by simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
118 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
119 |
assume "i \<noteq> 0" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
120 |
with isuccs |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
121 |
have "?isuccs p xs (1+n) (i - 1)" by auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
122 |
hence "p \<in> ?xs" unfolding succs_def by blast |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
123 |
thus ?thesis .. |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
124 |
qed |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
125 |
qed |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
126 |
thus "succs (x#xs) n \<subseteq> ?x \<union> ?xs" .. |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
127 |
|
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
128 |
have "p \<in> succs (x#xs) n" if assm: "p \<in> ?x \<or> p \<in> ?xs" for p |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
129 |
proof - |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
130 |
from assm show ?thesis |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
131 |
proof |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
132 |
assume "p \<in> ?x" thus ?thesis by (fastforce simp: succs_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
133 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
134 |
assume "p \<in> ?xs" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
135 |
then obtain i where "?isuccs p xs (1+n) i" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
136 |
unfolding succs_def by auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
137 |
hence "?isuccs p (x#xs) n (1+i)" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
138 |
by (simp add: algebra_simps) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
139 |
thus ?thesis unfolding succs_def by blast |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
140 |
qed |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
141 |
qed |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
142 |
thus "?x \<union> ?xs \<subseteq> succs (x#xs) n" by blast |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
143 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
144 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
145 |
lemma succs_iexec1: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
146 |
assumes "c' = iexec (P!!i) (i,s,stk)" "0 \<le> i" "i < size P" |
44004 | 147 |
shows "fst c' \<in> succs P 0" |
50060
43753eca324a
replaced relation by function - simplifies development
nipkow
parents:
47818
diff
changeset
|
148 |
using assms by (auto simp: succs_def isuccs_def split: instr.split) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
149 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
150 |
lemma succs_shift: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
151 |
"(p - n \<in> succs P 0) = (p \<in> succs P n)" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
152 |
by (fastforce simp: succs_def isuccs_def split: instr.split) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
153 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
154 |
lemma inj_op_plus [simp]: |
67399 | 155 |
"inj ((+) (i::int))" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
156 |
by (metis add_minus_cancel inj_on_inverseI) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
157 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
158 |
lemma succs_set_shift [simp]: |
67399 | 159 |
"(+) i ` succs xs 0 = succs xs i" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
160 |
by (force simp: succs_shift [where n=i, symmetric] intro: set_eqI) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
161 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
162 |
lemma succs_append [simp]: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
163 |
"succs (xs @ ys) n = succs xs n \<union> succs ys (n + size xs)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
164 |
by (induct xs arbitrary: n) (auto simp: succs_Cons algebra_simps) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
165 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
166 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
167 |
lemma exits_append [simp]: |
67399 | 168 |
"exits (xs @ ys) = exits xs \<union> ((+) (size xs)) ` exits ys - |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
169 |
{0..<size xs + size ys}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
170 |
by (auto simp: exits_def image_set_diff) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
171 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
172 |
lemma exits_single: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
173 |
"exits [x] = isuccs x 0 - {0}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
174 |
by (auto simp: exits_def succs_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
175 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
176 |
lemma exits_Cons: |
67399 | 177 |
"exits (x # xs) = (isuccs x 0 - {0}) \<union> ((+) 1) ` exits xs - |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
178 |
{0..<1 + size xs}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
179 |
using exits_append [of "[x]" xs] |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
180 |
by (simp add: exits_single) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
181 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
182 |
lemma exits_empty [iff]: "exits [] = {}" by (simp add: exits_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
183 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
184 |
lemma exits_simps [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
185 |
"exits [ADD] = {1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
186 |
"exits [LOADI v] = {1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
187 |
"exits [LOAD x] = {1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
188 |
"exits [STORE x] = {1}" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
189 |
"i \<noteq> -1 \<Longrightarrow> exits [JMP i] = {1 + i}" |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45218
diff
changeset
|
190 |
"i \<noteq> -1 \<Longrightarrow> exits [JMPGE i] = {1 + i, 1}" |
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45218
diff
changeset
|
191 |
"i \<noteq> -1 \<Longrightarrow> exits [JMPLESS i] = {1 + i, 1}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
192 |
by (auto simp: exits_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
193 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
194 |
lemma acomp_succs [simp]: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
195 |
"succs (acomp a) n = {n + 1 .. n + size (acomp a)}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
196 |
by (induct a arbitrary: n) auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
197 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
198 |
lemma acomp_size: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
199 |
"(1::int) \<le> size (acomp a)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
200 |
by (induct a) auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
201 |
|
44010 | 202 |
lemma acomp_exits [simp]: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
203 |
"exits (acomp a) = {size (acomp a)}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
204 |
by (auto simp: exits_def acomp_size) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
205 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
206 |
lemma bcomp_succs: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
207 |
"0 \<le> i \<Longrightarrow> |
53880 | 208 |
succs (bcomp b f i) n \<subseteq> {n .. n + size (bcomp b f i)} |
209 |
\<union> {n + i + size (bcomp b f i)}" |
|
210 |
proof (induction b arbitrary: f i n) |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
211 |
case (And b1 b2) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
212 |
from And.prems |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
213 |
show ?case |
53880 | 214 |
by (cases f) |
45015 | 215 |
(auto dest: And.IH(1) [THEN subsetD, rotated] |
216 |
And.IH(2) [THEN subsetD, rotated]) |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
217 |
qed auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
218 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
219 |
lemmas bcomp_succsD [dest!] = bcomp_succs [THEN subsetD, rotated] |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
220 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
221 |
lemma bcomp_exits: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
222 |
fixes i :: int |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
223 |
shows |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
224 |
"0 \<le> i \<Longrightarrow> |
53880 | 225 |
exits (bcomp b f i) \<subseteq> {size (bcomp b f i), i + size (bcomp b f i)}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
226 |
by (auto simp: exits_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
227 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
228 |
lemma bcomp_exitsD [dest!]: |
53880 | 229 |
"p \<in> exits (bcomp b f i) \<Longrightarrow> 0 \<le> i \<Longrightarrow> |
230 |
p = size (bcomp b f i) \<or> p = i + size (bcomp b f i)" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
231 |
using bcomp_exits by auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
232 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
233 |
lemma ccomp_succs: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
234 |
"succs (ccomp c) n \<subseteq> {n..n + size (ccomp c)}" |
45015 | 235 |
proof (induction c arbitrary: n) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
236 |
case SKIP thus ?case by simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
237 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
238 |
case Assign thus ?case by simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
239 |
next |
47818 | 240 |
case (Seq c1 c2) |
241 |
from Seq.prems |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
242 |
show ?case |
47818 | 243 |
by (fastforce dest: Seq.IH [THEN subsetD]) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
244 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
245 |
case (If b c1 c2) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
246 |
from If.prems |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
247 |
show ?case |
45015 | 248 |
by (auto dest!: If.IH [THEN subsetD] simp: isuccs_def succs_Cons) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
249 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
250 |
case (While b c) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
251 |
from While.prems |
45015 | 252 |
show ?case by (auto dest!: While.IH [THEN subsetD]) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
253 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
254 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
255 |
lemma ccomp_exits: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
256 |
"exits (ccomp c) \<subseteq> {size (ccomp c)}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
257 |
using ccomp_succs [of c 0] by (auto simp: exits_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
258 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
259 |
lemma ccomp_exitsD [dest!]: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
260 |
"p \<in> exits (ccomp c) \<Longrightarrow> p = size (ccomp c)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
261 |
using ccomp_exits by auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
262 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
263 |
|
67406 | 264 |
subsection \<open>Splitting up machine executions\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
265 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
266 |
lemma exec1_split: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
267 |
fixes i j :: int |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
268 |
shows |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
269 |
"P @ c @ P' \<turnstile> (size P + i, s) \<rightarrow> (j,s') \<Longrightarrow> 0 \<le> i \<Longrightarrow> i < size c \<Longrightarrow> |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
270 |
c \<turnstile> (i,s) \<rightarrow> (j - size P, s')" |
52915 | 271 |
by (auto split: instr.splits simp: exec1_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
272 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
273 |
lemma exec_n_split: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
274 |
fixes i j :: int |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
275 |
assumes "P @ c @ P' \<turnstile> (size P + i, s) \<rightarrow>^n (j, s')" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
276 |
"0 \<le> i" "i < size c" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
277 |
"j \<notin> {size P ..< size P + size c}" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
278 |
shows "\<exists>s'' (i'::int) k m. |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
279 |
c \<turnstile> (i, s) \<rightarrow>^k (i', s'') \<and> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
280 |
i' \<in> exits c \<and> |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
281 |
P @ c @ P' \<turnstile> (size P + i', s'') \<rightarrow>^m (j, s') \<and> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
282 |
n = k + m" |
45015 | 283 |
using assms proof (induction n arbitrary: i j s) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
284 |
case 0 |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
285 |
thus ?case by simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
286 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
287 |
case (Suc n) |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
288 |
have i: "0 \<le> i" "i < size c" by fact+ |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
289 |
from Suc.prems |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
290 |
have j: "\<not> (size P \<le> j \<and> j < size P + size c)" by simp |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
291 |
from Suc.prems |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
292 |
obtain i0 s0 where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
293 |
step: "P @ c @ P' \<turnstile> (size P + i, s) \<rightarrow> (i0,s0)" and |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
294 |
rest: "P @ c @ P' \<turnstile> (i0,s0) \<rightarrow>^n (j, s')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
295 |
by clarsimp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
296 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
297 |
from step i |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
298 |
have c: "c \<turnstile> (i,s) \<rightarrow> (i0 - size P, s0)" by (rule exec1_split) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
299 |
|
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
300 |
have "i0 = size P + (i0 - size P) " by simp |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
301 |
then obtain j0::int where j0: "i0 = size P + j0" .. |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
302 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
303 |
note split_paired_Ex [simp del] |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
304 |
|
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
305 |
have ?case if assm: "j0 \<in> {0 ..< size c}" |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
306 |
proof - |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
307 |
from assm j0 j rest c show ?case |
45015 | 308 |
by (fastforce dest!: Suc.IH intro!: exec_Suc) |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
309 |
qed |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
310 |
moreover |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
311 |
have ?case if assm: "j0 \<notin> {0 ..< size c}" |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
312 |
proof - |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
313 |
from c j0 have "j0 \<in> succs c 0" |
52915 | 314 |
by (auto dest: succs_iexec1 simp: exec1_def simp del: iexec.simps) |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
315 |
with assm have "j0 \<in> exits c" by (simp add: exits_def) |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
316 |
with c j0 rest show ?case by fastforce |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
317 |
qed |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
318 |
ultimately |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
319 |
show ?case by cases |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
320 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
321 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
322 |
lemma exec_n_drop_right: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
323 |
fixes j :: int |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
324 |
assumes "c @ P' \<turnstile> (0, s) \<rightarrow>^n (j, s')" "j \<notin> {0..<size c}" |
44004 | 325 |
shows "\<exists>s'' i' k m. |
326 |
(if c = [] then s'' = s \<and> i' = 0 \<and> k = 0 |
|
327 |
else c \<turnstile> (0, s) \<rightarrow>^k (i', s'') \<and> |
|
328 |
i' \<in> exits c) \<and> |
|
329 |
c @ P' \<turnstile> (i', s'') \<rightarrow>^m (j, s') \<and> |
|
330 |
n = k + m" |
|
331 |
using assms |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
332 |
by (cases "c = []") |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
333 |
(auto dest: exec_n_split [where P="[]", simplified]) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
334 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
335 |
|
67406 | 336 |
text \<open> |
69597 | 337 |
Dropping the left context of a potentially incomplete execution of \<^term>\<open>c\<close>. |
67406 | 338 |
\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
339 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
340 |
lemma exec1_drop_left: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
341 |
fixes i n :: int |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
342 |
assumes "P1 @ P2 \<turnstile> (i, s, stk) \<rightarrow> (n, s', stk')" and "size P1 \<le> i" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
343 |
shows "P2 \<turnstile> (i - size P1, s, stk) \<rightarrow> (n - size P1, s', stk')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
344 |
proof - |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
345 |
have "i = size P1 + (i - size P1)" by simp |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
346 |
then obtain i' :: int where "i = size P1 + i'" .. |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
347 |
moreover |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
348 |
have "n = size P1 + (n - size P1)" by simp |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
349 |
then obtain n' :: int where "n = size P1 + n'" .. |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
350 |
ultimately |
52915 | 351 |
show ?thesis using assms |
352 |
by (clarsimp simp: exec1_def simp del: iexec.simps) |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
353 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
354 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
355 |
lemma exec_n_drop_left: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
356 |
fixes i n :: int |
44004 | 357 |
assumes "P @ P' \<turnstile> (i, s, stk) \<rightarrow>^k (n, s', stk')" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
358 |
"size P \<le> i" "exits P' \<subseteq> {0..}" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
359 |
shows "P' \<turnstile> (i - size P, s, stk) \<rightarrow>^k (n - size P, s', stk')" |
45015 | 360 |
using assms proof (induction k arbitrary: i s stk) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
361 |
case 0 thus ?case by simp |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
362 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
363 |
case (Suc k) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
364 |
from Suc.prems |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
365 |
obtain i' s'' stk'' where |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
366 |
step: "P @ P' \<turnstile> (i, s, stk) \<rightarrow> (i', s'', stk'')" and |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
367 |
rest: "P @ P' \<turnstile> (i', s'', stk'') \<rightarrow>^k (n, s', stk')" |
53356 | 368 |
by auto |
67406 | 369 |
from step \<open>size P \<le> i\<close> |
63540 | 370 |
have *: "P' \<turnstile> (i - size P, s, stk) \<rightarrow> (i' - size P, s'', stk'')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
371 |
by (rule exec1_drop_left) |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
372 |
then have "i' - size P \<in> succs P' 0" |
52915 | 373 |
by (fastforce dest!: succs_iexec1 simp: exec1_def simp del: iexec.simps) |
67406 | 374 |
with \<open>exits P' \<subseteq> {0..}\<close> |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
375 |
have "size P \<le> i'" by (auto simp: exits_def) |
67406 | 376 |
from rest this \<open>exits P' \<subseteq> {0..}\<close> |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
377 |
have "P' \<turnstile> (i' - size P, s'', stk'') \<rightarrow>^k (n - size P, s', stk')" |
45015 | 378 |
by (rule Suc.IH) |
63540 | 379 |
with * show ?case by auto |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
380 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
381 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
382 |
lemmas exec_n_drop_Cons = |
45605 | 383 |
exec_n_drop_left [where P="[instr]", simplified] for instr |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
384 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
385 |
definition |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
386 |
"closed P \<longleftrightarrow> exits P \<subseteq> {size P}" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
387 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
388 |
lemma ccomp_closed [simp, intro!]: "closed (ccomp c)" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
389 |
using ccomp_exits by (auto simp: closed_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
390 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
391 |
lemma acomp_closed [simp, intro!]: "closed (acomp c)" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
392 |
by (simp add: closed_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
393 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
394 |
lemma exec_n_split_full: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
395 |
fixes j :: int |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
396 |
assumes exec: "P @ P' \<turnstile> (0,s,stk) \<rightarrow>^k (j, s', stk')" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
397 |
assumes P: "size P \<le> j" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
398 |
assumes closed: "closed P" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
399 |
assumes exits: "exits P' \<subseteq> {0..}" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
400 |
shows "\<exists>k1 k2 s'' stk''. P \<turnstile> (0,s,stk) \<rightarrow>^k1 (size P, s'', stk'') \<and> |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
401 |
P' \<turnstile> (0,s'',stk'') \<rightarrow>^k2 (j - size P, s', stk')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
402 |
proof (cases "P") |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
403 |
case Nil with exec |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
404 |
show ?thesis by fastforce |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
405 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
406 |
case Cons |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
407 |
hence "0 < size P" by simp |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
408 |
with exec P closed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
409 |
obtain k1 k2 s'' stk'' where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
410 |
1: "P \<turnstile> (0,s,stk) \<rightarrow>^k1 (size P, s'', stk'')" and |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
411 |
2: "P @ P' \<turnstile> (size P,s'',stk'') \<rightarrow>^k2 (j, s', stk')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
412 |
by (auto dest!: exec_n_split [where P="[]" and i=0, simplified] |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
413 |
simp: closed_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
414 |
moreover |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
415 |
have "j = size P + (j - size P)" by simp |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
416 |
then obtain j0 :: int where "j = size P + j0" .. |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
417 |
ultimately |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
418 |
show ?thesis using exits |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
419 |
by (fastforce dest: exec_n_drop_left) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
420 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
421 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
422 |
|
67406 | 423 |
subsection \<open>Correctness theorem\<close> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
424 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
425 |
lemma acomp_neq_Nil [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
426 |
"acomp a \<noteq> []" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
427 |
by (induct a) auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
428 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
429 |
lemma acomp_exec_n [dest!]: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
430 |
"acomp a \<turnstile> (0,s,stk) \<rightarrow>^n (size (acomp a),s',stk') \<Longrightarrow> |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
431 |
s' = s \<and> stk' = aval a s#stk" |
45015 | 432 |
proof (induction a arbitrary: n s' stk stk') |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
433 |
case (Plus a1 a2) |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
434 |
let ?sz = "size (acomp a1) + (size (acomp a2) + 1)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
435 |
from Plus.prems |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
436 |
have "acomp a1 @ acomp a2 @ [ADD] \<turnstile> (0,s,stk) \<rightarrow>^n (?sz, s', stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
437 |
by (simp add: algebra_simps) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
438 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
439 |
then obtain n1 s1 stk1 n2 s2 stk2 n3 where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
440 |
"acomp a1 \<turnstile> (0,s,stk) \<rightarrow>^n1 (size (acomp a1), s1, stk1)" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
441 |
"acomp a2 \<turnstile> (0,s1,stk1) \<rightarrow>^n2 (size (acomp a2), s2, stk2)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
442 |
"[ADD] \<turnstile> (0,s2,stk2) \<rightarrow>^n3 (1, s', stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
443 |
by (auto dest!: exec_n_split_full) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
444 |
|
52915 | 445 |
thus ?case by (fastforce dest: Plus.IH simp: exec_n_simps exec1_def) |
446 |
qed (auto simp: exec_n_simps exec1_def) |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
447 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
448 |
lemma bcomp_split: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
449 |
fixes i j :: int |
53880 | 450 |
assumes "bcomp b f i @ P' \<turnstile> (0, s, stk) \<rightarrow>^n (j, s', stk')" |
451 |
"j \<notin> {0..<size (bcomp b f i)}" "0 \<le> i" |
|
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
452 |
shows "\<exists>s'' stk'' (i'::int) k m. |
53880 | 453 |
bcomp b f i \<turnstile> (0, s, stk) \<rightarrow>^k (i', s'', stk'') \<and> |
454 |
(i' = size (bcomp b f i) \<or> i' = i + size (bcomp b f i)) \<and> |
|
455 |
bcomp b f i @ P' \<turnstile> (i', s'', stk'') \<rightarrow>^m (j, s', stk') \<and> |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
456 |
n = k + m" |
53880 | 457 |
using assms by (cases "bcomp b f i = []") (fastforce dest!: exec_n_drop_right)+ |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
458 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
459 |
lemma bcomp_exec_n [dest]: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
460 |
fixes i j :: int |
53880 | 461 |
assumes "bcomp b f j \<turnstile> (0, s, stk) \<rightarrow>^n (i, s', stk')" |
462 |
"size (bcomp b f j) \<le> i" "0 \<le> j" |
|
463 |
shows "i = size(bcomp b f j) + (if f = bval b s then j else 0) \<and> |
|
44004 | 464 |
s' = s \<and> stk' = stk" |
53880 | 465 |
using assms proof (induction b arbitrary: f j i n s' stk') |
45200 | 466 |
case Bc thus ?case |
62390 | 467 |
by (simp split: if_split_asm add: exec_n_simps exec1_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
468 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
469 |
case (Not b) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
470 |
from Not.prems show ?case |
45015 | 471 |
by (fastforce dest!: Not.IH) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
472 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
473 |
case (And b1 b2) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
474 |
|
53880 | 475 |
let ?b2 = "bcomp b2 f j" |
476 |
let ?m = "if f then size ?b2 else size ?b2 + j" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
477 |
let ?b1 = "bcomp b1 False ?m" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
478 |
|
53880 | 479 |
have j: "size (bcomp (And b1 b2) f j) \<le> i" "0 \<le> j" by fact+ |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
480 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
481 |
from And.prems |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
482 |
obtain s'' stk'' and i'::int and k m where |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
483 |
b1: "?b1 \<turnstile> (0, s, stk) \<rightarrow>^k (i', s'', stk'')" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
484 |
"i' = size ?b1 \<or> i' = ?m + size ?b1" and |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
485 |
b2: "?b2 \<turnstile> (i' - size ?b1, s'', stk'') \<rightarrow>^m (i - size ?b1, s', stk')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
486 |
by (auto dest!: bcomp_split dest: exec_n_drop_left) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
487 |
from b1 j |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
488 |
have "i' = size ?b1 + (if \<not>bval b1 s then ?m else 0) \<and> s'' = s \<and> stk'' = stk" |
45015 | 489 |
by (auto dest!: And.IH) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
490 |
with b2 j |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
491 |
show ?case |
62390 | 492 |
by (fastforce dest!: And.IH simp: exec_n_end split: if_split_asm) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
493 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
494 |
case Less |
52915 | 495 |
thus ?case by (auto dest!: exec_n_split_full simp: exec_n_simps exec1_def) (* takes time *) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
496 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
497 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
498 |
lemma ccomp_empty [elim!]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
499 |
"ccomp c = [] \<Longrightarrow> (c,s) \<Rightarrow> s" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
500 |
by (induct c) auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
501 |
|
44070 | 502 |
declare assign_simp [simp] |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
503 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
504 |
lemma ccomp_exec_n: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
505 |
"ccomp c \<turnstile> (0,s,stk) \<rightarrow>^n (size(ccomp c),t,stk') |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
506 |
\<Longrightarrow> (c,s) \<Rightarrow> t \<and> stk'=stk" |
45015 | 507 |
proof (induction c arbitrary: s t stk stk' n) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
508 |
case SKIP |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
509 |
thus ?case by auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
510 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
511 |
case (Assign x a) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
512 |
thus ?case |
52915 | 513 |
by simp (fastforce dest!: exec_n_split_full simp: exec_n_simps exec1_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
514 |
next |
47818 | 515 |
case (Seq c1 c2) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
516 |
thus ?case by (fastforce dest!: exec_n_split_full) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
517 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
518 |
case (If b c1 c2) |
45015 | 519 |
note If.IH [dest!] |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
520 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
521 |
let ?if = "IF b THEN c1 ELSE c2" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
522 |
let ?cs = "ccomp ?if" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
523 |
let ?bcomp = "bcomp b False (size (ccomp c1) + 1)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
524 |
|
67406 | 525 |
from \<open>?cs \<turnstile> (0,s,stk) \<rightarrow>^n (size ?cs,t,stk')\<close> |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
526 |
obtain i' :: int and k m s'' stk'' where |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
527 |
cs: "?cs \<turnstile> (i',s'',stk'') \<rightarrow>^m (size ?cs,t,stk')" and |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
528 |
"?bcomp \<turnstile> (0,s,stk) \<rightarrow>^k (i', s'', stk'')" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
529 |
"i' = size ?bcomp \<or> i' = size ?bcomp + size (ccomp c1) + 1" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
530 |
by (auto dest!: bcomp_split) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
531 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
532 |
hence i': |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
533 |
"s''=s" "stk'' = stk" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
534 |
"i' = (if bval b s then size ?bcomp else size ?bcomp+size(ccomp c1)+1)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
535 |
by auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
536 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
537 |
with cs have cs': |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
538 |
"ccomp c1@JMP (size (ccomp c2))#ccomp c2 \<turnstile> |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
539 |
(if bval b s then 0 else size (ccomp c1)+1, s, stk) \<rightarrow>^m |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
540 |
(1 + size (ccomp c1) + size (ccomp c2), t, stk')" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
541 |
by (fastforce dest: exec_n_drop_left simp: exits_Cons isuccs_def algebra_simps) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
542 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
543 |
show ?case |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
544 |
proof (cases "bval b s") |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
545 |
case True with cs' |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
546 |
show ?thesis |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
547 |
by simp |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
548 |
(fastforce dest: exec_n_drop_right |
62390 | 549 |
split: if_split_asm |
52915 | 550 |
simp: exec_n_simps exec1_def) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
551 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
552 |
case False with cs' |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
553 |
show ?thesis |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
554 |
by (auto dest!: exec_n_drop_Cons exec_n_drop_left |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
555 |
simp: exits_Cons isuccs_def) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
556 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
557 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
558 |
case (While b c) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
559 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
560 |
from While.prems |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
561 |
show ?case |
45015 | 562 |
proof (induction n arbitrary: s rule: nat_less_induct) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
563 |
case (1 n) |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
564 |
|
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
565 |
have ?case if assm: "\<not> bval b s" |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
566 |
proof - |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
567 |
from assm "1.prems" |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
568 |
show ?case |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
569 |
by simp (fastforce dest!: bcomp_split simp: exec_n_simps) |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
570 |
qed |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
571 |
moreover |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
572 |
have ?case if b: "bval b s" |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
573 |
proof - |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
574 |
let ?c0 = "WHILE b DO c" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
575 |
let ?cs = "ccomp ?c0" |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
576 |
let ?bs = "bcomp b False (size (ccomp c) + 1)" |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
577 |
let ?jmp = "[JMP (-((size ?bs + size (ccomp c) + 1)))]" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
578 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
579 |
from "1.prems" b |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
580 |
obtain k where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
581 |
cs: "?cs \<turnstile> (size ?bs, s, stk) \<rightarrow>^k (size ?cs, t, stk')" and |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
582 |
k: "k \<le> n" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44070
diff
changeset
|
583 |
by (fastforce dest!: bcomp_split) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
584 |
|
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
585 |
show ?case |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
586 |
proof cases |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
587 |
assume "ccomp c = []" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
588 |
with cs k |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
589 |
obtain m where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
590 |
"?cs \<turnstile> (0,s,stk) \<rightarrow>^m (size (ccomp ?c0), t, stk')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
591 |
"m < n" |
52915 | 592 |
by (auto simp: exec_n_step [where k=k] exec1_def) |
45015 | 593 |
with "1.IH" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
594 |
show ?case by blast |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
595 |
next |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
596 |
assume "ccomp c \<noteq> []" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
597 |
with cs |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
598 |
obtain m m' s'' stk'' where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
599 |
c: "ccomp c \<turnstile> (0, s, stk) \<rightarrow>^m' (size (ccomp c), s'', stk'')" and |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
600 |
rest: "?cs \<turnstile> (size ?bs + size (ccomp c), s'', stk'') \<rightarrow>^m |
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
601 |
(size ?cs, t, stk')" and |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
602 |
m: "k = m + m'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
603 |
by (auto dest: exec_n_split [where i=0, simplified]) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
604 |
from c |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
605 |
have "(c,s) \<Rightarrow> s''" and stk: "stk'' = stk" |
45015 | 606 |
by (auto dest!: While.IH) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
607 |
moreover |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
608 |
from rest m k stk |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
609 |
obtain k' where |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
610 |
"?cs \<turnstile> (0, s'', stk) \<rightarrow>^k' (size ?cs, t, stk')" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
611 |
"k' < n" |
52915 | 612 |
by (auto simp: exec_n_step [where k=m] exec1_def) |
45015 | 613 |
with "1.IH" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
614 |
have "(?c0, s'') \<Rightarrow> t \<and> stk' = stk" by blast |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
615 |
ultimately |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
616 |
show ?case using b by blast |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
617 |
qed |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63540
diff
changeset
|
618 |
qed |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
619 |
ultimately show ?case by cases |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
620 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
621 |
qed |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
622 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
623 |
theorem ccomp_exec: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
624 |
"ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk') \<Longrightarrow> (c,s) \<Rightarrow> t" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
625 |
by (auto dest: exec_exec_n ccomp_exec_n) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
626 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
627 |
corollary ccomp_sound: |
51259
1491459df114
eliminated isize in favour of size + type coercion
kleing
parents:
50061
diff
changeset
|
628 |
"ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk) \<longleftrightarrow> (c,s) \<Rightarrow> t" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
629 |
by (blast intro!: ccomp_exec ccomp_bigstep) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
630 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
diff
changeset
|
631 |
end |