src/HOL/NanoJava/State.thy
author wenzelm
Wed, 24 Feb 2010 22:09:50 +0100
changeset 35355 613e133966ea
parent 32960 69916a850301
child 35417 47ee18b6ae32
permissions -rw-r--r--
modernized syntax declarations, and make them actually work with authentic syntax;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     1
(*  Title:      HOL/NanoJava/State.thy
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     2
    Author:     David von Oheimb
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     3
    Copyright   2001 Technische Universitaet Muenchen
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     4
*)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     5
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     6
header "Program State"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14134
diff changeset
     8
theory State imports TypeRel begin
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     9
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    10
constdefs
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    11
11507
4b32a46ffd29 removed imname, uncurried Meth
oheimb
parents: 11499
diff changeset
    12
  body :: "cname \<times> mname => stmt"
11497
0e66e0114d9a corrected initialization of locals, streamlined Impl
oheimb
parents: 11376
diff changeset
    13
 "body \<equiv> \<lambda>(C,m). bdy (the (method C m))"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    14
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
    15
text {* Locations, i.e.\ abstract references to objects *}
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    16
typedecl loc 
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    17
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    18
datatype val
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    19
  = Null        --{* null reference *}
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    20
  | Addr loc    --{* address, i.e. location of object *}
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    21
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30235
diff changeset
    22
types   fields
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30235
diff changeset
    23
        = "(fname \<rightharpoonup> val)"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    24
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    25
        obj = "cname \<times> fields"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    26
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    27
translations
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    28
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    29
  "fields" \<leftharpoondown> (type)"fname => val option"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    30
  "obj"    \<leftharpoondown> (type)"cname \<times> fields"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    31
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    32
constdefs
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    33
14134
0fdf5708c7a8 Replaced \<leadsto> by \<rightharpoonup>
nipkow
parents: 13524
diff changeset
    34
  init_vars:: "('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> val)"
30235
58d147683393 Made Option a separate theory and renamed option_map to Option.map
nipkow
parents: 16417
diff changeset
    35
 "init_vars m == Option.map (\<lambda>T. Null) o m"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    36
  
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
    37
text {* private: *}
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30235
diff changeset
    38
types   heap   = "loc   \<rightharpoonup> obj"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30235
diff changeset
    39
        locals = "vname \<rightharpoonup> val"  
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    40
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
    41
text {* private: *}
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    42
record  state
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30235
diff changeset
    43
        = heap   :: heap
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    44
          locals :: locals
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    45
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    46
translations
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    47
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    48
  "heap"   \<leftharpoondown> (type)"loc   => obj option"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    49
  "locals" \<leftharpoondown> (type)"vname => val option"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    50
  "state" \<leftharpoondown> (type)"(|heap :: heap, locals :: locals|)"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    51
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    52
constdefs
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    53
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
    54
  del_locs     :: "state => state"
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
    55
 "del_locs s \<equiv> s (| locals := empty |)"
11497
0e66e0114d9a corrected initialization of locals, streamlined Impl
oheimb
parents: 11376
diff changeset
    56
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    57
  init_locs     :: "cname => mname => state => state"
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
    58
 "init_locs C m s \<equiv> s (| locals := locals s ++ 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
    59
                         init_vars (map_of (lcl (the (method C m)))) |)"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    60
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    61
text {* The first parameter of @{term set_locs} is of type @{typ state} 
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    62
        rather than @{typ locals} in order to keep @{typ locals} private.*}
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    63
constdefs
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    64
  set_locs  :: "state => state => state"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    65
 "set_locs s s' \<equiv> s' (| locals := locals s |)"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    66
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    67
  get_local     :: "state => vname => val" ("_<_>" [99,0] 99)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    68
 "get_local s x  \<equiv> the (locals s x)"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    69
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    70
--{* local function: *}
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    71
  get_obj       :: "state => loc => obj"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    72
 "get_obj s a \<equiv> the (heap s a)"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    73
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    74
  obj_class     :: "state => loc => cname"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    75
 "obj_class s a \<equiv> fst (get_obj s a)"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    76
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    77
  get_field     :: "state => loc => fname => val"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    78
 "get_field s a f \<equiv> the (snd (get_obj s a) f)"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    79
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    80
--{* local function: *}
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    81
  hupd       :: "loc => obj => state => state"   ("hupd'(_|->_')" [10,10] 1000)
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    82
 "hupd a obj s \<equiv> s (| heap   := ((heap   s)(a\<mapsto>obj))|)"
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    83
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    84
  lupd       :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000)
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    85
 "lupd x v s   \<equiv> s (| locals := ((locals s)(x\<mapsto>v  ))|)"
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    86
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 32960
diff changeset
    87
notation (xsymbols)
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 32960
diff changeset
    88
  hupd  ("hupd'(_\<mapsto>_')" [10,10] 1000) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 32960
diff changeset
    89
  lupd  ("lupd'(_\<mapsto>_')" [10,10] 1000)
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    90
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    91
constdefs
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    92
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    93
  new_obj    :: "loc => cname => state => state"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    94
 "new_obj a C   \<equiv> hupd(a\<mapsto>(C,init_vars (field C)))"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    95
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11507
diff changeset
    96
  upd_obj    :: "loc => fname => val => state => state"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    97
 "upd_obj a f v s \<equiv> let (C,fs) = the (heap s a) in hupd(a\<mapsto>(C,fs(f\<mapsto>v))) s"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    98
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30235
diff changeset
    99
  new_Addr      :: "state => val"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   100
 "new_Addr s == SOME v. (\<exists>a. v = Addr a \<and> (heap s) a = None) | v = Null"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   101
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   102
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   103
subsection "Properties not used in the meta theory"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   104
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   105
lemma locals_upd_id [simp]: "s\<lparr>locals := locals s\<rparr> = s" 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   106
by simp
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   107
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   108
lemma lupd_get_local_same [simp]: "lupd(x\<mapsto>v) s<x> = v"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   109
by (simp add: lupd_def get_local_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   110
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   111
lemma lupd_get_local_other [simp]: "x \<noteq> y \<Longrightarrow> lupd(x\<mapsto>v) s<y> = s<y>"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   112
apply (drule not_sym)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   113
by (simp add: lupd_def get_local_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   114
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   115
lemma get_field_lupd [simp]:
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   116
  "get_field (lupd(x\<mapsto>y) s) a f = get_field s a f"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   117
by (simp add: lupd_def get_field_def get_obj_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   118
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   119
lemma get_field_set_locs [simp]:
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   120
  "get_field (set_locs l s) a f = get_field s a f"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   121
by (simp add: lupd_def get_field_def set_locs_def get_obj_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   122
13524
604d0f3622d6 *** empty log message ***
wenzelm
parents: 11772
diff changeset
   123
lemma get_field_del_locs [simp]:
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   124
  "get_field (del_locs s) a f = get_field s a f"
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   125
by (simp add: lupd_def get_field_def del_locs_def get_obj_def)
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   126
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   127
lemma new_obj_get_local [simp]: "new_obj a C s <x> = s<x>"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   128
by (simp add: new_obj_def hupd_def get_local_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   129
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   130
lemma heap_lupd [simp]: "heap (lupd(x\<mapsto>y) s) = heap s"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   131
by (simp add: lupd_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   132
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   133
lemma heap_hupd_same [simp]: "heap (hupd(a\<mapsto>obj) s) a = Some obj"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   134
by (simp add: hupd_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   135
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   136
lemma heap_hupd_other [simp]: "aa \<noteq> a  \<Longrightarrow> heap (hupd(aa\<mapsto>obj) s) a = heap s a"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   137
apply (drule not_sym)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   138
by (simp add: hupd_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   139
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   140
lemma hupd_hupd [simp]: "hupd(a\<mapsto>obj) (hupd(a\<mapsto>obj') s) = hupd(a\<mapsto>obj) s"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   141
by (simp add: hupd_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   142
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   143
lemma heap_del_locs [simp]: "heap (del_locs s) = heap s"
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   144
by (simp add: del_locs_def)
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   145
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   146
lemma heap_set_locs [simp]: "heap (set_locs l s) = heap s"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   147
by (simp add: set_locs_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   148
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   149
lemma hupd_lupd [simp]: 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   150
  "hupd(a\<mapsto>obj) (lupd(x\<mapsto>y) s) = lupd(x\<mapsto>y) (hupd(a\<mapsto>obj) s)"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   151
by (simp add: hupd_def lupd_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   152
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   153
lemma hupd_del_locs [simp]: 
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   154
  "hupd(a\<mapsto>obj) (del_locs s) = del_locs (hupd(a\<mapsto>obj) s)"
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   155
by (simp add: hupd_def del_locs_def)
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   156
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   157
lemma new_obj_lupd [simp]: 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   158
  "new_obj a C (lupd(x\<mapsto>y) s) = lupd(x\<mapsto>y) (new_obj a C s)"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   159
by (simp add: new_obj_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   160
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   161
lemma new_obj_del_locs [simp]: 
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   162
  "new_obj a C (del_locs s) = del_locs (new_obj a C s)"
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   163
by (simp add: new_obj_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   164
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   165
lemma upd_obj_lupd [simp]: 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   166
  "upd_obj a f v (lupd(x\<mapsto>y) s) = lupd(x\<mapsto>y) (upd_obj a f v s)"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   167
by (simp add: upd_obj_def Let_def split_beta)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   168
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   169
lemma upd_obj_del_locs [simp]: 
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
   170
  "upd_obj a f v (del_locs s) = del_locs (upd_obj a f v s)"
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   171
by (simp add: upd_obj_def Let_def split_beta)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   172
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   173
lemma get_field_hupd_same [simp]:
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   174
 "get_field (hupd(a\<mapsto>(C, fs)) s) a = the \<circ> fs"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   175
apply (rule ext)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   176
by (simp add: get_field_def get_obj_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   177
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   178
lemma get_field_hupd_other [simp]:
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   179
 "aa \<noteq> a  \<Longrightarrow> get_field (hupd(aa\<mapsto>obj) s) a = get_field s a"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   180
apply (rule ext)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   181
by (simp add: get_field_def get_obj_def)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11558
diff changeset
   182
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   183
lemma new_AddrD: 
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   184
"new_Addr s = v \<Longrightarrow> (\<exists>a. v = Addr a \<and> heap s a = None) | v = Null"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   185
apply (unfold new_Addr_def)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   186
apply (erule subst)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   187
apply (rule someI)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   188
apply (rule disjI2)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   189
apply (rule HOL.refl)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   190
done
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   191
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   192
end