author | kleing |
Thu, 03 Nov 2011 16:22:29 +1100 | |
changeset 45322 | 654cc47f6115 |
parent 45200 | 1f1897ac7877 |
child 45323 | df7554ebe024 |
permissions | -rw-r--r-- |
43141 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
header "A Compiler for IMP" |
|
10343 | 4 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
5 |
theory Compiler imports Big_Step |
43141 | 6 |
begin |
12429 | 7 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
8 |
subsection "List setup" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
9 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
10 |
text {* |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
11 |
We are going to define a small machine language where programs are |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
12 |
lists of instructions. For nicer algebraic properties in our lemmas |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
13 |
later, we prefer @{typ int} to @{term nat} as program counter. |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
14 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
15 |
Therefore, we define notation for size and indexing for lists |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
16 |
on @{typ int}: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
17 |
*} |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
18 |
abbreviation "isize xs == int (length xs)" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
19 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
20 |
primrec |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
21 |
inth :: "'a list => int => 'a" (infixl "!!" 100) where |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
22 |
inth_Cons: "(x # xs) !! n = (if n = 0 then x else xs !! (n - 1))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
23 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
24 |
text {* |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
25 |
The only additional lemma we need is indexing over append: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
26 |
*} |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
27 |
lemma inth_append [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
28 |
"0 \<le> n \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
29 |
(xs @ ys) !! n = (if n < isize xs then xs !! n else ys !! (n - isize xs))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
30 |
by (induct xs arbitrary: n) (auto simp: algebra_simps) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
31 |
|
43141 | 32 |
subsection "Instructions and Stack Machine" |
10342 | 33 |
|
43141 | 34 |
datatype instr = |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
35 |
LOADI int | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
36 |
LOAD string | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
37 |
ADD | |
43141 | 38 |
STORE string | |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
39 |
JMP int | |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
40 |
JMPLESS int | |
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
41 |
JMPGE int |
10342 | 42 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
43 |
type_synonym stack = "val list" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
44 |
type_synonym config = "int\<times>state\<times>stack" |
43141 | 45 |
|
46 |
abbreviation "hd2 xs == hd(tl xs)" |
|
47 |
abbreviation "tl2 xs == tl(tl xs)" |
|
11275 | 48 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
49 |
inductive iexec1 :: "instr \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" |
44004 | 50 |
("(_/ \<turnstile>i (_ \<rightarrow>/ _))" [59,0,59] 60) |
43141 | 51 |
where |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
52 |
"LOADI n \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, n#stk)" | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
53 |
"LOAD x \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, s x # stk)" | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
54 |
"ADD \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk)" | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
55 |
"STORE n \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s(n := hd stk),tl stk)" | |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
56 |
"JMP n \<turnstile>i (i,s,stk) \<rightarrow> (i+1+n,s,stk)" | |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
57 |
"JMPLESS n \<turnstile>i (i,s,stk) \<rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk)" | |
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
58 |
"JMPGE n \<turnstile>i (i,s,stk) \<rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
59 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
60 |
code_pred iexec1 . |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
61 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
62 |
declare iexec1.intros |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
63 |
|
44000
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
64 |
definition |
44004 | 65 |
exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [59,0,59] 60) |
44000
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
66 |
where |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
67 |
"P \<turnstile> c \<rightarrow> c' = |
44004 | 68 |
(\<exists>i s stk. c = (i,s,stk) \<and> P!!i \<turnstile>i (i,s,stk) \<rightarrow> c' \<and> 0 \<le> i \<and> i < isize P)" |
44000
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
69 |
|
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
70 |
declare exec1_def [simp] |
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
71 |
|
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
72 |
lemma exec1I [intro, code_pred_intro]: |
44004 | 73 |
assumes "P!!i \<turnstile>i (i,s,stk) \<rightarrow> c'" "0 \<le> i" "i < isize P" |
74 |
shows "P \<turnstile> (i,s,stk) \<rightarrow> c'" |
|
75 |
using assms by simp |
|
43141 | 76 |
|
77 |
inductive exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("_/ \<turnstile> (_ \<rightarrow>*/ _)" 50) |
|
78 |
where |
|
79 |
refl: "P \<turnstile> c \<rightarrow>* c" | |
|
80 |
step: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''" |
|
81 |
||
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
82 |
declare refl[intro] step[intro] |
43141 | 83 |
|
84 |
lemmas exec_induct = exec.induct[split_format(complete)] |
|
85 |
||
44000
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
kleing
parents:
43438
diff
changeset
|
86 |
code_pred exec by force |
43141 | 87 |
|
88 |
values |
|
89 |
"{(i,map t [''x'',''y''],stk) | i t stk. |
|
90 |
[LOAD ''y'', STORE ''x''] \<turnstile> |
|
44036 | 91 |
(0, <''x'' := 3, ''y'' := 4>, []) \<rightarrow>* (i,t,stk)}" |
43141 | 92 |
|
93 |
||
94 |
subsection{* Verification infrastructure *} |
|
95 |
||
96 |
lemma exec_trans: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''" |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
97 |
by (induct rule: exec.induct) fastforce+ |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
98 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
99 |
inductive_cases iexec1_cases [elim!]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
100 |
"LOADI n \<turnstile>i c \<rightarrow> c'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
101 |
"LOAD x \<turnstile>i c \<rightarrow> c'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
102 |
"ADD \<turnstile>i c \<rightarrow> c'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
103 |
"STORE n \<turnstile>i c \<rightarrow> c'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
104 |
"JMP n \<turnstile>i c \<rightarrow> c'" |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
105 |
"JMPLESS n \<turnstile>i c \<rightarrow> c'" |
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
106 |
"JMPGE n \<turnstile>i c \<rightarrow> c'" |
43141 | 107 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
108 |
text {* Simplification rules for @{const iexec1}. *} |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
109 |
lemma iexec1_simps [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
110 |
"LOADI n \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, n # stk))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
111 |
"LOAD x \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, s x # stk))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
112 |
"ADD \<turnstile>i c \<rightarrow> c' = |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
113 |
(\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, (hd2 stk + hd stk) # tl2 stk))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
114 |
"STORE x \<turnstile>i c \<rightarrow> c' = |
44036 | 115 |
(\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s(x := hd stk), tl stk))" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
116 |
"JMP n \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1 + n, s, stk))" |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
117 |
"JMPLESS n \<turnstile>i c \<rightarrow> c' = |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
118 |
(\<exists>i s stk. c = (i, s, stk) \<and> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
119 |
c' = (if hd2 stk < hd stk then i + 1 + n else i + 1, s, tl2 stk))" |
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
120 |
"JMPGE n \<turnstile>i c \<rightarrow> c' = |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
121 |
(\<exists>i s stk. c = (i, s, stk) \<and> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
122 |
c' = (if hd stk \<le> hd2 stk then i + 1 + n else i + 1, s, tl2 stk))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
123 |
by (auto split del: split_if intro!: iexec1.intros) |
43141 | 124 |
|
125 |
||
126 |
text{* Below we need to argue about the execution of code that is embedded in |
|
127 |
larger programs. For this purpose we show that execution is preserved by |
|
128 |
appending code to the left or right of a program. *} |
|
129 |
||
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
130 |
lemma exec1_appendR: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow> c'" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
131 |
by auto |
11275 | 132 |
|
43141 | 133 |
lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
134 |
by (induct rule: exec.induct) (fastforce intro: exec1_appendR)+ |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
135 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
136 |
lemma iexec1_shiftI: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
137 |
assumes "X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
138 |
shows "X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
139 |
using assms by cases auto |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
140 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
141 |
lemma iexec1_shiftD: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
142 |
assumes "X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
143 |
shows "X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
144 |
using assms by cases auto |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
145 |
|
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
146 |
lemma iexec_shift [simp]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
147 |
"(X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')) = (X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk'))" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
148 |
by (blast intro: iexec1_shiftI dest: iexec1_shiftD) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
149 |
|
43141 | 150 |
lemma exec1_appendL: |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
151 |
"P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk') \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
152 |
P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow> (isize(P')+i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
153 |
by simp |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
154 |
|
43141 | 155 |
lemma exec_appendL: |
156 |
"P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow> |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
157 |
P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow>* (isize(P')+i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
158 |
by (induct rule: exec_induct) (blast intro!: exec1_appendL)+ |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
159 |
|
43141 | 160 |
text{* Now we specialise the above lemmas to enable automatic proofs of |
161 |
@{prop "P \<turnstile> c \<rightarrow>* c'"} where @{text P} is a mixture of concrete instructions and |
|
162 |
pieces of code that we already know how they execute (by induction), combined |
|
163 |
by @{text "@"} and @{text "#"}. Backward jumps are not supported. |
|
164 |
The details should be skipped on a first reading. |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
165 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
166 |
If we have just executed the first instruction of the program, drop it: *} |
43141 | 167 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
168 |
lemma exec_Cons_1 [intro]: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
169 |
"P \<turnstile> (0,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
170 |
instr#P \<turnstile> (1,s,stk) \<rightarrow>* (1+j,t,stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
171 |
by (drule exec_appendL[where P'="[instr]"]) simp |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
172 |
|
43141 | 173 |
lemma exec_appendL_if[intro]: |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
174 |
"isize P' <= i |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
175 |
\<Longrightarrow> P \<turnstile> (i - isize P',s,stk) \<rightarrow>* (i',s',stk') |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
176 |
\<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (isize P' + i',s',stk')" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
177 |
by (drule exec_appendL[where P'=P']) simp |
10342 | 178 |
|
43141 | 179 |
text{* Split the execution of a compound program up into the excution of its |
180 |
parts: *} |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
181 |
|
43141 | 182 |
lemma exec_append_trans[intro]: |
183 |
"P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow> |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
184 |
isize P \<le> i' \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
185 |
P' \<turnstile> (i' - isize P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
186 |
j'' = isize P + i'' |
43141 | 187 |
\<Longrightarrow> |
188 |
P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
189 |
by(metis exec_trans[OF exec_appendR exec_appendL_if]) |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
190 |
|
43141 | 191 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
192 |
declare Let_def[simp] |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
193 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
194 |
|
43141 | 195 |
subsection "Compilation" |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
196 |
|
43141 | 197 |
fun acomp :: "aexp \<Rightarrow> instr list" where |
198 |
"acomp (N n) = [LOADI n]" | |
|
199 |
"acomp (V x) = [LOAD x]" | |
|
200 |
"acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]" |
|
201 |
||
202 |
lemma acomp_correct[intro]: |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
203 |
"acomp a \<turnstile> (0,s,stk) \<rightarrow>* (isize(acomp a),s,aval a s#stk)" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
204 |
by (induct a arbitrary: stk) fastforce+ |
10342 | 205 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
206 |
fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> instr list" where |
45200 | 207 |
"bcomp (Bc v) c n = (if v=c then [JMP n] else [])" | |
43141 | 208 |
"bcomp (Not b) c n = bcomp b (\<not>c) n" | |
209 |
"bcomp (And b1 b2) c n = |
|
210 |
(let cb2 = bcomp b2 c n; |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
211 |
m = (if c then isize cb2 else isize cb2+n); |
43141 | 212 |
cb1 = bcomp b1 False m |
213 |
in cb1 @ cb2)" | |
|
214 |
"bcomp (Less a1 a2) c n = |
|
45322
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
kleing
parents:
45200
diff
changeset
|
215 |
acomp a1 @ acomp a2 @ (if c then [JMPLESS n] else [JMPGE n])" |
43141 | 216 |
|
217 |
value |
|
218 |
"bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v'')))) |
|
219 |
False 3" |
|
220 |
||
221 |
lemma bcomp_correct[intro]: |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
222 |
"0 \<le> n \<Longrightarrow> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
223 |
bcomp b c n \<turnstile> |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
224 |
(0,s,stk) \<rightarrow>* (isize(bcomp b c n) + (if c = bval b s then n else 0),s,stk)" |
45129
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
wenzelm
parents:
45114
diff
changeset
|
225 |
proof(induction b arbitrary: c n) |
43141 | 226 |
case Not |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
227 |
from Not(1)[where c="~c"] Not(2) show ?case by fastforce |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
228 |
next |
43141 | 229 |
case (And b1 b2) |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
230 |
from And(1)[of "if c then isize (bcomp b2 c n) else isize (bcomp b2 c n) + n" |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
231 |
"False"] |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
232 |
And(2)[of n "c"] And(3) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
233 |
show ?case by fastforce |
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
234 |
qed fastforce+ |
43141 | 235 |
|
236 |
fun ccomp :: "com \<Rightarrow> instr list" where |
|
237 |
"ccomp SKIP = []" | |
|
238 |
"ccomp (x ::= a) = acomp a @ [STORE x]" | |
|
239 |
"ccomp (c\<^isub>1;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" | |
|
240 |
"ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) = |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
241 |
(let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b False (isize cc\<^isub>1 + 1) |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
242 |
in cb @ cc\<^isub>1 @ JMP (isize cc\<^isub>2) # cc\<^isub>2)" | |
43141 | 243 |
"ccomp (WHILE b DO c) = |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
244 |
(let cc = ccomp c; cb = bcomp b False (isize cc + 1) |
44004 | 245 |
in cb @ cc @ [JMP (-(isize cb + isize cc + 1))])" |
246 |
||
43141 | 247 |
|
248 |
value "ccomp |
|
249 |
(IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1) |
|
250 |
ELSE ''v'' ::= V ''u'')" |
|
251 |
||
252 |
value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))" |
|
253 |
||
254 |
||
45114 | 255 |
subsection "Preservation of semantics" |
43141 | 256 |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
257 |
lemma ccomp_bigstep: |
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
258 |
"(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (isize(ccomp c),t,stk)" |
45015 | 259 |
proof(induction arbitrary: stk rule: big_step_induct) |
43141 | 260 |
case (Assign x a s) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
261 |
show ?case by (fastforce simp:fun_upd_def cong: if_cong) |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
262 |
next |
43141 | 263 |
case (Semi c1 s1 s2 c2 s3) |
264 |
let ?cc1 = "ccomp c1" let ?cc2 = "ccomp c2" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
265 |
have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cc1,s2,stk)" |
45015 | 266 |
using Semi.IH(1) by fastforce |
43141 | 267 |
moreover |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
268 |
have "?cc1 @ ?cc2 \<turnstile> (isize ?cc1,s2,stk) \<rightarrow>* (isize(?cc1 @ ?cc2),s3,stk)" |
45015 | 269 |
using Semi.IH(2) by fastforce |
43141 | 270 |
ultimately show ?case by simp (blast intro: exec_trans) |
271 |
next |
|
272 |
case (WhileTrue b s1 c s2 s3) |
|
273 |
let ?cc = "ccomp c" |
|
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
274 |
let ?cb = "bcomp b False (isize ?cc + 1)" |
43141 | 275 |
let ?cw = "ccomp(WHILE b DO c)" |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
276 |
have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cb + isize ?cc,s2,stk)" |
45015 | 277 |
using WhileTrue.IH(1) WhileTrue.hyps(1) by fastforce |
43141 | 278 |
moreover |
43438
a666b8d11252
IMP compiler with int, added reverse soundness direction
kleing
parents:
43158
diff
changeset
|
279 |
have "?cw \<turnstile> (isize ?cb + isize ?cc,s2,stk) \<rightarrow>* (0,s2,stk)" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
280 |
by fastforce |
43141 | 281 |
moreover |
45015 | 282 |
have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (isize ?cw,s3,stk)" by(rule WhileTrue.IH(2)) |
43141 | 283 |
ultimately show ?case by(blast intro: exec_trans) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44036
diff
changeset
|
284 |
qed fastforce+ |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
12566
diff
changeset
|
285 |
|
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
16417
diff
changeset
|
286 |
end |