62130
|
1 |
(* Author: Tobias Nipkow *)
|
|
2 |
|
|
3 |
section "AA Implementation of Maps"
|
|
4 |
|
|
5 |
theory AA_Map
|
|
6 |
imports
|
|
7 |
AA_Set
|
|
8 |
Lookup2
|
|
9 |
begin
|
|
10 |
|
|
11 |
fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
|
|
12 |
"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
|
|
13 |
"update x y (Node lv t1 (a,b) t2) =
|
|
14 |
(case cmp x a of
|
|
15 |
LT \<Rightarrow> split (skew (Node lv (update x y t1) (a,b) t2)) |
|
|
16 |
GT \<Rightarrow> split (skew (Node lv t1 (a,b) (update x y t2))) |
|
|
17 |
EQ \<Rightarrow> Node lv t1 (x,y) t2)"
|
|
18 |
|
|
19 |
fun delete :: "'a::cmp \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
|
|
20 |
"delete _ Leaf = Leaf" |
|
|
21 |
"delete x (Node lv l (a,b) r) =
|
|
22 |
(case cmp x a of
|
|
23 |
LT \<Rightarrow> adjust (Node lv (delete x l) (a,b) r) |
|
|
24 |
GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
|
|
25 |
EQ \<Rightarrow> (if l = Leaf then r
|
|
26 |
else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))"
|
|
27 |
|
|
28 |
|
|
29 |
subsection {* Functional Correctness Proofs *}
|
|
30 |
|
|
31 |
theorem inorder_update:
|
|
32 |
"sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
|
|
33 |
by (induct t) (auto simp: upd_list_simps inorder_split inorder_skew)
|
|
34 |
|
|
35 |
|
|
36 |
theorem inorder_delete:
|
|
37 |
"sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
|
|
38 |
by(induction t)
|
|
39 |
(auto simp: del_list_simps inorder_adjust del_maxD split: prod.splits)
|
|
40 |
|
|
41 |
interpretation Map_by_Ordered
|
|
42 |
where empty = Leaf and lookup = lookup and update = update and delete = delete
|
|
43 |
and inorder = inorder and inv = "\<lambda>_. True"
|
|
44 |
proof (standard, goal_cases)
|
|
45 |
case 1 show ?case by simp
|
|
46 |
next
|
|
47 |
case 2 thus ?case by(simp add: lookup_map_of)
|
|
48 |
next
|
|
49 |
case 3 thus ?case by(simp add: inorder_update)
|
|
50 |
next
|
|
51 |
case 4 thus ?case by(simp add: inorder_delete)
|
|
52 |
qed auto
|
|
53 |
|
|
54 |
end
|