src/HOL/Imperative_HOL/Heap.thy
author wenzelm
Tue, 03 Sep 2013 01:12:40 +0200
changeset 53374 a14d2a854c02
parent 42463 f270e3e18be5
child 58249 180f1b3508ed
permissions -rw-r--r--
tuned proofs -- clarified flow of facts wrt. calculation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37714
haftmann
parents: 37678
diff changeset
     1
(*  Title:      HOL/Imperative_HOL/Heap.thy
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     2
    Author:     John Matthews, Galois Connections; Alexander Krauss, TU Muenchen
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     3
*)
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     4
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     5
header {* A polymorphic heap based on cantor encodings *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     6
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     7
theory Heap
41413
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 39250
diff changeset
     8
imports Main "~~/src/HOL/Library/Countable"
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
     9
begin
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    10
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    11
subsection {* Representable types *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    12
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    13
text {* The type class of representable types *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    14
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28042
diff changeset
    15
class heap = typerep + countable
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    16
38408
721b4d6095b7 unit and bool are instances of heap
haftmann
parents: 37723
diff changeset
    17
instance unit :: heap ..
721b4d6095b7 unit and bool are instances of heap
haftmann
parents: 37723
diff changeset
    18
721b4d6095b7 unit and bool are instances of heap
haftmann
parents: 37723
diff changeset
    19
instance bool :: heap ..
721b4d6095b7 unit and bool are instances of heap
haftmann
parents: 37723
diff changeset
    20
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    21
instance nat :: heap ..
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    22
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 36640
diff changeset
    23
instance prod :: (heap, heap) heap ..
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    24
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 36640
diff changeset
    25
instance sum :: (heap, heap) heap ..
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    26
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    27
instance list :: (heap) heap ..
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    28
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    29
instance option :: (heap) heap ..
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    30
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    31
instance int :: heap ..
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    32
31205
98370b26c2ce String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents: 30738
diff changeset
    33
instance String.literal :: heap ..
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    34
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28042
diff changeset
    35
instance typerep :: heap ..
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    36
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    37
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    38
subsection {* A polymorphic heap with dynamic arrays and references *}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    39
37719
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    40
text {*
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    41
  References and arrays are developed in parallel,
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    42
  but keeping them separate makes some later proofs simpler.
271ecd4fb9f9 moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents: 37718
diff changeset
    43
*}
37714
haftmann
parents: 37678
diff changeset
    44
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 41413
diff changeset
    45
type_synonym addr = nat -- "untyped heap references"
f270e3e18be5 modernized specifications;
wenzelm
parents: 41413
diff changeset
    46
type_synonym heap_rep = nat -- "representable values"
37714
haftmann
parents: 37678
diff changeset
    47
haftmann
parents: 37678
diff changeset
    48
record heap =
haftmann
parents: 37678
diff changeset
    49
  arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list"
haftmann
parents: 37678
diff changeset
    50
  refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep"
haftmann
parents: 37678
diff changeset
    51
  lim  :: addr
haftmann
parents: 37678
diff changeset
    52
haftmann
parents: 37678
diff changeset
    53
definition empty :: heap where
37723
831b3eb7ed8e tuned empty heap
haftmann
parents: 37719
diff changeset
    54
  "empty = \<lparr>arrays = (\<lambda>_ _. []), refs = (\<lambda>_ _. 0), lim = 0\<rparr>"
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    55
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    56
datatype 'a array = Array addr
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    57
datatype 'a ref = Ref addr -- "note the phantom type 'a "
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    58
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    59
primrec addr_of_array :: "'a array \<Rightarrow> addr" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    60
  "addr_of_array (Array x) = x"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    61
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    62
primrec addr_of_ref :: "'a ref \<Rightarrow> addr" where
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    63
  "addr_of_ref (Ref x) = x"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    64
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    65
lemma addr_of_array_inj [simp]:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    66
  "addr_of_array a = addr_of_array a' \<longleftrightarrow> a = a'"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    67
  by (cases a, cases a') simp_all
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    68
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    69
lemma addr_of_ref_inj [simp]:
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    70
  "addr_of_ref r = addr_of_ref r' \<longleftrightarrow> r = r'"
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    71
  by (cases r, cases r') simp_all
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    72
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    73
instance array :: (type) countable
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    74
  by (rule countable_classI [of addr_of_array]) simp
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    75
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    76
instance ref :: (type) countable
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    77
  by (rule countable_classI [of addr_of_ref]) simp
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    78
37714
haftmann
parents: 37678
diff changeset
    79
text {* Syntactic convenience *}
haftmann
parents: 37678
diff changeset
    80
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    81
setup {*
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    82
  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
    83
  #> 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
    84
  #> 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
    85
  #> 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
    86
*}
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    87
37723
831b3eb7ed8e tuned empty heap
haftmann
parents: 37719
diff changeset
    88
hide_const (open) empty
831b3eb7ed8e tuned empty heap
haftmann
parents: 37719
diff changeset
    89
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents:
diff changeset
    90
end