author | paulson <lp15@cam.ac.uk> |
Tue, 18 May 2021 20:25:08 +0100 | |
changeset 73707 | 06aeb9054c07 |
parent 72586 | e3ba2578ad9d |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
57250 | 1 |
(* Author: Tobias Nipkow *) |
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
2 |
(* Todo: minimal ipl of almost complete trees *) |
57250 | 3 |
|
60500 | 4 |
section \<open>Binary Tree\<close> |
57250 | 5 |
|
6 |
theory Tree |
|
7 |
imports Main |
|
8 |
begin |
|
9 |
||
58424 | 10 |
datatype 'a tree = |
64887 | 11 |
Leaf ("\<langle>\<rangle>") | |
69655 | 12 |
Node "'a tree" ("value": 'a) "'a tree" ("(1\<langle>_,/ _,/ _\<rangle>)") |
57569
e20a999f7161
register tree with datatype_compat ot support QuickCheck
hoelzl
parents:
57530
diff
changeset
|
13 |
datatype_compat tree |
57250 | 14 |
|
69218 | 15 |
primrec left :: "'a tree \<Rightarrow> 'a tree" where |
16 |
"left (Node l v r) = l" | |
|
17 |
"left Leaf = Leaf" |
|
18 |
||
19 |
primrec right :: "'a tree \<Rightarrow> 'a tree" where |
|
20 |
"right (Node l v r) = r" | |
|
21 |
"right Leaf = Leaf" |
|
22 |
||
68998 | 23 |
text\<open>Counting the number of leaves rather than nodes:\<close> |
58438 | 24 |
|
68998 | 25 |
fun size1 :: "'a tree \<Rightarrow> nat" where |
26 |
"size1 \<langle>\<rangle> = 1" | |
|
27 |
"size1 \<langle>l, x, r\<rangle> = size1 l + size1 r" |
|
58438 | 28 |
|
63861 | 29 |
fun subtrees :: "'a tree \<Rightarrow> 'a tree set" where |
30 |
"subtrees \<langle>\<rangle> = {\<langle>\<rangle>}" | |
|
72586 | 31 |
"subtrees (\<langle>l, a, r\<rangle>) = {\<langle>l, a, r\<rangle>} \<union> subtrees l \<union> subtrees r" |
63861 | 32 |
|
33 |
fun mirror :: "'a tree \<Rightarrow> 'a tree" where |
|
34 |
"mirror \<langle>\<rangle> = Leaf" | |
|
35 |
"mirror \<langle>l,x,r\<rangle> = \<langle>mirror r, x, mirror l\<rangle>" |
|
36 |
||
37 |
class height = fixes height :: "'a \<Rightarrow> nat" |
|
38 |
||
39 |
instantiation tree :: (type)height |
|
40 |
begin |
|
41 |
||
42 |
fun height_tree :: "'a tree => nat" where |
|
43 |
"height Leaf = 0" | |
|
68999 | 44 |
"height (Node l a r) = max (height l) (height r) + 1" |
63861 | 45 |
|
46 |
instance .. |
|
47 |
||
48 |
end |
|
49 |
||
50 |
fun min_height :: "'a tree \<Rightarrow> nat" where |
|
51 |
"min_height Leaf = 0" | |
|
52 |
"min_height (Node l _ r) = min (min_height l) (min_height r) + 1" |
|
53 |
||
54 |
fun complete :: "'a tree \<Rightarrow> bool" where |
|
55 |
"complete Leaf = True" | |
|
72586 | 56 |
"complete (Node l x r) = (height l = height r \<and> complete l \<and> complete r)" |
63861 | 57 |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
58 |
text \<open>Almost complete:\<close> |
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
59 |
definition acomplete :: "'a tree \<Rightarrow> bool" where |
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
60 |
"acomplete t = (height t - min_height t \<le> 1)" |
63861 | 61 |
|
62 |
text \<open>Weight balanced:\<close> |
|
63 |
fun wbalanced :: "'a tree \<Rightarrow> bool" where |
|
64 |
"wbalanced Leaf = True" | |
|
65 |
"wbalanced (Node l x r) = (abs(int(size l) - int(size r)) \<le> 1 \<and> wbalanced l \<and> wbalanced r)" |
|
66 |
||
67 |
text \<open>Internal path length:\<close> |
|
64887 | 68 |
fun ipl :: "'a tree \<Rightarrow> nat" where |
69 |
"ipl Leaf = 0 " | |
|
70 |
"ipl (Node l _ r) = ipl l + size l + ipl r + size r" |
|
63861 | 71 |
|
72 |
fun preorder :: "'a tree \<Rightarrow> 'a list" where |
|
73 |
"preorder \<langle>\<rangle> = []" | |
|
74 |
"preorder \<langle>l, x, r\<rangle> = x # preorder l @ preorder r" |
|
75 |
||
76 |
fun inorder :: "'a tree \<Rightarrow> 'a list" where |
|
77 |
"inorder \<langle>\<rangle> = []" | |
|
78 |
"inorder \<langle>l, x, r\<rangle> = inorder l @ [x] @ inorder r" |
|
79 |
||
80 |
text\<open>A linear version avoiding append:\<close> |
|
81 |
fun inorder2 :: "'a tree \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
82 |
"inorder2 \<langle>\<rangle> xs = xs" | |
|
83 |
"inorder2 \<langle>l, x, r\<rangle> xs = inorder2 l (x # inorder2 r xs)" |
|
84 |
||
64925 | 85 |
fun postorder :: "'a tree \<Rightarrow> 'a list" where |
86 |
"postorder \<langle>\<rangle> = []" | |
|
87 |
"postorder \<langle>l, x, r\<rangle> = postorder l @ postorder r @ [x]" |
|
88 |
||
63861 | 89 |
text\<open>Binary Search Tree:\<close> |
66606 | 90 |
fun bst_wrt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a tree \<Rightarrow> bool" where |
91 |
"bst_wrt P \<langle>\<rangle> \<longleftrightarrow> True" | |
|
92 |
"bst_wrt P \<langle>l, a, r\<rangle> \<longleftrightarrow> |
|
72586 | 93 |
(\<forall>x\<in>set_tree l. P x a) \<and> (\<forall>x\<in>set_tree r. P a x) \<and> bst_wrt P l \<and> bst_wrt P r" |
63861 | 94 |
|
66606 | 95 |
abbreviation bst :: "('a::linorder) tree \<Rightarrow> bool" where |
67399 | 96 |
"bst \<equiv> bst_wrt (<)" |
63861 | 97 |
|
98 |
fun (in linorder) heap :: "'a tree \<Rightarrow> bool" where |
|
99 |
"heap Leaf = True" | |
|
100 |
"heap (Node l m r) = |
|
72586 | 101 |
((\<forall>x \<in> set_tree l \<union> set_tree r. m \<le> x) \<and> heap l \<and> heap r)" |
63861 | 102 |
|
103 |
||
69593 | 104 |
subsection \<open>\<^const>\<open>map_tree\<close>\<close> |
65339 | 105 |
|
65340 | 106 |
lemma eq_map_tree_Leaf[simp]: "map_tree f t = Leaf \<longleftrightarrow> t = Leaf" |
65339 | 107 |
by (rule tree.map_disc_iff) |
108 |
||
65340 | 109 |
lemma eq_Leaf_map_tree[simp]: "Leaf = map_tree f t \<longleftrightarrow> t = Leaf" |
65339 | 110 |
by (cases t) auto |
111 |
||
112 |
||
69593 | 113 |
subsection \<open>\<^const>\<open>size\<close>\<close> |
63861 | 114 |
|
68998 | 115 |
lemma size1_size: "size1 t = size t + 1" |
116 |
by (induction t) simp_all |
|
58438 | 117 |
|
62650 | 118 |
lemma size1_ge0[simp]: "0 < size1 t" |
68998 | 119 |
by (simp add: size1_size) |
62650 | 120 |
|
65340 | 121 |
lemma eq_size_0[simp]: "size t = 0 \<longleftrightarrow> t = Leaf" |
65339 | 122 |
by(cases t) auto |
123 |
||
65340 | 124 |
lemma eq_0_size[simp]: "0 = size t \<longleftrightarrow> t = Leaf" |
60505 | 125 |
by(cases t) auto |
126 |
||
58424 | 127 |
lemma neq_Leaf_iff: "(t \<noteq> \<langle>\<rangle>) = (\<exists>l a r. t = \<langle>l, a, r\<rangle>)" |
128 |
by (cases t) auto |
|
57530 | 129 |
|
59776 | 130 |
lemma size_map_tree[simp]: "size (map_tree f t) = size t" |
131 |
by (induction t) auto |
|
132 |
||
133 |
lemma size1_map_tree[simp]: "size1 (map_tree f t) = size1 t" |
|
68998 | 134 |
by (simp add: size1_size) |
59776 | 135 |
|
136 |
||
69593 | 137 |
subsection \<open>\<^const>\<open>set_tree\<close>\<close> |
65339 | 138 |
|
65340 | 139 |
lemma eq_set_tree_empty[simp]: "set_tree t = {} \<longleftrightarrow> t = Leaf" |
65339 | 140 |
by (cases t) auto |
141 |
||
65340 | 142 |
lemma eq_empty_set_tree[simp]: "{} = set_tree t \<longleftrightarrow> t = Leaf" |
65339 | 143 |
by (cases t) auto |
144 |
||
145 |
lemma finite_set_tree[simp]: "finite(set_tree t)" |
|
146 |
by(induction t) auto |
|
147 |
||
148 |
||
69593 | 149 |
subsection \<open>\<^const>\<open>subtrees\<close>\<close> |
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
150 |
|
65340 | 151 |
lemma neq_subtrees_empty[simp]: "subtrees t \<noteq> {}" |
152 |
by (cases t)(auto) |
|
153 |
||
154 |
lemma neq_empty_subtrees[simp]: "{} \<noteq> subtrees t" |
|
155 |
by (cases t)(auto) |
|
156 |
||
72313 | 157 |
lemma size_subtrees: "s \<in> subtrees t \<Longrightarrow> size s \<le> size t" |
158 |
by(induction t)(auto) |
|
159 |
||
63861 | 160 |
lemma set_treeE: "a \<in> set_tree t \<Longrightarrow> \<exists>l r. \<langle>l, a, r\<rangle> \<in> subtrees t" |
161 |
by (induction t)(auto) |
|
59776 | 162 |
|
63861 | 163 |
lemma Node_notin_subtrees_if[simp]: "a \<notin> set_tree t \<Longrightarrow> Node l a r \<notin> subtrees t" |
164 |
by (induction t) auto |
|
59776 | 165 |
|
63861 | 166 |
lemma in_set_tree_if: "\<langle>l, a, r\<rangle> \<in> subtrees t \<Longrightarrow> a \<in> set_tree t" |
167 |
by (metis Node_notin_subtrees_if) |
|
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
168 |
|
63861 | 169 |
|
69593 | 170 |
subsection \<open>\<^const>\<open>height\<close> and \<^const>\<open>min_height\<close>\<close> |
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
171 |
|
65340 | 172 |
lemma eq_height_0[simp]: "height t = 0 \<longleftrightarrow> t = Leaf" |
65339 | 173 |
by(cases t) auto |
174 |
||
65340 | 175 |
lemma eq_0_height[simp]: "0 = height t \<longleftrightarrow> t = Leaf" |
63665 | 176 |
by(cases t) auto |
177 |
||
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
178 |
lemma height_map_tree[simp]: "height (map_tree f t) = height t" |
59776 | 179 |
by (induction t) auto |
180 |
||
64414 | 181 |
lemma height_le_size_tree: "height t \<le> size (t::'a tree)" |
182 |
by (induction t) auto |
|
183 |
||
64533 | 184 |
lemma size1_height: "size1 t \<le> 2 ^ height (t::'a tree)" |
62202 | 185 |
proof(induction t) |
186 |
case (Node l a r) |
|
187 |
show ?case |
|
188 |
proof (cases "height l \<le> height r") |
|
189 |
case True |
|
64533 | 190 |
have "size1(Node l a r) = size1 l + size1 r" by simp |
64918 | 191 |
also have "\<dots> \<le> 2 ^ height l + 2 ^ height r" using Node.IH by arith |
192 |
also have "\<dots> \<le> 2 ^ height r + 2 ^ height r" using True by simp |
|
64922 | 193 |
also have "\<dots> = 2 ^ height (Node l a r)" |
64918 | 194 |
using True by (auto simp: max_def mult_2) |
195 |
finally show ?thesis . |
|
62202 | 196 |
next |
197 |
case False |
|
64533 | 198 |
have "size1(Node l a r) = size1 l + size1 r" by simp |
64918 | 199 |
also have "\<dots> \<le> 2 ^ height l + 2 ^ height r" using Node.IH by arith |
200 |
also have "\<dots> \<le> 2 ^ height l + 2 ^ height l" using False by simp |
|
62202 | 201 |
finally show ?thesis using False by (auto simp: max_def mult_2) |
202 |
qed |
|
203 |
qed simp |
|
204 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
205 |
corollary size_height: "size t \<le> 2 ^ height (t::'a tree) - 1" |
68998 | 206 |
using size1_height[of t, unfolded size1_size] by(arith) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
207 |
|
63861 | 208 |
lemma height_subtrees: "s \<in> subtrees t \<Longrightarrow> height s \<le> height t" |
209 |
by (induction t) auto |
|
57687 | 210 |
|
63598 | 211 |
|
64540 | 212 |
lemma min_height_le_height: "min_height t \<le> height t" |
63598 | 213 |
by(induction t) auto |
214 |
||
215 |
lemma min_height_map_tree[simp]: "min_height (map_tree f t) = min_height t" |
|
216 |
by (induction t) auto |
|
217 |
||
64533 | 218 |
lemma min_height_size1: "2 ^ min_height t \<le> size1 t" |
63598 | 219 |
proof(induction t) |
220 |
case (Node l a r) |
|
221 |
have "(2::nat) ^ min_height (Node l a r) \<le> 2 ^ min_height l + 2 ^ min_height r" |
|
222 |
by (simp add: min_def) |
|
64533 | 223 |
also have "\<dots> \<le> size1(Node l a r)" using Node.IH by simp |
63598 | 224 |
finally show ?case . |
225 |
qed simp |
|
226 |
||
227 |
||
69593 | 228 |
subsection \<open>\<^const>\<open>complete\<close>\<close> |
63036 | 229 |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
230 |
lemma complete_iff_height: "complete t \<longleftrightarrow> (min_height t = height t)" |
63598 | 231 |
apply(induction t) |
232 |
apply simp |
|
233 |
apply (simp add: min_def max_def) |
|
64540 | 234 |
by (metis le_antisym le_trans min_height_le_height) |
63598 | 235 |
|
63770 | 236 |
lemma size1_if_complete: "complete t \<Longrightarrow> size1 t = 2 ^ height t" |
63036 | 237 |
by (induction t) auto |
238 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
239 |
lemma size_if_complete: "complete t \<Longrightarrow> size t = 2 ^ height t - 1" |
68998 | 240 |
using size1_if_complete[simplified size1_size] by fastforce |
63770 | 241 |
|
69117 | 242 |
lemma size1_height_if_incomplete: |
243 |
"\<not> complete t \<Longrightarrow> size1 t < 2 ^ height t" |
|
244 |
proof(induction t) |
|
245 |
case Leaf thus ?case by simp |
|
63770 | 246 |
next |
69117 | 247 |
case (Node l x r) |
248 |
have 1: ?case if h: "height l < height r" |
|
249 |
using h size1_height[of l] size1_height[of r] power_strict_increasing[OF h, of "2::nat"] |
|
250 |
by(auto simp: max_def simp del: power_strict_increasing_iff) |
|
251 |
have 2: ?case if h: "height l > height r" |
|
252 |
using h size1_height[of l] size1_height[of r] power_strict_increasing[OF h, of "2::nat"] |
|
253 |
by(auto simp: max_def simp del: power_strict_increasing_iff) |
|
254 |
have 3: ?case if h: "height l = height r" and c: "\<not> complete l" |
|
255 |
using h size1_height[of r] Node.IH(1)[OF c] by(simp) |
|
256 |
have 4: ?case if h: "height l = height r" and c: "\<not> complete r" |
|
257 |
using h size1_height[of l] Node.IH(2)[OF c] by(simp) |
|
258 |
from 1 2 3 4 Node.prems show ?case apply (simp add: max_def) by linarith |
|
63770 | 259 |
qed |
260 |
||
69117 | 261 |
lemma complete_iff_min_height: "complete t \<longleftrightarrow> (height t = min_height t)" |
262 |
by(auto simp add: complete_iff_height) |
|
263 |
||
264 |
lemma min_height_size1_if_incomplete: |
|
265 |
"\<not> complete t \<Longrightarrow> 2 ^ min_height t < size1 t" |
|
266 |
proof(induction t) |
|
267 |
case Leaf thus ?case by simp |
|
268 |
next |
|
269 |
case (Node l x r) |
|
270 |
have 1: ?case if h: "min_height l < min_height r" |
|
271 |
using h min_height_size1[of l] min_height_size1[of r] power_strict_increasing[OF h, of "2::nat"] |
|
272 |
by(auto simp: max_def simp del: power_strict_increasing_iff) |
|
273 |
have 2: ?case if h: "min_height l > min_height r" |
|
274 |
using h min_height_size1[of l] min_height_size1[of r] power_strict_increasing[OF h, of "2::nat"] |
|
275 |
by(auto simp: max_def simp del: power_strict_increasing_iff) |
|
276 |
have 3: ?case if h: "min_height l = min_height r" and c: "\<not> complete l" |
|
277 |
using h min_height_size1[of r] Node.IH(1)[OF c] by(simp add: complete_iff_min_height) |
|
278 |
have 4: ?case if h: "min_height l = min_height r" and c: "\<not> complete r" |
|
279 |
using h min_height_size1[of l] Node.IH(2)[OF c] by(simp add: complete_iff_min_height) |
|
280 |
from 1 2 3 4 Node.prems show ?case |
|
281 |
by (fastforce simp: complete_iff_min_height[THEN iffD1]) |
|
282 |
qed |
|
283 |
||
284 |
lemma complete_if_size1_height: "size1 t = 2 ^ height t \<Longrightarrow> complete t" |
|
285 |
using size1_height_if_incomplete by fastforce |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
286 |
|
64533 | 287 |
lemma complete_if_size1_min_height: "size1 t = 2 ^ min_height t \<Longrightarrow> complete t" |
69117 | 288 |
using min_height_size1_if_incomplete by fastforce |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
289 |
|
64533 | 290 |
lemma complete_iff_size1: "complete t \<longleftrightarrow> size1 t = 2 ^ height t" |
291 |
using complete_if_size1_height size1_if_complete by blast |
|
292 |
||
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
293 |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
294 |
subsection \<open>\<^const>\<open>acomplete\<close>\<close> |
63861 | 295 |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
296 |
lemma acomplete_subtreeL: "acomplete (Node l x r) \<Longrightarrow> acomplete l" |
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
297 |
by(simp add: acomplete_def) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
298 |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
299 |
lemma acomplete_subtreeR: "acomplete (Node l x r) \<Longrightarrow> acomplete r" |
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
300 |
by(simp add: acomplete_def) |
63861 | 301 |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
302 |
lemma acomplete_subtrees: "\<lbrakk> acomplete t; s \<in> subtrees t \<rbrakk> \<Longrightarrow> acomplete s" |
63861 | 303 |
using [[simp_depth_limit=1]] |
304 |
by(induction t arbitrary: s) |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
305 |
(auto simp add: acomplete_subtreeL acomplete_subtreeR) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
306 |
|
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
307 |
text\<open>Balanced trees have optimal height:\<close> |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
308 |
|
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
309 |
lemma acomplete_optimal: |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
310 |
fixes t :: "'a tree" and t' :: "'b tree" |
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
311 |
assumes "acomplete t" "size t \<le> size t'" shows "height t \<le> height t'" |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
312 |
proof (cases "complete t") |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
313 |
case True |
64924 | 314 |
have "(2::nat) ^ height t \<le> 2 ^ height t'" |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
315 |
proof - |
64924 | 316 |
have "2 ^ height t = size1 t" |
69115 | 317 |
using True by (simp add: size1_if_complete) |
68998 | 318 |
also have "\<dots> \<le> size1 t'" using assms(2) by(simp add: size1_size) |
64924 | 319 |
also have "\<dots> \<le> 2 ^ height t'" by (rule size1_height) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
320 |
finally show ?thesis . |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
321 |
qed |
64924 | 322 |
thus ?thesis by (simp) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
323 |
next |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
324 |
case False |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
325 |
have "(2::nat) ^ min_height t < 2 ^ height t'" |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
326 |
proof - |
64533 | 327 |
have "(2::nat) ^ min_height t < size1 t" |
328 |
by(rule min_height_size1_if_incomplete[OF False]) |
|
68998 | 329 |
also have "\<dots> \<le> size1 t'" using assms(2) by (simp add: size1_size) |
64918 | 330 |
also have "\<dots> \<le> 2 ^ height t'" by(rule size1_height) |
331 |
finally have "(2::nat) ^ min_height t < (2::nat) ^ height t'" . |
|
64924 | 332 |
thus ?thesis . |
63755
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 |
hence *: "min_height t < height t'" by simp |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
335 |
have "min_height t + 1 = height t" |
64540 | 336 |
using min_height_le_height[of t] assms(1) False |
72566
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
nipkow
parents:
72313
diff
changeset
|
337 |
by (simp add: complete_iff_height acomplete_def) |
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
338 |
with * show ?thesis by arith |
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63665
diff
changeset
|
339 |
qed |
63036 | 340 |
|
341 |
||
69593 | 342 |
subsection \<open>\<^const>\<open>wbalanced\<close>\<close> |
63861 | 343 |
|
344 |
lemma wbalanced_subtrees: "\<lbrakk> wbalanced t; s \<in> subtrees t \<rbrakk> \<Longrightarrow> wbalanced s" |
|
345 |
using [[simp_depth_limit=1]] by(induction t arbitrary: s) auto |
|
346 |
||
347 |
||
69593 | 348 |
subsection \<open>\<^const>\<open>ipl\<close>\<close> |
63413 | 349 |
|
350 |
text \<open>The internal path length of a tree:\<close> |
|
351 |
||
64923 | 352 |
lemma ipl_if_complete_int: |
353 |
"complete t \<Longrightarrow> int(ipl t) = (int(height t) - 2) * 2^(height t) + 2" |
|
354 |
apply(induction t) |
|
355 |
apply simp |
|
356 |
apply simp |
|
357 |
apply (simp add: algebra_simps size_if_complete of_nat_diff) |
|
358 |
done |
|
63413 | 359 |
|
360 |
||
59776 | 361 |
subsection "List of entries" |
362 |
||
65340 | 363 |
lemma eq_inorder_Nil[simp]: "inorder t = [] \<longleftrightarrow> t = Leaf" |
65339 | 364 |
by (cases t) auto |
365 |
||
65340 | 366 |
lemma eq_Nil_inorder[simp]: "[] = inorder t \<longleftrightarrow> t = Leaf" |
65339 | 367 |
by (cases t) auto |
368 |
||
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
369 |
lemma set_inorder[simp]: "set (inorder t) = set_tree t" |
58424 | 370 |
by (induction t) auto |
57250 | 371 |
|
59776 | 372 |
lemma set_preorder[simp]: "set (preorder t) = set_tree t" |
373 |
by (induction t) auto |
|
374 |
||
64925 | 375 |
lemma set_postorder[simp]: "set (postorder t) = set_tree t" |
376 |
by (induction t) auto |
|
377 |
||
59776 | 378 |
lemma length_preorder[simp]: "length (preorder t) = size t" |
379 |
by (induction t) auto |
|
380 |
||
381 |
lemma length_inorder[simp]: "length (inorder t) = size t" |
|
382 |
by (induction t) auto |
|
383 |
||
64925 | 384 |
lemma length_postorder[simp]: "length (postorder t) = size t" |
385 |
by (induction t) auto |
|
386 |
||
59776 | 387 |
lemma preorder_map: "preorder (map_tree f t) = map f (preorder t)" |
388 |
by (induction t) auto |
|
389 |
||
390 |
lemma inorder_map: "inorder (map_tree f t) = map f (inorder t)" |
|
391 |
by (induction t) auto |
|
392 |
||
64925 | 393 |
lemma postorder_map: "postorder (map_tree f t) = map f (postorder t)" |
394 |
by (induction t) auto |
|
395 |
||
63765 | 396 |
lemma inorder2_inorder: "inorder2 t xs = inorder t @ xs" |
397 |
by (induction t arbitrary: xs) auto |
|
398 |
||
57687 | 399 |
|
63861 | 400 |
subsection \<open>Binary Search Tree\<close> |
59561 | 401 |
|
66606 | 402 |
lemma bst_wrt_mono: "(\<And>x y. P x y \<Longrightarrow> Q x y) \<Longrightarrow> bst_wrt P t \<Longrightarrow> bst_wrt Q t" |
59928 | 403 |
by (induction t) (auto) |
404 |
||
67399 | 405 |
lemma bst_wrt_le_if_bst: "bst t \<Longrightarrow> bst_wrt (\<le>) t" |
66606 | 406 |
using bst_wrt_mono less_imp_le by blast |
407 |
||
67399 | 408 |
lemma bst_wrt_le_iff_sorted: "bst_wrt (\<le>) t \<longleftrightarrow> sorted (inorder t)" |
59561 | 409 |
apply (induction t) |
410 |
apply(simp) |
|
68109 | 411 |
by (fastforce simp: sorted_append intro: less_imp_le less_trans) |
59561 | 412 |
|
67399 | 413 |
lemma bst_iff_sorted_wrt_less: "bst t \<longleftrightarrow> sorted_wrt (<) (inorder t)" |
59928 | 414 |
apply (induction t) |
415 |
apply simp |
|
68109 | 416 |
apply (fastforce simp: sorted_wrt_append) |
59928 | 417 |
done |
418 |
||
59776 | 419 |
|
69593 | 420 |
subsection \<open>\<^const>\<open>heap\<close>\<close> |
60505 | 421 |
|
422 |
||
69593 | 423 |
subsection \<open>\<^const>\<open>mirror\<close>\<close> |
59561 | 424 |
|
425 |
lemma mirror_Leaf[simp]: "mirror t = \<langle>\<rangle> \<longleftrightarrow> t = \<langle>\<rangle>" |
|
426 |
by (induction t) simp_all |
|
427 |
||
65339 | 428 |
lemma Leaf_mirror[simp]: "\<langle>\<rangle> = mirror t \<longleftrightarrow> t = \<langle>\<rangle>" |
429 |
using mirror_Leaf by fastforce |
|
430 |
||
59561 | 431 |
lemma size_mirror[simp]: "size(mirror t) = size t" |
432 |
by (induction t) simp_all |
|
433 |
||
434 |
lemma size1_mirror[simp]: "size1(mirror t) = size1 t" |
|
68998 | 435 |
by (simp add: size1_size) |
59561 | 436 |
|
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
437 |
lemma height_mirror[simp]: "height(mirror t) = height t" |
59776 | 438 |
by (induction t) simp_all |
439 |
||
66659 | 440 |
lemma min_height_mirror [simp]: "min_height (mirror t) = min_height t" |
441 |
by (induction t) simp_all |
|
442 |
||
443 |
lemma ipl_mirror [simp]: "ipl (mirror t) = ipl t" |
|
444 |
by (induction t) simp_all |
|
445 |
||
59776 | 446 |
lemma inorder_mirror: "inorder(mirror t) = rev(inorder t)" |
447 |
by (induction t) simp_all |
|
448 |
||
449 |
lemma map_mirror: "map_tree f (mirror t) = mirror (map_tree f t)" |
|
450 |
by (induction t) simp_all |
|
451 |
||
59561 | 452 |
lemma mirror_mirror[simp]: "mirror(mirror t) = t" |
453 |
by (induction t) simp_all |
|
454 |
||
57250 | 455 |
end |