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