| author | wenzelm | 
| Sat, 09 Mar 2024 19:57:08 +0100 | |
| changeset 79836 | c69ae2b8987e | 
| parent 68224 | 1f7308050349 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
1  | 
(* Title: HOL/Imperative_HOL/Ref.thy  | 
| 26170 | 2  | 
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen  | 
3  | 
*)  | 
|
4  | 
||
| 63167 | 5  | 
section \<open>Monadic references\<close>  | 
| 26170 | 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  | 
||
| 63167 | 11  | 
text \<open>  | 
| 26170 | 12  | 
Imperative reference operations; modeled after their ML counterparts.  | 
| 68224 | 13  | 
See \<^url>\<open>https://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html\<close>  | 
14  | 
and \<^url>\<open>https://www.smlnj.org/doc/Conversion/top-level-comparison.html\<close>.  | 
|
| 63167 | 15  | 
\<close>  | 
| 26170 | 16  | 
|
| 63167 | 17  | 
subsection \<open>Primitives\<close>  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
18  | 
|
| 61076 | 19  | 
definition present :: "heap \<Rightarrow> 'a::heap ref \<Rightarrow> bool" where  | 
| 37725 | 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  | 
|
| 61076 | 22  | 
definition get :: "heap \<Rightarrow> 'a::heap ref \<Rightarrow> 'a" where  | 
| 37725 | 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  | 
|
| 61076 | 25  | 
definition set :: "'a::heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where  | 
| 37725 | 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  | 
|
| 61076 | 29  | 
definition alloc :: "'a \<Rightarrow> heap \<Rightarrow> 'a::heap ref \<times> heap" where  | 
| 37725 | 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  | 
|
| 61076 | 35  | 
definition noteq :: "'a::heap ref \<Rightarrow> 'b::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  | 
|
| 63167 | 39  | 
subsection \<open>Monad operations\<close>  | 
| 37753 | 40  | 
|
| 61076 | 41  | 
definition ref :: "'a::heap \<Rightarrow> 'a ref Heap" where  | 
| 37753 | 42  | 
[code del]: "ref v = Heap_Monad.heap (alloc v)"  | 
43  | 
||
| 61076 | 44  | 
definition lookup :: "'a::heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
 | 
| 37758 | 45  | 
[code del]: "lookup r = Heap_Monad.tap (\<lambda>h. get h r)"  | 
| 37753 | 46  | 
|
| 61076 | 47  | 
definition update :: "'a ref \<Rightarrow> 'a::heap \<Rightarrow> unit Heap" ("_ := _" 62) where
 | 
| 37753 | 48  | 
[code del]: "update r v = Heap_Monad.heap (\<lambda>h. ((), set r v h))"  | 
49  | 
||
| 61076 | 50  | 
definition change :: "('a::heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
 | 
| 37792 | 51  | 
  "change f r = do {
 | 
| 37753 | 52  | 
x \<leftarrow> ! r;  | 
53  | 
let y = f x;  | 
|
54  | 
r := y;  | 
|
55  | 
return y  | 
|
| 37792 | 56  | 
}"  | 
| 37753 | 57  | 
|
58  | 
||
| 63167 | 59  | 
subsection \<open>Properties\<close>  | 
| 37753 | 60  | 
|
| 63167 | 61  | 
text \<open>Primitives\<close>  | 
| 37758 | 62  | 
|
| 37725 | 63  | 
lemma noteq_sym: "r =!= s \<Longrightarrow> s =!= r"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63680 
diff
changeset
 | 
64  | 
and unequal [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" \<comment> \<open>same types!\<close>  | 
| 37725 | 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  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
95  | 
lemma not_present_alloc [simp]:  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
96  | 
"\<not> present h (fst (alloc v h))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
97  | 
by (simp add: present_def alloc_def Let_def)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
98  | 
|
| 37725 | 99  | 
lemma set_set_swap:  | 
100  | 
"r =!= r' \<Longrightarrow> set r x (set r' x' h) = set r' x' (set r x h)"  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
101  | 
by (simp add: noteq_def set_def fun_eq_iff)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
102  | 
|
| 37725 | 103  | 
lemma alloc_set:  | 
104  | 
"fst (alloc x (set r x' h)) = fst (alloc x h)"  | 
|
105  | 
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
 | 
106  | 
|
| 37725 | 107  | 
lemma get_alloc [simp]:  | 
108  | 
"get (snd (alloc x h)) (fst (alloc x' h)) = x"  | 
|
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 set_alloc [simp]:  | 
112  | 
"set (fst (alloc v h)) v' (snd (alloc v h)) = snd (alloc v' h)"  | 
|
113  | 
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
 | 
114  | 
|
| 37725 | 115  | 
lemma get_alloc_neq: "r =!= fst (alloc v h) \<Longrightarrow>  | 
116  | 
get (snd (alloc v h)) r = get h r"  | 
|
117  | 
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
 | 
118  | 
|
| 37725 | 119  | 
lemma lim_set [simp]:  | 
120  | 
"lim (set r v h) = lim h"  | 
|
121  | 
by (simp add: set_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_alloc [simp]:  | 
124  | 
"present h r \<Longrightarrow> present (snd (alloc v h)) r"  | 
|
125  | 
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
 | 
126  | 
|
| 37725 | 127  | 
lemma present_set [simp]:  | 
128  | 
"present (set r v h) = present h"  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
129  | 
by (simp add: present_def fun_eq_iff)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
130  | 
|
| 37725 | 131  | 
lemma noteq_I:  | 
132  | 
"present h r \<Longrightarrow> \<not> present h r' \<Longrightarrow> r =!= r'"  | 
|
133  | 
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
 | 
134  | 
|
| 37758 | 135  | 
|
| 63167 | 136  | 
text \<open>Monad operations\<close>  | 
| 37758 | 137  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
138  | 
lemma execute_ref [execute_simps]:  | 
| 37758 | 139  | 
"execute (ref v) h = Some (alloc v h)"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
140  | 
by (simp add: ref_def execute_simps)  | 
| 26170 | 141  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
142  | 
lemma success_refI [success_intros]:  | 
| 37758 | 143  | 
"success (ref v) h"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
144  | 
by (auto intro: success_intros simp add: ref_def)  | 
| 37758 | 145  | 
|
| 40671 | 146  | 
lemma effect_refI [effect_intros]:  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
147  | 
assumes "(r, h') = alloc v h"  | 
| 40671 | 148  | 
shows "effect (ref v) h h' r"  | 
149  | 
by (rule effectI) (insert assms, simp add: execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
150  | 
|
| 40671 | 151  | 
lemma effect_refE [effect_elims]:  | 
152  | 
assumes "effect (ref v) h h' r"  | 
|
| 37796 | 153  | 
obtains "get h' r = v" and "present h' r" and "\<not> present h r"  | 
| 40671 | 154  | 
using assms by (rule effectE) (simp add: execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
155  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
156  | 
lemma execute_lookup [execute_simps]:  | 
| 37753 | 157  | 
"Heap_Monad.execute (lookup r) h = Some (get h r, h)"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
158  | 
by (simp add: lookup_def execute_simps)  | 
| 26182 | 159  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
160  | 
lemma success_lookupI [success_intros]:  | 
| 37758 | 161  | 
"success (lookup r) h"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
162  | 
by (auto intro: success_intros simp add: lookup_def)  | 
| 37758 | 163  | 
|
| 40671 | 164  | 
lemma effect_lookupI [effect_intros]:  | 
| 37796 | 165  | 
assumes "h' = h" "x = get h r"  | 
| 40671 | 166  | 
shows "effect (!r) h h' x"  | 
167  | 
by (rule effectI) (insert assms, simp add: execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
168  | 
|
| 40671 | 169  | 
lemma effect_lookupE [effect_elims]:  | 
170  | 
assumes "effect (!r) h h' x"  | 
|
| 37796 | 171  | 
obtains "h' = h" "x = get h r"  | 
| 40671 | 172  | 
using assms by (rule effectE) (simp add: execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
173  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
174  | 
lemma execute_update [execute_simps]:  | 
| 37753 | 175  | 
"Heap_Monad.execute (update r v) h = Some ((), set r v h)"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
176  | 
by (simp add: update_def execute_simps)  | 
| 26170 | 177  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
178  | 
lemma success_updateI [success_intros]:  | 
| 37758 | 179  | 
"success (update r v) h"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
180  | 
by (auto intro: success_intros simp add: update_def)  | 
| 37758 | 181  | 
|
| 40671 | 182  | 
lemma effect_updateI [effect_intros]:  | 
| 37796 | 183  | 
assumes "h' = set r v h"  | 
| 40671 | 184  | 
shows "effect (r := v) h h' x"  | 
185  | 
by (rule effectI) (insert assms, simp add: execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
186  | 
|
| 40671 | 187  | 
lemma effect_updateE [effect_elims]:  | 
188  | 
assumes "effect (r' := v) h h' r"  | 
|
| 37796 | 189  | 
obtains "h' = set r' v h"  | 
| 40671 | 190  | 
using assms by (rule effectE) (simp add: execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
191  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
192  | 
lemma execute_change [execute_simps]:  | 
| 37753 | 193  | 
"Heap_Monad.execute (change f r) h = Some (f (get h r), set r (f (get h r)) h)"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
194  | 
by (simp add: change_def bind_def Let_def execute_simps)  | 
| 37758 | 195  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37771 
diff
changeset
 | 
196  | 
lemma success_changeI [success_intros]:  | 
| 37758 | 197  | 
"success (change f r) h"  | 
| 40671 | 198  | 
by (auto intro!: success_intros effect_intros simp add: change_def)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
199  | 
|
| 40671 | 200  | 
lemma effect_changeI [effect_intros]:  | 
| 37796 | 201  | 
assumes "h' = set r (f (get h r)) h" "x = f (get h r)"  | 
| 40671 | 202  | 
shows "effect (change f r) h h' x"  | 
203  | 
by (rule effectI) (insert assms, simp add: execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
204  | 
|
| 40671 | 205  | 
lemma effect_changeE [effect_elims]:  | 
206  | 
assumes "effect (change f r') h h' r"  | 
|
| 37796 | 207  | 
obtains "h' = set r' (f (get h r')) h" "r = f (get h r')"  | 
| 40671 | 208  | 
using assms by (rule effectE) (simp add: execute_simps)  | 
| 26170 | 209  | 
|
210  | 
lemma lookup_chain:  | 
|
| 62026 | 211  | 
"(!r \<then> f) = f"  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
212  | 
by (rule Heap_eqI) (auto simp add: lookup_def execute_simps intro: execute_bind)  | 
| 26170 | 213  | 
|
| 28562 | 214  | 
lemma update_change [code]:  | 
| 62026 | 215  | 
"r := e = change (\<lambda>_. e) r \<then> return ()"  | 
| 37725 | 216  | 
by (rule Heap_eqI) (simp add: change_def lookup_chain)  | 
| 26170 | 217  | 
|
| 26182 | 218  | 
|
| 63167 | 219  | 
text \<open>Non-interaction between imperative arrays and imperative references\<close>  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
220  | 
|
| 
37806
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
221  | 
lemma array_get_set [simp]:  | 
| 
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
222  | 
"Array.get (set r v h) = Array.get h"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
223  | 
by (simp add: Array.get_def set_def fun_eq_iff)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
224  | 
|
| 37796 | 225  | 
lemma get_update [simp]:  | 
| 37805 | 226  | 
"get (Array.update a i v h) r = get h r"  | 
| 
37804
 
0145e59c1f6c
qualified names for (almost) all array operations
 
haftmann 
parents: 
37803 
diff
changeset
 | 
227  | 
by (simp add: get_def Array.update_def Array.set_def)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
228  | 
|
| 37796 | 229  | 
lemma alloc_update:  | 
230  | 
"fst (alloc v (Array.update a i v' h)) = fst (alloc v h)"  | 
|
| 
37806
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
231  | 
by (simp add: Array.update_def Array.get_def Array.set_def alloc_def Let_def)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
232  | 
|
| 37796 | 233  | 
lemma update_set_swap:  | 
234  | 
"Array.update a i v (set r v' h) = set r v' (Array.update a i v h)"  | 
|
| 
37806
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
235  | 
by (simp add: Array.update_def Array.get_def Array.set_def set_def)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
236  | 
|
| 37725 | 237  | 
lemma length_alloc [simp]:  | 
| 37802 | 238  | 
"Array.length (snd (alloc v h)) a = Array.length h a"  | 
| 
37806
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
239  | 
by (simp add: Array.length_def Array.get_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
 | 
240  | 
|
| 
37806
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
241  | 
lemma array_get_alloc [simp]:  | 
| 
 
a7679be14442
qualified names for (really) all array operations
 
haftmann 
parents: 
37805 
diff
changeset
 | 
242  | 
"Array.get (snd (alloc v h)) = Array.get h"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
243  | 
by (simp add: Array.get_def alloc_def set_def Let_def fun_eq_iff)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
244  | 
|
| 37796 | 245  | 
lemma present_update [simp]:  | 
246  | 
"present (Array.update a i v h) = present h"  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
247  | 
by (simp add: Array.update_def Array.set_def fun_eq_iff present_def)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
248  | 
|
| 37725 | 249  | 
lemma array_present_set [simp]:  | 
| 
37804
 
0145e59c1f6c
qualified names for (almost) all array operations
 
haftmann 
parents: 
37803 
diff
changeset
 | 
250  | 
"Array.present (set r v h) = Array.present h"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
251  | 
by (simp add: Array.present_def set_def fun_eq_iff)  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
252  | 
|
| 37725 | 253  | 
lemma array_present_alloc [simp]:  | 
| 
37804
 
0145e59c1f6c
qualified names for (almost) all array operations
 
haftmann 
parents: 
37803 
diff
changeset
 | 
254  | 
"Array.present h a \<Longrightarrow> Array.present (snd (alloc v h)) a"  | 
| 
 
0145e59c1f6c
qualified names for (almost) all array operations
 
haftmann 
parents: 
37803 
diff
changeset
 | 
255  | 
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
 | 
256  | 
|
| 37725 | 257  | 
lemma set_array_set_swap:  | 
| 
37804
 
0145e59c1f6c
qualified names for (almost) all array operations
 
haftmann 
parents: 
37803 
diff
changeset
 | 
258  | 
"Array.set a xs (set r x' h) = set r x' (Array.set a xs h)"  | 
| 
 
0145e59c1f6c
qualified names for (almost) all array operations
 
haftmann 
parents: 
37803 
diff
changeset
 | 
259  | 
by (simp add: Array.set_def set_def)  | 
| 37725 | 260  | 
|
| 37796 | 261  | 
hide_const (open) present get set alloc noteq lookup update change  | 
| 
37719
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
262  | 
|
| 
 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 
haftmann 
parents: 
37709 
diff
changeset
 | 
263  | 
|
| 63167 | 264  | 
subsection \<open>Code generator setup\<close>  | 
| 26182 | 265  | 
|
| 63167 | 266  | 
text \<open>Intermediate operation avoids invariance problem in \<open>Scala\<close> (similar to value restriction)\<close>  | 
| 
38068
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
267  | 
|
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
268  | 
definition ref' where  | 
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
269  | 
[code del]: "ref' = ref"  | 
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
270  | 
|
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
271  | 
lemma [code]:  | 
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
272  | 
"ref x = ref' x"  | 
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
273  | 
by (simp add: ref'_def)  | 
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
274  | 
|
| 
 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 
haftmann 
parents: 
37965 
diff
changeset
 | 
275  | 
|
| 63167 | 276  | 
text \<open>SML / Eval\<close>  | 
| 26182 | 277  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
278  | 
code_printing type_constructor ref \<rightharpoonup> (SML) "_/ ref"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
279  | 
code_printing type_constructor ref \<rightharpoonup> (Eval) "_/ Unsynchronized.ref"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
280  | 
code_printing constant Ref \<rightharpoonup> (SML) "raise/ (Fail/ \"bare Ref\")"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
281  | 
code_printing constant ref' \<rightharpoonup> (SML) "(fn/ ()/ =>/ ref/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
282  | 
code_printing constant ref' \<rightharpoonup> (Eval) "(fn/ ()/ =>/ Unsynchronized.ref/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
283  | 
code_printing constant Ref.lookup \<rightharpoonup> (SML) "(fn/ ()/ =>/ !/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
284  | 
code_printing constant Ref.update \<rightharpoonup> (SML) "(fn/ ()/ =>/ _/ :=/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
285  | 
code_printing constant "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" \<rightharpoonup> (SML) infixl 6 "="  | 
| 26182 | 286  | 
|
| 39607 | 287  | 
code_reserved Eval Unsynchronized  | 
| 26182 | 288  | 
|
289  | 
||
| 63167 | 290  | 
text \<open>OCaml\<close>  | 
| 26182 | 291  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
292  | 
code_printing type_constructor ref \<rightharpoonup> (OCaml) "_/ ref"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
293  | 
code_printing constant Ref \<rightharpoonup> (OCaml) "failwith/ \"bare Ref\""  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
294  | 
code_printing constant ref' \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ ref/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
295  | 
code_printing constant Ref.lookup \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ !/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
296  | 
code_printing constant Ref.update \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ _/ :=/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
297  | 
code_printing constant "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" \<rightharpoonup> (OCaml) infixl 4 "="  | 
| 26182 | 298  | 
|
299  | 
code_reserved OCaml ref  | 
|
300  | 
||
301  | 
||
| 63167 | 302  | 
text \<open>Haskell\<close>  | 
| 26182 | 303  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
304  | 
code_printing type_constructor ref \<rightharpoonup> (Haskell) "Heap.STRef/ Heap.RealWorld/ _"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
305  | 
code_printing constant Ref \<rightharpoonup> (Haskell) "error/ \"bare Ref\""  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
306  | 
code_printing constant ref' \<rightharpoonup> (Haskell) "Heap.newSTRef"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
307  | 
code_printing constant Ref.lookup \<rightharpoonup> (Haskell) "Heap.readSTRef"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
308  | 
code_printing constant Ref.update \<rightharpoonup> (Haskell) "Heap.writeSTRef"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
309  | 
code_printing constant "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" \<rightharpoonup> (Haskell) infix 4 "=="  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
310  | 
code_printing class_instance ref :: HOL.equal \<rightharpoonup> (Haskell) -  | 
| 26182 | 311  | 
|
| 37842 | 312  | 
|
| 63167 | 313  | 
text \<open>Scala\<close>  | 
| 37842 | 314  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
315  | 
code_printing type_constructor ref \<rightharpoonup> (Scala) "!Ref[_]"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
316  | 
code_printing constant Ref \<rightharpoonup> (Scala) "!sys.error(\"bare Ref\")"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
317  | 
code_printing constant ref' \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Ref((_))"
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
318  | 
code_printing constant Ref.lookup \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Ref.lookup((_))"
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
319  | 
code_printing constant Ref.update \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Ref.update((_), (_))"
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51090 
diff
changeset
 | 
320  | 
code_printing constant "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" \<rightharpoonup> (Scala) infixl 5 "=="  | 
| 37842 | 321  | 
|
| 37758 | 322  | 
end  |