author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 77106 | 5ef443fa4a5d |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
31870 | 1 |
(* Title: HOL/Imperative_HOL/Array.thy |
26170 | 2 |
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen |
3 |
*) |
|
4 |
||
63167 | 5 |
section \<open>Monadic arrays\<close> |
26170 | 6 |
|
7 |
theory Array |
|
31203
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
haftmann
parents:
29822
diff
changeset
|
8 |
imports Heap_Monad |
26170 | 9 |
begin |
10 |
||
63167 | 11 |
subsection \<open>Primitives\<close> |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
12 |
|
61076 | 13 |
definition present :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> bool" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
14 |
"present h a \<longleftrightarrow> addr_of_array a < lim h" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
15 |
|
61076 | 16 |
definition get :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> 'a list" where |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
17 |
"get h a = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
18 |
|
61076 | 19 |
definition set :: "'a::heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
20 |
"set a x = arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
21 |
|
61076 | 22 |
definition alloc :: "'a list \<Rightarrow> heap \<Rightarrow> 'a::heap array \<times> heap" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
23 |
"alloc xs h = (let |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
24 |
l = lim h; |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
25 |
r = Array l; |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
26 |
h'' = set r xs (h\<lparr>lim := l + 1\<rparr>) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
27 |
in (r, h''))" |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
28 |
|
61076 | 29 |
definition length :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> nat" where |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
30 |
"length h a = List.length (get h a)" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
31 |
|
61076 | 32 |
definition update :: "'a::heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
33 |
"update a i x h = set a ((get h a)[i:=x]) h" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
34 |
|
61076 | 35 |
definition noteq :: "'a::heap array \<Rightarrow> 'b::heap array \<Rightarrow> bool" (infix "=!!=" 70) where |
37752 | 36 |
"r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s" |
37 |
||
38 |
||
63167 | 39 |
subsection \<open>Monad operations\<close> |
37752 | 40 |
|
61076 | 41 |
definition new :: "nat \<Rightarrow> 'a::heap \<Rightarrow> 'a array Heap" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
42 |
[code del]: "new n x = Heap_Monad.heap (alloc (replicate n x))" |
37752 | 43 |
|
61076 | 44 |
definition of_list :: "'a::heap list \<Rightarrow> 'a array Heap" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
45 |
[code del]: "of_list xs = Heap_Monad.heap (alloc xs)" |
37752 | 46 |
|
61076 | 47 |
definition make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a::heap) \<Rightarrow> 'a array Heap" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
48 |
[code del]: "make n f = Heap_Monad.heap (alloc (map f [0 ..< n]))" |
37752 | 49 |
|
61076 | 50 |
definition len :: "'a::heap array \<Rightarrow> nat Heap" where |
37802 | 51 |
[code del]: "len a = Heap_Monad.tap (\<lambda>h. length h a)" |
37752 | 52 |
|
61076 | 53 |
definition nth :: "'a::heap array \<Rightarrow> nat \<Rightarrow> 'a Heap" where |
37802 | 54 |
[code del]: "nth a i = Heap_Monad.guard (\<lambda>h. i < length h a) |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
55 |
(\<lambda>h. (get h a ! i, h))" |
37752 | 56 |
|
61076 | 57 |
definition upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a::heap array \<Rightarrow> 'a::heap array Heap" where |
37802 | 58 |
[code del]: "upd i x a = Heap_Monad.guard (\<lambda>h. i < length h a) |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
59 |
(\<lambda>h. (a, update a i x h))" |
37752 | 60 |
|
61076 | 61 |
definition map_entry :: "nat \<Rightarrow> ('a::heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap" where |
37802 | 62 |
[code del]: "map_entry i f a = Heap_Monad.guard (\<lambda>h. i < length h a) |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
63 |
(\<lambda>h. (a, update a i (f (get h a ! i)) h))" |
37752 | 64 |
|
61076 | 65 |
definition swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a::heap array \<Rightarrow> 'a Heap" where |
37802 | 66 |
[code del]: "swap i x a = Heap_Monad.guard (\<lambda>h. i < length h a) |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
67 |
(\<lambda>h. (get h a ! i, update a i x h))" |
37752 | 68 |
|
61076 | 69 |
definition freeze :: "'a::heap array \<Rightarrow> 'a list Heap" where |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
70 |
[code del]: "freeze a = Heap_Monad.tap (\<lambda>h. get h a)" |
37752 | 71 |
|
72 |
||
63167 | 73 |
subsection \<open>Properties\<close> |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
74 |
|
63167 | 75 |
text \<open>FIXME: Does there exist a "canonical" array axiomatisation in |
76 |
the literature?\<close> |
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
77 |
|
63167 | 78 |
text \<open>Primitives\<close> |
37758 | 79 |
|
37807 | 80 |
lemma noteq_sym: "a =!!= b \<Longrightarrow> b =!!= a" |
81 |
and unequal [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'" |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
82 |
unfolding noteq_def by auto |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
83 |
|
37807 | 84 |
lemma noteq_irrefl: "r =!!= r \<Longrightarrow> False" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
85 |
unfolding noteq_def by auto |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
86 |
|
37807 | 87 |
lemma present_alloc_noteq: "present h a \<Longrightarrow> a =!!= fst (alloc xs h)" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
88 |
by (simp add: present_def noteq_def alloc_def Let_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
89 |
|
37807 | 90 |
lemma get_set_eq [simp]: "get (set r x h) r = x" |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
91 |
by (simp add: get_def set_def o_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
92 |
|
37807 | 93 |
lemma get_set_neq [simp]: "r =!!= s \<Longrightarrow> get (set s x h) r = get h r" |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
94 |
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:
37716
diff
changeset
|
95 |
|
37807 | 96 |
lemma set_same [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
97 |
"set r x (set r y h) = set r x h" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
98 |
by (simp add: set_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
99 |
|
37807 | 100 |
lemma set_set_swap: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
101 |
"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
|
102 |
by (simp add: Let_def fun_eq_iff noteq_def set_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
103 |
|
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
104 |
lemma get_update_eq [simp]: |
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
105 |
"get (update a i v h) a = (get h a) [i := v]" |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
106 |
by (simp add: update_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
107 |
|
37807 | 108 |
lemma nth_update_neq [simp]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
109 |
"a =!!= b \<Longrightarrow> get (update b j v h) a ! i = get h a ! i" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
110 |
by (simp add: update_def noteq_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
111 |
|
37807 | 112 |
lemma get_update_elem_neqIndex [simp]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
113 |
"i \<noteq> j \<Longrightarrow> get (update a j v h) a ! i = get h a ! i" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
114 |
by simp |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
115 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
116 |
lemma length_update [simp]: |
37802 | 117 |
"length (update b i v h) = length 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
|
118 |
by (simp add: update_def length_def set_def get_def fun_eq_iff) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
119 |
|
37807 | 120 |
lemma update_swap_neq: |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
121 |
"a =!!= a' \<Longrightarrow> |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
122 |
update a i v (update a' i' v' h) |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
123 |
= update a' i' v' (update a i v h)" |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
124 |
apply (unfold update_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
125 |
apply simp |
37807 | 126 |
apply (subst set_set_swap, assumption) |
127 |
apply (subst get_set_neq) |
|
128 |
apply (erule noteq_sym) |
|
129 |
apply simp |
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
130 |
done |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
131 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
132 |
lemma update_swap_neqIndex: |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
133 |
"\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> update a i v (update a i' v' h) = update a i' v' (update a i v h)" |
37807 | 134 |
by (auto simp add: update_def set_set_swap list_update_swap) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
135 |
|
37807 | 136 |
lemma get_alloc: |
40360 | 137 |
"get (snd (alloc xs h)) (fst (alloc ys h)) = xs" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
138 |
by (simp add: Let_def split_def alloc_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
139 |
|
40360 | 140 |
lemma length_alloc: |
141 |
"length (snd (alloc (xs :: 'a::heap list) h)) (fst (alloc (ys :: 'a list) h)) = List.length xs" |
|
142 |
by (simp add: Array.length_def get_alloc) |
|
143 |
||
37807 | 144 |
lemma set: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
145 |
"set (fst (alloc ls h)) |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
146 |
new_ls (snd (alloc ls h)) |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
147 |
= snd (alloc new_ls h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
148 |
by (simp add: Let_def split_def alloc_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
149 |
|
37807 | 150 |
lemma present_update [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
151 |
"present (update b 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
|
152 |
by (simp add: update_def present_def set_def get_def fun_eq_iff) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
153 |
|
37807 | 154 |
lemma present_alloc [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
155 |
"present (snd (alloc xs h)) (fst (alloc xs h))" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
156 |
by (simp add: present_def alloc_def set_def Let_def) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
157 |
|
37807 | 158 |
lemma not_present_alloc [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
159 |
"\<not> present h (fst (alloc xs h))" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
160 |
by (simp add: present_def alloc_def Let_def) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
161 |
|
37758 | 162 |
|
63167 | 163 |
text \<open>Monad operations\<close> |
37758 | 164 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
165 |
lemma execute_new [execute_simps]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
166 |
"execute (new n x) h = Some (alloc (replicate n x) h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
167 |
by (simp add: new_def execute_simps) |
37758 | 168 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
169 |
lemma success_newI [success_intros]: |
37758 | 170 |
"success (new n x) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
171 |
by (auto intro: success_intros simp add: new_def) |
26170 | 172 |
|
40671 | 173 |
lemma effect_newI [effect_intros]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
174 |
assumes "(a, h') = alloc (replicate n x) h" |
40671 | 175 |
shows "effect (new n x) h h' a" |
176 |
by (rule effectI) (simp add: assms execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
177 |
|
40671 | 178 |
lemma effect_newE [effect_elims]: |
179 |
assumes "effect (new n x) h h' r" |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
180 |
obtains "r = fst (alloc (replicate n x) h)" "h' = snd (alloc (replicate n x) h)" |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
181 |
"get h' r = replicate n x" "present h' r" "\<not> present h r" |
40671 | 182 |
using assms by (rule effectE) (simp add: get_alloc execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
183 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
184 |
lemma execute_of_list [execute_simps]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
185 |
"execute (of_list xs) h = Some (alloc xs h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
186 |
by (simp add: of_list_def execute_simps) |
37758 | 187 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
188 |
lemma success_of_listI [success_intros]: |
37758 | 189 |
"success (of_list xs) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
190 |
by (auto intro: success_intros simp add: of_list_def) |
26170 | 191 |
|
40671 | 192 |
lemma effect_of_listI [effect_intros]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
193 |
assumes "(a, h') = alloc xs h" |
40671 | 194 |
shows "effect (of_list xs) h h' a" |
195 |
by (rule effectI) (simp add: assms execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
196 |
|
40671 | 197 |
lemma effect_of_listE [effect_elims]: |
198 |
assumes "effect (of_list xs) h h' r" |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
199 |
obtains "r = fst (alloc xs h)" "h' = snd (alloc xs h)" |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
200 |
"get h' r = xs" "present h' r" "\<not> present h r" |
40671 | 201 |
using assms by (rule effectE) (simp add: get_alloc execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
202 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
203 |
lemma execute_make [execute_simps]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
204 |
"execute (make n f) h = Some (alloc (map f [0 ..< n]) h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
205 |
by (simp add: make_def execute_simps) |
26170 | 206 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
207 |
lemma success_makeI [success_intros]: |
37758 | 208 |
"success (make n f) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
209 |
by (auto intro: success_intros simp add: make_def) |
37758 | 210 |
|
40671 | 211 |
lemma effect_makeI [effect_intros]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
212 |
assumes "(a, h') = alloc (map f [0 ..< n]) h" |
40671 | 213 |
shows "effect (make n f) h h' a" |
214 |
by (rule effectI) (simp add: assms execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
215 |
|
40671 | 216 |
lemma effect_makeE [effect_elims]: |
217 |
assumes "effect (make n f) h h' r" |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
218 |
obtains "r = fst (alloc (map f [0 ..< n]) h)" "h' = snd (alloc (map f [0 ..< n]) h)" |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
219 |
"get h' r = map f [0 ..< n]" "present h' r" "\<not> present h r" |
40671 | 220 |
using assms by (rule effectE) (simp add: get_alloc execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
221 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
222 |
lemma execute_len [execute_simps]: |
37802 | 223 |
"execute (len a) h = Some (length h a, h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
224 |
by (simp add: len_def execute_simps) |
37758 | 225 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
226 |
lemma success_lenI [success_intros]: |
37758 | 227 |
"success (len a) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
228 |
by (auto intro: success_intros simp add: len_def) |
37752 | 229 |
|
40671 | 230 |
lemma effect_lengthI [effect_intros]: |
37802 | 231 |
assumes "h' = h" "r = length h a" |
40671 | 232 |
shows "effect (len a) h h' r" |
233 |
by (rule effectI) (simp add: assms execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
234 |
|
40671 | 235 |
lemma effect_lengthE [effect_elims]: |
236 |
assumes "effect (len a) h h' r" |
|
37802 | 237 |
obtains "r = length h' a" "h' = h" |
40671 | 238 |
using assms by (rule effectE) (simp add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
239 |
|
37758 | 240 |
lemma execute_nth [execute_simps]: |
37802 | 241 |
"i < length h a \<Longrightarrow> |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
242 |
execute (nth a i) h = Some (get h a ! i, h)" |
37802 | 243 |
"i \<ge> length h a \<Longrightarrow> execute (nth a i) h = None" |
37758 | 244 |
by (simp_all add: nth_def execute_simps) |
245 |
||
246 |
lemma success_nthI [success_intros]: |
|
37802 | 247 |
"i < length h a \<Longrightarrow> success (nth a i) h" |
37758 | 248 |
by (auto intro: success_intros simp add: nth_def) |
26170 | 249 |
|
40671 | 250 |
lemma effect_nthI [effect_intros]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
251 |
assumes "i < length h a" "h' = h" "r = get h a ! i" |
40671 | 252 |
shows "effect (nth a i) h h' r" |
253 |
by (rule effectI) (insert assms, simp add: execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
254 |
|
40671 | 255 |
lemma effect_nthE [effect_elims]: |
256 |
assumes "effect (nth a i) h h' r" |
|
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
257 |
obtains "i < length h a" "r = get h a ! i" "h' = h" |
58510 | 258 |
using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
259 |
|
37758 | 260 |
lemma execute_upd [execute_simps]: |
37802 | 261 |
"i < length h a \<Longrightarrow> |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
262 |
execute (upd i x a) h = Some (a, update a i x h)" |
37802 | 263 |
"i \<ge> length h a \<Longrightarrow> execute (upd i x a) h = None" |
37758 | 264 |
by (simp_all add: upd_def execute_simps) |
26170 | 265 |
|
37758 | 266 |
lemma success_updI [success_intros]: |
37802 | 267 |
"i < length h a \<Longrightarrow> success (upd i x a) h" |
37758 | 268 |
by (auto intro: success_intros simp add: upd_def) |
269 |
||
40671 | 270 |
lemma effect_updI [effect_intros]: |
37802 | 271 |
assumes "i < length h a" "h' = update a i v h" |
40671 | 272 |
shows "effect (upd i v a) h h' a" |
273 |
by (rule effectI) (insert assms, simp add: execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
274 |
|
40671 | 275 |
lemma effect_updE [effect_elims]: |
276 |
assumes "effect (upd i v a) h h' r" |
|
37802 | 277 |
obtains "r = a" "h' = update a i v h" "i < length h a" |
58510 | 278 |
using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
279 |
|
37758 | 280 |
lemma execute_map_entry [execute_simps]: |
37802 | 281 |
"i < length h a \<Longrightarrow> |
37758 | 282 |
execute (map_entry i f a) h = |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
283 |
Some (a, update a i (f (get h a ! i)) h)" |
37802 | 284 |
"i \<ge> length h a \<Longrightarrow> execute (map_entry i f a) h = None" |
37758 | 285 |
by (simp_all add: map_entry_def execute_simps) |
37752 | 286 |
|
37758 | 287 |
lemma success_map_entryI [success_intros]: |
37802 | 288 |
"i < length h a \<Longrightarrow> success (map_entry i f a) h" |
37758 | 289 |
by (auto intro: success_intros simp add: map_entry_def) |
290 |
||
40671 | 291 |
lemma effect_map_entryI [effect_intros]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
292 |
assumes "i < length h a" "h' = update a i (f (get h a ! i)) h" "r = a" |
40671 | 293 |
shows "effect (map_entry i f a) h h' r" |
294 |
by (rule effectI) (insert assms, simp add: execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
295 |
|
40671 | 296 |
lemma effect_map_entryE [effect_elims]: |
297 |
assumes "effect (map_entry i f a) h h' r" |
|
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
298 |
obtains "r = a" "h' = update a i (f (get h a ! i)) h" "i < length h a" |
58510 | 299 |
using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
300 |
|
37758 | 301 |
lemma execute_swap [execute_simps]: |
37802 | 302 |
"i < length h a \<Longrightarrow> |
37758 | 303 |
execute (swap i x a) h = |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
304 |
Some (get h a ! i, update a i x h)" |
37802 | 305 |
"i \<ge> length h a \<Longrightarrow> execute (swap i x a) h = None" |
37758 | 306 |
by (simp_all add: swap_def execute_simps) |
307 |
||
308 |
lemma success_swapI [success_intros]: |
|
37802 | 309 |
"i < length h a \<Longrightarrow> success (swap i x a) h" |
37758 | 310 |
by (auto intro: success_intros simp add: swap_def) |
37752 | 311 |
|
40671 | 312 |
lemma effect_swapI [effect_intros]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
313 |
assumes "i < length h a" "h' = update a i x h" "r = get h a ! i" |
40671 | 314 |
shows "effect (swap i x a) h h' r" |
315 |
by (rule effectI) (insert assms, simp add: execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
316 |
|
40671 | 317 |
lemma effect_swapE [effect_elims]: |
318 |
assumes "effect (swap i x a) h h' r" |
|
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
319 |
obtains "r = get h a ! i" "h' = update a i x h" "i < length h a" |
58510 | 320 |
using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
321 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
322 |
lemma execute_freeze [execute_simps]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
323 |
"execute (freeze a) h = Some (get h a, h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
324 |
by (simp add: freeze_def execute_simps) |
37758 | 325 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
326 |
lemma success_freezeI [success_intros]: |
37758 | 327 |
"success (freeze a) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
328 |
by (auto intro: success_intros simp add: freeze_def) |
26170 | 329 |
|
40671 | 330 |
lemma effect_freezeI [effect_intros]: |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
331 |
assumes "h' = h" "r = get h a" |
40671 | 332 |
shows "effect (freeze a) h h' r" |
333 |
by (rule effectI) (insert assms, simp add: execute_simps) |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
334 |
|
40671 | 335 |
lemma effect_freezeE [effect_elims]: |
336 |
assumes "effect (freeze a) h h' r" |
|
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
337 |
obtains "h' = h" "r = get h a" |
40671 | 338 |
using assms by (rule effectE) (simp add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
339 |
|
26170 | 340 |
lemma upd_return: |
62026 | 341 |
"upd i x a \<then> return a = upd i x a" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
342 |
by (rule Heap_eqI) (simp add: bind_def guard_def upd_def execute_simps) |
26170 | 343 |
|
37752 | 344 |
lemma array_make: |
345 |
"new n x = make n (\<lambda>_. x)" |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
346 |
by (rule Heap_eqI) (simp add: map_replicate_trivial execute_simps) |
26170 | 347 |
|
37845
b70d7a347964
first roughly working version of Imperative HOL for Scala
haftmann
parents:
37842
diff
changeset
|
348 |
lemma array_of_list_make [code]: |
37752 | 349 |
"of_list xs = make (List.length xs) (\<lambda>n. xs ! n)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
350 |
by (rule Heap_eqI) (simp add: map_nth execute_simps) |
26170 | 351 |
|
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
352 |
hide_const (open) present get set alloc length update noteq new of_list make len nth upd map_entry swap freeze |
26170 | 353 |
|
26182 | 354 |
|
63167 | 355 |
subsection \<open>Code generator setup\<close> |
26182 | 356 |
|
63167 | 357 |
subsubsection \<open>Logical intermediate layer\<close> |
26182 | 358 |
|
359 |
definition new' where |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48073
diff
changeset
|
360 |
[code del]: "new' = Array.new o nat_of_integer" |
37752 | 361 |
|
28562 | 362 |
lemma [code]: |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48073
diff
changeset
|
363 |
"Array.new = new' o of_nat" |
26182 | 364 |
by (simp add: new'_def o_def) |
365 |
||
366 |
definition make' where |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48073
diff
changeset
|
367 |
[code del]: "make' i f = Array.make (nat_of_integer i) (f o of_nat)" |
37752 | 368 |
|
28562 | 369 |
lemma [code]: |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48073
diff
changeset
|
370 |
"Array.make n f = make' (of_nat n) (f o nat_of_integer)" |
26182 | 371 |
by (simp add: make'_def o_def) |
372 |
||
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
373 |
definition len' where |
62026 | 374 |
[code del]: "len' a = Array.len a \<bind> (\<lambda>n. return (of_nat n))" |
37752 | 375 |
|
28562 | 376 |
lemma [code]: |
62026 | 377 |
"Array.len a = len' a \<bind> (\<lambda>i. return (nat_of_integer i))" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
378 |
by (simp add: len'_def) |
26182 | 379 |
|
380 |
definition nth' where |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48073
diff
changeset
|
381 |
[code del]: "nth' a = Array.nth a o nat_of_integer" |
37752 | 382 |
|
28562 | 383 |
lemma [code]: |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
48073
diff
changeset
|
384 |
"Array.nth a n = nth' a (of_nat n)" |
26182 | 385 |
by (simp add: nth'_def) |
386 |
||
387 |
definition upd' where |
|
62026 | 388 |
[code del]: "upd' a i x = Array.upd (nat_of_integer i) x a \<then> return ()" |
37752 | 389 |
|
28562 | 390 |
lemma [code]: |
62026 | 391 |
"Array.upd i x a = upd' a (of_nat i) x \<then> return a" |
37709 | 392 |
by (simp add: upd'_def upd_return) |
26182 | 393 |
|
37752 | 394 |
lemma [code]: |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
395 |
"Array.map_entry i f a = do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
396 |
x \<leftarrow> Array.nth a i; |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
397 |
Array.upd i (f x) a |
37792 | 398 |
}" |
37758 | 399 |
by (rule Heap_eqI) (simp add: bind_def guard_def map_entry_def execute_simps) |
26182 | 400 |
|
37752 | 401 |
lemma [code]: |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
402 |
"Array.swap i x a = do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
403 |
y \<leftarrow> Array.nth a i; |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
404 |
Array.upd i x a; |
37752 | 405 |
return y |
37792 | 406 |
}" |
37758 | 407 |
by (rule Heap_eqI) (simp add: bind_def guard_def swap_def execute_simps) |
37752 | 408 |
|
409 |
lemma [code]: |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
410 |
"Array.freeze a = do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
411 |
n \<leftarrow> Array.len a; |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
412 |
Heap_Monad.fold_map (\<lambda>i. Array.nth a i) [0..<n] |
37792 | 413 |
}" |
37752 | 414 |
proof (rule Heap_eqI) |
415 |
fix h |
|
416 |
have *: "List.map |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
417 |
(\<lambda>x. fst (the (if x < Array.length h a |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
418 |
then Some (Array.get h a ! x, h) else None))) |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
419 |
[0..<Array.length h a] = |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
420 |
List.map (List.nth (Array.get h a)) [0..<Array.length h a]" |
37752 | 421 |
by simp |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
422 |
have "execute (Heap_Monad.fold_map (Array.nth a) [0..<Array.length h a]) h = |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
423 |
Some (Array.get h a, h)" |
37756
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
haftmann
parents:
37752
diff
changeset
|
424 |
apply (subst execute_fold_map_unchanged_heap) |
37752 | 425 |
apply (simp_all add: nth_def guard_def *) |
426 |
apply (simp add: length_def map_nth) |
|
427 |
done |
|
37792 | 428 |
then have "execute (do { |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
429 |
n \<leftarrow> Array.len a; |
37756
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
haftmann
parents:
37752
diff
changeset
|
430 |
Heap_Monad.fold_map (Array.nth a) [0..<n] |
37806
a7679be14442
qualified names for (really) all array operations
haftmann
parents:
37805
diff
changeset
|
431 |
}) h = Some (Array.get h a, h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
432 |
by (auto intro: execute_bind_eq_SomeI simp add: execute_simps) |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
433 |
then show "execute (Array.freeze a) h = execute (do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
434 |
n \<leftarrow> Array.len a; |
37756
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
haftmann
parents:
37752
diff
changeset
|
435 |
Heap_Monad.fold_map (Array.nth a) [0..<n] |
37792 | 436 |
}) h" by (simp add: execute_simps) |
37752 | 437 |
qed |
438 |
||
37831
fa3a2e35c4f1
repaired some implementations of imperative operations
haftmann
parents:
37827
diff
changeset
|
439 |
hide_const (open) new' make' len' nth' upd' |
37752 | 440 |
|
441 |
||
63167 | 442 |
text \<open>SML\<close> |
26182 | 443 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
444 |
code_printing type_constructor array \<rightharpoonup> (SML) "_/ array" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
445 |
code_printing constant Array \<rightharpoonup> (SML) "raise/ (Fail/ \"bare Array\")" |
66003
5b2fab45db92
Fixed to also work with non-IntInf default int type
lammich <lammich@in.tum.de>
parents:
63167
diff
changeset
|
446 |
code_printing constant Array.new' \<rightharpoonup> (SML) "(fn/ ()/ =>/ Array.array/ (IntInf.toInt _,/ (_)))" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
447 |
code_printing constant Array.of_list \<rightharpoonup> (SML) "(fn/ ()/ =>/ Array.fromList/ _)" |
66003
5b2fab45db92
Fixed to also work with non-IntInf default int type
lammich <lammich@in.tum.de>
parents:
63167
diff
changeset
|
448 |
code_printing constant Array.make' \<rightharpoonup> (SML) "(fn/ ()/ =>/ Array.tabulate/ (IntInf.toInt _,/ _ o IntInf.fromInt))" |
5b2fab45db92
Fixed to also work with non-IntInf default int type
lammich <lammich@in.tum.de>
parents:
63167
diff
changeset
|
449 |
code_printing constant Array.len' \<rightharpoonup> (SML) "(fn/ ()/ =>/ IntInf.fromInt (Array.length/ _))" |
5b2fab45db92
Fixed to also work with non-IntInf default int type
lammich <lammich@in.tum.de>
parents:
63167
diff
changeset
|
450 |
code_printing constant Array.nth' \<rightharpoonup> (SML) "(fn/ ()/ =>/ Array.sub/ ((_),/ IntInf.toInt _))" |
5b2fab45db92
Fixed to also work with non-IntInf default int type
lammich <lammich@in.tum.de>
parents:
63167
diff
changeset
|
451 |
code_printing constant Array.upd' \<rightharpoonup> (SML) "(fn/ ()/ =>/ Array.update/ ((_),/ IntInf.toInt _,/ (_)))" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
452 |
code_printing constant "HOL.equal :: 'a array \<Rightarrow> 'a array \<Rightarrow> bool" \<rightharpoonup> (SML) infixl 6 "=" |
77106 | 453 |
|
26182 | 454 |
code_reserved SML Array |
455 |
||
456 |
||
63167 | 457 |
text \<open>OCaml\<close> |
26182 | 458 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
459 |
code_printing type_constructor array \<rightharpoonup> (OCaml) "_/ array" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
460 |
code_printing constant Array \<rightharpoonup> (OCaml) "failwith/ \"bare Array\"" |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
66003
diff
changeset
|
461 |
code_printing constant Array.new' \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ Array.make/ (Z.to'_int/ _)/ _)" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
462 |
code_printing constant Array.of_list \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ Array.of'_list/ _)" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
463 |
code_printing constant Array.make' \<rightharpoonup> (OCaml) |
69906
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
66003
diff
changeset
|
464 |
"(fun/ ()/ ->/ Array.init/ (Z.to'_int/ _)/ (fun k'_ ->/ _/ (Z.of'_int/ k'_)))" |
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
66003
diff
changeset
|
465 |
code_printing constant Array.len' \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ Z.of'_int/ (Array.length/ _))" |
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
66003
diff
changeset
|
466 |
code_printing constant Array.nth' \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ Array.get/ _/ (Z.to'_int/ _))" |
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
haftmann
parents:
66003
diff
changeset
|
467 |
code_printing constant Array.upd' \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ Array.set/ _/ (Z.to'_int/ _)/ _)" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
468 |
code_printing constant "HOL.equal :: 'a array \<Rightarrow> 'a array \<Rightarrow> bool" \<rightharpoonup> (OCaml) infixl 4 "=" |
26182 | 469 |
|
470 |
code_reserved OCaml Array |
|
471 |
||
472 |
||
63167 | 473 |
text \<open>Haskell\<close> |
26182 | 474 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
475 |
code_printing type_constructor array \<rightharpoonup> (Haskell) "Heap.STArray/ Heap.RealWorld/ _" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
476 |
code_printing constant Array \<rightharpoonup> (Haskell) "error/ \"bare Array\"" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
477 |
code_printing constant Array.new' \<rightharpoonup> (Haskell) "Heap.newArray" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
478 |
code_printing constant Array.of_list \<rightharpoonup> (Haskell) "Heap.newListArray" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
479 |
code_printing constant Array.make' \<rightharpoonup> (Haskell) "Heap.newFunArray" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
480 |
code_printing constant Array.len' \<rightharpoonup> (Haskell) "Heap.lengthArray" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
481 |
code_printing constant Array.nth' \<rightharpoonup> (Haskell) "Heap.readArray" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
482 |
code_printing constant Array.upd' \<rightharpoonup> (Haskell) "Heap.writeArray" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
483 |
code_printing constant "HOL.equal :: 'a array \<Rightarrow> 'a array \<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:
51143
diff
changeset
|
484 |
code_printing class_instance array :: HOL.equal \<rightharpoonup> (Haskell) - |
26182 | 485 |
|
37842 | 486 |
|
63167 | 487 |
text \<open>Scala\<close> |
37842 | 488 |
|
73137
ca450d902198
updated to scala-2.13: its ArraySeq implementation is not usable here (requires scala.relection.ClassTag);
wenzelm
parents:
69906
diff
changeset
|
489 |
code_printing type_constructor array \<rightharpoonup> (Scala) "!Array.T[_]" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
490 |
code_printing constant Array \<rightharpoonup> (Scala) "!sys.error(\"bare Array\")" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
491 |
code_printing constant Array.new' \<rightharpoonup> (Scala) "('_: Unit)/ => / Array.alloc((_))((_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
492 |
code_printing constant Array.make' \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Array.make((_))((_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
493 |
code_printing constant Array.len' \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Array.len((_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
494 |
code_printing constant Array.nth' \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Array.nth((_), (_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
495 |
code_printing constant Array.upd' \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Array.upd((_), (_), (_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
496 |
code_printing constant Array.freeze \<rightharpoonup> (Scala) "('_: Unit)/ =>/ Array.freeze((_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
497 |
code_printing constant "HOL.equal :: 'a array \<Rightarrow> 'a array \<Rightarrow> bool" \<rightharpoonup> (Scala) infixl 5 "==" |
37842 | 498 |
|
26170 | 499 |
end |