author | haftmann |
Mon, 20 Oct 2014 07:45:58 +0200 | |
changeset 58710 | 7216a10d69ba |
parent 58310 | 91ea607a34d8 |
child 58889 | 5b7a9633cfa8 |
permissions | -rw-r--r-- |
45266 | 1 |
header "Stack Machine and Compilation" |
43141 | 2 |
|
3 |
theory ASM imports AExp begin |
|
4 |
||
45266 | 5 |
subsection "Stack Machine" |
43141 | 6 |
|
49191 | 7 |
text_raw{*\snip{ASMinstrdef}{0}{1}{% *} |
58310 | 8 |
datatype instr = LOADI val | LOAD vname | ADD |
49191 | 9 |
text_raw{*}%endsnip*} |
43141 | 10 |
|
49191 | 11 |
text_raw{*\snip{ASMstackdef}{1}{2}{% *} |
43141 | 12 |
type_synonym stack = "val list" |
49191 | 13 |
text_raw{*}%endsnip*} |
43141 | 14 |
|
15 |
abbreviation "hd2 xs == hd(tl xs)" |
|
16 |
abbreviation "tl2 xs == tl(tl xs)" |
|
17 |
||
18 |
text{* \noindent Abbreviations are transparent: they are unfolded after |
|
19 |
parsing and folded back again before printing. Internally, they do not |
|
49191 | 20 |
exist.*} |
43141 | 21 |
|
49191 | 22 |
text_raw{*\snip{ASMexeconedef}{0}{1}{% *} |
45257 | 23 |
fun exec1 :: "instr \<Rightarrow> state \<Rightarrow> stack \<Rightarrow> stack" where |
24 |
"exec1 (LOADI n) _ stk = n # stk" | |
|
25 |
"exec1 (LOAD x) s stk = s(x) # stk" | |
|
26 |
"exec1 ADD _ stk = (hd2 stk + hd stk) # tl2 stk" |
|
49191 | 27 |
text_raw{*}%endsnip*} |
43141 | 28 |
|
49191 | 29 |
text_raw{*\snip{ASMexecdef}{1}{2}{% *} |
45257 | 30 |
fun exec :: "instr list \<Rightarrow> state \<Rightarrow> stack \<Rightarrow> stack" where |
31 |
"exec [] _ stk = stk" | |
|
32 |
"exec (i#is) s stk = exec is s (exec1 i s stk)" |
|
49191 | 33 |
text_raw{*}%endsnip*} |
43141 | 34 |
|
50007
56f269baae76
executable true liveness analysis incl an approximating version
nipkow
parents:
49191
diff
changeset
|
35 |
value "exec [LOADI 5, LOAD ''y'', ADD] <''x'' := 42, ''y'' := 43> [50]" |
43141 | 36 |
|
45257 | 37 |
lemma exec_append[simp]: |
38 |
"exec (is1@is2) s stk = exec is2 s (exec is1 s stk)" |
|
45015 | 39 |
apply(induction is1 arbitrary: stk) |
43141 | 40 |
apply (auto) |
41 |
done |
|
42 |
||
43 |
||
44 |
subsection "Compilation" |
|
45 |
||
49191 | 46 |
text_raw{*\snip{ASMcompdef}{0}{2}{% *} |
45257 | 47 |
fun comp :: "aexp \<Rightarrow> instr list" where |
48 |
"comp (N n) = [LOADI n]" | |
|
49 |
"comp (V x) = [LOAD x]" | |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
50007
diff
changeset
|
50 |
"comp (Plus e\<^sub>1 e\<^sub>2) = comp e\<^sub>1 @ comp e\<^sub>2 @ [ADD]" |
49191 | 51 |
text_raw{*}%endsnip*} |
43141 | 52 |
|
45257 | 53 |
value "comp (Plus (Plus (V ''x'') (N 1)) (V ''z''))" |
43141 | 54 |
|
45257 | 55 |
theorem exec_comp: "exec (comp a) s stk = aval a s # stk" |
45015 | 56 |
apply(induction a arbitrary: stk) |
43141 | 57 |
apply (auto) |
58 |
done |
|
59 |
||
60 |
end |