| author | wenzelm | 
| Thu, 10 Apr 2008 14:53:25 +0200 | |
| changeset 26609 | 53754d9ee31f | 
| parent 24783 | 5a3e336a2e37 | 
| child 30235 | 58d147683393 | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/J/State.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: David von Oheimb | |
| 4 | Copyright 1999 Technische Universitaet Muenchen | |
| 11070 | 5 | *) | 
| 8011 | 6 | |
| 12911 | 7 | header {* \isaheader{Program State} *}
 | 
| 8011 | 8 | |
| 16417 | 9 | theory State imports TypeRel Value begin | 
| 8011 | 10 | |
| 12517 | 11 | types | 
| 24783 | 12 | fields' = "(vname \<times> cname \<rightharpoonup> val)" -- "field name, defining class, value" | 
| 8011 | 13 | |
| 24783 | 14 | obj = "cname \<times> fields'" -- "class instance with class name and fields" | 
| 8011 | 15 | |
| 16 | constdefs | |
| 12517 | 17 | obj_ty :: "obj => ty" | 
| 10042 | 18 | "obj_ty obj == Class (fst obj)" | 
| 8011 | 19 | |
| 14134 | 20 |   init_vars :: "('a \<times> ty) list => ('a \<rightharpoonup> val)"
 | 
| 12517 | 21 | "init_vars == map_of o map (\<lambda>(n,T). (n,default_val T))" | 
| 8011 | 22 | |
| 12517 | 23 | |
| 14134 | 24 | types aheap  = "loc \<rightharpoonup> obj"    -- {* "@{text heap}" used in a translation below *}
 | 
| 25 | locals = "vname \<rightharpoonup> val" -- "simple state, i.e. variable contents" | |
| 8011 | 26 | |
| 12517 | 27 | state = "aheap \<times> locals" -- "heap, local parameter including This" | 
| 13672 | 28 | xstate = "val option \<times> state" -- "state including exception information" | 
| 12517 | 29 | |
| 8011 | 30 | syntax | 
| 12517 | 31 | heap :: "state => aheap" | 
| 32 | locals :: "state => locals" | |
| 33 | Norm :: "state => xstate" | |
| 13672 | 34 | abrupt :: "xstate \<Rightarrow> val option" | 
| 35 | store :: "xstate \<Rightarrow> state" | |
| 36 | lookup_obj :: "state \<Rightarrow> val \<Rightarrow> obj" | |
| 8011 | 37 | |
| 38 | translations | |
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 39 | "heap" => "fst" | 
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 40 | "locals" => "snd" | 
| 12517 | 41 | "Norm s" == "(None,s)" | 
| 13672 | 42 | "abrupt" => "fst" | 
| 43 | "store" => "snd" | |
| 44 | "lookup_obj s a'" == "the (heap s (the_Addr a'))" | |
| 45 | ||
| 8011 | 46 | |
| 47 | constdefs | |
| 13672 | 48 | raise_if :: "bool \<Rightarrow> xcpt \<Rightarrow> val option \<Rightarrow> val option" | 
| 49 | "raise_if b x xo \<equiv> if b \<and> (xo = None) then Some (Addr (XcptRef x)) else xo" | |
| 8011 | 50 | |
| 13672 | 51 | new_Addr :: "aheap => loc \<times> val option" | 
| 52 | "new_Addr h \<equiv> SOME (a,x). (h a = None \<and> x = None) | x = Some (Addr (XcptRef OutOfMemory))" | |
| 8011 | 53 | |
| 13672 | 54 | np :: "val => val option => val option" | 
| 10042 | 55 | "np v == raise_if (v = Null) NullPointer" | 
| 8011 | 56 | |
| 12517 | 57 | c_hupd :: "aheap => xstate => xstate" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 58 | "c_hupd h'== \<lambda>(xo,(h,l)). if xo = None then (None,(h',l)) else (xo,(h,l))" | 
| 8011 | 59 | |
| 12517 | 60 | cast_ok :: "'c prog => cname => aheap => val => bool" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 61 | "cast_ok G C h v == v = Null \<or> G\<turnstile>obj_ty (the (h (the_Addr v)))\<preceq> Class C" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 62 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 63 | lemma obj_ty_def2 [simp]: "obj_ty (C,fs) = Class C" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 64 | apply (unfold obj_ty_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 65 | apply (simp (no_asm)) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 66 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 67 | |
| 13672 | 68 | |
| 69 | lemma new_AddrD: "new_Addr hp = (ref, xcp) \<Longrightarrow> | |
| 70 | hp ref = None \<and> xcp = None \<or> xcp = Some (Addr (XcptRef OutOfMemory))" | |
| 71 | apply (drule sym) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 72 | apply (unfold new_Addr_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 73 | apply(simp add: Pair_fst_snd_eq Eps_split) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 74 | apply(rule someI) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 75 | apply(rule disjI2) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
14144diff
changeset | 76 | apply(rule_tac r = "snd (?a,Some (Addr (XcptRef OutOfMemory)))" in trans) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 77 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 78 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 79 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 80 | lemma raise_if_True [simp]: "raise_if True x y \<noteq> None" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 81 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 82 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 83 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 84 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 85 | lemma raise_if_False [simp]: "raise_if False x y = y" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 86 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 87 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 88 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 89 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 90 | lemma raise_if_Some [simp]: "raise_if c x (Some y) \<noteq> None" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 91 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 92 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 93 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 94 | |
| 12517 | 95 | lemma raise_if_Some2 [simp]: | 
| 96 | "raise_if c z (if x = None then Some y else x) \<noteq> None" | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 97 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 98 | apply(induct_tac "x") | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 99 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 100 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 101 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 102 | lemma raise_if_SomeD [rule_format (no_asm)]: | 
| 13672 | 103 | "raise_if c x y = Some z \<longrightarrow> c \<and> Some z = Some (Addr (XcptRef x)) | y = Some z" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 104 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 105 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 106 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 107 | |
| 12517 | 108 | lemma raise_if_NoneD [rule_format (no_asm)]: | 
| 109 | "raise_if c x y = None --> \<not> c \<and> y = None" | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 110 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 111 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 112 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 113 | |
| 12517 | 114 | lemma np_NoneD [rule_format (no_asm)]: | 
| 115 | "np a' x' = None --> x' = None \<and> a' \<noteq> Null" | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 116 | apply (unfold np_def raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 117 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 118 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 119 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 120 | lemma np_None [rule_format (no_asm), simp]: "a' \<noteq> Null --> np a' x' = x'" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 121 | apply (unfold np_def raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 122 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 123 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 124 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 125 | lemma np_Some [simp]: "np a' (Some xc) = Some xc" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 126 | apply (unfold np_def raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 127 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 128 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 129 | |
| 13672 | 130 | lemma np_Null [simp]: "np Null None = Some (Addr (XcptRef NullPointer))" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 131 | apply (unfold np_def raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 132 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 133 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 134 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 135 | lemma np_Addr [simp]: "np (Addr a) None = None" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 136 | apply (unfold np_def raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 137 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 138 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 139 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 140 | lemma np_raise_if [simp]: "(np Null (raise_if c xc None)) = | 
| 13672 | 141 | Some (Addr (XcptRef (if c then xc else NullPointer)))" | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 142 | apply (unfold raise_if_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 143 | apply (simp (no_asm)) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10061diff
changeset | 144 | done | 
| 8011 | 145 | |
| 14144 | 146 | lemma c_hupd_fst [simp]: "fst (c_hupd h (x, s)) = x" | 
| 147 | by (simp add: c_hupd_def split_beta) | |
| 148 | ||
| 8011 | 149 | end |