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