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