| author | wenzelm | 
| Tue, 19 Sep 2006 23:01:52 +0200 | |
| changeset 20618 | 3f763be47c2f | 
| parent 20503 | 503ac4c5ef91 | 
| child 21404 | eb85850d3eb7 | 
| permissions | -rw-r--r-- | 
| 19234 | 1  | 
(* Title: HOL/Library/Library.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Norbert Schirmer, Tobias Nipkow, Martin Wildmoser  | 
|
4  | 
*)  | 
|
5  | 
||
6  | 
header {* Map operations implemented on association lists*}
 | 
|
7  | 
||
8  | 
theory AssocList  | 
|
9  | 
imports Map  | 
|
10  | 
||
11  | 
begin  | 
|
12  | 
||
13  | 
text {* The operations preserve distinctness of keys and 
 | 
|
| 
19332
 
bb71a64e1263
added map_val, superseding map_at and substitute
 
schirmer 
parents: 
19323 
diff
changeset
 | 
14  | 
        function @{term "clearjunk"} distributes over them. Since 
 | 
| 
 
bb71a64e1263
added map_val, superseding map_at and substitute
 
schirmer 
parents: 
19323 
diff
changeset
 | 
15  | 
        @{term clearjunk} enforces distinctness of keys it can be used
 | 
| 
 
bb71a64e1263
added map_val, superseding map_at and substitute
 
schirmer 
parents: 
19323 
diff
changeset
 | 
16  | 
to establish the invariant, e.g. for inductive proofs.*}  | 
| 19234 | 17  | 
consts  | 
18  | 
  delete     :: "'key \<Rightarrow> ('key * 'val)list \<Rightarrow>  ('key * 'val)list"
 | 
|
19  | 
  update     :: "'key \<Rightarrow> 'val \<Rightarrow> ('key * 'val)list \<Rightarrow>  ('key * 'val)list"
 | 
|
20  | 
  updates    :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key * 'val)list \<Rightarrow>  ('key * 'val)list"
 | 
|
21  | 
  merge      :: "('key * 'val)list \<Rightarrow> ('key * 'val)list \<Rightarrow> ('key * 'val)list"
 | 
|
22  | 
  compose    :: "('key * 'a)list \<Rightarrow> ('a * 'b)list \<Rightarrow> ('key * 'b)list"
 | 
|
23  | 
  restrict   :: "('key set) \<Rightarrow> ('key * 'val)list \<Rightarrow> ('key * 'val)list"
 | 
|
| 19333 | 24  | 
  map_ran    :: "('key \<Rightarrow> 'val \<Rightarrow> 'val) \<Rightarrow> ('key * 'val)list \<Rightarrow> ('key * 'val)list"
 | 
| 19234 | 25  | 
|
26  | 
  clearjunk  :: "('key * 'val)list \<Rightarrow> ('key * 'val)list"
 | 
|
27  | 
||
| 19323 | 28  | 
|
| 19234 | 29  | 
defs  | 
30  | 
delete_def: "delete k \<equiv> filter (\<lambda>p. fst p \<noteq> k)"  | 
|
31  | 
||
32  | 
primrec  | 
|
33  | 
"update k v [] = [(k,v)]"  | 
|
34  | 
"update k v (p#ps) = (if fst p = k then (k,v)#ps else p # update k v ps)"  | 
|
35  | 
primrec  | 
|
36  | 
"updates [] vs al = al"  | 
|
37  | 
"updates (k#ks) vs al = (case vs of [] \<Rightarrow> al  | 
|
38  | 
| (v#vs') \<Rightarrow> updates ks vs' (update k v al))"  | 
|
39  | 
primrec  | 
|
| 19323 | 40  | 
"merge xs [] = xs"  | 
41  | 
"merge xs (p#ps) = update (fst p) (snd p) (merge xs ps)"  | 
|
42  | 
||
43  | 
primrec  | 
|
| 19333 | 44  | 
"map_ran f [] = []"  | 
45  | 
"map_ran f (p#ps) = (fst p, f (fst p) (snd p))#map_ran f ps"  | 
|
| 
19332
 
bb71a64e1263
added map_val, superseding map_at and substitute
 
schirmer 
parents: 
19323 
diff
changeset
 | 
46  | 
|
| 19234 | 47  | 
|
48  | 
lemma length_delete_le: "length (delete k al) \<le> length al"  | 
|
49  | 
proof (induct al)  | 
|
50  | 
case Nil thus ?case by (simp add: delete_def)  | 
|
51  | 
next  | 
|
52  | 
case (Cons a al)  | 
|
53  | 
note length_filter_le [of "\<lambda>p. fst p \<noteq> fst a" al]  | 
|
54  | 
also have "\<And>n. n \<le> Suc n"  | 
|
55  | 
by simp  | 
|
56  | 
finally have "length [p\<in>al . fst p \<noteq> fst a] \<le> Suc (length al)" .  | 
|
57  | 
with Cons show ?case  | 
|
58  | 
by (auto simp add: delete_def)  | 
|
59  | 
qed  | 
|
60  | 
||
61  | 
lemma compose_hint: "length (delete k al) < Suc (length al)"  | 
|
62  | 
proof -  | 
|
63  | 
note length_delete_le  | 
|
64  | 
also have "\<And>n. n < Suc n"  | 
|
65  | 
by simp  | 
|
66  | 
finally show ?thesis .  | 
|
67  | 
qed  | 
|
68  | 
||
69  | 
recdef compose "measure size"  | 
|
70  | 
"compose [] = (\<lambda>ys. [])"  | 
|
71  | 
"compose (x#xs) = (\<lambda>ys. (case (map_of ys (snd x)) of  | 
|
72  | 
None \<Rightarrow> compose (delete (fst x) xs) ys  | 
|
73  | 
| Some v \<Rightarrow> (fst x,v)#compose xs ys))"  | 
|
74  | 
(hints intro: compose_hint)  | 
|
75  | 
||
76  | 
defs  | 
|
77  | 
restrict_def: "restrict A \<equiv> filter (\<lambda>(k,v). k \<in> A)"  | 
|
78  | 
||
79  | 
recdef clearjunk "measure size"  | 
|
80  | 
"clearjunk [] = []"  | 
|
81  | 
"clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)"  | 
|
82  | 
(hints intro: compose_hint)  | 
|
83  | 
||
84  | 
||
85  | 
(* ******************************************************************************** *)  | 
|
86  | 
subsection {* Lookup *}
 | 
|
87  | 
(* ******************************************************************************** *)  | 
|
88  | 
||
89  | 
lemma lookup_simps:  | 
|
90  | 
"map_of [] k = None"  | 
|
91  | 
"map_of (p#ps) k = (if fst p = k then Some (snd p) else map_of ps k)"  | 
|
92  | 
by simp_all  | 
|
93  | 
||
94  | 
(* ******************************************************************************** *)  | 
|
95  | 
subsection {* @{const delete} *}
 | 
|
96  | 
(* ******************************************************************************** *)  | 
|
97  | 
||
98  | 
lemma delete_simps [simp]:  | 
|
99  | 
"delete k [] = []"  | 
|
100  | 
"delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)"  | 
|
101  | 
by (simp_all add: delete_def)  | 
|
102  | 
||
103  | 
lemma delete_id[simp]: "k \<notin> fst ` set al \<Longrightarrow> delete k al = al"  | 
|
104  | 
by(induct al, auto)  | 
|
105  | 
||
106  | 
lemma delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'"  | 
|
107  | 
by (induct al) auto  | 
|
108  | 
||
109  | 
lemma delete_conv': "map_of (delete k al) = ((map_of al)(k := None))"  | 
|
110  | 
by (rule ext) (rule delete_conv)  | 
|
111  | 
||
112  | 
lemma delete_idem: "delete k (delete k al) = delete k al"  | 
|
113  | 
by (induct al) auto  | 
|
114  | 
||
115  | 
lemma map_of_delete[simp]:  | 
|
116  | 
"k' \<noteq> k \<Longrightarrow> map_of (delete k al) k' = map_of al k'"  | 
|
117  | 
by(induct al, auto)  | 
|
118  | 
||
119  | 
lemma delete_notin_dom: "k \<notin> fst ` set (delete k al)"  | 
|
120  | 
by (induct al) auto  | 
|
121  | 
||
122  | 
lemma dom_delete_subset: "fst ` set (delete k al) \<subseteq> fst ` set al"  | 
|
123  | 
by (induct al) auto  | 
|
124  | 
||
125  | 
lemma distinct_delete:  | 
|
126  | 
assumes "distinct (map fst al)"  | 
|
127  | 
shows "distinct (map fst (delete k al))"  | 
|
128  | 
using prems  | 
|
129  | 
proof (induct al)  | 
|
130  | 
case Nil thus ?case by simp  | 
|
131  | 
next  | 
|
132  | 
case (Cons a al)  | 
|
133  | 
from Cons.prems obtain  | 
|
134  | 
a_notin_al: "fst a \<notin> fst ` set al" and  | 
|
135  | 
dist_al: "distinct (map fst al)"  | 
|
136  | 
by auto  | 
|
137  | 
show ?case  | 
|
138  | 
proof (cases "fst a = k")  | 
|
139  | 
case True  | 
|
140  | 
from True dist_al show ?thesis by simp  | 
|
141  | 
next  | 
|
142  | 
case False  | 
|
143  | 
from dist_al  | 
|
144  | 
have "distinct (map fst (delete k al))"  | 
|
145  | 
by (rule Cons.hyps)  | 
|
146  | 
moreover from a_notin_al dom_delete_subset [of k al]  | 
|
147  | 
have "fst a \<notin> fst ` set (delete k al)"  | 
|
148  | 
by blast  | 
|
149  | 
ultimately show ?thesis using False by simp  | 
|
150  | 
qed  | 
|
151  | 
qed  | 
|
152  | 
||
153  | 
lemma delete_twist: "delete x (delete y al) = delete y (delete x al)"  | 
|
154  | 
by (induct al) auto  | 
|
155  | 
||
156  | 
lemma clearjunk_delete: "clearjunk (delete x al) = delete x (clearjunk al)"  | 
|
157  | 
by (induct al rule: clearjunk.induct) (auto simp add: delete_idem delete_twist)  | 
|
158  | 
||
159  | 
(* ******************************************************************************** *)  | 
|
160  | 
subsection {* @{const clearjunk} *}
 | 
|
161  | 
(* ******************************************************************************** *)  | 
|
162  | 
||
163  | 
lemma insert_fst_filter:  | 
|
164  | 
  "insert a(fst ` {x \<in> set ps. fst x \<noteq> a}) = insert a (fst ` set ps)"
 | 
|
165  | 
by (induct ps) auto  | 
|
166  | 
||
167  | 
lemma dom_clearjunk: "fst ` set (clearjunk al) = fst ` set al"  | 
|
168  | 
by (induct al rule: clearjunk.induct) (simp_all add: insert_fst_filter delete_def)  | 
|
169  | 
||
170  | 
lemma notin_filter_fst: "a \<notin> fst ` {x \<in> set ps. fst x \<noteq> a}"
 | 
|
171  | 
by (induct ps) auto  | 
|
172  | 
||
173  | 
lemma distinct_clearjunk [simp]: "distinct (map fst (clearjunk al))"  | 
|
174  | 
by (induct al rule: clearjunk.induct)  | 
|
175  | 
(simp_all add: dom_clearjunk notin_filter_fst delete_def)  | 
|
176  | 
||
177  | 
lemma map_of_filter: "k \<noteq> a \<Longrightarrow> map_of [q\<in>ps . fst q \<noteq> a] k = map_of ps k"  | 
|
178  | 
by (induct ps) auto  | 
|
179  | 
||
180  | 
lemma map_of_clearjunk: "map_of (clearjunk al) = map_of al"  | 
|
181  | 
apply (rule ext)  | 
|
182  | 
apply (induct al rule: clearjunk.induct)  | 
|
183  | 
apply simp  | 
|
184  | 
apply (simp add: map_of_filter)  | 
|
185  | 
done  | 
|
186  | 
||
187  | 
lemma length_clearjunk: "length (clearjunk al) \<le> length al"  | 
|
188  | 
proof (induct al rule: clearjunk.induct [case_names Nil Cons])  | 
|
189  | 
case Nil thus ?case by simp  | 
|
190  | 
next  | 
|
191  | 
case (Cons k v ps)  | 
|
192  | 
from Cons have "length (clearjunk [q\<in>ps . fst q \<noteq> k]) \<le> length [q\<in>ps . fst q \<noteq> k]"  | 
|
193  | 
by (simp add: delete_def)  | 
|
194  | 
also have "\<dots> \<le> length ps"  | 
|
195  | 
by simp  | 
|
196  | 
finally show ?case  | 
|
197  | 
by (simp add: delete_def)  | 
|
198  | 
qed  | 
|
199  | 
||
200  | 
lemma notin_fst_filter: "a \<notin> fst ` set ps \<Longrightarrow> [q\<in>ps . fst q \<noteq> a] = ps"  | 
|
201  | 
by (induct ps) auto  | 
|
202  | 
||
203  | 
lemma distinct_clearjunk_id [simp]: "distinct (map fst al) \<Longrightarrow> clearjunk al = al"  | 
|
204  | 
by (induct al rule: clearjunk.induct) (auto simp add: notin_fst_filter)  | 
|
205  | 
||
206  | 
lemma clearjunk_idem: "clearjunk (clearjunk al) = clearjunk al"  | 
|
207  | 
by simp  | 
|
208  | 
||
209  | 
(* ******************************************************************************** *)  | 
|
210  | 
subsection {* @{const dom} and @{term "ran"} *}
 | 
|
211  | 
(* ******************************************************************************** *)  | 
|
212  | 
||
213  | 
lemma dom_map_of': "fst ` set al = dom (map_of al)"  | 
|
214  | 
by (induct al) auto  | 
|
215  | 
||
216  | 
lemmas dom_map_of = dom_map_of' [symmetric]  | 
|
217  | 
||
218  | 
lemma ran_clearjunk: "ran (map_of (clearjunk al)) = ran (map_of al)"  | 
|
219  | 
by (simp add: map_of_clearjunk)  | 
|
220  | 
||
221  | 
lemma ran_distinct:  | 
|
222  | 
assumes dist: "distinct (map fst al)"  | 
|
223  | 
shows "ran (map_of al) = snd ` set al"  | 
|
224  | 
using dist  | 
|
225  | 
proof (induct al)  | 
|
226  | 
case Nil  | 
|
227  | 
thus ?case by simp  | 
|
228  | 
next  | 
|
229  | 
case (Cons a al)  | 
|
230  | 
hence hyp: "snd ` set al = ran (map_of al)"  | 
|
231  | 
by simp  | 
|
232  | 
||
233  | 
  have "ran (map_of (a # al)) = {snd a} \<union> ran (map_of al)"
 | 
|
234  | 
proof  | 
|
235  | 
    show "ran (map_of (a # al)) \<subseteq> {snd a} \<union> ran (map_of al)"
 | 
|
236  | 
proof  | 
|
237  | 
fix v  | 
|
238  | 
assume "v \<in> ran (map_of (a#al))"  | 
|
239  | 
then obtain x where "map_of (a#al) x = Some v"  | 
|
240  | 
by (auto simp add: ran_def)  | 
|
241  | 
      then show "v \<in> {snd a} \<union> ran (map_of al)"
 | 
|
242  | 
by (auto split: split_if_asm simp add: ran_def)  | 
|
243  | 
qed  | 
|
244  | 
next  | 
|
245  | 
    show "{snd a} \<union> ran (map_of al) \<subseteq> ran (map_of (a # al))"
 | 
|
246  | 
proof  | 
|
247  | 
fix v  | 
|
248  | 
      assume v_in: "v \<in> {snd a} \<union> ran (map_of al)"
 | 
|
249  | 
show "v \<in> ran (map_of (a#al))"  | 
|
250  | 
proof (cases "v=snd a")  | 
|
251  | 
case True  | 
|
252  | 
with v_in show ?thesis  | 
|
253  | 
by (auto simp add: ran_def)  | 
|
254  | 
next  | 
|
255  | 
case False  | 
|
256  | 
with v_in have "v \<in> ran (map_of al)" by auto  | 
|
257  | 
then obtain x where al_x: "map_of al x = Some v"  | 
|
258  | 
by (auto simp add: ran_def)  | 
|
259  | 
from map_of_SomeD [OF this]  | 
|
260  | 
have "x \<in> fst ` set al"  | 
|
261  | 
by (force simp add: image_def)  | 
|
262  | 
with Cons.prems have "x\<noteq>fst a"  | 
|
263  | 
by - (rule ccontr,simp)  | 
|
264  | 
with al_x  | 
|
265  | 
show ?thesis  | 
|
266  | 
by (auto simp add: ran_def)  | 
|
267  | 
qed  | 
|
268  | 
qed  | 
|
269  | 
qed  | 
|
270  | 
with hyp show ?case  | 
|
271  | 
by (simp only:) auto  | 
|
272  | 
qed  | 
|
273  | 
||
274  | 
lemma ran_map_of: "ran (map_of al) = snd ` set (clearjunk al)"  | 
|
275  | 
proof -  | 
|
276  | 
have "ran (map_of al) = ran (map_of (clearjunk al))"  | 
|
277  | 
by (simp add: ran_clearjunk)  | 
|
278  | 
also have "\<dots> = snd ` set (clearjunk al)"  | 
|
279  | 
by (simp add: ran_distinct)  | 
|
280  | 
finally show ?thesis .  | 
|
281  | 
qed  | 
|
282  | 
||
283  | 
(* ******************************************************************************** *)  | 
|
284  | 
subsection {* @{const update} *}
 | 
|
285  | 
(* ******************************************************************************** *)  | 
|
286  | 
||
287  | 
lemma update_conv: "map_of (update k v al) k' = ((map_of al)(k\<mapsto>v)) k'"  | 
|
288  | 
by (induct al) auto  | 
|
289  | 
||
290  | 
lemma update_conv': "map_of (update k v al) = ((map_of al)(k\<mapsto>v))"  | 
|
291  | 
by (rule ext) (rule update_conv)  | 
|
292  | 
||
293  | 
lemma dom_update: "fst ` set (update k v al) = {k} \<union> fst ` set al"
 | 
|
294  | 
by (induct al) auto  | 
|
295  | 
||
296  | 
lemma distinct_update:  | 
|
297  | 
assumes "distinct (map fst al)"  | 
|
298  | 
shows "distinct (map fst (update k v al))"  | 
|
299  | 
using prems  | 
|
300  | 
proof (induct al)  | 
|
301  | 
case Nil thus ?case by simp  | 
|
302  | 
next  | 
|
303  | 
case (Cons a al)  | 
|
304  | 
from Cons.prems obtain  | 
|
305  | 
a_notin_al: "fst a \<notin> fst ` set al" and  | 
|
306  | 
dist_al: "distinct (map fst al)"  | 
|
307  | 
by auto  | 
|
308  | 
show ?case  | 
|
309  | 
proof (cases "fst a = k")  | 
|
310  | 
case True  | 
|
311  | 
from True dist_al a_notin_al show ?thesis by simp  | 
|
312  | 
next  | 
|
313  | 
case False  | 
|
314  | 
from dist_al  | 
|
315  | 
have "distinct (map fst (update k v al))"  | 
|
316  | 
by (rule Cons.hyps)  | 
|
317  | 
with False a_notin_al show ?thesis by (simp add: dom_update)  | 
|
318  | 
qed  | 
|
319  | 
qed  | 
|
320  | 
||
321  | 
lemma update_filter:  | 
|
322  | 
"a\<noteq>k \<Longrightarrow> update k v [q\<in>ps . fst q \<noteq> a] = [q\<in>update k v ps . fst q \<noteq> a]"  | 
|
323  | 
by (induct ps) auto  | 
|
324  | 
||
325  | 
lemma clearjunk_update: "clearjunk (update k v al) = update k v (clearjunk al)"  | 
|
326  | 
by (induct al rule: clearjunk.induct) (auto simp add: update_filter delete_def)  | 
|
327  | 
||
328  | 
lemma update_triv: "map_of al k = Some v \<Longrightarrow> update k v al = al"  | 
|
329  | 
by (induct al) auto  | 
|
330  | 
||
331  | 
lemma update_nonempty [simp]: "update k v al \<noteq> []"  | 
|
332  | 
by (induct al) auto  | 
|
333  | 
||
334  | 
lemma update_eqD: "update k v al = update k v' al' \<Longrightarrow> v=v'"  | 
|
| 20503 | 335  | 
proof (induct al arbitrary: al')  | 
| 19234 | 336  | 
case Nil thus ?case  | 
337  | 
by (cases al') (auto split: split_if_asm)  | 
|
338  | 
next  | 
|
339  | 
case Cons thus ?case  | 
|
340  | 
by (cases al') (auto split: split_if_asm)  | 
|
341  | 
qed  | 
|
342  | 
||
343  | 
lemma update_last [simp]: "update k v (update k v' al) = update k v al"  | 
|
344  | 
by (induct al) auto  | 
|
345  | 
||
346  | 
text {* Note that the lists are not necessarily the same:
 | 
|
347  | 
        @{term "update k v (update k' v' []) = [(k',v'),(k,v)]"} and 
 | 
|
348  | 
        @{term "update k' v' (update k v []) = [(k,v),(k',v')]"}.*}
 | 
|
349  | 
lemma update_swap: "k\<noteq>k'  | 
|
350  | 
\<Longrightarrow> map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))"  | 
|
351  | 
by (auto simp add: update_conv' intro: ext)  | 
|
352  | 
||
353  | 
lemma update_Some_unfold:  | 
|
354  | 
"(map_of (update k v al) x = Some y) =  | 
|
355  | 
(x = k \<and> v = y \<or> x \<noteq> k \<and> map_of al x = Some y)"  | 
|
356  | 
by (simp add: update_conv' map_upd_Some_unfold)  | 
|
357  | 
||
358  | 
lemma image_update[simp]: "x \<notin> A \<Longrightarrow> map_of (update x y al) ` A = map_of al ` A"  | 
|
359  | 
by (simp add: update_conv' image_map_upd)  | 
|
360  | 
||
361  | 
||
362  | 
(* ******************************************************************************** *)  | 
|
363  | 
subsection {* @{const updates} *}
 | 
|
364  | 
(* ******************************************************************************** *)  | 
|
365  | 
||
366  | 
lemma updates_conv: "map_of (updates ks vs al) k = ((map_of al)(ks[\<mapsto>]vs)) k"  | 
|
| 20503 | 367  | 
proof (induct ks arbitrary: vs al)  | 
| 19234 | 368  | 
case Nil  | 
369  | 
thus ?case by simp  | 
|
370  | 
next  | 
|
371  | 
case (Cons k ks)  | 
|
372  | 
show ?case  | 
|
373  | 
proof (cases vs)  | 
|
374  | 
case Nil  | 
|
375  | 
with Cons show ?thesis by simp  | 
|
376  | 
next  | 
|
377  | 
case (Cons k ks')  | 
|
378  | 
with Cons.hyps show ?thesis  | 
|
379  | 
by (simp add: update_conv fun_upd_def)  | 
|
380  | 
qed  | 
|
381  | 
qed  | 
|
382  | 
||
383  | 
lemma updates_conv': "map_of (updates ks vs al) = ((map_of al)(ks[\<mapsto>]vs))"  | 
|
384  | 
by (rule ext) (rule updates_conv)  | 
|
385  | 
||
386  | 
lemma distinct_updates:  | 
|
387  | 
assumes "distinct (map fst al)"  | 
|
388  | 
shows "distinct (map fst (updates ks vs al))"  | 
|
389  | 
using prems  | 
|
| 20503 | 390  | 
by (induct ks arbitrary: vs al) (auto simp add: distinct_update split: list.splits)  | 
| 19234 | 391  | 
|
392  | 
lemma clearjunk_updates:  | 
|
393  | 
"clearjunk (updates ks vs al) = updates ks vs (clearjunk al)"  | 
|
| 20503 | 394  | 
by (induct ks arbitrary: vs al) (auto simp add: clearjunk_update split: list.splits)  | 
| 19234 | 395  | 
|
396  | 
lemma updates_empty[simp]: "updates vs [] al = al"  | 
|
397  | 
by (induct vs) auto  | 
|
398  | 
||
399  | 
lemma updates_Cons: "updates (k#ks) (v#vs) al = updates ks vs (update k v al)"  | 
|
400  | 
by simp  | 
|
401  | 
||
402  | 
lemma updates_append1[simp]: "size ks < size vs \<Longrightarrow>  | 
|
403  | 
updates (ks@[k]) vs al = update k (vs!size ks) (updates ks vs al)"  | 
|
| 20503 | 404  | 
by (induct ks arbitrary: vs al) (auto split: list.splits)  | 
| 19234 | 405  | 
|
406  | 
lemma updates_list_update_drop[simp]:  | 
|
407  | 
"\<lbrakk>size ks \<le> i; i < size vs\<rbrakk>  | 
|
408  | 
\<Longrightarrow> updates ks (vs[i:=v]) al = updates ks vs al"  | 
|
| 20503 | 409  | 
by (induct ks arbitrary: al vs i) (auto split:list.splits nat.splits)  | 
| 19234 | 410  | 
|
411  | 
lemma update_updates_conv_if: "  | 
|
412  | 
map_of (updates xs ys (update x y al)) =  | 
|
413  | 
map_of (if x \<in> set(take (length ys) xs) then updates xs ys al  | 
|
414  | 
else (update x y (updates xs ys al)))"  | 
|
415  | 
by (simp add: updates_conv' update_conv' map_upd_upds_conv_if)  | 
|
416  | 
||
417  | 
lemma updates_twist [simp]:  | 
|
418  | 
"k \<notin> set ks \<Longrightarrow>  | 
|
419  | 
map_of (updates ks vs (update k v al)) = map_of (update k v (updates ks vs al))"  | 
|
420  | 
by (simp add: updates_conv' update_conv' map_upds_twist)  | 
|
421  | 
||
422  | 
lemma updates_apply_notin[simp]:  | 
|
423  | 
"k \<notin> set ks ==> map_of (updates ks vs al) k = map_of al k"  | 
|
424  | 
by (simp add: updates_conv)  | 
|
425  | 
||
426  | 
lemma updates_append_drop[simp]:  | 
|
427  | 
"size xs = size ys \<Longrightarrow> updates (xs@zs) ys al = updates xs ys al"  | 
|
| 20503 | 428  | 
by (induct xs arbitrary: ys al) (auto split: list.splits)  | 
| 19234 | 429  | 
|
430  | 
lemma updates_append2_drop[simp]:  | 
|
431  | 
"size xs = size ys \<Longrightarrow> updates xs (ys@zs) al = updates xs ys al"  | 
|
| 20503 | 432  | 
by (induct xs arbitrary: ys al) (auto split: list.splits)  | 
| 19234 | 433  | 
|
434  | 
(* ******************************************************************************** *)  | 
|
| 19333 | 435  | 
subsection {* @{const map_ran} *}
 | 
| 19234 | 436  | 
(* ******************************************************************************** *)  | 
437  | 
||
| 19333 | 438  | 
lemma map_ran_conv: "map_of (map_ran f al) k = option_map (f k) (map_of al k)"  | 
| 19234 | 439  | 
by (induct al) auto  | 
440  | 
||
| 19333 | 441  | 
lemma dom_map_ran: "fst ` set (map_ran f al) = fst ` set al"  | 
| 19234 | 442  | 
by (induct al) auto  | 
443  | 
||
| 19333 | 444  | 
lemma distinct_map_ran: "distinct (map fst al) \<Longrightarrow> distinct (map fst (map_ran f al))"  | 
445  | 
by (induct al) (auto simp add: dom_map_ran)  | 
|
| 19234 | 446  | 
|
| 19333 | 447  | 
lemma map_ran_filter: "map_ran f [p\<in>ps. fst p \<noteq> a] = [p\<in>map_ran f ps. fst p \<noteq> a]"  | 
| 19234 | 448  | 
by (induct ps) auto  | 
449  | 
||
| 19333 | 450  | 
lemma clearjunk_map_ran: "clearjunk (map_ran f al) = map_ran f (clearjunk al)"  | 
451  | 
by (induct al rule: clearjunk.induct) (auto simp add: delete_def map_ran_filter)  | 
|
| 19234 | 452  | 
|
453  | 
(* ******************************************************************************** *)  | 
|
454  | 
subsection {* @{const merge} *}
 | 
|
455  | 
(* ******************************************************************************** *)  | 
|
456  | 
||
457  | 
lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \<union> fst ` set ys"  | 
|
| 20503 | 458  | 
by (induct ys arbitrary: xs) (auto simp add: dom_update)  | 
| 19234 | 459  | 
|
460  | 
lemma distinct_merge:  | 
|
461  | 
assumes "distinct (map fst xs)"  | 
|
462  | 
shows "distinct (map fst (merge xs ys))"  | 
|
463  | 
using prems  | 
|
| 20503 | 464  | 
by (induct ys arbitrary: xs) (auto simp add: dom_merge distinct_update)  | 
| 19234 | 465  | 
|
466  | 
lemma clearjunk_merge:  | 
|
467  | 
"clearjunk (merge xs ys) = merge (clearjunk xs) ys"  | 
|
468  | 
by (induct ys) (auto simp add: clearjunk_update)  | 
|
469  | 
||
470  | 
lemma merge_conv: "map_of (merge xs ys) k = (map_of xs ++ map_of ys) k"  | 
|
471  | 
proof (induct ys)  | 
|
472  | 
case Nil thus ?case by simp  | 
|
473  | 
next  | 
|
474  | 
case (Cons y ys)  | 
|
475  | 
show ?case  | 
|
476  | 
proof (cases "k = fst y")  | 
|
477  | 
case True  | 
|
478  | 
from True show ?thesis  | 
|
479  | 
by (simp add: update_conv)  | 
|
480  | 
next  | 
|
481  | 
case False  | 
|
482  | 
from False show ?thesis  | 
|
483  | 
by (auto simp add: update_conv Cons.hyps map_add_def)  | 
|
484  | 
qed  | 
|
485  | 
qed  | 
|
486  | 
||
487  | 
lemma merge_conv': "map_of (merge xs ys) = (map_of xs ++ map_of ys)"  | 
|
488  | 
by (rule ext) (rule merge_conv)  | 
|
489  | 
||
490  | 
lemma merge_emty: "map_of (merge [] ys) = map_of ys"  | 
|
491  | 
by (simp add: merge_conv')  | 
|
492  | 
||
493  | 
lemma merge_assoc[simp]: "map_of (merge m1 (merge m2 m3)) =  | 
|
494  | 
map_of (merge (merge m1 m2) m3)"  | 
|
495  | 
by (simp add: merge_conv')  | 
|
496  | 
||
497  | 
lemma merge_Some_iff:  | 
|
498  | 
"(map_of (merge m n) k = Some x) =  | 
|
499  | 
(map_of n k = Some x \<or> map_of n k = None \<and> map_of m k = Some x)"  | 
|
500  | 
by (simp add: merge_conv' map_add_Some_iff)  | 
|
501  | 
||
502  | 
lemmas merge_SomeD = merge_Some_iff [THEN iffD1, standard]  | 
|
503  | 
declare merge_SomeD [dest!]  | 
|
504  | 
||
505  | 
lemma merge_find_right[simp]: "map_of n k = Some v \<Longrightarrow> map_of (merge m n) k = Some v"  | 
|
506  | 
by (simp add: merge_conv')  | 
|
507  | 
||
508  | 
lemma merge_None [iff]:  | 
|
509  | 
"(map_of (merge m n) k = None) = (map_of n k = None \<and> map_of m k = None)"  | 
|
510  | 
by (simp add: merge_conv')  | 
|
511  | 
||
512  | 
lemma merge_upd[simp]:  | 
|
513  | 
"map_of (merge m (update k v n)) = map_of (update k v (merge m n))"  | 
|
514  | 
by (simp add: update_conv' merge_conv')  | 
|
515  | 
||
516  | 
lemma merge_updatess[simp]:  | 
|
517  | 
"map_of (merge m (updates xs ys n)) = map_of (updates xs ys (merge m n))"  | 
|
518  | 
by (simp add: updates_conv' merge_conv')  | 
|
519  | 
||
520  | 
lemma merge_append: "map_of (xs@ys) = map_of (merge ys xs)"  | 
|
521  | 
by (simp add: merge_conv')  | 
|
522  | 
||
523  | 
(* ******************************************************************************** *)  | 
|
524  | 
subsection {* @{const compose} *}
 | 
|
525  | 
(* ******************************************************************************** *)  | 
|
526  | 
||
527  | 
lemma compose_induct [case_names Nil Cons]:  | 
|
528  | 
assumes Nil: "P [] ys"  | 
|
529  | 
assumes Cons: "\<And>x xs.  | 
|
530  | 
\<lbrakk>\<And>v. map_of ys (snd x) = Some v \<Longrightarrow> P xs ys;  | 
|
531  | 
map_of ys (snd x) = None \<Longrightarrow> P (delete (fst x) xs) ys\<rbrakk>  | 
|
532  | 
\<Longrightarrow> P (x # xs) ys"  | 
|
533  | 
shows "P xs ys"  | 
|
534  | 
apply (rule compose.induct [where ?P="\<lambda>xs. P xs ys"])  | 
|
535  | 
apply (rule Nil)  | 
|
536  | 
apply (rule Cons)  | 
|
537  | 
apply (erule allE, erule allE, erule impE, assumption,assumption)  | 
|
538  | 
apply (erule allE, erule impE,assumption,assumption)  | 
|
539  | 
done  | 
|
540  | 
||
541  | 
lemma compose_first_None [simp]:  | 
|
542  | 
assumes "map_of xs k = None"  | 
|
543  | 
shows "map_of (compose xs ys) k = None"  | 
|
544  | 
using prems  | 
|
545  | 
by (induct xs ys rule: compose_induct) (auto split: option.splits split_if_asm)  | 
|
546  | 
||
547  | 
||
548  | 
lemma compose_conv:  | 
|
549  | 
shows "map_of (compose xs ys) k = (map_of ys \<circ>\<^sub>m map_of xs) k"  | 
|
550  | 
proof (induct xs ys rule: compose_induct )  | 
|
551  | 
case Nil thus ?case by simp  | 
|
552  | 
next  | 
|
553  | 
case (Cons x xs)  | 
|
554  | 
show ?case  | 
|
555  | 
proof (cases "map_of ys (snd x)")  | 
|
556  | 
case None  | 
|
557  | 
with Cons  | 
|
558  | 
have hyp: "map_of (compose (delete (fst x) xs) ys) k =  | 
|
559  | 
(map_of ys \<circ>\<^sub>m map_of (delete (fst x) xs)) k"  | 
|
560  | 
by simp  | 
|
561  | 
show ?thesis  | 
|
562  | 
proof (cases "fst x = k")  | 
|
563  | 
case True  | 
|
564  | 
from True delete_notin_dom [of k xs]  | 
|
565  | 
have "map_of (delete (fst x) xs) k = None"  | 
|
566  | 
by (simp add: map_of_eq_None_iff)  | 
|
567  | 
with hyp show ?thesis  | 
|
568  | 
using True None  | 
|
569  | 
by simp  | 
|
570  | 
next  | 
|
571  | 
case False  | 
|
572  | 
from False have "map_of (delete (fst x) xs) k = map_of xs k"  | 
|
573  | 
by simp  | 
|
574  | 
with hyp show ?thesis  | 
|
575  | 
using False None  | 
|
576  | 
by (simp add: map_comp_def)  | 
|
577  | 
qed  | 
|
578  | 
next  | 
|
579  | 
case (Some v)  | 
|
580  | 
with Cons  | 
|
581  | 
have "map_of (compose xs ys) k = (map_of ys \<circ>\<^sub>m map_of xs) k"  | 
|
582  | 
by simp  | 
|
583  | 
with Some show ?thesis  | 
|
584  | 
by (auto simp add: map_comp_def)  | 
|
585  | 
qed  | 
|
586  | 
qed  | 
|
587  | 
||
588  | 
lemma compose_conv':  | 
|
589  | 
shows "map_of (compose xs ys) = (map_of ys \<circ>\<^sub>m map_of xs)"  | 
|
590  | 
by (rule ext) (rule compose_conv)  | 
|
591  | 
||
592  | 
lemma compose_first_Some [simp]:  | 
|
593  | 
assumes "map_of xs k = Some v"  | 
|
594  | 
shows "map_of (compose xs ys) k = map_of ys v"  | 
|
595  | 
using prems by (simp add: compose_conv)  | 
|
596  | 
||
597  | 
lemma dom_compose: "fst ` set (compose xs ys) \<subseteq> fst ` set xs"  | 
|
598  | 
proof (induct xs ys rule: compose_induct )  | 
|
599  | 
case Nil thus ?case by simp  | 
|
600  | 
next  | 
|
601  | 
case (Cons x xs)  | 
|
602  | 
show ?case  | 
|
603  | 
proof (cases "map_of ys (snd x)")  | 
|
604  | 
case None  | 
|
605  | 
with Cons.hyps  | 
|
606  | 
have "fst ` set (compose (delete (fst x) xs) ys) \<subseteq> fst ` set (delete (fst x) xs)"  | 
|
607  | 
by simp  | 
|
608  | 
also  | 
|
609  | 
have "\<dots> \<subseteq> fst ` set xs"  | 
|
610  | 
by (rule dom_delete_subset)  | 
|
611  | 
finally show ?thesis  | 
|
612  | 
using None  | 
|
613  | 
by auto  | 
|
614  | 
next  | 
|
615  | 
case (Some v)  | 
|
616  | 
with Cons.hyps  | 
|
617  | 
have "fst ` set (compose xs ys) \<subseteq> fst ` set xs"  | 
|
618  | 
by simp  | 
|
619  | 
with Some show ?thesis  | 
|
620  | 
by auto  | 
|
621  | 
qed  | 
|
622  | 
qed  | 
|
623  | 
||
624  | 
lemma distinct_compose:  | 
|
625  | 
assumes "distinct (map fst xs)"  | 
|
626  | 
shows "distinct (map fst (compose xs ys))"  | 
|
627  | 
using prems  | 
|
628  | 
proof (induct xs ys rule: compose_induct)  | 
|
629  | 
case Nil thus ?case by simp  | 
|
630  | 
next  | 
|
631  | 
case (Cons x xs)  | 
|
632  | 
show ?case  | 
|
633  | 
proof (cases "map_of ys (snd x)")  | 
|
634  | 
case None  | 
|
635  | 
with Cons show ?thesis by simp  | 
|
636  | 
next  | 
|
637  | 
case (Some v)  | 
|
638  | 
with Cons dom_compose [of xs ys] show ?thesis  | 
|
639  | 
by (auto)  | 
|
640  | 
qed  | 
|
641  | 
qed  | 
|
642  | 
||
643  | 
lemma compose_delete_twist: "(compose (delete k xs) ys) = delete k (compose xs ys)"  | 
|
644  | 
proof (induct xs ys rule: compose_induct)  | 
|
645  | 
case Nil thus ?case by simp  | 
|
646  | 
next  | 
|
647  | 
case (Cons x xs)  | 
|
648  | 
show ?case  | 
|
649  | 
proof (cases "map_of ys (snd x)")  | 
|
650  | 
case None  | 
|
651  | 
with Cons have  | 
|
652  | 
hyp: "compose (delete k (delete (fst x) xs)) ys =  | 
|
653  | 
delete k (compose (delete (fst x) xs) ys)"  | 
|
654  | 
by simp  | 
|
655  | 
show ?thesis  | 
|
656  | 
proof (cases "fst x = k")  | 
|
657  | 
case True  | 
|
658  | 
with None hyp  | 
|
659  | 
show ?thesis  | 
|
660  | 
by (simp add: delete_idem)  | 
|
661  | 
next  | 
|
662  | 
case False  | 
|
663  | 
from None False hyp  | 
|
664  | 
show ?thesis  | 
|
665  | 
by (simp add: delete_twist)  | 
|
666  | 
qed  | 
|
667  | 
next  | 
|
668  | 
case (Some v)  | 
|
669  | 
with Cons have hyp: "compose (delete k xs) ys = delete k (compose xs ys)" by simp  | 
|
670  | 
with Some show ?thesis  | 
|
671  | 
by simp  | 
|
672  | 
qed  | 
|
673  | 
qed  | 
|
674  | 
||
675  | 
lemma compose_clearjunk: "compose xs (clearjunk ys) = compose xs ys"  | 
|
676  | 
by (induct xs ys rule: compose_induct)  | 
|
677  | 
(auto simp add: map_of_clearjunk split: option.splits)  | 
|
678  | 
||
679  | 
lemma clearjunk_compose: "clearjunk (compose xs ys) = compose (clearjunk xs) ys"  | 
|
680  | 
by (induct xs rule: clearjunk.induct)  | 
|
681  | 
(auto split: option.splits simp add: clearjunk_delete delete_idem  | 
|
682  | 
compose_delete_twist)  | 
|
683  | 
||
684  | 
lemma compose_empty [simp]:  | 
|
685  | 
"compose xs [] = []"  | 
|
686  | 
by (induct xs rule: compose_induct [where ys="[]"]) auto  | 
|
687  | 
||
688  | 
||
689  | 
lemma compose_Some_iff:  | 
|
690  | 
"(map_of (compose xs ys) k = Some v) =  | 
|
691  | 
(\<exists>k'. map_of xs k = Some k' \<and> map_of ys k' = Some v)"  | 
|
692  | 
by (simp add: compose_conv map_comp_Some_iff)  | 
|
693  | 
||
694  | 
lemma map_comp_None_iff:  | 
|
695  | 
"(map_of (compose xs ys) k = None) =  | 
|
696  | 
(map_of xs k = None \<or> (\<exists>k'. map_of xs k = Some k' \<and> map_of ys k' = None)) "  | 
|
697  | 
by (simp add: compose_conv map_comp_None_iff)  | 
|
698  | 
||
699  | 
||
700  | 
(* ******************************************************************************** *)  | 
|
701  | 
subsection {* @{const restrict} *}
 | 
|
702  | 
(* ******************************************************************************** *)  | 
|
703  | 
||
704  | 
lemma restrict_simps [simp]:  | 
|
705  | 
"restrict A [] = []"  | 
|
706  | 
"restrict A (p#ps) = (if fst p \<in> A then p#restrict A ps else restrict A ps)"  | 
|
707  | 
by (auto simp add: restrict_def)  | 
|
708  | 
||
709  | 
lemma distinct_restr: "distinct (map fst al) \<Longrightarrow> distinct (map fst (restrict A al))"  | 
|
710  | 
by (induct al) (auto simp add: restrict_def)  | 
|
711  | 
||
712  | 
lemma restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k"  | 
|
713  | 
apply (induct al)  | 
|
714  | 
apply (simp add: restrict_def)  | 
|
715  | 
apply (cases "k\<in>A")  | 
|
716  | 
apply (auto simp add: restrict_def)  | 
|
717  | 
done  | 
|
718  | 
||
719  | 
lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)"  | 
|
720  | 
by (rule ext) (rule restr_conv)  | 
|
721  | 
||
722  | 
lemma restr_empty [simp]:  | 
|
723  | 
  "restrict {} al = []" 
 | 
|
724  | 
"restrict A [] = []"  | 
|
725  | 
by (induct al) (auto simp add: restrict_def)  | 
|
726  | 
||
727  | 
lemma restr_in [simp]: "x \<in> A \<Longrightarrow> map_of (restrict A al) x = map_of al x"  | 
|
728  | 
by (simp add: restr_conv')  | 
|
729  | 
||
730  | 
lemma restr_out [simp]: "x \<notin> A \<Longrightarrow> map_of (restrict A al) x = None"  | 
|
731  | 
by (simp add: restr_conv')  | 
|
732  | 
||
733  | 
lemma dom_restr [simp]: "fst ` set (restrict A al) = fst ` set al \<inter> A"  | 
|
734  | 
by (induct al) (auto simp add: restrict_def)  | 
|
735  | 
||
736  | 
lemma restr_upd_same [simp]: "restrict (-{x}) (update x y al) = restrict (-{x}) al"
 | 
|
737  | 
by (induct al) (auto simp add: restrict_def)  | 
|
738  | 
||
739  | 
lemma restr_restr [simp]: "restrict A (restrict B al) = restrict (A\<inter>B) al"  | 
|
740  | 
by (induct al) (auto simp add: restrict_def)  | 
|
741  | 
||
742  | 
lemma restr_update[simp]:  | 
|
743  | 
"map_of (restrict D (update x y al)) =  | 
|
744  | 
  map_of ((if x \<in> D then (update x y (restrict (D-{x}) al)) else restrict D al))"
 | 
|
745  | 
by (simp add: restr_conv' update_conv')  | 
|
746  | 
||
747  | 
lemma restr_delete [simp]:  | 
|
748  | 
"(delete x (restrict D al)) =  | 
|
749  | 
    (if x\<in> D then restrict (D - {x}) al else restrict D al)"
 | 
|
750  | 
proof (induct al)  | 
|
751  | 
case Nil thus ?case by simp  | 
|
752  | 
next  | 
|
753  | 
case (Cons a al)  | 
|
754  | 
show ?case  | 
|
755  | 
proof (cases "x \<in> D")  | 
|
756  | 
case True  | 
|
757  | 
note x_D = this  | 
|
758  | 
    with Cons have hyp: "delete x (restrict D al) = restrict (D - {x}) al"
 | 
|
759  | 
by simp  | 
|
760  | 
show ?thesis  | 
|
761  | 
proof (cases "fst a = x")  | 
|
762  | 
case True  | 
|
763  | 
from Cons.hyps  | 
|
764  | 
show ?thesis  | 
|
765  | 
using x_D True  | 
|
766  | 
by simp  | 
|
767  | 
next  | 
|
768  | 
case False  | 
|
769  | 
note not_fst_a_x = this  | 
|
770  | 
show ?thesis  | 
|
771  | 
proof (cases "fst a \<in> D")  | 
|
772  | 
case True  | 
|
773  | 
with not_fst_a_x  | 
|
774  | 
have "delete x (restrict D (a#al)) = a#(delete x (restrict D al))"  | 
|
775  | 
by (cases a) (simp add: restrict_def)  | 
|
776  | 
	also from not_fst_a_x True hyp have "\<dots> = restrict (D - {x}) (a # al)"
 | 
|
777  | 
by (cases a) (simp add: restrict_def)  | 
|
778  | 
finally show ?thesis  | 
|
779  | 
using x_D by simp  | 
|
780  | 
next  | 
|
781  | 
case False  | 
|
782  | 
hence "delete x (restrict D (a#al)) = delete x (restrict D al)"  | 
|
783  | 
by (cases a) (simp add: restrict_def)  | 
|
784  | 
moreover from False not_fst_a_x  | 
|
785  | 
	have "restrict (D - {x}) (a # al) = restrict (D - {x}) al"
 | 
|
786  | 
by (cases a) (simp add: restrict_def)  | 
|
787  | 
ultimately  | 
|
788  | 
show ?thesis using x_D hyp by simp  | 
|
789  | 
qed  | 
|
790  | 
qed  | 
|
791  | 
next  | 
|
792  | 
case False  | 
|
793  | 
from False Cons show ?thesis  | 
|
794  | 
by simp  | 
|
795  | 
qed  | 
|
796  | 
qed  | 
|
797  | 
||
798  | 
lemma update_restr:  | 
|
799  | 
 "map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
 | 
|
800  | 
by (simp add: update_conv' restr_conv') (rule fun_upd_restrict)  | 
|
801  | 
||
802  | 
lemma upate_restr_conv[simp]:  | 
|
803  | 
"x \<in> D \<Longrightarrow>  | 
|
804  | 
 map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
 | 
|
805  | 
by (simp add: update_conv' restr_conv')  | 
|
806  | 
||
807  | 
lemma restr_updates[simp]: "  | 
|
808  | 
\<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>  | 
|
809  | 
\<Longrightarrow> map_of (restrict D (updates xs ys al)) =  | 
|
810  | 
map_of (updates xs ys (restrict (D - set xs) al))"  | 
|
811  | 
by (simp add: updates_conv' restr_conv')  | 
|
812  | 
||
813  | 
lemma restr_delete_twist: "(restrict A (delete a ps)) = delete a (restrict A ps)"  | 
|
814  | 
by (induct ps) auto  | 
|
815  | 
||
816  | 
lemma clearjunk_restrict:  | 
|
817  | 
"clearjunk (restrict A al) = restrict A (clearjunk al)"  | 
|
818  | 
by (induct al rule: clearjunk.induct) (auto simp add: restr_delete_twist)  | 
|
819  | 
||
820  | 
end  |