12519
|
1 |
(* Title: HOL/MicroJava/JVM/JVMExceptions.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Gerwin Klein, Martin Strecker
|
|
4 |
Copyright 2001 Technische Universitaet Muenchen
|
|
5 |
*)
|
|
6 |
|
|
7 |
header {* Exception handling in the JVM *}
|
|
8 |
|
|
9 |
theory JVMExceptions = JVMInstructions:
|
|
10 |
|
|
11 |
constdefs
|
|
12 |
match_exception_entry :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_entry \<Rightarrow> bool"
|
|
13 |
"match_exception_entry G cn pc ee ==
|
|
14 |
let (start_pc, end_pc, handler_pc, catch_type) = ee in
|
|
15 |
start_pc <= pc \<and> pc < end_pc \<and> G\<turnstile> cn \<preceq>C catch_type"
|
|
16 |
|
|
17 |
|
|
18 |
consts
|
|
19 |
match_exception_table :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_table
|
|
20 |
\<Rightarrow> p_count option"
|
|
21 |
primrec
|
|
22 |
"match_exception_table G cn pc [] = None"
|
|
23 |
"match_exception_table G cn pc (e#es) = (if match_exception_entry G cn pc e
|
|
24 |
then Some (fst (snd (snd e)))
|
|
25 |
else match_exception_table G cn pc es)"
|
|
26 |
|
|
27 |
|
|
28 |
consts
|
|
29 |
cname_of :: "aheap \<Rightarrow> val \<Rightarrow> cname"
|
|
30 |
ex_table_of :: "jvm_method \<Rightarrow> exception_table"
|
|
31 |
|
|
32 |
translations
|
|
33 |
"cname_of hp v" == "fst (the (hp (the_Addr v)))"
|
|
34 |
"ex_table_of m" == "snd (snd (snd m))"
|
|
35 |
|
|
36 |
|
|
37 |
consts
|
|
38 |
find_handler :: "jvm_prog \<Rightarrow> val option \<Rightarrow> aheap \<Rightarrow> frame list \<Rightarrow> jvm_state"
|
|
39 |
primrec
|
|
40 |
"find_handler G xcpt hp [] = (xcpt, hp, [])"
|
|
41 |
"find_handler G xcpt hp (fr#frs) =
|
|
42 |
(case xcpt of
|
|
43 |
None \<Rightarrow> (None, hp, fr#frs)
|
|
44 |
| Some xc \<Rightarrow>
|
|
45 |
let (stk,loc,C,sig,pc) = fr in
|
|
46 |
(case match_exception_table G (cname_of hp xc) pc
|
|
47 |
(ex_table_of (snd(snd(the(method (G,C) sig))))) of
|
|
48 |
None \<Rightarrow> find_handler G (Some xc) hp frs
|
|
49 |
| Some handler_pc \<Rightarrow> (None, hp, ([xc], loc, C, sig, handler_pc)#frs)))"
|
|
50 |
|
|
51 |
|
|
52 |
|
|
53 |
lemma cname_of_xcp:
|
|
54 |
"raise_system_xcpt b x = Some xcp \<Longrightarrow> cname_of (hp::aheap) xcp = Xcpt x"
|
|
55 |
proof -
|
|
56 |
assume "raise_system_xcpt b x = Some xcp"
|
|
57 |
hence "xcp = Addr (XcptRef x)"
|
|
58 |
by (simp add: raise_system_xcpt_def split: split_if_asm)
|
|
59 |
moreover
|
|
60 |
have "hp (XcptRef x) = Some (Xcpt x, empty)"
|
|
61 |
by (rule system_xcpt_allocated)
|
|
62 |
ultimately
|
|
63 |
show ?thesis by simp
|
|
64 |
qed
|
|
65 |
|
|
66 |
end
|