| author | kleing | 
| Tue, 30 Nov 2004 06:50:03 +0100 | |
| changeset 15343 | 444bb25d3da0 | 
| parent 13674 | f4c64597fb02 | 
| child 15860 | a344c4284972 | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/JVM/JVMState.thy | 
| 2 | ID: $Id$ | |
| 12519 | 3 | Author: Cornelia Pusch, Gerwin Klein | 
| 8011 | 4 | Copyright 1999 Technische Universitaet Muenchen | 
| 5 | *) | |
| 6 | ||
| 12911 | 7 | header {* 
 | 
| 8 |   \chapter{Java Virtual Machine}\label{cha:jvm}
 | |
| 9 |   \isaheader{State of the JVM} 
 | |
| 10 | *} | |
| 8011 | 11 | |
| 10922 
f1209aff9517
Store.thy is obsolete (newref isn't used any more)
 kleing parents: 
10057diff
changeset | 12 | theory JVMState = Conform: | 
| 8011 | 13 | |
| 12519 | 14 | section {* Frame Stack *}
 | 
| 8011 | 15 | types | 
| 12519 | 16 | opstack = "val list" | 
| 17 | locvars = "val list" | |
| 18 | p_count = nat | |
| 8011 | 19 | |
| 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 *}
 | |
| 34 | constdefs | |
| 35 | raise_system_xcpt :: "bool \<Rightarrow> xcpt \<Rightarrow> val option" | |
| 13674 | 36 | "raise_system_xcpt b x \<equiv> raise_if b x None" | 
| 12519 | 37 | |
| 38 | section {* Runtime State *}
 | |
| 39 | types | |
| 40 | jvm_state = "val option \<times> aheap \<times> frame list" -- "exception flag, heap, frames" | |
| 8011 | 41 | |
| 42 | ||
| 12519 | 43 | section {* Lemmas *}
 | 
| 8011 | 44 | |
| 12519 | 45 | lemma new_Addr_OutOfMemory: | 
| 46 | "snd (new_Addr hp) = Some xcp \<Longrightarrow> xcp = Addr (XcptRef OutOfMemory)" | |
| 47 | proof - | |
| 48 | obtain ref xp where "new_Addr hp = (ref, xp)" by (cases "new_Addr hp") | |
| 49 | moreover | |
| 50 | assume "snd (new_Addr hp) = Some xcp" | |
| 51 | ultimately | |
| 52 | show ?thesis by (auto dest: new_AddrD) | |
| 13052 | 53 | qed | 
| 54 | ||
| 13674 | 55 | end |