author | haftmann |
Tue, 13 Jul 2010 16:00:56 +0200 | |
changeset 37804 | 0145e59c1f6c |
parent 37803 | 582d0fbd201e |
child 37805 | 0f797d586ce5 |
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 |
||
5 |
header {* Monadic arrays *} |
|
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 |
||
37752 | 11 |
subsection {* Primitives *} |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
12 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
13 |
definition present :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> bool" where |
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 |
|
37752 | 16 |
definition (*FIXME get :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a list" where*) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
17 |
get_array :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> 'a list" where |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
18 |
"get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))" |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
19 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
20 |
definition set :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
21 |
"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
|
22 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
23 |
definition alloc :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
24 |
"alloc xs h = (let |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
25 |
l = lim h; |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
26 |
r = Array l; |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
27 |
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
|
28 |
in (r, h''))" |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
29 |
|
37802 | 30 |
definition length :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> nat" where |
31 |
"length h a = List.length (get_array a h)" |
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
32 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
33 |
definition update :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
34 |
"update a i x h = set a ((get_array a h)[i:=x]) h" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
35 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
36 |
definition noteq :: "'a\<Colon>heap array \<Rightarrow> 'b\<Colon>heap array \<Rightarrow> bool" (infix "=!!=" 70) where |
37752 | 37 |
"r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s" |
38 |
||
39 |
||
40 |
subsection {* Monad operations *} |
|
41 |
||
42 |
definition new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
43 |
[code del]: "new n x = Heap_Monad.heap (alloc (replicate n x))" |
37752 | 44 |
|
45 |
definition of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
46 |
[code del]: "of_list xs = Heap_Monad.heap (alloc xs)" |
37752 | 47 |
|
48 |
definition make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap" where |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
49 |
[code del]: "make n f = Heap_Monad.heap (alloc (map f [0 ..< n]))" |
37752 | 50 |
|
51 |
definition len :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where |
|
37802 | 52 |
[code del]: "len a = Heap_Monad.tap (\<lambda>h. length h a)" |
37752 | 53 |
|
54 |
definition nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap" where |
|
37802 | 55 |
[code del]: "nth a i = Heap_Monad.guard (\<lambda>h. i < length h a) |
37752 | 56 |
(\<lambda>h. (get_array a h ! i, h))" |
57 |
||
58 |
definition upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap" where |
|
37802 | 59 |
[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
|
60 |
(\<lambda>h. (a, update a i x h))" |
37752 | 61 |
|
62 |
definition map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap" where |
|
37802 | 63 |
[code del]: "map_entry i f 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
|
64 |
(\<lambda>h. (a, update a i (f (get_array a h ! i)) h))" |
37752 | 65 |
|
66 |
definition swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap" where |
|
37802 | 67 |
[code del]: "swap 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
|
68 |
(\<lambda>h. (get_array a h ! i, update a i x h))" |
37752 | 69 |
|
70 |
definition freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap" where |
|
37758 | 71 |
[code del]: "freeze a = Heap_Monad.tap (\<lambda>h. get_array a h)" |
37752 | 72 |
|
73 |
||
74 |
subsection {* Properties *} |
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
75 |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
76 |
text {* FIXME: Does there exist a "canonical" array axiomatisation in |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
77 |
the literature? *} |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
78 |
|
37758 | 79 |
text {* Primitives *} |
80 |
||
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
81 |
lemma noteq_arrs_sym: "a =!!= b \<Longrightarrow> b =!!= a" |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
82 |
and unequal_arrs [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
83 |
unfolding noteq_def by auto |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
84 |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
85 |
lemma noteq_arrs_irrefl: "r =!!= r \<Longrightarrow> False" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
86 |
unfolding noteq_def by auto |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
87 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
88 |
lemma present_new_arr: "present h a \<Longrightarrow> a =!!= fst (alloc xs h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
89 |
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
|
90 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
91 |
lemma array_get_set_eq [simp]: "get_array r (set r x h) = x" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
92 |
by (simp add: get_array_def set_def o_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
93 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
94 |
lemma array_get_set_neq [simp]: "r =!!= s \<Longrightarrow> get_array r (set s x h) = get_array r h" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
95 |
by (simp add: noteq_def get_array_def set_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
96 |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
97 |
lemma set_array_same [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
98 |
"set r x (set r y h) = set r x h" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
99 |
by (simp add: set_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
100 |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
101 |
lemma array_set_set_swap: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
102 |
"r =!!= r' \<Longrightarrow> set r x (set r' x' h) = set r' x' (set r x h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
103 |
by (simp add: Let_def expand_fun_eq noteq_def set_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
104 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
105 |
lemma get_array_update_eq [simp]: |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
106 |
"get_array a (update a i v h) = (get_array a h) [i := v]" |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
107 |
by (simp add: update_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
108 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
109 |
lemma nth_update_array_neq_array [simp]: |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
110 |
"a =!!= b \<Longrightarrow> get_array a (update b j v h) ! i = get_array a h ! i" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
111 |
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
|
112 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
113 |
lemma get_arry_array_update_elem_neqIndex [simp]: |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
114 |
"i \<noteq> j \<Longrightarrow> get_array a (update a j v h) ! i = get_array a h ! i" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
115 |
by simp |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
116 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
117 |
lemma length_update [simp]: |
37802 | 118 |
"length (update b i v h) = length h" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
119 |
by (simp add: update_def length_def set_def get_array_def expand_fun_eq) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
120 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
121 |
lemma update_swap_neqArray: |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
122 |
"a =!!= a' \<Longrightarrow> |
37798
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 |
= update a' i' v' (update a i v h)" |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
125 |
apply (unfold update_def) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
126 |
apply simp |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
127 |
apply (subst array_set_set_swap, assumption) |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
128 |
apply (subst array_get_set_neq) |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
129 |
apply (erule noteq_arrs_sym) |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
130 |
apply (simp) |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
131 |
done |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
132 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
133 |
lemma update_swap_neqIndex: |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
134 |
"\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> update a i v (update a i' v' h) = update a i' v' (update a i v h)" |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
135 |
by (auto simp add: update_def array_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
|
136 |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
137 |
lemma get_array_init_array_list: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
138 |
"get_array (fst (alloc ls h)) (snd (alloc ls' h)) = ls'" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
139 |
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
|
140 |
|
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
141 |
lemma set_array: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
142 |
"set (fst (alloc ls h)) |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
143 |
new_ls (snd (alloc ls h)) |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
144 |
= snd (alloc new_ls h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
145 |
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
|
146 |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
147 |
lemma array_present_update [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
148 |
"present (update b i v h) = present h" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
149 |
by (simp add: update_def present_def set_def get_array_def expand_fun_eq) |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
150 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
151 |
lemma array_present_array [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
152 |
"present (snd (alloc xs h)) (fst (alloc xs h))" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
153 |
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
|
154 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
155 |
lemma not_array_present_array [simp]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
156 |
"\<not> present h (fst (alloc xs h))" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
157 |
by (simp add: present_def alloc_def Let_def) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
158 |
|
37758 | 159 |
|
160 |
text {* Monad operations *} |
|
161 |
||
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
162 |
lemma execute_new [execute_simps]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
163 |
"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
|
164 |
by (simp add: new_def execute_simps) |
37758 | 165 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
166 |
lemma success_newI [success_intros]: |
37758 | 167 |
"success (new n x) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
168 |
by (auto intro: success_intros simp add: new_def) |
26170 | 169 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
170 |
lemma crel_newI [crel_intros]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
171 |
assumes "(a, h') = alloc (replicate n x) h" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
172 |
shows "crel (new n x) h h' a" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
173 |
by (rule crelI) (simp add: assms execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
174 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
175 |
lemma crel_newE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
176 |
assumes "crel (new n x) h h' r" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
177 |
obtains "r = fst (alloc (replicate n x) h)" "h' = snd (alloc (replicate n x) h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
178 |
"get_array r h' = replicate n x" "present h' r" "\<not> present h r" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
179 |
using assms by (rule crelE) (simp add: get_array_init_array_list execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
180 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
181 |
lemma execute_of_list [execute_simps]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
182 |
"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
|
183 |
by (simp add: of_list_def execute_simps) |
37758 | 184 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
185 |
lemma success_of_listI [success_intros]: |
37758 | 186 |
"success (of_list xs) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
187 |
by (auto intro: success_intros simp add: of_list_def) |
26170 | 188 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
189 |
lemma crel_of_listI [crel_intros]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
190 |
assumes "(a, h') = alloc xs h" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
191 |
shows "crel (of_list xs) h h' a" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
192 |
by (rule crelI) (simp add: assms execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
193 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
194 |
lemma crel_of_listE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
195 |
assumes "crel (of_list xs) h h' r" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
196 |
obtains "r = fst (alloc xs h)" "h' = snd (alloc xs h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
197 |
"get_array r h' = xs" "present h' r" "\<not> present h r" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
198 |
using assms by (rule crelE) (simp add: get_array_init_array_list execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
199 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
200 |
lemma execute_make [execute_simps]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
201 |
"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
|
202 |
by (simp add: make_def execute_simps) |
26170 | 203 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
204 |
lemma success_makeI [success_intros]: |
37758 | 205 |
"success (make n f) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
206 |
by (auto intro: success_intros simp add: make_def) |
37758 | 207 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
208 |
lemma crel_makeI [crel_intros]: |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
209 |
assumes "(a, h') = alloc (map f [0 ..< n]) h" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
210 |
shows "crel (make n f) h h' a" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
211 |
by (rule crelI) (simp add: assms execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
212 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
213 |
lemma crel_makeE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
214 |
assumes "crel (make n f) h h' r" |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
215 |
obtains "r = fst (alloc (map f [0 ..< n]) h)" "h' = snd (alloc (map f [0 ..< n]) h)" |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
216 |
"get_array r h' = map f [0 ..< n]" "present h' r" "\<not> present h r" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
217 |
using assms by (rule crelE) (simp add: get_array_init_array_list execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
218 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
219 |
lemma execute_len [execute_simps]: |
37802 | 220 |
"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
|
221 |
by (simp add: len_def execute_simps) |
37758 | 222 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
223 |
lemma success_lenI [success_intros]: |
37758 | 224 |
"success (len a) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
225 |
by (auto intro: success_intros simp add: len_def) |
37752 | 226 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
227 |
lemma crel_lengthI [crel_intros]: |
37802 | 228 |
assumes "h' = h" "r = length h a" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
229 |
shows "crel (len a) h h' r" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
230 |
by (rule crelI) (simp add: assms execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
231 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
232 |
lemma crel_lengthE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
233 |
assumes "crel (len a) h h' r" |
37802 | 234 |
obtains "r = length h' a" "h' = h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
235 |
using assms by (rule crelE) (simp add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
236 |
|
37758 | 237 |
lemma execute_nth [execute_simps]: |
37802 | 238 |
"i < length h a \<Longrightarrow> |
37758 | 239 |
execute (nth a i) h = Some (get_array a h ! i, h)" |
37802 | 240 |
"i \<ge> length h a \<Longrightarrow> execute (nth a i) h = None" |
37758 | 241 |
by (simp_all add: nth_def execute_simps) |
242 |
||
243 |
lemma success_nthI [success_intros]: |
|
37802 | 244 |
"i < length h a \<Longrightarrow> success (nth a i) h" |
37758 | 245 |
by (auto intro: success_intros simp add: nth_def) |
26170 | 246 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
247 |
lemma crel_nthI [crel_intros]: |
37802 | 248 |
assumes "i < length h a" "h' = h" "r = get_array a h ! i" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
249 |
shows "crel (nth a i) h h' r" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
250 |
by (rule crelI) (insert assms, simp add: execute_simps) |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
251 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
252 |
lemma crel_nthE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
253 |
assumes "crel (nth a i) h h' r" |
37802 | 254 |
obtains "i < length h a" "r = get_array a h ! i" "h' = h" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
255 |
using assms by (rule crelE) |
37802 | 256 |
(erule successE, cases "i < length h a", simp_all add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
257 |
|
37758 | 258 |
lemma execute_upd [execute_simps]: |
37802 | 259 |
"i < length h a \<Longrightarrow> |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
260 |
execute (upd i x a) h = Some (a, update a i x h)" |
37802 | 261 |
"i \<ge> length h a \<Longrightarrow> execute (upd i x a) h = None" |
37758 | 262 |
by (simp_all add: upd_def execute_simps) |
26170 | 263 |
|
37758 | 264 |
lemma success_updI [success_intros]: |
37802 | 265 |
"i < length h a \<Longrightarrow> success (upd i x a) h" |
37758 | 266 |
by (auto intro: success_intros simp add: upd_def) |
267 |
||
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
268 |
lemma crel_updI [crel_intros]: |
37802 | 269 |
assumes "i < length h a" "h' = update a i v h" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
270 |
shows "crel (upd i v a) h h' a" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
271 |
by (rule crelI) (insert assms, simp add: execute_simps) |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
272 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
273 |
lemma crel_updE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
274 |
assumes "crel (upd i v a) h h' r" |
37802 | 275 |
obtains "r = a" "h' = update a i v h" "i < length h a" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
276 |
using assms by (rule crelE) |
37802 | 277 |
(erule successE, cases "i < length h a", simp_all add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
278 |
|
37758 | 279 |
lemma execute_map_entry [execute_simps]: |
37802 | 280 |
"i < length h a \<Longrightarrow> |
37758 | 281 |
execute (map_entry i f a) h = |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
282 |
Some (a, update a i (f (get_array a h ! i)) h)" |
37802 | 283 |
"i \<ge> length h a \<Longrightarrow> execute (map_entry i f a) h = None" |
37758 | 284 |
by (simp_all add: map_entry_def execute_simps) |
37752 | 285 |
|
37758 | 286 |
lemma success_map_entryI [success_intros]: |
37802 | 287 |
"i < length h a \<Longrightarrow> success (map_entry i f a) h" |
37758 | 288 |
by (auto intro: success_intros simp add: map_entry_def) |
289 |
||
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
290 |
lemma crel_map_entryI [crel_intros]: |
37802 | 291 |
assumes "i < length h a" "h' = update a i (f (get_array a h ! i)) h" "r = a" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
292 |
shows "crel (map_entry i f a) h h' r" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
293 |
by (rule crelI) (insert assms, simp add: execute_simps) |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
294 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
295 |
lemma crel_map_entryE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
296 |
assumes "crel (map_entry i f a) h h' r" |
37802 | 297 |
obtains "r = a" "h' = update a i (f (get_array a h ! i)) h" "i < length h a" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
298 |
using assms by (rule crelE) |
37802 | 299 |
(erule successE, cases "i < length h a", simp_all add: execute_simps) |
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 = |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
304 |
Some (get_array a h ! 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 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
312 |
lemma crel_swapI [crel_intros]: |
37802 | 313 |
assumes "i < length h a" "h' = update a i x h" "r = get_array a h ! i" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
314 |
shows "crel (swap i x a) h h' r" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
315 |
by (rule crelI) (insert assms, simp add: execute_simps) |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
316 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
317 |
lemma crel_swapE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
318 |
assumes "crel (swap i x a) h h' r" |
37802 | 319 |
obtains "r = get_array a h ! i" "h' = update a i x h" "i < length h a" |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
320 |
using assms by (rule crelE) |
37802 | 321 |
(erule successE, cases "i < length h a", simp_all add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
322 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
323 |
lemma execute_freeze [execute_simps]: |
37758 | 324 |
"execute (freeze a) h = Some (get_array a h, h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
325 |
by (simp add: freeze_def execute_simps) |
37758 | 326 |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
327 |
lemma success_freezeI [success_intros]: |
37758 | 328 |
"success (freeze a) h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
329 |
by (auto intro: success_intros simp add: freeze_def) |
26170 | 330 |
|
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
331 |
lemma crel_freezeI [crel_intros]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
332 |
assumes "h' = h" "r = get_array a h" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
333 |
shows "crel (freeze a) h h' r" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
334 |
by (rule crelI) (insert assms, simp add: execute_simps) |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
335 |
|
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
336 |
lemma crel_freezeE [crel_elims]: |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
337 |
assumes "crel (freeze a) h h' r" |
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
338 |
obtains "h' = h" "r = get_array a h" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
339 |
using assms by (rule crelE) (simp add: execute_simps) |
37771
1bec64044b5e
spelt out relational framework in a consistent way
haftmann
parents:
37758
diff
changeset
|
340 |
|
26170 | 341 |
lemma upd_return: |
342 |
"upd i x a \<guillemotright> return a = upd i x a" |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
343 |
by (rule Heap_eqI) (simp add: bind_def guard_def upd_def execute_simps) |
26170 | 344 |
|
37752 | 345 |
lemma array_make: |
346 |
"new n x = make n (\<lambda>_. x)" |
|
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
347 |
by (rule Heap_eqI) (simp add: map_replicate_trivial execute_simps) |
26170 | 348 |
|
37752 | 349 |
lemma array_of_list_make: |
350 |
"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
|
351 |
by (rule Heap_eqI) (simp add: map_nth execute_simps) |
26170 | 352 |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
353 |
hide_const (open) present (*get*) set alloc length update noteq new of_list make len nth upd map_entry swap freeze |
26170 | 354 |
|
26182 | 355 |
|
356 |
subsection {* Code generator setup *} |
|
357 |
||
358 |
subsubsection {* Logical intermediate layer *} |
|
359 |
||
360 |
definition new' where |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
361 |
[code del]: "new' = Array.new o Code_Numeral.nat_of" |
37752 | 362 |
|
28562 | 363 |
lemma [code]: |
37752 | 364 |
"Array.new = new' o Code_Numeral.of_nat" |
26182 | 365 |
by (simp add: new'_def o_def) |
366 |
||
367 |
definition of_list' where |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
368 |
[code del]: "of_list' i xs = Array.of_list (take (Code_Numeral.nat_of i) xs)" |
37752 | 369 |
|
28562 | 370 |
lemma [code]: |
37752 | 371 |
"Array.of_list xs = of_list' (Code_Numeral.of_nat (List.length xs)) xs" |
26182 | 372 |
by (simp add: of_list'_def) |
373 |
||
374 |
definition make' where |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
375 |
[code del]: "make' i f = Array.make (Code_Numeral.nat_of i) (f o Code_Numeral.of_nat)" |
37752 | 376 |
|
28562 | 377 |
lemma [code]: |
37752 | 378 |
"Array.make n f = make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)" |
26182 | 379 |
by (simp add: make'_def o_def) |
380 |
||
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
381 |
definition len' where |
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
382 |
[code del]: "len' a = Array.len a \<guillemotright>= (\<lambda>n. return (Code_Numeral.of_nat n))" |
37752 | 383 |
|
28562 | 384 |
lemma [code]: |
37752 | 385 |
"Array.len a = len' a \<guillemotright>= (\<lambda>i. return (Code_Numeral.nat_of i))" |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
386 |
by (simp add: len'_def) |
26182 | 387 |
|
388 |
definition nth' where |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
389 |
[code del]: "nth' a = Array.nth a o Code_Numeral.nat_of" |
37752 | 390 |
|
28562 | 391 |
lemma [code]: |
37752 | 392 |
"Array.nth a n = nth' a (Code_Numeral.of_nat n)" |
26182 | 393 |
by (simp add: nth'_def) |
394 |
||
395 |
definition upd' where |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
396 |
[code del]: "upd' a i x = Array.upd (Code_Numeral.nat_of i) x a \<guillemotright> return ()" |
37752 | 397 |
|
28562 | 398 |
lemma [code]: |
37752 | 399 |
"Array.upd i x a = upd' a (Code_Numeral.of_nat i) x \<guillemotright> return a" |
37709 | 400 |
by (simp add: upd'_def upd_return) |
26182 | 401 |
|
37752 | 402 |
lemma [code]: |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
403 |
"Array.map_entry i f a = do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
404 |
x \<leftarrow> Array.nth a i; |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
405 |
Array.upd i (f x) a |
37792 | 406 |
}" |
37758 | 407 |
by (rule Heap_eqI) (simp add: bind_def guard_def map_entry_def execute_simps) |
26182 | 408 |
|
37752 | 409 |
lemma [code]: |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
410 |
"Array.swap i x a = do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
411 |
y \<leftarrow> Array.nth a i; |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
412 |
Array.upd i x a; |
37752 | 413 |
return y |
37792 | 414 |
}" |
37758 | 415 |
by (rule Heap_eqI) (simp add: bind_def guard_def swap_def execute_simps) |
37752 | 416 |
|
417 |
lemma [code]: |
|
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
418 |
"Array.freeze a = do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
419 |
n \<leftarrow> Array.len a; |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
420 |
Heap_Monad.fold_map (\<lambda>i. Array.nth a i) [0..<n] |
37792 | 421 |
}" |
37752 | 422 |
proof (rule Heap_eqI) |
423 |
fix h |
|
424 |
have *: "List.map |
|
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
425 |
(\<lambda>x. fst (the (if x < Array.length h a |
37752 | 426 |
then Some (get_array a h ! x, h) else None))) |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
427 |
[0..<Array.length h a] = |
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
428 |
List.map (List.nth (get_array a h)) [0..<Array.length h a]" |
37752 | 429 |
by simp |
37804
0145e59c1f6c
qualified names for (almost) all array operations
haftmann
parents:
37803
diff
changeset
|
430 |
have "execute (Heap_Monad.fold_map (Array.nth a) [0..<Array.length h a]) h = |
37752 | 431 |
Some (get_array a h, 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
|
432 |
apply (subst execute_fold_map_unchanged_heap) |
37752 | 433 |
apply (simp_all add: nth_def guard_def *) |
434 |
apply (simp add: length_def map_nth) |
|
435 |
done |
|
37792 | 436 |
then have "execute (do { |
37798
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
437 |
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
|
438 |
Heap_Monad.fold_map (Array.nth a) [0..<n] |
37792 | 439 |
}) h = Some (get_array a h, h)" |
37787
30dc3abf4a58
theorem collections do not contain default rules any longer
haftmann
parents:
37771
diff
changeset
|
440 |
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
|
441 |
then show "execute (Array.freeze a) h = execute (do { |
0b0570445a2a
proper merge of operation changes and generic do-syntax
haftmann
parents:
37797
diff
changeset
|
442 |
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
|
443 |
Heap_Monad.fold_map (Array.nth a) [0..<n] |
37792 | 444 |
}) h" by (simp add: execute_simps) |
37752 | 445 |
qed |
446 |
||
447 |
hide_const (open) new' of_list' make' len' nth' upd' |
|
448 |
||
449 |
||
450 |
text {* SML *} |
|
26182 | 451 |
|
452 |
code_type array (SML "_/ array") |
|
453 |
code_const Array (SML "raise/ (Fail/ \"bare Array\")") |
|
26752 | 454 |
code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))") |
35846 | 455 |
code_const Array.of_list' (SML "(fn/ ()/ =>/ Array.fromList/ _)") |
26752 | 456 |
code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))") |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
457 |
code_const Array.len' (SML "(fn/ ()/ =>/ Array.length/ _)") |
26752 | 458 |
code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))") |
459 |
code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))") |
|
26182 | 460 |
|
461 |
code_reserved SML Array |
|
462 |
||
463 |
||
37752 | 464 |
text {* OCaml *} |
26182 | 465 |
|
466 |
code_type array (OCaml "_/ array") |
|
467 |
code_const Array (OCaml "failwith/ \"bare Array\"") |
|
32580 | 468 |
code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ (Big'_int.int'_of'_big'_int/ _)/ _)") |
35846 | 469 |
code_const Array.of_list' (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)") |
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
470 |
code_const Array.len' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))") |
32580 | 471 |
code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ (Big'_int.int'_of'_big'_int/ _))") |
472 |
code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ (Big'_int.int'_of'_big'_int/ _)/ _)") |
|
26182 | 473 |
|
474 |
code_reserved OCaml Array |
|
475 |
||
476 |
||
37752 | 477 |
text {* Haskell *} |
26182 | 478 |
|
29793 | 479 |
code_type array (Haskell "Heap.STArray/ Heap.RealWorld/ _") |
26182 | 480 |
code_const Array (Haskell "error/ \"bare Array\"") |
29793 | 481 |
code_const Array.new' (Haskell "Heap.newArray/ (0,/ _)") |
482 |
code_const Array.of_list' (Haskell "Heap.newListArray/ (0,/ _)") |
|
37719
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
haftmann
parents:
37716
diff
changeset
|
483 |
code_const Array.len' (Haskell "Heap.lengthArray") |
29793 | 484 |
code_const Array.nth' (Haskell "Heap.readArray") |
485 |
code_const Array.upd' (Haskell "Heap.writeArray") |
|
26182 | 486 |
|
26170 | 487 |
end |