author | haftmann |
Fri, 25 Sep 2009 09:50:31 +0200 | |
changeset 32705 | 04ce6bb14d85 |
parent 31758 | 3edd5f813f01 |
permissions | -rw-r--r-- |
31758 | 1 |
(* Title: HOL/Isar_examples/Expr_Compiler.thy |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
2 |
Author: Markus Wenzel, TU Muenchen |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
3 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
4 |
Correctness of a simple expression/stack-machine compiler. |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
5 |
*) |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
6 |
|
10007 | 7 |
header {* Correctness of a simple expression compiler *} |
7748 | 8 |
|
31758 | 9 |
theory Expr_Compiler |
10 |
imports Main |
|
11 |
begin |
|
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
12 |
|
7869 | 13 |
text {* |
7874 | 14 |
This is a (rather trivial) example of program verification. We model |
15 |
a compiler for translating expressions to stack machine instructions, |
|
16 |
and prove its correctness wrt.\ some evaluation semantics. |
|
10007 | 17 |
*} |
7869 | 18 |
|
19 |
||
10007 | 20 |
subsection {* Binary operations *} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
21 |
|
6744 | 22 |
text {* |
7869 | 23 |
Binary operations are just functions over some type of values. This |
7982 | 24 |
is both for abstract syntax and semantics, i.e.\ we use a ``shallow |
7874 | 25 |
embedding'' here. |
10007 | 26 |
*} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
27 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
28 |
types |
10007 | 29 |
'val binop = "'val => 'val => 'val" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
30 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
31 |
|
10007 | 32 |
subsection {* Expressions *} |
7869 | 33 |
|
34 |
text {* |
|
35 |
The language of expressions is defined as an inductive type, |
|
36 |
consisting of variables, constants, and binary operations on |
|
37 |
expressions. |
|
10007 | 38 |
*} |
7869 | 39 |
|
40 |
datatype ('adr, 'val) expr = |
|
41 |
Variable 'adr | |
|
42 |
Constant 'val | |
|
10007 | 43 |
Binop "'val binop" "('adr, 'val) expr" "('adr, 'val) expr" |
7869 | 44 |
|
45 |
text {* |
|
7874 | 46 |
Evaluation (wrt.\ some environment of variable assignments) is |
47 |
defined by primitive recursion over the structure of expressions. |
|
10007 | 48 |
*} |
7869 | 49 |
|
50 |
consts |
|
10007 | 51 |
eval :: "('adr, 'val) expr => ('adr => 'val) => 'val" |
7869 | 52 |
|
53 |
primrec |
|
54 |
"eval (Variable x) env = env x" |
|
55 |
"eval (Constant c) env = c" |
|
10007 | 56 |
"eval (Binop f e1 e2) env = f (eval e1 env) (eval e2 env)" |
7869 | 57 |
|
58 |
||
10007 | 59 |
subsection {* Machine *} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
60 |
|
6744 | 61 |
text {* |
6792 | 62 |
Next we model a simple stack machine, with three instructions. |
10007 | 63 |
*} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
64 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
65 |
datatype ('adr, 'val) instr = |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
66 |
Const 'val | |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
67 |
Load 'adr | |
10007 | 68 |
Apply "'val binop" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
69 |
|
6744 | 70 |
text {* |
7869 | 71 |
Execution of a list of stack machine instructions is easily defined |
72 |
as follows. |
|
10007 | 73 |
*} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
74 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
75 |
consts |
7761 | 76 |
exec :: "(('adr, 'val) instr) list |
10007 | 77 |
=> 'val list => ('adr => 'val) => 'val list" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
78 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
79 |
primrec |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
80 |
"exec [] stack env = stack" |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
81 |
"exec (instr # instrs) stack env = |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
82 |
(case instr of |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
83 |
Const c => exec instrs (c # stack) env |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
84 |
| Load x => exec instrs (env x # stack) env |
7761 | 85 |
| Apply f => exec instrs (f (hd stack) (hd (tl stack)) |
10007 | 86 |
# (tl (tl stack))) env)" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
87 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
88 |
constdefs |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
89 |
execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val" |
10007 | 90 |
"execute instrs env == hd (exec instrs [] env)" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
91 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
92 |
|
10007 | 93 |
subsection {* Compiler *} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
94 |
|
6744 | 95 |
text {* |
7869 | 96 |
We are ready to define the compilation function of expressions to |
6792 | 97 |
lists of stack machine instructions. |
10007 | 98 |
*} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
99 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
100 |
consts |
10007 | 101 |
compile :: "('adr, 'val) expr => (('adr, 'val) instr) list" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
102 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
103 |
primrec |
8031
826c6222cca2
renamed comp to compile (avoids clash with Relation.comp);
wenzelm
parents:
7982
diff
changeset
|
104 |
"compile (Variable x) = [Load x]" |
826c6222cca2
renamed comp to compile (avoids clash with Relation.comp);
wenzelm
parents:
7982
diff
changeset
|
105 |
"compile (Constant c) = [Const c]" |
10007 | 106 |
"compile (Binop f e1 e2) = compile e2 @ compile e1 @ [Apply f]" |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
107 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
108 |
|
6744 | 109 |
text {* |
8051 | 110 |
The main result of this development is the correctness theorem for |
111 |
$\idt{compile}$. We first establish a lemma about $\idt{exec}$ and |
|
112 |
list append. |
|
10007 | 113 |
*} |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
114 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
115 |
lemma exec_append: |
18153 | 116 |
"exec (xs @ ys) stack env = |
117 |
exec ys (exec xs stack env) env" |
|
20503 | 118 |
proof (induct xs arbitrary: stack) |
18153 | 119 |
case Nil |
120 |
show ?case by simp |
|
11809 | 121 |
next |
18153 | 122 |
case (Cons x xs) |
123 |
show ?case |
|
11809 | 124 |
proof (induct x) |
23373 | 125 |
case Const |
126 |
from Cons show ?case by simp |
|
18153 | 127 |
next |
23373 | 128 |
case Load |
129 |
from Cons show ?case by simp |
|
18153 | 130 |
next |
23373 | 131 |
case Apply |
132 |
from Cons show ?case by simp |
|
10007 | 133 |
qed |
134 |
qed |
|
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
135 |
|
10007 | 136 |
theorem correctness: "execute (compile e) env = eval e env" |
137 |
proof - |
|
18193 | 138 |
have "\<And>stack. exec (compile e) stack env = eval e env # stack" |
11809 | 139 |
proof (induct e) |
18153 | 140 |
case Variable show ?case by simp |
141 |
next |
|
142 |
case Constant show ?case by simp |
|
143 |
next |
|
144 |
case Binop then show ?case by (simp add: exec_append) |
|
10007 | 145 |
qed |
23373 | 146 |
then show ?thesis by (simp add: execute_def) |
10007 | 147 |
qed |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
148 |
|
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
149 |
|
8051 | 150 |
text {* |
151 |
\bigskip In the proofs above, the \name{simp} method does quite a lot |
|
152 |
of work behind the scenes (mostly ``functional program execution''). |
|
153 |
Subsequently, the same reasoning is elaborated in detail --- at most |
|
154 |
one recursive function definition is used at a time. Thus we get a |
|
155 |
better idea of what is actually going on. |
|
10007 | 156 |
*} |
8051 | 157 |
|
13524 | 158 |
lemma exec_append': |
18153 | 159 |
"exec (xs @ ys) stack env = exec ys (exec xs stack env) env" |
20503 | 160 |
proof (induct xs arbitrary: stack) |
18153 | 161 |
case (Nil s) |
162 |
have "exec ([] @ ys) s env = exec ys s env" by simp |
|
163 |
also have "... = exec ys (exec [] s env) env" by simp |
|
164 |
finally show ?case . |
|
165 |
next |
|
166 |
case (Cons x xs s) |
|
167 |
show ?case |
|
10007 | 168 |
proof (induct x) |
18153 | 169 |
case (Const val) |
170 |
have "exec ((Const val # xs) @ ys) s env = exec (Const val # xs @ ys) s env" |
|
171 |
by simp |
|
172 |
also have "... = exec (xs @ ys) (val # s) env" by simp |
|
173 |
also from Cons have "... = exec ys (exec xs (val # s) env) env" . |
|
174 |
also have "... = exec ys (exec (Const val # xs) s env) env" by simp |
|
175 |
finally show ?case . |
|
10007 | 176 |
next |
18153 | 177 |
case (Load adr) |
178 |
from Cons show ?case by simp -- {* same as above *} |
|
179 |
next |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
180 |
case (Apply fn) |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
181 |
have "exec ((Apply fn # xs) @ ys) s env = |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
182 |
exec (Apply fn # xs @ ys) s env" by simp |
18153 | 183 |
also have "... = |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
184 |
exec (xs @ ys) (fn (hd s) (hd (tl s)) # (tl (tl s))) env" by simp |
18153 | 185 |
also from Cons have "... = |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
186 |
exec ys (exec xs (fn (hd s) (hd (tl s)) # tl (tl s)) env) env" . |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
187 |
also have "... = exec ys (exec (Apply fn # xs) s env) env" by simp |
18153 | 188 |
finally show ?case . |
10007 | 189 |
qed |
190 |
qed |
|
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
191 |
|
13537 | 192 |
theorem correctness': "execute (compile e) env = eval e env" |
10007 | 193 |
proof - |
18193 | 194 |
have exec_compile: "\<And>stack. exec (compile e) stack env = eval e env # stack" |
10007 | 195 |
proof (induct e) |
18153 | 196 |
case (Variable adr s) |
197 |
have "exec (compile (Variable adr)) s env = exec [Load adr] s env" |
|
198 |
by simp |
|
199 |
also have "... = env adr # s" by simp |
|
200 |
also have "env adr = eval (Variable adr) env" by simp |
|
201 |
finally show ?case . |
|
10007 | 202 |
next |
18153 | 203 |
case (Constant val s) |
204 |
show ?case by simp -- {* same as above *} |
|
10007 | 205 |
next |
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
206 |
case (Binop fn e1 e2 s) |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
207 |
have "exec (compile (Binop fn e1 e2)) s env = |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
208 |
exec (compile e2 @ compile e1 @ [Apply fn]) s env" by simp |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
209 |
also have "... = exec [Apply fn] |
18153 | 210 |
(exec (compile e1) (exec (compile e2) s env) env) env" |
211 |
by (simp only: exec_append) |
|
212 |
also have "exec (compile e2) s env = eval e2 env # s" by fact |
|
213 |
also have "exec (compile e1) ... env = eval e1 env # ..." by fact |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
214 |
also have "exec [Apply fn] ... env = |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
215 |
fn (hd ...) (hd (tl ...)) # (tl (tl ...))" by simp |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
216 |
also have "... = fn (eval e1 env) (eval e2 env) # s" by simp |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
217 |
also have "fn (eval e1 env) (eval e2 env) = |
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
20503
diff
changeset
|
218 |
eval (Binop fn e1 e2) env" |
18153 | 219 |
by simp |
220 |
finally show ?case . |
|
10007 | 221 |
qed |
8051 | 222 |
|
10007 | 223 |
have "execute (compile e) env = hd (exec (compile e) [] env)" |
224 |
by (simp add: execute_def) |
|
225 |
also from exec_compile |
|
18153 | 226 |
have "exec (compile e) [] env = [eval e env]" . |
10007 | 227 |
also have "hd ... = eval e env" by simp |
228 |
finally show ?thesis . |
|
229 |
qed |
|
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
230 |
|
10007 | 231 |
end |