| author | wenzelm | 
| Thu, 06 Oct 2016 11:13:12 +0200 | |
| changeset 64062 | a7352cbde7d7 | 
| parent 63861 | 90360390a916 | 
| child 64018 | c6eb691770d8 | 
| permissions | -rw-r--r-- | 
| 63829 | 1 | (* Author: Tobias Nipkow *) | 
| 63643 | 2 | |
| 63829 | 3 | section \<open>Creating Balanced Trees\<close> | 
| 63643 | 4 | |
| 63829 | 5 | theory Balance | 
| 63643 | 6 | imports | 
| 7 | "~~/src/HOL/Library/Tree" | |
| 63663 
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
 nipkow parents: 
63643diff
changeset | 8 | "~~/src/HOL/Library/Log_Nat" | 
| 63643 | 9 | begin | 
| 10 | ||
| 11 | fun bal :: "'a list \<Rightarrow> nat \<Rightarrow> 'a tree * 'a list" where | |
| 12 | "bal xs n = (if n=0 then (Leaf,xs) else | |
| 13 | (let m = n div 2; | |
| 14 | (l, ys) = bal xs m; | |
| 15 | (r, zs) = bal (tl ys) (n-1-m) | |
| 16 | in (Node l (hd ys) r, zs)))" | |
| 17 | ||
| 18 | declare bal.simps[simp del] | |
| 19 | ||
| 63829 | 20 | definition balance_list :: "'a list \<Rightarrow> 'a tree" where | 
| 21 | "balance_list xs = fst (bal xs (length xs))" | |
| 22 | ||
| 23 | definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where | |
| 24 | "balance_tree = balance_list o inorder" | |
| 25 | ||
| 63843 | 26 | lemma bal_simps: | 
| 27 | "bal xs 0 = (Leaf, xs)" | |
| 28 | "n > 0 \<Longrightarrow> | |
| 29 | bal xs n = | |
| 30 | (let m = n div 2; | |
| 31 | (l, ys) = Balance.bal xs m; | |
| 32 | (r, zs) = Balance.bal (tl ys) (n-1-m) | |
| 33 | in (Node l (hd ys) r, zs))" | |
| 34 | by(simp_all add: bal.simps) | |
| 63643 | 35 | |
| 63861 | 36 | text\<open>The following lemmas take advantage of the fact | 
| 37 | that \<open>bal xs n\<close> yields a result even if \<open>n > length xs\<close>.\<close> | |
| 38 | ||
| 39 | lemma size_bal: "bal xs n = (t,ys) \<Longrightarrow> size t = n" | |
| 40 | proof(induction xs n arbitrary: t ys rule: bal.induct) | |
| 41 | case (1 xs n) | |
| 42 | thus ?case | |
| 43 | by(cases "n=0") | |
| 44 | (auto simp add: bal_simps Let_def split: prod.splits) | |
| 45 | qed | |
| 46 | ||
| 63643 | 47 | lemma bal_inorder: | 
| 63755 
182c111190e5
Renamed balanced to complete; added balanced; more about both
 nipkow parents: 
63663diff
changeset | 48 | "\<lbrakk> bal xs n = (t,ys); n \<le> length xs \<rbrakk> | 
| 
182c111190e5
Renamed balanced to complete; added balanced; more about both
 nipkow parents: 
63663diff
changeset | 49 | \<Longrightarrow> inorder t = take n xs \<and> ys = drop n xs" | 
| 63643 | 50 | proof(induction xs n arbitrary: t ys rule: bal.induct) | 
| 51 | case (1 xs n) show ?case | |
| 52 | proof cases | |
| 63843 | 53 | assume "n = 0" thus ?thesis using 1 by (simp add: bal_simps) | 
| 63643 | 54 | next | 
| 55 | assume [arith]: "n \<noteq> 0" | |
| 56 | let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1" | |
| 57 | from "1.prems" obtain l r xs' where | |
| 58 | b1: "bal xs ?n1 = (l,xs')" and | |
| 59 | b2: "bal (tl xs') ?n2 = (r,ys)" and | |
| 60 | t: "t = \<langle>l, hd xs', r\<rangle>" | |
| 63843 | 61 | by(auto simp: Let_def bal_simps split: prod.splits) | 
| 63643 | 62 | have IH1: "inorder l = take ?n1 xs \<and> xs' = drop ?n1 xs" | 
| 63 | using b1 "1.prems" by(intro "1.IH"(1)) auto | |
| 64 | have IH2: "inorder r = take ?n2 (tl xs') \<and> ys = drop ?n2 (tl xs')" | |
| 65 | using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto | |
| 66 | have "drop (n div 2) xs \<noteq> []" using "1.prems"(2) by simp | |
| 67 | hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)" | |
| 68 | by (metis Suc_eq_plus1 take_Suc) | |
| 69 | hence *: "inorder t = take n xs" using t IH1 IH2 | |
| 70 | using take_add[of ?n1 "?n2+1" xs] by(simp) | |
| 71 | have "n - n div 2 + n div 2 = n" by simp | |
| 72 | hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) | |
| 73 | thus ?thesis using * by blast | |
| 74 | qed | |
| 75 | qed | |
| 76 | ||
| 63829 | 77 | corollary inorder_balance_list: "inorder(balance_list xs) = xs" | 
| 63643 | 78 | using bal_inorder[of xs "length xs"] | 
| 63829 | 79 | by (metis balance_list_def order_refl prod.collapse take_all) | 
| 63643 | 80 | |
| 81 | lemma bal_height: "bal xs n = (t,ys) \<Longrightarrow> height t = floorlog 2 n" | |
| 82 | proof(induction xs n arbitrary: t ys rule: bal.induct) | |
| 83 | case (1 xs n) show ?case | |
| 84 | proof cases | |
| 85 | assume "n = 0" thus ?thesis | |
| 63843 | 86 | using "1.prems" by (simp add: floorlog_def bal_simps) | 
| 63643 | 87 | next | 
| 88 | assume [arith]: "n \<noteq> 0" | |
| 89 | from "1.prems" obtain l r xs' where | |
| 90 | b1: "bal xs (n div 2) = (l,xs')" and | |
| 91 | b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and | |
| 92 | t: "t = \<langle>l, hd xs', r\<rangle>" | |
| 63843 | 93 | by(auto simp: bal_simps Let_def split: prod.splits) | 
| 63643 | 94 | let ?log1 = "floorlog 2 (n div 2)" | 
| 95 | let ?log2 = "floorlog 2 (n - 1 - n div 2)" | |
| 96 | have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp | |
| 97 | have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp | |
| 98 | have "n div 2 \<ge> n - 1 - n div 2" by arith | |
| 99 | hence le: "?log2 \<le> ?log1" by(simp add:floorlog_mono) | |
| 100 | have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) | |
| 101 | also have "\<dots> = ?log1 + 1" using le by (simp add: max_absorb1) | |
| 63663 
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
 nipkow parents: 
63643diff
changeset | 102 | also have "\<dots> = floorlog 2 n" by (simp add: compute_floorlog) | 
| 63643 | 103 | finally show ?thesis . | 
| 104 | qed | |
| 105 | qed | |
| 106 | ||
| 107 | lemma bal_min_height: | |
| 108 | "bal xs n = (t,ys) \<Longrightarrow> min_height t = floorlog 2 (n + 1) - 1" | |
| 109 | proof(induction xs n arbitrary: t ys rule: bal.induct) | |
| 110 | case (1 xs n) show ?case | |
| 111 | proof cases | |
| 112 | assume "n = 0" thus ?thesis | |
| 63843 | 113 | using "1.prems" by (simp add: floorlog_def bal_simps) | 
| 63643 | 114 | next | 
| 115 | assume [arith]: "n \<noteq> 0" | |
| 116 | from "1.prems" obtain l r xs' where | |
| 117 | b1: "bal xs (n div 2) = (l,xs')" and | |
| 118 | b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and | |
| 119 | t: "t = \<langle>l, hd xs', r\<rangle>" | |
| 63843 | 120 | by(auto simp: bal_simps Let_def split: prod.splits) | 
| 63643 | 121 | let ?log1 = "floorlog 2 (n div 2 + 1) - 1" | 
| 122 | let ?log2 = "floorlog 2 (n - 1 - n div 2 + 1) - 1" | |
| 123 | let ?log2' = "floorlog 2 (n - n div 2) - 1" | |
| 124 | have "n - 1 - n div 2 + 1 = n - n div 2" by arith | |
| 125 | hence IH2: "min_height r = ?log2'" using "1.IH"(2) b1 b2 by simp | |
| 126 | have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp | |
| 127 | have *: "floorlog 2 (n - n div 2) \<ge> 1" by (simp add: floorlog_def) | |
| 128 | have "n div 2 + 1 \<ge> n - n div 2" by arith | |
| 129 | with * have le: "?log2' \<le> ?log1" by(simp add: floorlog_mono diff_le_mono) | |
| 130 | have "min_height t = min ?log1 ?log2' + 1" by (simp add: t IH1 IH2) | |
| 131 | also have "\<dots> = ?log2' + 1" using le by (simp add: min_absorb2) | |
| 132 | also have "\<dots> = floorlog 2 (n - n div 2)" by(simp add: floorlog_def) | |
| 133 | also have "n - n div 2 = (n+1) div 2" by arith | |
| 134 | also have "floorlog 2 \<dots> = floorlog 2 (n+1) - 1" | |
| 63663 
28d1deca302e
Extracted floorlog and bitlen to separate theory Log_Nat
 nipkow parents: 
63643diff
changeset | 135 | by (simp add: compute_floorlog) | 
| 63643 | 136 | finally show ?thesis . | 
| 137 | qed | |
| 138 | qed | |
| 139 | ||
| 140 | lemma balanced_bal: | |
| 63755 
182c111190e5
Renamed balanced to complete; added balanced; more about both
 nipkow parents: 
63663diff
changeset | 141 | assumes "bal xs n = (t,ys)" shows "balanced t" | 
| 63643 | 142 | proof - | 
| 143 | have "floorlog 2 n \<le> floorlog 2 (n+1)" by (rule floorlog_mono) auto | |
| 63829 | 144 | thus ?thesis unfolding balanced_def | 
| 145 | using bal_height[OF assms] bal_min_height[OF assms] by linarith | |
| 63643 | 146 | qed | 
| 147 | ||
| 63829 | 148 | corollary size_balance_list[simp]: "size(balance_list xs) = length xs" | 
| 149 | by (metis inorder_balance_list length_inorder) | |
| 150 | ||
| 151 | corollary balanced_balance_list[simp]: "balanced (balance_list xs)" | |
| 152 | by (metis balance_list_def balanced_bal prod.collapse) | |
| 153 | ||
| 154 | lemma height_balance_list: "height(balance_list xs) = floorlog 2 (length xs)" | |
| 155 | by (metis bal_height balance_list_def prod.collapse) | |
| 156 | ||
| 157 | lemma inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t" | |
| 158 | by(simp add: balance_tree_def inorder_balance_list) | |
| 159 | ||
| 160 | lemma size_balance_tree[simp]: "size(balance_tree t) = size t" | |
| 161 | by(simp add: balance_tree_def inorder_balance_list) | |
| 162 | ||
| 163 | corollary balanced_balance_tree[simp]: "balanced (balance_tree t)" | |
| 164 | by (simp add: balance_tree_def) | |
| 165 | ||
| 166 | lemma height_balance_tree: "height(balance_tree t) = floorlog 2 (size t)" | |
| 167 | by(simp add: balance_tree_def height_balance_list) | |
| 168 | ||
| 63861 | 169 | lemma wbalanced_bal: "bal xs n = (t,ys) \<Longrightarrow> wbalanced t" | 
| 170 | proof(induction xs n arbitrary: t ys rule: bal.induct) | |
| 171 | case (1 xs n) | |
| 172 | show ?case | |
| 173 | proof cases | |
| 174 | assume "n = 0" | |
| 175 | thus ?thesis | |
| 176 | using "1.prems" by(simp add: bal_simps) | |
| 177 | next | |
| 178 | assume "n \<noteq> 0" | |
| 179 | with "1.prems" obtain l ys r zs where | |
| 180 | rec1: "bal xs (n div 2) = (l, ys)" and | |
| 181 | rec2: "bal (tl ys) (n - 1 - n div 2) = (r, zs)" and | |
| 182 | t: "t = \<langle>l, hd ys, r\<rangle>" | |
| 183 | by(auto simp add: bal_simps Let_def split: prod.splits) | |
| 184 | have l: "wbalanced l" using "1.IH"(1)[OF \<open>n\<noteq>0\<close> refl rec1] . | |
| 185 | have "wbalanced r" using "1.IH"(2)[OF \<open>n\<noteq>0\<close> refl rec1[symmetric] refl rec2] . | |
| 186 | with l t size_bal[OF rec1] size_bal[OF rec2] | |
| 187 | show ?thesis by auto | |
| 188 | qed | |
| 189 | qed | |
| 190 | ||
| 191 | lemma wbalanced_balance_tree: "wbalanced (balance_tree t)" | |
| 192 | by(simp add: balance_tree_def balance_list_def) | |
| 193 | (metis prod.collapse wbalanced_bal) | |
| 194 | ||
| 63829 | 195 | hide_const (open) bal | 
| 63643 | 196 | |
| 197 | end |