author | nipkow |
Mon, 05 Dec 2016 18:14:41 +0100 | |
changeset 64540 | f1f4ba6d02c9 |
parent 64533 | 172f3a047f4a |
child 64771 | 23c56f483775 |
permissions | -rw-r--r-- |
57250 | 1 |
(* Author: Tobias Nipkow *) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
2 |
(* Todo: |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
3 |
(min_)height of balanced trees via floorlog |
63770 | 4 |
minimal path_len of balanced trees |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
5 |
*) |
57250 | 6 |
|
60500 | 7 |
section \<open>Binary Tree\<close> |
57250 | 8 |
|
9 |
theory Tree |
|
10 |
imports Main |
|
11 |
begin |
|
12 |
||
58424 | 13 |
datatype 'a tree = |
62160 | 14 |
is_Leaf: Leaf ("\<langle>\<rangle>") | |
15 |
Node (left: "'a tree") (val: 'a) (right: "'a tree") ("(1\<langle>_,/ _,/ _\<rangle>)") |
|
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
16 |
where |
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
17 |
"left Leaf = Leaf" |
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
18 |
| "right Leaf = Leaf" |
57569
e20a999f7161
register tree with datatype_compat ot support QuickCheck
hoelzl
parents:
57530
diff
changeset
|
19 |
datatype_compat tree |
57250 | 20 |
|
60500 | 21 |
text\<open>Can be seen as counting the number of leaves rather than nodes:\<close> |
58438 | 22 |
|
23 |
definition size1 :: "'a tree \<Rightarrow> nat" where |
|
24 |
"size1 t = size t + 1" |
|
25 |
||
63861 | 26 |
fun subtrees :: "'a tree \<Rightarrow> 'a tree set" where |
27 |
"subtrees \<langle>\<rangle> = {\<langle>\<rangle>}" | |
|
28 |
"subtrees (\<langle>l, a, r\<rangle>) = insert \<langle>l, a, r\<rangle> (subtrees l \<union> subtrees r)" |
|
29 |
||
30 |
fun mirror :: "'a tree \<Rightarrow> 'a tree" where |
|
31 |
"mirror \<langle>\<rangle> = Leaf" | |
|
32 |
"mirror \<langle>l,x,r\<rangle> = \<langle>mirror r, x, mirror l\<rangle>" |
|
33 |
||
34 |
class height = fixes height :: "'a \<Rightarrow> nat" |
|
35 |
||
36 |
instantiation tree :: (type)height |
|
37 |
begin |
|
38 |
||
39 |
fun height_tree :: "'a tree => nat" where |
|
40 |
"height Leaf = 0" | |
|
41 |
"height (Node t1 a t2) = max (height t1) (height t2) + 1" |
|
42 |
||
43 |
instance .. |
|
44 |
||
45 |
end |
|
46 |
||
47 |
fun min_height :: "'a tree \<Rightarrow> nat" where |
|
48 |
"min_height Leaf = 0" | |
|
49 |
"min_height (Node l _ r) = min (min_height l) (min_height r) + 1" |
|
50 |
||
51 |
fun complete :: "'a tree \<Rightarrow> bool" where |
|
52 |
"complete Leaf = True" | |
|
53 |
"complete (Node l x r) = (complete l \<and> complete r \<and> height l = height r)" |
|
54 |
||
55 |
definition balanced :: "'a tree \<Rightarrow> bool" where |
|
56 |
"balanced t = (height t - min_height t \<le> 1)" |
|
57 |
||
58 |
text \<open>Weight balanced:\<close> |
|
59 |
fun wbalanced :: "'a tree \<Rightarrow> bool" where |
|
60 |
"wbalanced Leaf = True" | |
|
61 |
"wbalanced (Node l x r) = (abs(int(size l) - int(size r)) \<le> 1 \<and> wbalanced l \<and> wbalanced r)" |
|
62 |
||
63 |
text \<open>Internal path length:\<close> |
|
64 |
fun path_len :: "'a tree \<Rightarrow> nat" where |
|
65 |
"path_len Leaf = 0 " | |
|
66 |
"path_len (Node l _ r) = path_len l + size l + path_len r + size r" |
|
67 |
||
68 |
fun preorder :: "'a tree \<Rightarrow> 'a list" where |
|
69 |
"preorder \<langle>\<rangle> = []" | |
|
70 |
"preorder \<langle>l, x, r\<rangle> = x # preorder l @ preorder r" |
|
71 |
||
72 |
fun inorder :: "'a tree \<Rightarrow> 'a list" where |
|
73 |
"inorder \<langle>\<rangle> = []" | |
|
74 |
"inorder \<langle>l, x, r\<rangle> = inorder l @ [x] @ inorder r" |
|
75 |
||
76 |
text\<open>A linear version avoiding append:\<close> |
|
77 |
fun inorder2 :: "'a tree \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
78 |
"inorder2 \<langle>\<rangle> xs = xs" | |
|
79 |
"inorder2 \<langle>l, x, r\<rangle> xs = inorder2 l (x # inorder2 r xs)" |
|
80 |
||
81 |
text\<open>Binary Search Tree:\<close> |
|
82 |
fun (in linorder) bst :: "'a tree \<Rightarrow> bool" where |
|
83 |
"bst \<langle>\<rangle> \<longleftrightarrow> True" | |
|
84 |
"bst \<langle>l, a, r\<rangle> \<longleftrightarrow> bst l \<and> bst r \<and> (\<forall>x\<in>set_tree l. x < a) \<and> (\<forall>x\<in>set_tree r. a < x)" |
|
85 |
||
86 |
text\<open>Binary Search Tree with duplicates:\<close> |
|
87 |
fun (in linorder) bst_eq :: "'a tree \<Rightarrow> bool" where |
|
88 |
"bst_eq \<langle>\<rangle> \<longleftrightarrow> True" | |
|
89 |
"bst_eq \<langle>l,a,r\<rangle> \<longleftrightarrow> |
|
90 |
bst_eq l \<and> bst_eq r \<and> (\<forall>x\<in>set_tree l. x \<le> a) \<and> (\<forall>x\<in>set_tree r. a \<le> x)" |
|
91 |
||
92 |
fun (in linorder) heap :: "'a tree \<Rightarrow> bool" where |
|
93 |
"heap Leaf = True" | |
|
94 |
"heap (Node l m r) = |
|
95 |
(heap l \<and> heap r \<and> (\<forall>x \<in> set_tree l \<union> set_tree r. m \<le> x))" |
|
96 |
||
97 |
||
98 |
subsection \<open>@{const size}\<close> |
|
99 |
||
58438 | 100 |
lemma size1_simps[simp]: |
101 |
"size1 \<langle>\<rangle> = 1" |
|
102 |
"size1 \<langle>l, x, r\<rangle> = size1 l + size1 r" |
|
103 |
by (simp_all add: size1_def) |
|
104 |
||
62650 | 105 |
lemma size1_ge0[simp]: "0 < size1 t" |
106 |
by (simp add: size1_def) |
|
107 |
||
60507 | 108 |
lemma size_0_iff_Leaf: "size t = 0 \<longleftrightarrow> t = Leaf" |
60505 | 109 |
by(cases t) auto |
110 |
||
58424 | 111 |
lemma neq_Leaf_iff: "(t \<noteq> \<langle>\<rangle>) = (\<exists>l a r. t = \<langle>l, a, r\<rangle>)" |
112 |
by (cases t) auto |
|
57530 | 113 |
|
57687 | 114 |
lemma finite_set_tree[simp]: "finite(set_tree t)" |
115 |
by(induction t) auto |
|
116 |
||
59776 | 117 |
lemma size_map_tree[simp]: "size (map_tree f t) = size t" |
118 |
by (induction t) auto |
|
119 |
||
120 |
lemma size1_map_tree[simp]: "size1 (map_tree f t) = size1 t" |
|
121 |
by (simp add: size1_def) |
|
122 |
||
123 |
||
63861 | 124 |
subsection \<open>@{const subtrees}\<close> |
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
125 |
|
63861 | 126 |
lemma set_treeE: "a \<in> set_tree t \<Longrightarrow> \<exists>l r. \<langle>l, a, r\<rangle> \<in> subtrees t" |
127 |
by (induction t)(auto) |
|
59776 | 128 |
|
63861 | 129 |
lemma Node_notin_subtrees_if[simp]: "a \<notin> set_tree t \<Longrightarrow> Node l a r \<notin> subtrees t" |
130 |
by (induction t) auto |
|
59776 | 131 |
|
63861 | 132 |
lemma in_set_tree_if: "\<langle>l, a, r\<rangle> \<in> subtrees t \<Longrightarrow> a \<in> set_tree t" |
133 |
by (metis Node_notin_subtrees_if) |
|
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
134 |
|
63861 | 135 |
|
136 |
subsection \<open>@{const height} and @{const min_height}\<close> |
|
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
137 |
|
63665 | 138 |
lemma height_0_iff_Leaf: "height t = 0 \<longleftrightarrow> t = Leaf" |
139 |
by(cases t) auto |
|
140 |
||
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
141 |
lemma height_map_tree[simp]: "height (map_tree f t) = height t" |
59776 | 142 |
by (induction t) auto |
143 |
||
64414 | 144 |
lemma height_le_size_tree: "height t \<le> size (t::'a tree)" |
145 |
by (induction t) auto |
|
146 |
||
64533 | 147 |
lemma size1_height: "size1 t \<le> 2 ^ height (t::'a tree)" |
62202 | 148 |
proof(induction t) |
149 |
case (Node l a r) |
|
150 |
show ?case |
|
151 |
proof (cases "height l \<le> height r") |
|
152 |
case True |
|
64533 | 153 |
have "size1(Node l a r) = size1 l + size1 r" by simp |
154 |
also have "size1 l \<le> 2 ^ height l" by(rule Node.IH(1)) |
|
155 |
also have "size1 r \<le> 2 ^ height r" by(rule Node.IH(2)) |
|
62202 | 156 |
also have "(2::nat) ^ height l \<le> 2 ^ height r" using True by simp |
157 |
finally show ?thesis using True by (auto simp: max_def mult_2) |
|
158 |
next |
|
159 |
case False |
|
64533 | 160 |
have "size1(Node l a r) = size1 l + size1 r" by simp |
161 |
also have "size1 l \<le> 2 ^ height l" by(rule Node.IH(1)) |
|
162 |
also have "size1 r \<le> 2 ^ height r" by(rule Node.IH(2)) |
|
62202 | 163 |
also have "(2::nat) ^ height r \<le> 2 ^ height l" using False by simp |
164 |
finally show ?thesis using False by (auto simp: max_def mult_2) |
|
165 |
qed |
|
166 |
qed simp |
|
167 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
168 |
corollary size_height: "size t \<le> 2 ^ height (t::'a tree) - 1" |
64533 | 169 |
using size1_height[of t, unfolded size1_def] by(arith) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
170 |
|
63861 | 171 |
lemma height_subtrees: "s \<in> subtrees t \<Longrightarrow> height s \<le> height t" |
172 |
by (induction t) auto |
|
57687 | 173 |
|
63598 | 174 |
|
64540 | 175 |
lemma min_height_le_height: "min_height t \<le> height t" |
63598 | 176 |
by(induction t) auto |
177 |
||
178 |
lemma min_height_map_tree[simp]: "min_height (map_tree f t) = min_height t" |
|
179 |
by (induction t) auto |
|
180 |
||
64533 | 181 |
lemma min_height_size1: "2 ^ min_height t \<le> size1 t" |
63598 | 182 |
proof(induction t) |
183 |
case (Node l a r) |
|
184 |
have "(2::nat) ^ min_height (Node l a r) \<le> 2 ^ min_height l + 2 ^ min_height r" |
|
185 |
by (simp add: min_def) |
|
64533 | 186 |
also have "\<dots> \<le> size1(Node l a r)" using Node.IH by simp |
63598 | 187 |
finally show ?case . |
188 |
qed simp |
|
189 |
||
190 |
||
63861 | 191 |
subsection \<open>@{const complete}\<close> |
63036 | 192 |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
193 |
lemma complete_iff_height: "complete t \<longleftrightarrow> (min_height t = height t)" |
63598 | 194 |
apply(induction t) |
195 |
apply simp |
|
196 |
apply (simp add: min_def max_def) |
|
64540 | 197 |
by (metis le_antisym le_trans min_height_le_height) |
63598 | 198 |
|
63770 | 199 |
lemma size1_if_complete: "complete t \<Longrightarrow> size1 t = 2 ^ height t" |
63036 | 200 |
by (induction t) auto |
201 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
202 |
lemma size_if_complete: "complete t \<Longrightarrow> size t = 2 ^ height t - 1" |
63770 | 203 |
using size1_if_complete[simplified size1_def] by fastforce |
204 |
||
64533 | 205 |
lemma complete_if_size1_height: "size1 t = 2 ^ height t \<Longrightarrow> complete t" |
63770 | 206 |
proof (induct "height t" arbitrary: t) |
64533 | 207 |
case 0 thus ?case by (simp add: height_0_iff_Leaf) |
63770 | 208 |
next |
209 |
case (Suc h) |
|
210 |
hence "t \<noteq> Leaf" by auto |
|
211 |
then obtain l a r where [simp]: "t = Node l a r" |
|
212 |
by (auto simp: neq_Leaf_iff) |
|
213 |
have 1: "height l \<le> h" and 2: "height r \<le> h" using Suc(2) by(auto) |
|
64533 | 214 |
have 3: "\<not> height l < h" |
63770 | 215 |
proof |
216 |
assume 0: "height l < h" |
|
64533 | 217 |
have "size1 t = size1 l + size1 r" by simp |
218 |
also note size1_height[of l] |
|
63770 | 219 |
also note size1_height[of r] |
64533 | 220 |
also have "(2::nat) ^ height l < 2 ^ h" |
63770 | 221 |
using 0 by (simp add: diff_less_mono) |
222 |
also have "(2::nat) ^ height r \<le> 2 ^ h" using 2 by simp |
|
64533 | 223 |
also have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp) |
224 |
also have "\<dots> = size1 t" using Suc(2,3) by simp |
|
63770 | 225 |
finally show False by (simp add: diff_le_mono) |
226 |
qed |
|
227 |
have 4: "~ height r < h" |
|
228 |
proof |
|
229 |
assume 0: "height r < h" |
|
64533 | 230 |
have "size1 t = size1 l + size1 r" by simp |
231 |
also note size1_height[of r] |
|
63770 | 232 |
also note size1_height[of l] |
64533 | 233 |
also have "(2::nat) ^ height r < 2 ^ h" |
63770 | 234 |
using 0 by (simp add: diff_less_mono) |
235 |
also have "(2::nat) ^ height l \<le> 2 ^ h" using 1 by simp |
|
64533 | 236 |
also have "(2::nat) ^ h +2 ^ h = 2 ^ (Suc h)" by (simp) |
237 |
also have "\<dots> = size1 t" using Suc(2,3) by simp |
|
63770 | 238 |
finally show False by (simp add: diff_le_mono) |
239 |
qed |
|
240 |
from 1 2 3 4 have *: "height l = h" "height r = h" by linarith+ |
|
64533 | 241 |
hence "size1 l = 2 ^ height l" "size1 r = 2 ^ height r" |
242 |
using Suc(3) size1_height[of l] size1_height[of r] by (auto) |
|
63770 | 243 |
with * Suc(1) show ?case by simp |
244 |
qed |
|
245 |
||
64533 | 246 |
text\<open>The following proof involves \<open>\<ge>\<close>/\<open>>\<close> chains rather than the standard |
247 |
\<open>\<le>\<close>/\<open><\<close> chains. To chain the elements together the transitivity rules \<open>xtrans\<close> |
|
248 |
are used.\<close> |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
249 |
|
64533 | 250 |
lemma complete_if_size1_min_height: "size1 t = 2 ^ min_height t \<Longrightarrow> complete t" |
251 |
proof (induct "min_height t" arbitrary: t) |
|
252 |
case 0 thus ?case by (simp add: size_0_iff_Leaf size1_def) |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
253 |
next |
64533 | 254 |
case (Suc h) |
255 |
hence "t \<noteq> Leaf" by auto |
|
256 |
then obtain l a r where [simp]: "t = Node l a r" |
|
257 |
by (auto simp: neq_Leaf_iff) |
|
258 |
have 1: "h \<le> min_height l" and 2: "h \<le> min_height r" using Suc(2) by(auto) |
|
259 |
have 3: "\<not> h < min_height l" |
|
260 |
proof |
|
261 |
assume 0: "h < min_height l" |
|
262 |
have "size1 t = size1 l + size1 r" by simp |
|
263 |
also note min_height_size1[of l] |
|
264 |
also(xtrans) note min_height_size1[of r] |
|
265 |
also(xtrans) have "(2::nat) ^ min_height l > 2 ^ h" |
|
266 |
using 0 by (simp add: diff_less_mono) |
|
267 |
also(xtrans) have "(2::nat) ^ min_height r \<ge> 2 ^ h" using 2 by simp |
|
268 |
also(xtrans) have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp) |
|
269 |
also have "\<dots> = size1 t" using Suc(2,3) by simp |
|
270 |
finally show False by (simp add: diff_le_mono) |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
271 |
qed |
64533 | 272 |
have 4: "\<not> h < min_height r" |
273 |
proof |
|
274 |
assume 0: "h < min_height r" |
|
275 |
have "size1 t = size1 l + size1 r" by simp |
|
276 |
also note min_height_size1[of l] |
|
277 |
also(xtrans) note min_height_size1[of r] |
|
278 |
also(xtrans) have "(2::nat) ^ min_height r > 2 ^ h" |
|
279 |
using 0 by (simp add: diff_less_mono) |
|
280 |
also(xtrans) have "(2::nat) ^ min_height l \<ge> 2 ^ h" using 1 by simp |
|
281 |
also(xtrans) have "(2::nat) ^ h + 2 ^ h = 2 ^ (Suc h)" by (simp) |
|
282 |
also have "\<dots> = size1 t" using Suc(2,3) by simp |
|
283 |
finally show False by (simp add: diff_le_mono) |
|
284 |
qed |
|
285 |
from 1 2 3 4 have *: "min_height l = h" "min_height r = h" by linarith+ |
|
286 |
hence "size1 l = 2 ^ min_height l" "size1 r = 2 ^ min_height r" |
|
287 |
using Suc(3) min_height_size1[of l] min_height_size1[of r] by (auto) |
|
288 |
with * Suc(1) show ?case |
|
289 |
by (simp add: complete_iff_height) |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
290 |
qed |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
291 |
|
64533 | 292 |
lemma complete_iff_size1: "complete t \<longleftrightarrow> size1 t = 2 ^ height t" |
293 |
using complete_if_size1_height size1_if_complete by blast |
|
294 |
||
295 |
text\<open>Better bounds for incomplete trees:\<close> |
|
296 |
||
297 |
lemma size1_height_if_incomplete: |
|
298 |
"\<not> complete t \<Longrightarrow> size1 t < 2 ^ height t" |
|
299 |
by (meson antisym_conv complete_iff_size1 not_le size1_height) |
|
300 |
||
301 |
lemma min_height_size1_if_incomplete: |
|
302 |
"\<not> complete t \<Longrightarrow> 2 ^ min_height t < size1 t" |
|
303 |
by (metis complete_if_size1_min_height le_less min_height_size1) |
|
304 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
305 |
|
63861 | 306 |
subsection \<open>@{const balanced}\<close> |
307 |
||
308 |
lemma balanced_subtreeL: "balanced (Node l x r) \<Longrightarrow> balanced l" |
|
309 |
by(simp add: balanced_def) |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
310 |
|
63861 | 311 |
lemma balanced_subtreeR: "balanced (Node l x r) \<Longrightarrow> balanced r" |
312 |
by(simp add: balanced_def) |
|
313 |
||
314 |
lemma balanced_subtrees: "\<lbrakk> balanced t; s \<in> subtrees t \<rbrakk> \<Longrightarrow> balanced s" |
|
315 |
using [[simp_depth_limit=1]] |
|
316 |
by(induction t arbitrary: s) |
|
317 |
(auto simp add: balanced_subtreeL balanced_subtreeR) |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
318 |
|
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
319 |
text\<open>Balanced trees have optimal height:\<close> |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
320 |
|
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
321 |
lemma balanced_optimal: |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
322 |
fixes t :: "'a tree" and t' :: "'b tree" |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
323 |
assumes "balanced t" "size t \<le> size t'" shows "height t \<le> height t'" |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
324 |
proof (cases "complete t") |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
325 |
case True |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
326 |
have "(2::nat) ^ height t - 1 \<le> 2 ^ height t' - 1" |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
327 |
proof - |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
328 |
have "(2::nat) ^ height t - 1 = size t" |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
329 |
using True by (simp add: complete_iff_height size_if_complete) |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
330 |
also note assms(2) |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
331 |
also have "size t' \<le> 2 ^ height t' - 1" by (rule size_height) |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
332 |
finally show ?thesis . |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
333 |
qed |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
334 |
thus ?thesis by (simp add: le_diff_iff) |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
335 |
next |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
336 |
case False |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
337 |
have "(2::nat) ^ min_height t < 2 ^ height t'" |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
338 |
proof - |
64533 | 339 |
have "(2::nat) ^ min_height t < size1 t" |
340 |
by(rule min_height_size1_if_incomplete[OF False]) |
|
341 |
also have "size1 t \<le> size1 t'" using assms(2) by (simp add: size1_def) |
|
342 |
also have "size1 t' \<le> 2 ^ height t'" by(rule size1_height) |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
343 |
finally show ?thesis |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
344 |
using power_eq_0_iff[of "2::nat" "height t'"] by linarith |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
345 |
qed |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
346 |
hence *: "min_height t < height t'" by simp |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
347 |
have "min_height t + 1 = height t" |
64540 | 348 |
using min_height_le_height[of t] assms(1) False |
63829 | 349 |
by (simp add: complete_iff_height balanced_def) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
350 |
with * show ?thesis by arith |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
351 |
qed |
63036 | 352 |
|
353 |
||
63861 | 354 |
subsection \<open>@{const wbalanced}\<close> |
355 |
||
356 |
lemma wbalanced_subtrees: "\<lbrakk> wbalanced t; s \<in> subtrees t \<rbrakk> \<Longrightarrow> wbalanced s" |
|
357 |
using [[simp_depth_limit=1]] by(induction t arbitrary: s) auto |
|
358 |
||
359 |
(* show wbalanced \<Longrightarrow> balanced and use that in Balanced.thy *) |
|
360 |
||
361 |
||
362 |
subsection \<open>@{const path_len}\<close> |
|
63413 | 363 |
|
364 |
text \<open>The internal path length of a tree:\<close> |
|
365 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
366 |
lemma path_len_if_bal: "complete t |
63413 | 367 |
\<Longrightarrow> path_len t = (let n = height t in 2 + n*2^n - 2^(n+1))" |
368 |
proof(induction t) |
|
369 |
case (Node l x r) |
|
370 |
have *: "2^(n+1) \<le> 2 + n*2^n" for n :: nat |
|
371 |
by(induction n) auto |
|
372 |
have **: "(0::nat) < 2^n" for n :: nat by simp |
|
373 |
let ?h = "height r" |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
374 |
show ?case using Node *[of ?h] **[of ?h] by (simp add: size_if_complete Let_def) |
63413 | 375 |
qed simp |
376 |
||
377 |
||
59776 | 378 |
subsection "List of entries" |
379 |
||
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
380 |
lemma set_inorder[simp]: "set (inorder t) = set_tree t" |
58424 | 381 |
by (induction t) auto |
57250 | 382 |
|
59776 | 383 |
lemma set_preorder[simp]: "set (preorder t) = set_tree t" |
384 |
by (induction t) auto |
|
385 |
||
386 |
lemma length_preorder[simp]: "length (preorder t) = size t" |
|
387 |
by (induction t) auto |
|
388 |
||
389 |
lemma length_inorder[simp]: "length (inorder t) = size t" |
|
390 |
by (induction t) auto |
|
391 |
||
392 |
lemma preorder_map: "preorder (map_tree f t) = map f (preorder t)" |
|
393 |
by (induction t) auto |
|
394 |
||
395 |
lemma inorder_map: "inorder (map_tree f t) = map f (inorder t)" |
|
396 |
by (induction t) auto |
|
397 |
||
63765 | 398 |
lemma inorder2_inorder: "inorder2 t xs = inorder t @ xs" |
399 |
by (induction t arbitrary: xs) auto |
|
400 |
||
57687 | 401 |
|
63861 | 402 |
subsection \<open>Binary Search Tree\<close> |
59561 | 403 |
|
59928 | 404 |
lemma (in linorder) bst_eq_if_bst: "bst t \<Longrightarrow> bst_eq t" |
405 |
by (induction t) (auto) |
|
406 |
||
59561 | 407 |
lemma (in linorder) bst_eq_imp_sorted: "bst_eq t \<Longrightarrow> sorted (inorder t)" |
408 |
apply (induction t) |
|
409 |
apply(simp) |
|
410 |
by (fastforce simp: sorted_append sorted_Cons intro: less_imp_le less_trans) |
|
411 |
||
59928 | 412 |
lemma (in linorder) distinct_preorder_if_bst: "bst t \<Longrightarrow> distinct (preorder t)" |
413 |
apply (induction t) |
|
414 |
apply simp |
|
415 |
apply(fastforce elim: order.asym) |
|
416 |
done |
|
417 |
||
418 |
lemma (in linorder) distinct_inorder_if_bst: "bst t \<Longrightarrow> distinct (inorder t)" |
|
419 |
apply (induction t) |
|
420 |
apply simp |
|
421 |
apply(fastforce elim: order.asym) |
|
422 |
done |
|
423 |
||
59776 | 424 |
|
63861 | 425 |
subsection \<open>@{const heap}\<close> |
60505 | 426 |
|
427 |
||
63861 | 428 |
subsection \<open>@{const mirror}\<close> |
59561 | 429 |
|
430 |
lemma mirror_Leaf[simp]: "mirror t = \<langle>\<rangle> \<longleftrightarrow> t = \<langle>\<rangle>" |
|
431 |
by (induction t) simp_all |
|
432 |
||
433 |
lemma size_mirror[simp]: "size(mirror t) = size t" |
|
434 |
by (induction t) simp_all |
|
435 |
||
436 |
lemma size1_mirror[simp]: "size1(mirror t) = size1 t" |
|
437 |
by (simp add: size1_def) |
|
438 |
||
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
439 |
lemma height_mirror[simp]: "height(mirror t) = height t" |
59776 | 440 |
by (induction t) simp_all |
441 |
||
442 |
lemma inorder_mirror: "inorder(mirror t) = rev(inorder t)" |
|
443 |
by (induction t) simp_all |
|
444 |
||
445 |
lemma map_mirror: "map_tree f (mirror t) = mirror (map_tree f t)" |
|
446 |
by (induction t) simp_all |
|
447 |
||
59561 | 448 |
lemma mirror_mirror[simp]: "mirror(mirror t) = t" |
449 |
by (induction t) simp_all |
|
450 |
||
57250 | 451 |
end |