| author | wenzelm | 
| Sun, 05 Jul 2020 11:06:09 +0200 | |
| changeset 71999 | 720b72513ae5 | 
| parent 67406 | 23307fd33906 | 
| permissions | -rw-r--r-- | 
| 58889 | 1  | 
section "Stack Machine and Compilation"  | 
| 43141 | 2  | 
|
3  | 
theory ASM imports AExp begin  | 
|
4  | 
||
| 45266 | 5  | 
subsection "Stack Machine"  | 
| 43141 | 6  | 
|
| 67406 | 7  | 
text_raw\<open>\snip{ASMinstrdef}{0}{1}{%\<close>
 | 
| 58310 | 8  | 
datatype instr = LOADI val | LOAD vname | ADD  | 
| 67406 | 9  | 
text_raw\<open>}%endsnip\<close>  | 
| 43141 | 10  | 
|
| 67406 | 11  | 
text_raw\<open>\snip{ASMstackdef}{1}{2}{%\<close>
 | 
| 43141 | 12  | 
type_synonym stack = "val list"  | 
| 67406 | 13  | 
text_raw\<open>}%endsnip\<close>  | 
| 43141 | 14  | 
|
| 67406 | 15  | 
text\<open>\noindent Abbreviations are transparent: they are unfolded after  | 
| 43141 | 16  | 
parsing and folded back again before printing. Internally, they do not  | 
| 67406 | 17  | 
exist.\<close>  | 
| 43141 | 18  | 
|
| 67406 | 19  | 
text_raw\<open>\snip{ASMexeconedef}{0}{1}{%\<close>
 | 
| 45257 | 20  | 
fun exec1 :: "instr \<Rightarrow> state \<Rightarrow> stack \<Rightarrow> stack" where  | 
21  | 
"exec1 (LOADI n) _ stk = n # stk" |  | 
|
22  | 
"exec1 (LOAD x) s stk = s(x) # stk" |  | 
|
| 67079 | 23  | 
"exec1 ADD _ (j # i # stk) = (i + j) # stk"  | 
| 67406 | 24  | 
text_raw\<open>}%endsnip\<close>  | 
| 43141 | 25  | 
|
| 67406 | 26  | 
text_raw\<open>\snip{ASMexecdef}{1}{2}{%\<close>
 | 
| 45257 | 27  | 
fun exec :: "instr list \<Rightarrow> state \<Rightarrow> stack \<Rightarrow> stack" where  | 
28  | 
"exec [] _ stk = stk" |  | 
|
29  | 
"exec (i#is) s stk = exec is s (exec1 i s stk)"  | 
|
| 67406 | 30  | 
text_raw\<open>}%endsnip\<close>  | 
| 43141 | 31  | 
|
| 
50007
 
56f269baae76
executable true liveness analysis incl an approximating version
 
nipkow 
parents: 
49191 
diff
changeset
 | 
32  | 
value "exec [LOADI 5, LOAD ''y'', ADD] <''x'' := 42, ''y'' := 43> [50]"  | 
| 43141 | 33  | 
|
| 45257 | 34  | 
lemma exec_append[simp]:  | 
35  | 
"exec (is1@is2) s stk = exec is2 s (exec is1 s stk)"  | 
|
| 45015 | 36  | 
apply(induction is1 arbitrary: stk)  | 
| 43141 | 37  | 
apply (auto)  | 
38  | 
done  | 
|
39  | 
||
40  | 
||
41  | 
subsection "Compilation"  | 
|
42  | 
||
| 67406 | 43  | 
text_raw\<open>\snip{ASMcompdef}{0}{2}{%\<close>
 | 
| 45257 | 44  | 
fun comp :: "aexp \<Rightarrow> instr list" where  | 
45  | 
"comp (N n) = [LOADI n]" |  | 
|
46  | 
"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
 | 
47  | 
"comp (Plus e\<^sub>1 e\<^sub>2) = comp e\<^sub>1 @ comp e\<^sub>2 @ [ADD]"  | 
| 67406 | 48  | 
text_raw\<open>}%endsnip\<close>  | 
| 43141 | 49  | 
|
| 45257 | 50  | 
value "comp (Plus (Plus (V ''x'') (N 1)) (V ''z''))"  | 
| 43141 | 51  | 
|
| 45257 | 52  | 
theorem exec_comp: "exec (comp a) s stk = aval a s # stk"  | 
| 45015 | 53  | 
apply(induction a arbitrary: stk)  | 
| 43141 | 54  | 
apply (auto)  | 
55  | 
done  | 
|
56  | 
||
57  | 
end  |