| author | blanchet | 
| Wed, 23 Jun 2010 09:40:06 +0200 | |
| changeset 37511 | 26afa11a1fb2 | 
| parent 36640 | 7eadf5acdaf4 | 
| child 37678 | 0040bafffdef | 
| permissions | -rw-r--r-- | 
| 26170 | 1 | (* Title: HOL/Library/Heap.thy | 
| 2 | Author: John Matthews, Galois Connections; Alexander Krauss, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* A polymorphic heap based on cantor encodings *}
 | |
| 6 | ||
| 7 | theory Heap | |
| 30738 | 8 | imports Main Countable | 
| 26170 | 9 | begin | 
| 10 | ||
| 11 | subsection {* Representable types *}
 | |
| 12 | ||
| 13 | text {* The type class of representable types *}
 | |
| 14 | ||
| 28335 | 15 | class heap = typerep + countable | 
| 26170 | 16 | |
| 17 | text {* Instances for common HOL types *}
 | |
| 18 | ||
| 19 | instance nat :: heap .. | |
| 20 | ||
| 21 | instance "*" :: (heap, heap) heap .. | |
| 22 | ||
| 23 | instance "+" :: (heap, heap) heap .. | |
| 24 | ||
| 25 | instance list :: (heap) heap .. | |
| 26 | ||
| 27 | instance option :: (heap) heap .. | |
| 28 | ||
| 29 | instance int :: heap .. | |
| 30 | ||
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
30738diff
changeset | 31 | instance String.literal :: countable | 
| 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
30738diff
changeset | 32 | by (rule countable_classI [of "String.literal_case to_nat"]) | 
| 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
30738diff
changeset | 33 | (auto split: String.literal.splits) | 
| 26170 | 34 | |
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
30738diff
changeset | 35 | instance String.literal :: heap .. | 
| 26170 | 36 | |
| 37 | text {* Reflected types themselves are heap-representable *}
 | |
| 38 | ||
| 28335 | 39 | instantiation typerep :: countable | 
| 26170 | 40 | begin | 
| 41 | ||
| 28335 | 42 | fun to_nat_typerep :: "typerep \<Rightarrow> nat" where | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28524diff
changeset | 43 | "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))" | 
| 26170 | 44 | |
| 26932 | 45 | instance | 
| 46 | proof (rule countable_classI) | |
| 28335 | 47 | fix t t' :: typerep and ts | 
| 48 | have "(\<forall>t'. to_nat_typerep t = to_nat_typerep t' \<longrightarrow> t = t') | |
| 49 | \<and> (\<forall>ts'. map to_nat_typerep ts = map to_nat_typerep ts' \<longrightarrow> ts = ts')" | |
| 50 | proof (induct rule: typerep.induct) | |
| 51 | case (Typerep c ts) show ?case | |
| 26170 | 52 | proof (rule allI, rule impI) | 
| 53 | fix t' | |
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28524diff
changeset | 54 | assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'" | 
| 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28524diff
changeset | 55 | then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')" | 
| 26170 | 56 | by (cases t') auto | 
| 28335 | 57 | with Typerep hyp have "c = c'" and "ts = ts'" by simp_all | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28524diff
changeset | 58 | with t' show "Typerep.Typerep c ts = t'" by simp | 
| 26170 | 59 | qed | 
| 60 | next | |
| 28335 | 61 | case Nil_typerep then show ?case by simp | 
| 26170 | 62 | next | 
| 28335 | 63 | case (Cons_typerep t ts) then show ?case by auto | 
| 26170 | 64 | qed | 
| 28335 | 65 | then have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'" by auto | 
| 66 | moreover assume "to_nat_typerep t = to_nat_typerep t'" | |
| 26170 | 67 | ultimately show "t = t'" by simp | 
| 68 | qed | |
| 69 | ||
| 70 | end | |
| 71 | ||
| 28335 | 72 | instance typerep :: heap .. | 
| 26170 | 73 | |
| 74 | ||
| 75 | subsection {* A polymorphic heap with dynamic arrays and references *}
 | |
| 76 | ||
| 77 | types addr = nat -- "untyped heap references" | |
| 78 | ||
| 79 | datatype 'a array = Array addr | |
| 80 | datatype 'a ref = Ref addr -- "note the phantom type 'a " | |
| 81 | ||
| 82 | primrec addr_of_array :: "'a array \<Rightarrow> addr" where | |
| 83 | "addr_of_array (Array x) = x" | |
| 84 | ||
| 85 | primrec addr_of_ref :: "'a ref \<Rightarrow> addr" where | |
| 86 | "addr_of_ref (Ref x) = x" | |
| 87 | ||
| 88 | lemma addr_of_array_inj [simp]: | |
| 89 | "addr_of_array a = addr_of_array a' \<longleftrightarrow> a = a'" | |
| 90 | by (cases a, cases a') simp_all | |
| 91 | ||
| 92 | lemma addr_of_ref_inj [simp]: | |
| 93 | "addr_of_ref r = addr_of_ref r' \<longleftrightarrow> r = r'" | |
| 94 | by (cases r, cases r') simp_all | |
| 95 | ||
| 96 | instance array :: (type) countable | |
| 97 | by (rule countable_classI [of addr_of_array]) simp | |
| 98 | ||
| 99 | instance ref :: (type) countable | |
| 100 | by (rule countable_classI [of addr_of_ref]) simp | |
| 101 | ||
| 102 | setup {*
 | |
| 103 |   Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap array"})
 | |
| 104 |   #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap ref"})
 | |
| 105 |   #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\<Colon>heap array \<Rightarrow> nat"})
 | |
| 106 |   #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\<Colon>heap ref \<Rightarrow> nat"})
 | |
| 107 | *} | |
| 108 | ||
| 109 | types heap_rep = nat -- "representable values" | |
| 110 | ||
| 111 | record heap = | |
| 28335 | 112 | arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list" | 
| 113 | refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep" | |
| 26170 | 114 | lim :: addr | 
| 115 | ||
| 116 | definition empty :: heap where | |
| 28524 | 117 | "empty = \<lparr>arrays = (\<lambda>_. undefined), refs = (\<lambda>_. undefined), lim = 0\<rparr>" -- "why undefined?" | 
| 26170 | 118 | |
| 119 | ||
| 120 | subsection {* Imperative references and arrays *}
 | |
| 121 | ||
| 122 | text {*
 | |
| 123 | References and arrays are developed in parallel, | |
| 26586 | 124 | but keeping them separate makes some later proofs simpler. | 
| 26170 | 125 | *} | 
| 126 | ||
| 127 | subsubsection {* Primitive operations *}
 | |
| 128 | ||
| 129 | definition | |
| 130 |   new_ref :: "heap \<Rightarrow> ('a\<Colon>heap) ref \<times> heap" where
 | |
| 131 | "new_ref h = (let l = lim h in (Ref l, h\<lparr>lim := l + 1\<rparr>))" | |
| 132 | ||
| 133 | definition | |
| 134 |   new_array :: "heap \<Rightarrow> ('a\<Colon>heap) array \<times> heap" where
 | |
| 135 | "new_array h = (let l = lim h in (Array l, h\<lparr>lim := l + 1\<rparr>))" | |
| 136 | ||
| 137 | definition | |
| 138 | ref_present :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> bool" where | |
| 139 | "ref_present r h \<longleftrightarrow> addr_of_ref r < lim h" | |
| 140 | ||
| 141 | definition | |
| 142 | array_present :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> bool" where | |
| 143 | "array_present a h \<longleftrightarrow> addr_of_array a < lim h" | |
| 144 | ||
| 145 | definition | |
| 146 | get_ref :: "'a\<Colon>heap ref \<Rightarrow> heap \<Rightarrow> 'a" where | |
| 28335 | 147 |   "get_ref r h = from_nat (refs h (TYPEREP('a)) (addr_of_ref r))"
 | 
| 26170 | 148 | |
| 149 | definition | |
| 150 | get_array :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> 'a list" where | |
| 28335 | 151 |   "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))"
 | 
| 26170 | 152 | |
| 153 | definition | |
| 154 | set_ref :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where | |
| 155 | "set_ref r x = | |
| 28335 | 156 |   refs_update (\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r:=to_nat x))))"
 | 
| 26170 | 157 | |
| 158 | definition | |
| 159 | set_array :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where | |
| 160 | "set_array a x = | |
| 28335 | 161 |   arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))"
 | 
| 26170 | 162 | |
| 163 | subsubsection {* Interface operations *}
 | |
| 164 | ||
| 165 | definition | |
| 166 | ref :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where | |
| 167 | "ref x h = (let (r, h') = new_ref h; | |
| 168 | h'' = set_ref r x h' | |
| 169 | in (r, h''))" | |
| 170 | ||
| 171 | definition | |
| 172 | array :: "nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where | |
| 173 | "array n x h = (let (r, h') = new_array h; | |
| 174 | h'' = set_array r (replicate n x) h' | |
| 175 | in (r, h''))" | |
| 176 | ||
| 177 | definition | |
| 178 | array_of_list :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where | |
| 179 | "array_of_list xs h = (let (r, h') = new_array h; | |
| 180 | h'' = set_array r xs h' | |
| 181 | in (r, h''))" | |
| 182 | ||
| 183 | definition | |
| 184 | upd :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where | |
| 185 | "upd a i x h = set_array a ((get_array a h)[i:=x]) h" | |
| 186 | ||
| 187 | definition | |
| 188 | length :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> nat" where | |
| 189 | "length a h = size (get_array a h)" | |
| 190 | ||
| 191 | definition | |
| 192 |   array_ran :: "('a\<Colon>heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
 | |
| 193 |   "array_ran a h = {e. Some e \<in> set (get_array a h)}"
 | |
| 194 |     -- {*FIXME*}
 | |
| 195 | ||
| 196 | ||
| 197 | subsubsection {* Reference equality *}
 | |
| 198 | ||
| 199 | text {* 
 | |
| 200 | The following relations are useful for comparing arrays and references. | |
| 201 | *} | |
| 202 | ||
| 203 | definition | |
| 204 |   noteq_refs :: "('a\<Colon>heap) ref \<Rightarrow> ('b\<Colon>heap) ref \<Rightarrow> bool" (infix "=!=" 70)
 | |
| 205 | where | |
| 28335 | 206 |   "r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
 | 
| 26170 | 207 | |
| 208 | definition | |
| 209 |   noteq_arrs :: "('a\<Colon>heap) array \<Rightarrow> ('b\<Colon>heap) array \<Rightarrow> bool" (infix "=!!=" 70)
 | |
| 210 | where | |
| 28335 | 211 |   "r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
 | 
| 26170 | 212 | |
| 213 | lemma noteq_refs_sym: "r =!= s \<Longrightarrow> s =!= r" | |
| 214 | and noteq_arrs_sym: "a =!!= b \<Longrightarrow> b =!!= a" | |
| 215 | and unequal_refs [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!" | |
| 216 | and unequal_arrs [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'" | |
| 217 | unfolding noteq_refs_def noteq_arrs_def by auto | |
| 218 | ||
| 36640 
7eadf5acdaf4
added lemma about irreflexivity of pointer inequality in Imperative_HOL
 bulwahn parents: 
36176diff
changeset | 219 | lemma noteq_refs_irrefl: "r =!= r \<Longrightarrow> False" | 
| 
7eadf5acdaf4
added lemma about irreflexivity of pointer inequality in Imperative_HOL
 bulwahn parents: 
36176diff
changeset | 220 | unfolding noteq_refs_def by auto | 
| 
7eadf5acdaf4
added lemma about irreflexivity of pointer inequality in Imperative_HOL
 bulwahn parents: 
36176diff
changeset | 221 | |
| 26170 | 222 | lemma present_new_ref: "ref_present r h \<Longrightarrow> r =!= fst (ref v h)" | 
| 223 | by (simp add: ref_present_def new_ref_def ref_def Let_def noteq_refs_def) | |
| 224 | ||
| 225 | lemma present_new_arr: "array_present a h \<Longrightarrow> a =!!= fst (array v x h)" | |
| 226 | by (simp add: array_present_def noteq_arrs_def new_array_def array_def Let_def) | |
| 227 | ||
| 228 | ||
| 229 | subsubsection {* Properties of heap containers *}
 | |
| 230 | ||
| 231 | text {* Properties of imperative arrays *}
 | |
| 232 | ||
| 233 | text {* FIXME: Does there exist a "canonical" array axiomatisation in
 | |
| 234 | the literature? *} | |
| 235 | ||
| 236 | lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x" | |
| 33639 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 hoelzl parents: 
31205diff
changeset | 237 | by (simp add: get_array_def set_array_def o_def) | 
| 26170 | 238 | |
| 239 | lemma array_get_set_neq [simp]: "r =!!= s \<Longrightarrow> get_array r (set_array s x h) = get_array r h" | |
| 240 | by (simp add: noteq_arrs_def get_array_def set_array_def) | |
| 241 | ||
| 242 | lemma set_array_same [simp]: | |
| 243 | "set_array r x (set_array r y h) = set_array r x h" | |
| 244 | by (simp add: set_array_def) | |
| 245 | ||
| 246 | lemma array_set_set_swap: | |
| 247 | "r =!!= r' \<Longrightarrow> set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)" | |
| 248 | by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def) | |
| 249 | ||
| 250 | lemma array_ref_set_set_swap: | |
| 251 | "set_array r x (set_ref r' x' h) = set_ref r' x' (set_array r x h)" | |
| 252 | by (simp add: Let_def expand_fun_eq set_array_def set_ref_def) | |
| 253 | ||
| 254 | lemma get_array_upd_eq [simp]: | |
| 255 | "get_array a (upd a i v h) = (get_array a h) [i := v]" | |
| 256 | by (simp add: upd_def) | |
| 257 | ||
| 258 | lemma nth_upd_array_neq_array [simp]: | |
| 259 | "a =!!= b \<Longrightarrow> get_array a (upd b j v h) ! i = get_array a h ! i" | |
| 260 | by (simp add: upd_def noteq_arrs_def) | |
| 261 | ||
| 262 | lemma get_arry_array_upd_elem_neqIndex [simp]: | |
| 263 | "i \<noteq> j \<Longrightarrow> get_array a (upd a j v h) ! i = get_array a h ! i" | |
| 264 | by simp | |
| 265 | ||
| 266 | lemma length_upd_eq [simp]: | |
| 267 | "length a (upd a i v h) = length a h" | |
| 268 | by (simp add: length_def upd_def) | |
| 269 | ||
| 270 | lemma length_upd_neq [simp]: | |
| 271 | "length a (upd b i v h) = length a h" | |
| 272 | by (simp add: upd_def length_def set_array_def get_array_def) | |
| 273 | ||
| 274 | lemma upd_swap_neqArray: | |
| 275 | "a =!!= a' \<Longrightarrow> | |
| 276 | upd a i v (upd a' i' v' h) | |
| 277 | = upd a' i' v' (upd a i v h)" | |
| 278 | apply (unfold upd_def) | |
| 279 | apply simp | |
| 280 | apply (subst array_set_set_swap, assumption) | |
| 281 | apply (subst array_get_set_neq) | |
| 282 | apply (erule noteq_arrs_sym) | |
| 283 | apply (simp) | |
| 284 | done | |
| 285 | ||
| 286 | lemma upd_swap_neqIndex: | |
| 287 | "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> upd a i v (upd a i' v' h) = upd a i' v' (upd a i v h)" | |
| 288 | by (auto simp add: upd_def array_set_set_swap list_update_swap) | |
| 289 | ||
| 290 | lemma get_array_init_array_list: | |
| 291 | "get_array (fst (array_of_list ls h)) (snd (array_of_list ls' h)) = ls'" | |
| 292 | by (simp add: Let_def split_def array_of_list_def) | |
| 293 | ||
| 294 | lemma set_array: | |
| 295 | "set_array (fst (array_of_list ls h)) | |
| 296 | new_ls (snd (array_of_list ls h)) | |
| 297 | = snd (array_of_list new_ls h)" | |
| 298 | by (simp add: Let_def split_def array_of_list_def) | |
| 299 | ||
| 300 | lemma array_present_upd [simp]: | |
| 301 | "array_present a (upd b i v h) = array_present a h" | |
| 302 | by (simp add: upd_def array_present_def set_array_def get_array_def) | |
| 303 | ||
| 304 | lemma array_of_list_replicate: | |
| 305 | "array_of_list (replicate n x) = array n x" | |
| 306 | by (simp add: expand_fun_eq array_of_list_def array_def) | |
| 307 | ||
| 308 | ||
| 309 | text {* Properties of imperative references *}
 | |
| 310 | ||
| 311 | lemma next_ref_fresh [simp]: | |
| 312 | assumes "(r, h') = new_ref h" | |
| 313 | shows "\<not> ref_present r h" | |
| 314 | using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def) | |
| 315 | ||
| 316 | lemma next_ref_present [simp]: | |
| 317 | assumes "(r, h') = new_ref h" | |
| 318 | shows "ref_present r h'" | |
| 319 | using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def) | |
| 320 | ||
| 321 | lemma ref_get_set_eq [simp]: "get_ref r (set_ref r x h) = x" | |
| 322 | by (simp add: get_ref_def set_ref_def) | |
| 323 | ||
| 324 | lemma ref_get_set_neq [simp]: "r =!= s \<Longrightarrow> get_ref r (set_ref s x h) = get_ref r h" | |
| 325 | by (simp add: noteq_refs_def get_ref_def set_ref_def) | |
| 326 | ||
| 327 | (* FIXME: We need some infrastructure to infer that locally generated | |
| 328 |   new refs (by new_ref(_no_init), new_array(')) are distinct
 | |
| 329 | from all existing refs. | |
| 330 | *) | |
| 331 | ||
| 332 | lemma ref_set_get: "set_ref r (get_ref r h) h = h" | |
| 333 | apply (simp add: set_ref_def get_ref_def) | |
| 334 | oops | |
| 335 | ||
| 336 | lemma set_ref_same[simp]: | |
| 337 | "set_ref r x (set_ref r y h) = set_ref r x h" | |
| 338 | by (simp add: set_ref_def) | |
| 339 | ||
| 340 | lemma ref_set_set_swap: | |
| 341 | "r =!= r' \<Longrightarrow> set_ref r x (set_ref r' x' h) = set_ref r' x' (set_ref r x h)" | |
| 342 | by (simp add: Let_def expand_fun_eq noteq_refs_def set_ref_def) | |
| 343 | ||
| 344 | lemma ref_new_set: "fst (ref v (set_ref r v' h)) = fst (ref v h)" | |
| 345 | by (simp add: ref_def new_ref_def set_ref_def Let_def) | |
| 346 | ||
| 347 | lemma ref_get_new [simp]: | |
| 348 | "get_ref (fst (ref v h)) (snd (ref v' h)) = v'" | |
| 349 | by (simp add: ref_def Let_def split_def) | |
| 350 | ||
| 351 | lemma ref_set_new [simp]: | |
| 352 | "set_ref (fst (ref v h)) new_v (snd (ref v h)) = snd (ref new_v h)" | |
| 353 | by (simp add: ref_def Let_def split_def) | |
| 354 | ||
| 355 | lemma ref_get_new_neq: "r =!= (fst (ref v h)) \<Longrightarrow> | |
| 356 | get_ref r (snd (ref v h)) = get_ref r h" | |
| 357 | by (simp add: get_ref_def set_ref_def ref_def Let_def new_ref_def noteq_refs_def) | |
| 358 | ||
| 359 | lemma lim_set_ref [simp]: | |
| 360 | "lim (set_ref r v h) = lim h" | |
| 361 | by (simp add: set_ref_def) | |
| 362 | ||
| 363 | lemma ref_present_new_ref [simp]: | |
| 364 | "ref_present r h \<Longrightarrow> ref_present r (snd (ref v h))" | |
| 365 | by (simp add: new_ref_def ref_present_def ref_def Let_def) | |
| 366 | ||
| 367 | lemma ref_present_set_ref [simp]: | |
| 368 | "ref_present r (set_ref r' v h) = ref_present r h" | |
| 369 | by (simp add: set_ref_def ref_present_def) | |
| 370 | ||
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: 
33639diff
changeset | 371 | lemma noteq_refsI: "\<lbrakk> ref_present r h; \<not>ref_present r' h \<rbrakk> \<Longrightarrow> r =!= r'" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: 
33639diff
changeset | 372 | unfolding noteq_refs_def ref_present_def | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: 
33639diff
changeset | 373 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: 
33639diff
changeset | 374 | |
| 26170 | 375 | lemma array_ranI: "\<lbrakk> Some b = get_array a h ! i; i < Heap.length a h \<rbrakk> \<Longrightarrow> b \<in> array_ran a h" | 
| 376 | unfolding array_ran_def Heap.length_def by simp | |
| 377 | ||
| 378 | lemma array_ran_upd_array_Some: | |
| 379 | assumes "cl \<in> array_ran a (Heap.upd a i (Some b) h)" | |
| 380 | shows "cl \<in> array_ran a h \<or> cl = b" | |
| 381 | proof - | |
| 382 | have "set (get_array a h[i := Some b]) \<subseteq> insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert) | |
| 383 | with assms show ?thesis | |
| 384 | unfolding array_ran_def Heap.upd_def by fastsimp | |
| 385 | qed | |
| 386 | ||
| 387 | lemma array_ran_upd_array_None: | |
| 388 | assumes "cl \<in> array_ran a (Heap.upd a i None h)" | |
| 389 | shows "cl \<in> array_ran a h" | |
| 390 | proof - | |
| 391 | have "set (get_array a h[i := None]) \<subseteq> insert None (set (get_array a h))" by (rule set_update_subset_insert) | |
| 392 | with assms show ?thesis | |
| 393 | unfolding array_ran_def Heap.upd_def by auto | |
| 394 | qed | |
| 395 | ||
| 396 | ||
| 397 | text {* Non-interaction between imperative array and imperative references *}
 | |
| 398 | ||
| 399 | lemma get_array_set_ref [simp]: "get_array a (set_ref r v h) = get_array a h" | |
| 400 | by (simp add: get_array_def set_ref_def) | |
| 401 | ||
| 402 | lemma nth_set_ref [simp]: "get_array a (set_ref r v h) ! i = get_array a h ! i" | |
| 403 | by simp | |
| 404 | ||
| 405 | lemma get_ref_upd [simp]: "get_ref r (upd a i v h) = get_ref r h" | |
| 406 | by (simp add: get_ref_def set_array_def upd_def) | |
| 407 | ||
| 408 | lemma new_ref_upd: "fst (ref v (upd a i v' h)) = fst (ref v h)" | |
| 409 | by (simp add: set_array_def get_array_def Let_def ref_new_set upd_def ref_def new_ref_def) | |
| 410 | ||
| 26300 | 411 | text {*not actually true ???*}
 | 
| 26170 | 412 | lemma upd_set_ref_swap: "upd a i v (set_ref r v' h) = set_ref r v' (upd a i v h)" | 
| 413 | apply (case_tac a) | |
| 414 | apply (simp add: Let_def upd_def) | |
| 415 | apply auto | |
| 26300 | 416 | oops | 
| 26170 | 417 | |
| 418 | lemma length_new_ref[simp]: | |
| 419 | "length a (snd (ref v h)) = length a h" | |
| 420 | by (simp add: get_array_def set_ref_def length_def new_ref_def ref_def Let_def) | |
| 421 | ||
| 422 | lemma get_array_new_ref [simp]: | |
| 423 | "get_array a (snd (ref v h)) = get_array a h" | |
| 424 | by (simp add: new_ref_def ref_def set_ref_def get_array_def Let_def) | |
| 425 | ||
| 426 | lemma ref_present_upd [simp]: | |
| 427 | "ref_present r (upd a i v h) = ref_present r h" | |
| 428 | by (simp add: upd_def ref_present_def set_array_def get_array_def) | |
| 429 | ||
| 430 | lemma array_present_set_ref [simp]: | |
| 431 | "array_present a (set_ref r v h) = array_present a h" | |
| 432 | by (simp add: array_present_def set_ref_def) | |
| 433 | ||
| 434 | lemma array_present_new_ref [simp]: | |
| 435 | "array_present a h \<Longrightarrow> array_present a (snd (ref v h))" | |
| 436 | by (simp add: array_present_def new_ref_def ref_def Let_def) | |
| 437 | ||
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
34051diff
changeset | 438 | hide_const (open) empty array array_of_list upd length ref | 
| 26170 | 439 | |
| 440 | end |