author | nipkow |
Fri, 12 Aug 2016 18:08:40 +0200 | |
changeset 63665 | 15f48ce7ec23 |
parent 63598 | 025d6e52d86f |
child 63755 | 182c111190e5 |
permissions | -rw-r--r-- |
57250 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
60500 | 3 |
section \<open>Binary Tree\<close> |
57250 | 4 |
|
5 |
theory Tree |
|
6 |
imports Main |
|
7 |
begin |
|
8 |
||
58424 | 9 |
datatype 'a tree = |
62160 | 10 |
is_Leaf: Leaf ("\<langle>\<rangle>") | |
11 |
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
|
12 |
where |
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
13 |
"left Leaf = Leaf" |
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
14 |
| "right Leaf = Leaf" |
57569
e20a999f7161
register tree with datatype_compat ot support QuickCheck
hoelzl
parents:
57530
diff
changeset
|
15 |
datatype_compat tree |
57250 | 16 |
|
60500 | 17 |
text\<open>Can be seen as counting the number of leaves rather than nodes:\<close> |
58438 | 18 |
|
19 |
definition size1 :: "'a tree \<Rightarrow> nat" where |
|
20 |
"size1 t = size t + 1" |
|
21 |
||
22 |
lemma size1_simps[simp]: |
|
23 |
"size1 \<langle>\<rangle> = 1" |
|
24 |
"size1 \<langle>l, x, r\<rangle> = size1 l + size1 r" |
|
25 |
by (simp_all add: size1_def) |
|
26 |
||
62650 | 27 |
lemma size1_ge0[simp]: "0 < size1 t" |
28 |
by (simp add: size1_def) |
|
29 |
||
60507 | 30 |
lemma size_0_iff_Leaf: "size t = 0 \<longleftrightarrow> t = Leaf" |
60505 | 31 |
by(cases t) auto |
32 |
||
58424 | 33 |
lemma neq_Leaf_iff: "(t \<noteq> \<langle>\<rangle>) = (\<exists>l a r. t = \<langle>l, a, r\<rangle>)" |
34 |
by (cases t) auto |
|
57530 | 35 |
|
57687 | 36 |
lemma finite_set_tree[simp]: "finite(set_tree t)" |
37 |
by(induction t) auto |
|
38 |
||
59776 | 39 |
lemma size_map_tree[simp]: "size (map_tree f t) = size t" |
40 |
by (induction t) auto |
|
41 |
||
42 |
lemma size1_map_tree[simp]: "size1 (map_tree f t) = size1 t" |
|
43 |
by (simp add: size1_def) |
|
44 |
||
45 |
||
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
46 |
subsection "The Height" |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
47 |
|
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
48 |
class height = fixes height :: "'a \<Rightarrow> nat" |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
49 |
|
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
50 |
instantiation tree :: (type)height |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
51 |
begin |
59776 | 52 |
|
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
53 |
fun height_tree :: "'a tree => nat" where |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
54 |
"height Leaf = 0" | |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
55 |
"height (Node t1 a t2) = max (height t1) (height t2) + 1" |
59776 | 56 |
|
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
57 |
instance .. |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
58 |
|
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
59 |
end |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
60 |
|
63665 | 61 |
lemma height_0_iff_Leaf: "height t = 0 \<longleftrightarrow> t = Leaf" |
62 |
by(cases t) auto |
|
63 |
||
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
64 |
lemma height_map_tree[simp]: "height (map_tree f t) = height t" |
59776 | 65 |
by (induction t) auto |
66 |
||
62202 | 67 |
lemma size1_height: "size t + 1 \<le> 2 ^ height (t::'a tree)" |
68 |
proof(induction t) |
|
69 |
case (Node l a r) |
|
70 |
show ?case |
|
71 |
proof (cases "height l \<le> height r") |
|
72 |
case True |
|
73 |
have "size(Node l a r) + 1 = (size l + 1) + (size r + 1)" by simp |
|
74 |
also have "size l + 1 \<le> 2 ^ height l" by(rule Node.IH(1)) |
|
75 |
also have "size r + 1 \<le> 2 ^ height r" by(rule Node.IH(2)) |
|
76 |
also have "(2::nat) ^ height l \<le> 2 ^ height r" using True by simp |
|
77 |
finally show ?thesis using True by (auto simp: max_def mult_2) |
|
78 |
next |
|
79 |
case False |
|
80 |
have "size(Node l a r) + 1 = (size l + 1) + (size r + 1)" by simp |
|
81 |
also have "size l + 1 \<le> 2 ^ height l" by(rule Node.IH(1)) |
|
82 |
also have "size r + 1 \<le> 2 ^ height r" by(rule Node.IH(2)) |
|
83 |
also have "(2::nat) ^ height r \<le> 2 ^ height l" using False by simp |
|
84 |
finally show ?thesis using False by (auto simp: max_def mult_2) |
|
85 |
qed |
|
86 |
qed simp |
|
87 |
||
57687 | 88 |
|
63598 | 89 |
fun min_height :: "'a tree \<Rightarrow> nat" where |
90 |
"min_height Leaf = 0" | |
|
91 |
"min_height (Node l _ r) = min (min_height l) (min_height r) + 1" |
|
92 |
||
93 |
lemma min_hight_le_height: "min_height t \<le> height t" |
|
94 |
by(induction t) auto |
|
95 |
||
96 |
lemma min_height_map_tree[simp]: "min_height (map_tree f t) = min_height t" |
|
97 |
by (induction t) auto |
|
98 |
||
99 |
lemma min_height_le_size1: "2 ^ min_height t \<le> size t + 1" |
|
100 |
proof(induction t) |
|
101 |
case (Node l a r) |
|
102 |
have "(2::nat) ^ min_height (Node l a r) \<le> 2 ^ min_height l + 2 ^ min_height r" |
|
103 |
by (simp add: min_def) |
|
104 |
also have "\<dots> \<le> size(Node l a r) + 1" using Node.IH by simp |
|
105 |
finally show ?case . |
|
106 |
qed simp |
|
107 |
||
108 |
||
63036 | 109 |
subsection "Balanced" |
110 |
||
111 |
fun balanced :: "'a tree \<Rightarrow> bool" where |
|
112 |
"balanced Leaf = True" | |
|
113 |
"balanced (Node l x r) = (balanced l \<and> balanced r \<and> height l = height r)" |
|
114 |
||
63598 | 115 |
lemma balanced_iff_min_height: "balanced t \<longleftrightarrow> (min_height t = height t)" |
116 |
apply(induction t) |
|
117 |
apply simp |
|
118 |
apply (simp add: min_def max_def) |
|
119 |
by (metis le_antisym le_trans min_hight_le_height) |
|
120 |
||
63036 | 121 |
lemma balanced_size1: "balanced t \<Longrightarrow> size1 t = 2 ^ height t" |
122 |
by (induction t) auto |
|
123 |
||
124 |
lemma balanced_size: "balanced t \<Longrightarrow> size t = 2 ^ height t - 1" |
|
125 |
using balanced_size1[simplified size1_def] by fastforce |
|
126 |
||
127 |
||
63413 | 128 |
subsection \<open>Path length\<close> |
129 |
||
130 |
text \<open>The internal path length of a tree:\<close> |
|
131 |
||
132 |
fun path_len :: "'a tree \<Rightarrow> nat" where |
|
133 |
"path_len Leaf = 0 " | |
|
134 |
"path_len (Node l _ r) = path_len l + size l + path_len r + size r" |
|
135 |
||
136 |
lemma path_len_if_bal: "balanced t |
|
137 |
\<Longrightarrow> path_len t = (let n = height t in 2 + n*2^n - 2^(n+1))" |
|
138 |
proof(induction t) |
|
139 |
case (Node l x r) |
|
140 |
have *: "2^(n+1) \<le> 2 + n*2^n" for n :: nat |
|
141 |
by(induction n) auto |
|
142 |
have **: "(0::nat) < 2^n" for n :: nat by simp |
|
143 |
let ?h = "height r" |
|
144 |
show ?case using Node *[of ?h] **[of ?h] by (simp add: balanced_size Let_def) |
|
145 |
qed simp |
|
146 |
||
147 |
||
57687 | 148 |
subsection "The set of subtrees" |
149 |
||
57250 | 150 |
fun subtrees :: "'a tree \<Rightarrow> 'a tree set" where |
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
151 |
"subtrees \<langle>\<rangle> = {\<langle>\<rangle>}" | |
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
152 |
"subtrees (\<langle>l, a, r\<rangle>) = insert \<langle>l, a, r\<rangle> (subtrees l \<union> subtrees r)" |
57250 | 153 |
|
58424 | 154 |
lemma set_treeE: "a \<in> set_tree t \<Longrightarrow> \<exists>l r. \<langle>l, a, r\<rangle> \<in> subtrees t" |
155 |
by (induction t)(auto) |
|
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
156 |
|
57450 | 157 |
lemma Node_notin_subtrees_if[simp]: "a \<notin> set_tree t \<Longrightarrow> Node l a r \<notin> subtrees t" |
58424 | 158 |
by (induction t) auto |
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
159 |
|
58424 | 160 |
lemma in_set_tree_if: "\<langle>l, a, r\<rangle> \<in> subtrees t \<Longrightarrow> a \<in> set_tree t" |
161 |
by (metis Node_notin_subtrees_if) |
|
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
162 |
|
57687 | 163 |
|
59776 | 164 |
subsection "List of entries" |
165 |
||
166 |
fun preorder :: "'a tree \<Rightarrow> 'a list" where |
|
167 |
"preorder \<langle>\<rangle> = []" | |
|
168 |
"preorder \<langle>l, x, r\<rangle> = x # preorder l @ preorder r" |
|
57687 | 169 |
|
57250 | 170 |
fun inorder :: "'a tree \<Rightarrow> 'a list" where |
58424 | 171 |
"inorder \<langle>\<rangle> = []" | |
172 |
"inorder \<langle>l, x, r\<rangle> = inorder l @ [x] @ inorder r" |
|
57250 | 173 |
|
57449
f81da03b9ebd
Library/Tree: use datatype_new, bst is an inductive predicate
hoelzl
parents:
57250
diff
changeset
|
174 |
lemma set_inorder[simp]: "set (inorder t) = set_tree t" |
58424 | 175 |
by (induction t) auto |
57250 | 176 |
|
59776 | 177 |
lemma set_preorder[simp]: "set (preorder t) = set_tree t" |
178 |
by (induction t) auto |
|
179 |
||
180 |
lemma length_preorder[simp]: "length (preorder t) = size t" |
|
181 |
by (induction t) auto |
|
182 |
||
183 |
lemma length_inorder[simp]: "length (inorder t) = size t" |
|
184 |
by (induction t) auto |
|
185 |
||
186 |
lemma preorder_map: "preorder (map_tree f t) = map f (preorder t)" |
|
187 |
by (induction t) auto |
|
188 |
||
189 |
lemma inorder_map: "inorder (map_tree f t) = map f (inorder t)" |
|
190 |
by (induction t) auto |
|
191 |
||
57687 | 192 |
|
60500 | 193 |
subsection \<open>Binary Search Tree predicate\<close> |
57250 | 194 |
|
57450 | 195 |
fun (in linorder) bst :: "'a tree \<Rightarrow> bool" where |
58424 | 196 |
"bst \<langle>\<rangle> \<longleftrightarrow> True" | |
197 |
"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)" |
|
57250 | 198 |
|
60500 | 199 |
text\<open>In case there are duplicates:\<close> |
59561 | 200 |
|
201 |
fun (in linorder) bst_eq :: "'a tree \<Rightarrow> bool" where |
|
202 |
"bst_eq \<langle>\<rangle> \<longleftrightarrow> True" | |
|
203 |
"bst_eq \<langle>l,a,r\<rangle> \<longleftrightarrow> |
|
204 |
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)" |
|
205 |
||
59928 | 206 |
lemma (in linorder) bst_eq_if_bst: "bst t \<Longrightarrow> bst_eq t" |
207 |
by (induction t) (auto) |
|
208 |
||
59561 | 209 |
lemma (in linorder) bst_eq_imp_sorted: "bst_eq t \<Longrightarrow> sorted (inorder t)" |
210 |
apply (induction t) |
|
211 |
apply(simp) |
|
212 |
by (fastforce simp: sorted_append sorted_Cons intro: less_imp_le less_trans) |
|
213 |
||
59928 | 214 |
lemma (in linorder) distinct_preorder_if_bst: "bst t \<Longrightarrow> distinct (preorder t)" |
215 |
apply (induction t) |
|
216 |
apply simp |
|
217 |
apply(fastforce elim: order.asym) |
|
218 |
done |
|
219 |
||
220 |
lemma (in linorder) distinct_inorder_if_bst: "bst t \<Longrightarrow> distinct (inorder t)" |
|
221 |
apply (induction t) |
|
222 |
apply simp |
|
223 |
apply(fastforce elim: order.asym) |
|
224 |
done |
|
225 |
||
59776 | 226 |
|
60505 | 227 |
subsection "The heap predicate" |
228 |
||
229 |
fun heap :: "'a::linorder tree \<Rightarrow> bool" where |
|
230 |
"heap Leaf = True" | |
|
231 |
"heap (Node l m r) = |
|
232 |
(heap l \<and> heap r \<and> (\<forall>x \<in> set_tree l \<union> set_tree r. m \<le> x))" |
|
233 |
||
234 |
||
61585 | 235 |
subsection "Function \<open>mirror\<close>" |
59561 | 236 |
|
237 |
fun mirror :: "'a tree \<Rightarrow> 'a tree" where |
|
238 |
"mirror \<langle>\<rangle> = Leaf" | |
|
239 |
"mirror \<langle>l,x,r\<rangle> = \<langle>mirror r, x, mirror l\<rangle>" |
|
240 |
||
241 |
lemma mirror_Leaf[simp]: "mirror t = \<langle>\<rangle> \<longleftrightarrow> t = \<langle>\<rangle>" |
|
242 |
by (induction t) simp_all |
|
243 |
||
244 |
lemma size_mirror[simp]: "size(mirror t) = size t" |
|
245 |
by (induction t) simp_all |
|
246 |
||
247 |
lemma size1_mirror[simp]: "size1(mirror t) = size1 t" |
|
248 |
by (simp add: size1_def) |
|
249 |
||
60808
fd26519b1a6a
depth -> height; removed del_rightmost (too specifi)
nipkow
parents:
60507
diff
changeset
|
250 |
lemma height_mirror[simp]: "height(mirror t) = height t" |
59776 | 251 |
by (induction t) simp_all |
252 |
||
253 |
lemma inorder_mirror: "inorder(mirror t) = rev(inorder t)" |
|
254 |
by (induction t) simp_all |
|
255 |
||
256 |
lemma map_mirror: "map_tree f (mirror t) = mirror (map_tree f t)" |
|
257 |
by (induction t) simp_all |
|
258 |
||
59561 | 259 |
lemma mirror_mirror[simp]: "mirror(mirror t) = t" |
260 |
by (induction t) simp_all |
|
261 |
||
57250 | 262 |
end |