| author | wenzelm | 
| Mon, 08 Jun 2020 22:31:36 +0200 | |
| changeset 71928 | ae643fb4ca30 | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 12519 | 1 | (* Title: HOL/MicroJava/JVM/JVMExceptions.thy | 
| 2 | Author: Gerwin Klein, Martin Strecker | |
| 3 | Copyright 2001 Technische Universitaet Muenchen | |
| 4 | *) | |
| 5 | ||
| 61361 | 6 | section \<open>Exception handling in the JVM\<close> | 
| 12519 | 7 | |
| 16417 | 8 | theory JVMExceptions imports JVMInstructions begin | 
| 12519 | 9 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 10 | definition match_exception_entry :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_entry \<Rightarrow> bool" where | 
| 12519 | 11 | "match_exception_entry G cn pc ee == | 
| 12 | let (start_pc, end_pc, handler_pc, catch_type) = ee in | |
| 13 | start_pc <= pc \<and> pc < end_pc \<and> G\<turnstile> cn \<preceq>C catch_type" | |
| 14 | ||
| 15 | ||
| 39758 | 16 | primrec match_exception_table :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_table | 
| 17 | \<Rightarrow> p_count option" | |
| 18 | where | |
| 12519 | 19 | "match_exception_table G cn pc [] = None" | 
| 39758 | 20 | | "match_exception_table G cn pc (e#es) = (if match_exception_entry G cn pc e | 
| 12519 | 21 | then Some (fst (snd (snd e))) | 
| 22 | else match_exception_table G cn pc es)" | |
| 23 | ||
| 35102 | 24 | abbreviation | 
| 12519 | 25 | ex_table_of :: "jvm_method \<Rightarrow> exception_table" | 
| 35102 | 26 | where "ex_table_of m == snd (snd (snd m))" | 
| 12519 | 27 | |
| 28 | ||
| 39758 | 29 | primrec find_handler :: "jvm_prog \<Rightarrow> val option \<Rightarrow> aheap \<Rightarrow> frame list | 
| 30 | \<Rightarrow> jvm_state" | |
| 31 | where | |
| 12519 | 32 | "find_handler G xcpt hp [] = (xcpt, hp, [])" | 
| 39758 | 33 | | "find_handler G xcpt hp (fr#frs) = | 
| 12519 | 34 | (case xcpt of | 
| 35 | None \<Rightarrow> (None, hp, fr#frs) | |
| 36 | | Some xc \<Rightarrow> | |
| 37 | let (stk,loc,C,sig,pc) = fr in | |
| 38 | (case match_exception_table G (cname_of hp xc) pc | |
| 39 | (ex_table_of (snd(snd(the(method (G,C) sig))))) of | |
| 40 | None \<Rightarrow> find_handler G (Some xc) hp frs | |
| 41 | | Some handler_pc \<Rightarrow> (None, hp, ([xc], loc, C, sig, handler_pc)#frs)))" | |
| 42 | ||
| 43 | ||
| 61361 | 44 | text \<open> | 
| 13052 | 45 | System exceptions are allocated in all heaps: | 
| 61361 | 46 | \<close> | 
| 13052 | 47 | |
| 13065 | 48 | |
| 61361 | 49 | text \<open> | 
| 13065 | 50 | Only program counters that are mentioned in the exception table | 
| 69597 | 51 | can be returned by \<^term>\<open>match_exception_table\<close>: | 
| 61361 | 52 | \<close> | 
| 13065 | 53 | lemma match_exception_table_in_et: | 
| 54 | "match_exception_table G C pc et = Some pc' \<Longrightarrow> \<exists>e \<in> set et. pc' = fst (snd (snd e))" | |
| 62390 | 55 | by (induct et) (auto split: if_split_asm) | 
| 13065 | 56 | |
| 57 | ||
| 12519 | 58 | end |