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