| author | haftmann |
| Mon, 12 Jul 2010 08:58:12 +0200 | |
| changeset 37764 | 3489daf839d5 |
| parent 37758 | bf86a65403a8 |
| child 37771 | 1bec64044b5e |
| permissions | -rw-r--r-- |
| 26170 | 1 |
(* Title: HOL/Library/Ref.thy |
2 |
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen |
|
3 |
*) |
|
4 |
||
5 |
header {* Monadic references *}
|
|
6 |
||
7 |
theory Ref |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
8 |
imports Array |
| 26170 | 9 |
begin |
10 |
||
11 |
text {*
|
|
12 |
Imperative reference operations; modeled after their ML counterparts. |
|
13 |
See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html |
|
14 |
and http://www.smlnj.org/doc/Conversion/top-level-comparison.html |
|
15 |
*} |
|
16 |
||
| 37753 | 17 |
subsection {* Primitives *}
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
18 |
|
| 37725 | 19 |
definition present :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> bool" where |
20 |
"present h r \<longleftrightarrow> addr_of_ref r < lim h" |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
21 |
|
| 37725 | 22 |
definition get :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> 'a" where |
23 |
"get h = from_nat \<circ> refs h TYPEREP('a) \<circ> addr_of_ref"
|
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
24 |
|
| 37725 | 25 |
definition set :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where |
26 |
"set r x = refs_update |
|
27 |
(\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))"
|
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
28 |
|
| 37725 | 29 |
definition alloc :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where |
30 |
"alloc x h = (let |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
31 |
l = lim h; |
| 37725 | 32 |
r = Ref l |
33 |
in (r, set r x (h\<lparr>lim := l + 1\<rparr>)))" |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
34 |
|
| 37725 | 35 |
definition noteq :: "'a\<Colon>heap ref \<Rightarrow> 'b\<Colon>heap ref \<Rightarrow> bool" (infix "=!=" 70) where |
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
36 |
"r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
|
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
37 |
|
| 37753 | 38 |
|
39 |
subsection {* Monad operations *}
|
|
40 |
||
41 |
definition ref :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where |
|
42 |
[code del]: "ref v = Heap_Monad.heap (alloc v)" |
|
43 |
||
44 |
definition lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
|
|
| 37758 | 45 |
[code del]: "lookup r = Heap_Monad.tap (\<lambda>h. get h r)" |
| 37753 | 46 |
|
47 |
definition update :: "'a ref \<Rightarrow> 'a\<Colon>heap \<Rightarrow> unit Heap" ("_ := _" 62) where
|
|
48 |
[code del]: "update r v = Heap_Monad.heap (\<lambda>h. ((), set r v h))" |
|
49 |
||
50 |
definition change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
|
|
51 |
"change f r = (do |
|
52 |
x \<leftarrow> ! r; |
|
53 |
let y = f x; |
|
54 |
r := y; |
|
55 |
return y |
|
56 |
done)" |
|
57 |
||
58 |
||
59 |
subsection {* Properties *}
|
|
60 |
||
| 37758 | 61 |
text {* Primitives *}
|
62 |
||
| 37725 | 63 |
lemma noteq_sym: "r =!= s \<Longrightarrow> s =!= r" |
64 |
and unequal [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!" |
|
65 |
by (auto simp add: noteq_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
66 |
|
| 37725 | 67 |
lemma noteq_irrefl: "r =!= r \<Longrightarrow> False" |
68 |
by (auto simp add: noteq_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
69 |
|
| 37725 | 70 |
lemma present_alloc_neq: "present h r \<Longrightarrow> r =!= fst (alloc v h)" |
71 |
by (simp add: present_def alloc_def noteq_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
72 |
|
| 37725 | 73 |
lemma next_fresh [simp]: |
74 |
assumes "(r, h') = alloc x h" |
|
75 |
shows "\<not> present h r" |
|
76 |
using assms by (cases h) (auto simp add: alloc_def present_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
77 |
|
| 37725 | 78 |
lemma next_present [simp]: |
79 |
assumes "(r, h') = alloc x h" |
|
80 |
shows "present h' r" |
|
81 |
using assms by (cases h) (auto simp add: alloc_def set_def present_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
82 |
|
| 37725 | 83 |
lemma get_set_eq [simp]: |
84 |
"get (set r x h) r = x" |
|
85 |
by (simp add: get_def set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
86 |
|
| 37725 | 87 |
lemma get_set_neq [simp]: |
88 |
"r =!= s \<Longrightarrow> get (set s x h) r = get h r" |
|
89 |
by (simp add: noteq_def get_def set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
90 |
|
| 37725 | 91 |
lemma set_same [simp]: |
92 |
"set r x (set r y h) = set r x h" |
|
93 |
by (simp add: set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
94 |
|
| 37725 | 95 |
lemma set_set_swap: |
96 |
"r =!= r' \<Longrightarrow> set r x (set r' x' h) = set r' x' (set r x h)" |
|
97 |
by (simp add: noteq_def set_def expand_fun_eq) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
98 |
|
| 37725 | 99 |
lemma alloc_set: |
100 |
"fst (alloc x (set r x' h)) = fst (alloc x h)" |
|
101 |
by (simp add: alloc_def set_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
102 |
|
| 37725 | 103 |
lemma get_alloc [simp]: |
104 |
"get (snd (alloc x h)) (fst (alloc x' h)) = x" |
|
105 |
by (simp add: alloc_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
106 |
|
| 37725 | 107 |
lemma set_alloc [simp]: |
108 |
"set (fst (alloc v h)) v' (snd (alloc v h)) = snd (alloc v' h)" |
|
109 |
by (simp add: alloc_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
110 |
|
| 37725 | 111 |
lemma get_alloc_neq: "r =!= fst (alloc v h) \<Longrightarrow> |
112 |
get (snd (alloc v h)) r = get h r" |
|
113 |
by (simp add: get_def set_def alloc_def Let_def noteq_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
114 |
|
| 37725 | 115 |
lemma lim_set [simp]: |
116 |
"lim (set r v h) = lim h" |
|
117 |
by (simp add: set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
118 |
|
| 37725 | 119 |
lemma present_alloc [simp]: |
120 |
"present h r \<Longrightarrow> present (snd (alloc v h)) r" |
|
121 |
by (simp add: present_def alloc_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
122 |
|
| 37725 | 123 |
lemma present_set [simp]: |
124 |
"present (set r v h) = present h" |
|
125 |
by (simp add: present_def expand_fun_eq) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
126 |
|
| 37725 | 127 |
lemma noteq_I: |
128 |
"present h r \<Longrightarrow> \<not> present h r' \<Longrightarrow> r =!= r'" |
|
129 |
by (auto simp add: noteq_def present_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
130 |
|
| 37758 | 131 |
|
132 |
text {* Monad operations *}
|
|
133 |
||
134 |
lemma execute_ref [simp, execute_simps]: |
|
135 |
"execute (ref v) h = Some (alloc v h)" |
|
| 37753 | 136 |
by (simp add: ref_def) |
| 26170 | 137 |
|
| 37758 | 138 |
lemma success_refI [iff, success_intros]: |
139 |
"success (ref v) h" |
|
140 |
by (auto simp add: ref_def) |
|
141 |
||
142 |
lemma execute_lookup [simp, execute_simps]: |
|
| 37753 | 143 |
"Heap_Monad.execute (lookup r) h = Some (get h r, h)" |
144 |
by (simp add: lookup_def) |
|
| 26182 | 145 |
|
| 37758 | 146 |
lemma success_lookupI [iff, success_intros]: |
147 |
"success (lookup r) h" |
|
148 |
by (auto simp add: lookup_def) |
|
149 |
||
150 |
lemma execute_update [simp, execute_simps]: |
|
| 37753 | 151 |
"Heap_Monad.execute (update r v) h = Some ((), set r v h)" |
152 |
by (simp add: update_def) |
|
| 26170 | 153 |
|
| 37758 | 154 |
lemma success_updateI [iff, success_intros]: |
155 |
"success (update r v) h" |
|
156 |
by (auto simp add: update_def) |
|
157 |
||
158 |
lemma execute_change [simp, execute_simps]: |
|
| 37753 | 159 |
"Heap_Monad.execute (change f r) h = Some (f (get h r), set r (f (get h r)) h)" |
| 37758 | 160 |
by (simp add: change_def bind_def Let_def) |
161 |
||
162 |
lemma success_changeI [iff, success_intros]: |
|
163 |
"success (change f r) h" |
|
164 |
by (auto intro!: success_intros simp add: change_def) |
|
| 26170 | 165 |
|
166 |
lemma lookup_chain: |
|
167 |
"(!r \<guillemotright> f) = f" |
|
| 37758 | 168 |
by (rule Heap_eqI) (auto simp add: lookup_def execute_simps) |
| 26170 | 169 |
|
| 28562 | 170 |
lemma update_change [code]: |
| 37725 | 171 |
"r := e = change (\<lambda>_. e) r \<guillemotright> return ()" |
172 |
by (rule Heap_eqI) (simp add: change_def lookup_chain) |
|
| 26170 | 173 |
|
| 26182 | 174 |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
175 |
text {* Non-interaction between imperative array and imperative references *}
|
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
176 |
|
| 37725 | 177 |
lemma get_array_set [simp]: |
178 |
"get_array a (set r v h) = get_array a h" |
|
179 |
by (simp add: get_array_def set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
180 |
|
| 37725 | 181 |
lemma nth_set [simp]: |
182 |
"get_array a (set r v h) ! i = get_array a h ! i" |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
183 |
by simp |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
184 |
|
| 37725 | 185 |
lemma get_change [simp]: |
186 |
"get (Array.change a i v h) r = get h r" |
|
187 |
by (simp add: get_def Array.change_def set_array_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
188 |
|
| 37725 | 189 |
lemma alloc_change: |
190 |
"fst (alloc v (Array.change a i v' h)) = fst (alloc v h)" |
|
191 |
by (simp add: Array.change_def get_array_def set_array_def alloc_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
192 |
|
| 37725 | 193 |
lemma change_set_swap: |
194 |
"Array.change a i v (set r v' h) = set r v' (Array.change a i v h)" |
|
195 |
by (simp add: Array.change_def get_array_def set_array_def set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
196 |
|
| 37725 | 197 |
lemma length_alloc [simp]: |
198 |
"Array.length a (snd (alloc v h)) = Array.length a h" |
|
199 |
by (simp add: Array.length_def get_array_def alloc_def set_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
200 |
|
| 37725 | 201 |
lemma get_array_alloc [simp]: |
202 |
"get_array a (snd (alloc v h)) = get_array a h" |
|
203 |
by (simp add: get_array_def alloc_def set_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
204 |
|
| 37725 | 205 |
lemma present_change [simp]: |
206 |
"present (Array.change a i v h) = present h" |
|
207 |
by (simp add: Array.change_def set_array_def expand_fun_eq present_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
208 |
|
| 37725 | 209 |
lemma array_present_set [simp]: |
210 |
"array_present a (set r v h) = array_present a h" |
|
211 |
by (simp add: array_present_def set_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
212 |
|
| 37725 | 213 |
lemma array_present_alloc [simp]: |
214 |
"array_present a h \<Longrightarrow> array_present a (snd (alloc v h))" |
|
215 |
by (simp add: array_present_def alloc_def Let_def) |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
216 |
|
| 37725 | 217 |
lemma set_array_set_swap: |
218 |
"set_array a xs (set r x' h) = set r x' (set_array a xs h)" |
|
219 |
by (simp add: set_array_def set_def) |
|
220 |
||
221 |
hide_const (open) present get set alloc lookup update change |
|
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
222 |
|
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37709
diff
changeset
|
223 |
|
| 26182 | 224 |
subsection {* Code generator setup *}
|
225 |
||
| 37753 | 226 |
text {* SML *}
|
| 26182 | 227 |
|
|
34051
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
bulwahn
parents:
29793
diff
changeset
|
228 |
code_type ref (SML "_/ Unsynchronized.ref") |
| 26182 | 229 |
code_const Ref (SML "raise/ (Fail/ \"bare Ref\")") |
| 37725 | 230 |
code_const ref (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)") |
| 26752 | 231 |
code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)") |
232 |
code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)") |
|
| 26182 | 233 |
|
234 |
code_reserved SML ref |
|
235 |
||
236 |
||
| 37753 | 237 |
text {* OCaml *}
|
| 26182 | 238 |
|
239 |
code_type ref (OCaml "_/ ref") |
|
| 26752 | 240 |
code_const Ref (OCaml "failwith/ \"bare Ref\")") |
| 37725 | 241 |
code_const ref (OCaml "(fn/ ()/ =>/ ref/ _)") |
| 26752 | 242 |
code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)") |
243 |
code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)") |
|
| 26182 | 244 |
|
245 |
code_reserved OCaml ref |
|
246 |
||
247 |
||
| 37753 | 248 |
text {* Haskell *}
|
| 26182 | 249 |
|
| 29793 | 250 |
code_type ref (Haskell "Heap.STRef/ Heap.RealWorld/ _") |
| 26182 | 251 |
code_const Ref (Haskell "error/ \"bare Ref\"") |
| 37725 | 252 |
code_const ref (Haskell "Heap.newSTRef") |
| 29793 | 253 |
code_const Ref.lookup (Haskell "Heap.readSTRef") |
254 |
code_const Ref.update (Haskell "Heap.writeSTRef") |
|
| 26182 | 255 |
|
| 37758 | 256 |
end |