# HG changeset patch # User nipkow # Date 1473423316 -7200 # Node ID 6a05c8cbf7de06b0cff01f2670862c78749de0a6 # Parent ca467e73f9122f28d425ae94c0976127b20f943f More on balancing; renamed theory to Balance diff -r ca467e73f912 -r 6a05c8cbf7de src/HOL/Data_Structures/Balance.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Data_Structures/Balance.thy Fri Sep 09 14:15:16 2016 +0200 @@ -0,0 +1,151 @@ +(* Author: Tobias Nipkow *) + +section \Creating Balanced Trees\ + +theory Balance +imports + "~~/src/HOL/Library/Tree" + "~~/src/HOL/Library/Log_Nat" +begin + +fun bal :: "'a list \ nat \ 'a tree * 'a list" where +"bal xs n = (if n=0 then (Leaf,xs) else + (let m = n div 2; + (l, ys) = bal xs m; + (r, zs) = bal (tl ys) (n-1-m) + in (Node l (hd ys) r, zs)))" + +declare bal.simps[simp del] + +definition balance_list :: "'a list \ 'a tree" where +"balance_list xs = fst (bal xs (length xs))" + +definition balance_tree :: "'a tree \ 'a tree" where +"balance_tree = balance_list o inorder" + + +lemma bal_inorder: + "\ bal xs n = (t,ys); n \ length xs \ + \ inorder t = take n xs \ ys = drop n xs" +proof(induction xs n arbitrary: t ys rule: bal.induct) + case (1 xs n) show ?case + proof cases + assume "n = 0" thus ?thesis using 1 by (simp add: bal.simps) + next + assume [arith]: "n \ 0" + let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1" + from "1.prems" obtain l r xs' where + b1: "bal xs ?n1 = (l,xs')" and + b2: "bal (tl xs') ?n2 = (r,ys)" and + t: "t = \l, hd xs', r\" + using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits) + have IH1: "inorder l = take ?n1 xs \ xs' = drop ?n1 xs" + using b1 "1.prems" by(intro "1.IH"(1)) auto + have IH2: "inorder r = take ?n2 (tl xs') \ ys = drop ?n2 (tl xs')" + using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto + have "drop (n div 2) xs \ []" using "1.prems"(2) by simp + hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)" + by (metis Suc_eq_plus1 take_Suc) + hence *: "inorder t = take n xs" using t IH1 IH2 + using take_add[of ?n1 "?n2+1" xs] by(simp) + have "n - n div 2 + n div 2 = n" by simp + hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) + thus ?thesis using * by blast + qed +qed + +corollary inorder_balance_list: "inorder(balance_list xs) = xs" +using bal_inorder[of xs "length xs"] +by (metis balance_list_def order_refl prod.collapse take_all) + +lemma bal_height: "bal xs n = (t,ys) \ height t = floorlog 2 n" +proof(induction xs n arbitrary: t ys rule: bal.induct) + case (1 xs n) show ?case + proof cases + assume "n = 0" thus ?thesis + using "1.prems" by (simp add: floorlog_def bal.simps) + next + assume [arith]: "n \ 0" + from "1.prems" obtain l r xs' where + b1: "bal xs (n div 2) = (l,xs')" and + b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and + t: "t = \l, hd xs', r\" + using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits) + let ?log1 = "floorlog 2 (n div 2)" + let ?log2 = "floorlog 2 (n - 1 - n div 2)" + have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp + have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp + have "n div 2 \ n - 1 - n div 2" by arith + hence le: "?log2 \ ?log1" by(simp add:floorlog_mono) + have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) + also have "\ = ?log1 + 1" using le by (simp add: max_absorb1) + also have "\ = floorlog 2 n" by (simp add: compute_floorlog) + finally show ?thesis . + qed +qed + +lemma bal_min_height: + "bal xs n = (t,ys) \ min_height t = floorlog 2 (n + 1) - 1" +proof(induction xs n arbitrary: t ys rule: bal.induct) + case (1 xs n) show ?case + proof cases + assume "n = 0" thus ?thesis + using "1.prems" by (simp add: floorlog_def bal.simps) + next + assume [arith]: "n \ 0" + from "1.prems" obtain l r xs' where + b1: "bal xs (n div 2) = (l,xs')" and + b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and + t: "t = \l, hd xs', r\" + using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits) + let ?log1 = "floorlog 2 (n div 2 + 1) - 1" + let ?log2 = "floorlog 2 (n - 1 - n div 2 + 1) - 1" + let ?log2' = "floorlog 2 (n - n div 2) - 1" + have "n - 1 - n div 2 + 1 = n - n div 2" by arith + hence IH2: "min_height r = ?log2'" using "1.IH"(2) b1 b2 by simp + have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp + have *: "floorlog 2 (n - n div 2) \ 1" by (simp add: floorlog_def) + have "n div 2 + 1 \ n - n div 2" by arith + with * have le: "?log2' \ ?log1" by(simp add: floorlog_mono diff_le_mono) + have "min_height t = min ?log1 ?log2' + 1" by (simp add: t IH1 IH2) + also have "\ = ?log2' + 1" using le by (simp add: min_absorb2) + also have "\ = floorlog 2 (n - n div 2)" by(simp add: floorlog_def) + also have "n - n div 2 = (n+1) div 2" by arith + also have "floorlog 2 \ = floorlog 2 (n+1) - 1" + by (simp add: compute_floorlog) + finally show ?thesis . + qed +qed + +lemma balanced_bal: + assumes "bal xs n = (t,ys)" shows "balanced t" +proof - + have "floorlog 2 n \ floorlog 2 (n+1)" by (rule floorlog_mono) auto + thus ?thesis unfolding balanced_def + using bal_height[OF assms] bal_min_height[OF assms] by linarith +qed + +corollary size_balance_list[simp]: "size(balance_list xs) = length xs" +by (metis inorder_balance_list length_inorder) + +corollary balanced_balance_list[simp]: "balanced (balance_list xs)" +by (metis balance_list_def balanced_bal prod.collapse) + +lemma height_balance_list: "height(balance_list xs) = floorlog 2 (length xs)" +by (metis bal_height balance_list_def prod.collapse) + +lemma inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t" +by(simp add: balance_tree_def inorder_balance_list) + +lemma size_balance_tree[simp]: "size(balance_tree t) = size t" +by(simp add: balance_tree_def inorder_balance_list) + +corollary balanced_balance_tree[simp]: "balanced (balance_tree t)" +by (simp add: balance_tree_def) + +lemma height_balance_tree: "height(balance_tree t) = floorlog 2 (size t)" +by(simp add: balance_tree_def height_balance_list) + +hide_const (open) bal + +end diff -r ca467e73f912 -r 6a05c8cbf7de src/HOL/Data_Structures/Balance_List.thy --- a/src/HOL/Data_Structures/Balance_List.thy Fri Sep 09 13:39:21 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,126 +0,0 @@ -(* Tobias Nipkow *) - -section \Creating a Balanced Tree from a List\ - -theory Balance_List -imports - "~~/src/HOL/Library/Tree" - "~~/src/HOL/Library/Log_Nat" -begin - -fun bal :: "'a list \ nat \ 'a tree * 'a list" where -"bal xs n = (if n=0 then (Leaf,xs) else - (let m = n div 2; - (l, ys) = bal xs m; - (r, zs) = bal (tl ys) (n-1-m) - in (Node l (hd ys) r, zs)))" - -declare bal.simps[simp del] - -definition "balance xs = fst (bal xs (length xs))" - -lemma bal_inorder: - "\ bal xs n = (t,ys); n \ length xs \ - \ inorder t = take n xs \ ys = drop n xs" -proof(induction xs n arbitrary: t ys rule: bal.induct) - case (1 xs n) show ?case - proof cases - assume "n = 0" thus ?thesis using 1 by (simp add: bal.simps) - next - assume [arith]: "n \ 0" - let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1" - from "1.prems" obtain l r xs' where - b1: "bal xs ?n1 = (l,xs')" and - b2: "bal (tl xs') ?n2 = (r,ys)" and - t: "t = \l, hd xs', r\" - using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits) - have IH1: "inorder l = take ?n1 xs \ xs' = drop ?n1 xs" - using b1 "1.prems" by(intro "1.IH"(1)) auto - have IH2: "inorder r = take ?n2 (tl xs') \ ys = drop ?n2 (tl xs')" - using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto - have "drop (n div 2) xs \ []" using "1.prems"(2) by simp - hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)" - by (metis Suc_eq_plus1 take_Suc) - hence *: "inorder t = take n xs" using t IH1 IH2 - using take_add[of ?n1 "?n2+1" xs] by(simp) - have "n - n div 2 + n div 2 = n" by simp - hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) - thus ?thesis using * by blast - qed -qed - -corollary balance_inorder: "inorder(balance xs) = xs" -using bal_inorder[of xs "length xs"] -by (metis balance_def order_refl prod.collapse take_all) - -lemma bal_height: "bal xs n = (t,ys) \ height t = floorlog 2 n" -proof(induction xs n arbitrary: t ys rule: bal.induct) - case (1 xs n) show ?case - proof cases - assume "n = 0" thus ?thesis - using "1.prems" by (simp add: floorlog_def bal.simps) - next - assume [arith]: "n \ 0" - from "1.prems" obtain l r xs' where - b1: "bal xs (n div 2) = (l,xs')" and - b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and - t: "t = \l, hd xs', r\" - using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits) - let ?log1 = "floorlog 2 (n div 2)" - let ?log2 = "floorlog 2 (n - 1 - n div 2)" - have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp - have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp - have "n div 2 \ n - 1 - n div 2" by arith - hence le: "?log2 \ ?log1" by(simp add:floorlog_mono) - have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) - also have "\ = ?log1 + 1" using le by (simp add: max_absorb1) - also have "\ = floorlog 2 n" by (simp add: compute_floorlog) - finally show ?thesis . - qed -qed - -lemma bal_min_height: - "bal xs n = (t,ys) \ min_height t = floorlog 2 (n + 1) - 1" -proof(induction xs n arbitrary: t ys rule: bal.induct) - case (1 xs n) show ?case - proof cases - assume "n = 0" thus ?thesis - using "1.prems" by (simp add: floorlog_def bal.simps) - next - assume [arith]: "n \ 0" - from "1.prems" obtain l r xs' where - b1: "bal xs (n div 2) = (l,xs')" and - b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and - t: "t = \l, hd xs', r\" - using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits) - let ?log1 = "floorlog 2 (n div 2 + 1) - 1" - let ?log2 = "floorlog 2 (n - 1 - n div 2 + 1) - 1" - let ?log2' = "floorlog 2 (n - n div 2) - 1" - have "n - 1 - n div 2 + 1 = n - n div 2" by arith - hence IH2: "min_height r = ?log2'" using "1.IH"(2) b1 b2 by simp - have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp - have *: "floorlog 2 (n - n div 2) \ 1" by (simp add: floorlog_def) - have "n div 2 + 1 \ n - n div 2" by arith - with * have le: "?log2' \ ?log1" by(simp add: floorlog_mono diff_le_mono) - have "min_height t = min ?log1 ?log2' + 1" by (simp add: t IH1 IH2) - also have "\ = ?log2' + 1" using le by (simp add: min_absorb2) - also have "\ = floorlog 2 (n - n div 2)" by(simp add: floorlog_def) - also have "n - n div 2 = (n+1) div 2" by arith - also have "floorlog 2 \ = floorlog 2 (n+1) - 1" - by (simp add: compute_floorlog) - finally show ?thesis . - qed -qed - -lemma balanced_bal: - assumes "bal xs n = (t,ys)" shows "balanced t" -proof - - have "floorlog 2 n \ floorlog 2 (n+1)" by (rule floorlog_mono) auto - thus ?thesis - using bal_height[OF assms] bal_min_height[OF assms] by arith -qed - -corollary balanced_balance: "balanced (balance xs)" -by (metis balance_def balanced_bal prod.collapse) - -end diff -r ca467e73f912 -r 6a05c8cbf7de src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Fri Sep 09 13:39:21 2016 +0200 +++ b/src/HOL/Library/Tree.thy Fri Sep 09 14:15:16 2016 +0200 @@ -230,7 +230,8 @@ subsection \Balanced\ -abbreviation "balanced t \ (height t - min_height t \ 1)" +definition balanced :: "'a tree \ bool" where +"balanced t = (height t - min_height t \ 1)" text\Balanced trees have optimal height:\ @@ -262,7 +263,7 @@ hence *: "min_height t < height t'" by simp have "min_height t + 1 = height t" using min_hight_le_height[of t] assms(1) False - by (simp add: complete_iff_height) + by (simp add: complete_iff_height balanced_def) with * show ?thesis by arith qed diff -r ca467e73f912 -r 6a05c8cbf7de src/HOL/ROOT --- a/src/HOL/ROOT Fri Sep 09 13:39:21 2016 +0200 +++ b/src/HOL/ROOT Fri Sep 09 14:15:16 2016 +0200 @@ -183,7 +183,7 @@ "~~/src/HOL/Library/Multiset" "~~/src/HOL/Library/Float" theories - Balance_List + Balance Tree_Map AVL_Map RBT_Map