author | blanchet |
Tue, 22 Mar 2011 19:04:32 +0100 | |
changeset 42064 | f4e53c8630c0 |
parent 41413 | 64cd30d6b0b8 |
child 42463 | f270e3e18be5 |
permissions | -rw-r--r-- |
37714 | 1 |
(* Title: HOL/Imperative_HOL/Heap.thy |
26170 | 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 |
|
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
39250
diff
changeset
|
8 |
imports Main "~~/src/HOL/Library/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 |
|
38408 | 17 |
instance unit :: heap .. |
18 |
||
19 |
instance bool :: heap .. |
|
20 |
||
26170 | 21 |
instance nat :: heap .. |
22 |
||
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
36640
diff
changeset
|
23 |
instance prod :: (heap, heap) heap .. |
26170 | 24 |
|
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
36640
diff
changeset
|
25 |
instance sum :: (heap, heap) heap .. |
26170 | 26 |
|
27 |
instance list :: (heap) heap .. |
|
28 |
||
29 |
instance option :: (heap) heap .. |
|
30 |
||
31 |
instance int :: heap .. |
|
32 |
||
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
30738
diff
changeset
|
33 |
instance String.literal :: heap .. |
26170 | 34 |
|
28335 | 35 |
instance typerep :: heap .. |
26170 | 36 |
|
37 |
||
38 |
subsection {* A polymorphic heap with dynamic arrays and references *} |
|
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 | 44 |
|
26170 | 45 |
types addr = nat -- "untyped heap references" |
37714 | 46 |
types heap_rep = nat -- "representable values" |
47 |
||
48 |
record heap = |
|
49 |
arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list" |
|
50 |
refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep" |
|
51 |
lim :: addr |
|
52 |
||
53 |
definition empty :: heap where |
|
37723 | 54 |
"empty = \<lparr>arrays = (\<lambda>_ _. []), refs = (\<lambda>_ _. 0), lim = 0\<rparr>" |
26170 | 55 |
|
56 |
datatype 'a array = Array addr |
|
57 |
datatype 'a ref = Ref addr -- "note the phantom type 'a " |
|
58 |
||
59 |
primrec addr_of_array :: "'a array \<Rightarrow> addr" where |
|
60 |
"addr_of_array (Array x) = x" |
|
61 |
||
62 |
primrec addr_of_ref :: "'a ref \<Rightarrow> addr" where |
|
63 |
"addr_of_ref (Ref x) = x" |
|
64 |
||
65 |
lemma addr_of_array_inj [simp]: |
|
66 |
"addr_of_array a = addr_of_array a' \<longleftrightarrow> a = a'" |
|
67 |
by (cases a, cases a') simp_all |
|
68 |
||
69 |
lemma addr_of_ref_inj [simp]: |
|
70 |
"addr_of_ref r = addr_of_ref r' \<longleftrightarrow> r = r'" |
|
71 |
by (cases r, cases r') simp_all |
|
72 |
||
73 |
instance array :: (type) countable |
|
74 |
by (rule countable_classI [of addr_of_array]) simp |
|
75 |
||
76 |
instance ref :: (type) countable |
|
77 |
by (rule countable_classI [of addr_of_ref]) simp |
|
78 |
||
37714 | 79 |
text {* Syntactic convenience *} |
80 |
||
26170 | 81 |
setup {* |
82 |
Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap array"}) |
|
83 |
#> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap ref"}) |
|
84 |
#> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\<Colon>heap array \<Rightarrow> nat"}) |
|
85 |
#> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\<Colon>heap ref \<Rightarrow> nat"}) |
|
86 |
*} |
|
87 |
||
37723 | 88 |
hide_const (open) empty |
89 |
||
26170 | 90 |
end |