| author | wenzelm | 
| Thu, 09 Jun 2011 16:34:49 +0200 | |
| changeset 43324 | 2b47822868e4 | 
| parent 42463 | f270e3e18be5 | 
| child 58886 | 8a6cac7c7247 | 
| permissions | -rw-r--r-- | 
| 8011 | 1  | 
(* Title: HOL/MicroJava/JVM/JVMState.thy  | 
| 41589 | 2  | 
Author: Cornelia Pusch, Gerwin Klein, Technische Universitaet Muenchen  | 
| 8011 | 3  | 
*)  | 
4  | 
||
| 12911 | 5  | 
header {* 
 | 
6  | 
  \chapter{Java Virtual Machine}\label{cha:jvm}
 | 
|
7  | 
  \isaheader{State of the JVM} 
 | 
|
8  | 
*}  | 
|
| 8011 | 9  | 
|
| 15860 | 10  | 
theory JVMState  | 
11  | 
imports "../J/Conform"  | 
|
12  | 
begin  | 
|
| 8011 | 13  | 
|
| 12519 | 14  | 
section {* Frame Stack *}
 | 
| 42463 | 15  | 
type_synonym opstack = "val list"  | 
16  | 
type_synonym locvars = "val list"  | 
|
17  | 
type_synonym p_count = nat  | 
|
| 8011 | 18  | 
|
| 42463 | 19  | 
type_synonym  | 
| 12519 | 20  | 
frame = "opstack \<times>  | 
21  | 
locvars \<times>  | 
|
22  | 
cname \<times>  | 
|
23  | 
sig \<times>  | 
|
| 10057 | 24  | 
p_count"  | 
| 8011 | 25  | 
|
| 12519 | 26  | 
-- "operand stack"  | 
27  | 
-- "local variables (including this pointer and method parameters)"  | 
|
28  | 
-- "name of class where current method is defined"  | 
|
29  | 
-- "method name + parameter types"  | 
|
30  | 
-- "program counter within frame"  | 
|
31  | 
||
32  | 
||
33  | 
section {* Exceptions *}
 | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
15860 
diff
changeset
 | 
34  | 
definition raise_system_xcpt :: "bool \<Rightarrow> xcpt \<Rightarrow> val option" where  | 
| 13674 | 35  | 
"raise_system_xcpt b x \<equiv> raise_if b x None"  | 
| 12519 | 36  | 
|
37  | 
section {* Runtime State *}
 | 
|
| 42463 | 38  | 
type_synonym  | 
| 12519 | 39  | 
jvm_state = "val option \<times> aheap \<times> frame list" -- "exception flag, heap, frames"  | 
| 8011 | 40  | 
|
41  | 
||
| 12519 | 42  | 
section {* Lemmas *}
 | 
| 8011 | 43  | 
|
| 12519 | 44  | 
lemma new_Addr_OutOfMemory:  | 
45  | 
"snd (new_Addr hp) = Some xcp \<Longrightarrow> xcp = Addr (XcptRef OutOfMemory)"  | 
|
46  | 
proof -  | 
|
47  | 
obtain ref xp where "new_Addr hp = (ref, xp)" by (cases "new_Addr hp")  | 
|
48  | 
moreover  | 
|
49  | 
assume "snd (new_Addr hp) = Some xcp"  | 
|
50  | 
ultimately  | 
|
51  | 
show ?thesis by (auto dest: new_AddrD)  | 
|
| 13052 | 52  | 
qed  | 
53  | 
||
| 13674 | 54  | 
end  |