71844
|
1 |
(* Author: Tobias Nipkow *)
|
71814
|
2 |
|
71844
|
3 |
section "AVL Tree with Balance Factors (1)"
|
71814
|
4 |
|
|
5 |
theory AVL_Bal_Set
|
|
6 |
imports
|
|
7 |
Cmp
|
|
8 |
Isin2
|
|
9 |
begin
|
|
10 |
|
71844
|
11 |
text \<open>This version detects height increase/decrease from above via the change in balance factors.\<close>
|
|
12 |
|
71814
|
13 |
datatype bal = Lh | Bal | Rh
|
|
14 |
|
|
15 |
type_synonym 'a tree_bal = "('a * bal) tree"
|
|
16 |
|
|
17 |
text \<open>Invariant:\<close>
|
|
18 |
|
|
19 |
fun avl :: "'a tree_bal \<Rightarrow> bool" where
|
|
20 |
"avl Leaf = True" |
|
|
21 |
"avl (Node l (a,b) r) =
|
|
22 |
((case b of
|
|
23 |
Bal \<Rightarrow> height r = height l |
|
|
24 |
Lh \<Rightarrow> height l = height r + 1 |
|
|
25 |
Rh \<Rightarrow> height r = height l + 1)
|
|
26 |
\<and> avl l \<and> avl r)"
|
|
27 |
|
|
28 |
|
|
29 |
subsection \<open>Code\<close>
|
|
30 |
|
71844
|
31 |
fun is_bal where
|
|
32 |
"is_bal (Node l (a,b) r) = (b = Bal)"
|
71828
|
33 |
|
71844
|
34 |
fun incr where
|
|
35 |
"incr t t' = (t = Leaf \<or> is_bal t \<and> \<not> is_bal t')"
|
71814
|
36 |
|
71820
|
37 |
fun rot2 where
|
|
38 |
"rot2 A a B c C = (case B of
|
|
39 |
(Node B\<^sub>1 (b, bb) B\<^sub>2) \<Rightarrow>
|
|
40 |
let b\<^sub>1 = if bb = Rh then Lh else Bal;
|
|
41 |
b\<^sub>2 = if bb = Lh then Rh else Bal
|
|
42 |
in Node (Node A (a,b\<^sub>1) B\<^sub>1) (b,Bal) (Node B\<^sub>2 (c,b\<^sub>2) C))"
|
71814
|
43 |
|
71844
|
44 |
fun balL :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
|
|
45 |
"balL AB c bc C = (case bc of
|
|
46 |
Bal \<Rightarrow> Node AB (c,Lh) C |
|
|
47 |
Rh \<Rightarrow> Node AB (c,Bal) C |
|
71824
|
48 |
Lh \<Rightarrow> (case AB of
|
71844
|
49 |
Node A (a,Lh) B \<Rightarrow> Node A (a,Bal) (Node B (c,Bal) C) |
|
|
50 |
Node A (a,Bal) B \<Rightarrow> Node A (a,Rh) (Node B (c,Lh) C) |
|
|
51 |
Node A (a,Rh) B \<Rightarrow> rot2 A a B c C))"
|
71814
|
52 |
|
71844
|
53 |
fun balR :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
|
|
54 |
"balR A a ba BC = (case ba of
|
|
55 |
Bal \<Rightarrow> Node A (a,Rh) BC |
|
|
56 |
Lh \<Rightarrow> Node A (a,Bal) BC |
|
71824
|
57 |
Rh \<Rightarrow> (case BC of
|
71844
|
58 |
Node B (c,Rh) C \<Rightarrow> Node (Node A (a,Bal) B) (c,Bal) C |
|
|
59 |
Node B (c,Bal) C \<Rightarrow> Node (Node A (a,Rh) B) (c,Lh) C |
|
|
60 |
Node B (c,Lh) C \<Rightarrow> rot2 A a B c C))"
|
71814
|
61 |
|
71844
|
62 |
fun insert :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
|
|
63 |
"insert x Leaf = Node Leaf (x, Bal) Leaf" |
|
|
64 |
"insert x (Node l (a, b) r) = (case cmp x a of
|
|
65 |
EQ \<Rightarrow> Node l (a, b) r |
|
|
66 |
LT \<Rightarrow> let l' = insert x l in if incr l l' then balL l' a b r else Node l' (a,b) r |
|
|
67 |
GT \<Rightarrow> let r' = insert x r in if incr r r' then balR l a b r' else Node l (a,b) r')"
|
|
68 |
|
|
69 |
fun decr where
|
|
70 |
"decr t t' = (t \<noteq> Leaf \<and> (t' = Leaf \<or> \<not> is_bal t \<and> is_bal t'))"
|
71814
|
71 |
|
71844
|
72 |
fun split_max :: "'a tree_bal \<Rightarrow> 'a tree_bal * 'a" where
|
71814
|
73 |
"split_max (Node l (a, ba) r) =
|
71844
|
74 |
(if r = Leaf then (l,a)
|
|
75 |
else let (r',a') = split_max r;
|
|
76 |
t' = if decr r r' then balL l a ba r' else Node l (a,ba) r'
|
|
77 |
in (t', a'))"
|
71814
|
78 |
|
71844
|
79 |
fun delete :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
|
|
80 |
"delete _ Leaf = Leaf" |
|
|
81 |
"delete x (Node l (a, ba) r) =
|
71814
|
82 |
(case cmp x a of
|
71844
|
83 |
EQ \<Rightarrow> if l = Leaf then r
|
|
84 |
else let (l', a') = split_max l in
|
|
85 |
if decr l l' then balR l' a' ba r else Node l' (a',ba) r |
|
|
86 |
LT \<Rightarrow> let l' = delete x l in if decr l l' then balR l' a ba r else Node l' (a,ba) r |
|
|
87 |
GT \<Rightarrow> let r' = delete x r in if decr r r' then balL l a ba r' else Node l (a,ba) r')"
|
71814
|
88 |
|
71846
|
89 |
|
|
90 |
subsection \<open>Proofs\<close>
|
|
91 |
|
71814
|
92 |
lemmas split_max_induct = split_max.induct[case_names Node Leaf]
|
|
93 |
|
71844
|
94 |
lemmas splits = if_splits tree.splits bal.splits
|
71814
|
95 |
|
71846
|
96 |
declare Let_def [simp]
|
71814
|
97 |
|
71828
|
98 |
subsubsection "Proofs about insertion"
|
71814
|
99 |
|
71844
|
100 |
lemma avl_insert: "avl t \<Longrightarrow>
|
|
101 |
avl(insert x t) \<and>
|
|
102 |
height(insert x t) = height t + (if incr t (insert x t) then 1 else 0)"
|
|
103 |
apply(induction x t rule: insert.induct)
|
|
104 |
apply(auto split!: splits)
|
71814
|
105 |
done
|
|
106 |
|
71844
|
107 |
text \<open>The following two auxiliary lemma merely simplify the proof of \<open>inorder_insert\<close>.\<close>
|
71814
|
108 |
|
71844
|
109 |
lemma [simp]: "[] \<noteq> ins_list x xs"
|
|
110 |
by(cases xs) auto
|
71824
|
111 |
|
71844
|
112 |
lemma [simp]: "avl t \<Longrightarrow> insert x t \<noteq> \<langle>l, (a, Rh), \<langle>\<rangle>\<rangle> \<and> insert x t \<noteq> \<langle>\<langle>\<rangle>, (a, Lh), r\<rangle>"
|
|
113 |
by(drule avl_insert[of _ x]) (auto split: splits)
|
71824
|
114 |
|
71844
|
115 |
theorem inorder_insert:
|
|
116 |
"\<lbrakk> avl t; sorted(inorder t) \<rbrakk> \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
|
71824
|
117 |
apply(induction t)
|
71846
|
118 |
apply (auto simp: ins_list_simps split!: splits)
|
71824
|
119 |
done
|
|
120 |
|
71814
|
121 |
|
71828
|
122 |
subsubsection "Proofs about deletion"
|
71814
|
123 |
|
71844
|
124 |
lemma inorder_balR:
|
71814
|
125 |
"\<lbrakk> ba = Rh \<longrightarrow> r \<noteq> Leaf; avl r \<rbrakk>
|
71844
|
126 |
\<Longrightarrow> inorder (balR l a ba r) = inorder l @ a # inorder r"
|
71814
|
127 |
by (auto split: splits)
|
|
128 |
|
71844
|
129 |
lemma inorder_balL:
|
71814
|
130 |
"\<lbrakk> ba = Lh \<longrightarrow> l \<noteq> Leaf; avl l \<rbrakk>
|
71844
|
131 |
\<Longrightarrow> inorder (balL l a ba r) = inorder l @ a # inorder r"
|
71814
|
132 |
by (auto split: splits)
|
|
133 |
|
71844
|
134 |
lemma height_1_iff: "avl t \<Longrightarrow> height t = Suc 0 \<longleftrightarrow> (\<exists>x. t = Node Leaf (x,Bal) Leaf)"
|
|
135 |
by(cases t) (auto split: splits prod.splits)
|
|
136 |
|
71814
|
137 |
lemma avl_split_max:
|
71844
|
138 |
"\<lbrakk> split_max t = (t',a); avl t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
|
|
139 |
avl t' \<and> height t = height t' + (if decr t t' then 1 else 0)"
|
71814
|
140 |
apply(induction t arbitrary: t' a rule: split_max_induct)
|
71844
|
141 |
apply(auto simp: max_absorb1 max_absorb2 height_1_iff split!: splits prod.splits)
|
71814
|
142 |
done
|
|
143 |
|
71844
|
144 |
lemma avl_delete: "avl t \<Longrightarrow>
|
|
145 |
avl (delete x t) \<and>
|
|
146 |
height t = height (delete x t) + (if decr t (delete x t) then 1 else 0)"
|
|
147 |
apply(induction x t rule: delete.induct)
|
71846
|
148 |
apply(auto simp: max_absorb1 max_absorb2 height_1_iff dest: avl_split_max split!: splits prod.splits)
|
71814
|
149 |
done
|
|
150 |
|
|
151 |
lemma inorder_split_maxD:
|
|
152 |
"\<lbrakk> split_max t = (t',a); t \<noteq> Leaf; avl t \<rbrakk> \<Longrightarrow>
|
71844
|
153 |
inorder t' @ [a] = inorder t"
|
71814
|
154 |
apply(induction t arbitrary: t' rule: split_max.induct)
|
|
155 |
apply(fastforce split!: splits prod.splits)
|
|
156 |
apply simp
|
|
157 |
done
|
|
158 |
|
71844
|
159 |
lemma neq_Leaf_if_height_neq_0: "height t \<noteq> 0 \<Longrightarrow> t \<noteq> Leaf"
|
71814
|
160 |
by auto
|
|
161 |
|
71844
|
162 |
lemma split_max_Leaf: "\<lbrakk> t \<noteq> Leaf; avl t \<rbrakk> \<Longrightarrow> split_max t = (\<langle>\<rangle>, x) \<longleftrightarrow> t = Node Leaf (x,Bal) Leaf"
|
|
163 |
by(cases t) (auto split: splits prod.splits)
|
|
164 |
|
|
165 |
theorem inorder_delete:
|
|
166 |
"\<lbrakk> avl t; sorted(inorder t) \<rbrakk> \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
|
71814
|
167 |
apply(induction t rule: tree2_induct)
|
71846
|
168 |
apply(auto simp: del_list_simps inorder_balR inorder_balL avl_delete inorder_split_maxD
|
71844
|
169 |
split_max_Leaf neq_Leaf_if_height_neq_0
|
|
170 |
simp del: balL.simps balR.simps split!: splits prod.splits)
|
71814
|
171 |
done
|
|
172 |
|
|
173 |
|
|
174 |
subsubsection \<open>Set Implementation\<close>
|
|
175 |
|
|
176 |
interpretation S: Set_by_Ordered
|
|
177 |
where empty = Leaf and isin = isin
|
71828
|
178 |
and insert = insert
|
|
179 |
and delete = delete
|
71814
|
180 |
and inorder = inorder and inv = avl
|
|
181 |
proof (standard, goal_cases)
|
|
182 |
case 1 show ?case by (simp)
|
|
183 |
next
|
|
184 |
case 2 thus ?case by(simp add: isin_set_inorder)
|
|
185 |
next
|
71844
|
186 |
case 3 thus ?case by(simp add: inorder_insert)
|
71814
|
187 |
next
|
71844
|
188 |
case 4 thus ?case by(simp add: inorder_delete)
|
71814
|
189 |
next
|
71828
|
190 |
case 5 thus ?case by (simp)
|
71814
|
191 |
next
|
|
192 |
case 6 thus ?case by (simp add: avl_insert)
|
|
193 |
next
|
|
194 |
case 7 thus ?case by (simp add: avl_delete)
|
|
195 |
qed
|
|
196 |
|
|
197 |
end
|