author | kleing |
Sun, 03 Mar 2002 16:59:08 +0100 | |
changeset 13006 | 51c5f3f11d16 |
parent 12911 | 704713ca07ea |
child 13052 | 3bf41c474a88 |
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:
10057
diff
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" |
|
36 |
"raise_system_xcpt b x == if b then Some (Addr (XcptRef x)) else None" |
|
37 |
||
38 |
-- "redefines State.new\\_Addr:" |
|
39 |
new_Addr :: "aheap => loc \<times> val option" |
|
40 |
"new_Addr h == SOME (a,x). (h a = None \<and> x = None) | |
|
41 |
x = raise_system_xcpt True OutOfMemory" |
|
42 |
||
43 |
||
44 |
section {* Runtime State *} |
|
45 |
types |
|
46 |
jvm_state = "val option \<times> aheap \<times> frame list" -- "exception flag, heap, frames" |
|
8011 | 47 |
|
48 |
||
12519 | 49 |
section {* Lemmas *} |
8011 | 50 |
|
12519 | 51 |
lemma new_AddrD: |
52 |
"new_Addr hp = (ref, xcp) \<Longrightarrow> hp ref = None \<and> xcp = None \<or> xcp = Some (Addr (XcptRef OutOfMemory))" |
|
53 |
apply (drule sym) |
|
54 |
apply (unfold new_Addr_def) |
|
55 |
apply (simp add: raise_system_xcpt_def) |
|
56 |
apply (simp add: Pair_fst_snd_eq Eps_split) |
|
57 |
apply (rule someI) |
|
58 |
apply (rule disjI2) |
|
59 |
apply (rule_tac "r" = "snd (?a,Some (Addr (XcptRef OutOfMemory)))" in trans) |
|
60 |
apply auto |
|
61 |
done |
|
8011 | 62 |
|
12519 | 63 |
lemma new_Addr_OutOfMemory: |
64 |
"snd (new_Addr hp) = Some xcp \<Longrightarrow> xcp = Addr (XcptRef OutOfMemory)" |
|
65 |
proof - |
|
66 |
obtain ref xp where "new_Addr hp = (ref, xp)" by (cases "new_Addr hp") |
|
67 |
moreover |
|
68 |
assume "snd (new_Addr hp) = Some xcp" |
|
69 |
ultimately |
|
70 |
show ?thesis by (auto dest: new_AddrD) |
|
71 |
qed |
|
72 |
||
73 |
end |