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