src/HOL/Data_Structures/Balance.thy
author nipkow
Fri, 09 Sep 2016 14:15:16 +0200
changeset 63829 6a05c8cbf7de
parent 63755 src/HOL/Data_Structures/Balance_List.thy@182c111190e5
child 63843 ade7c3a20917
permissions -rw-r--r--
More on balancing; renamed theory to Balance
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
     1
(* Author: Tobias Nipkow *)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     2
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
     3
section \<open>Creating Balanced Trees\<close>
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     4
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
     5
theory Balance
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     6
imports
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     7
  "~~/src/HOL/Library/Tree"
63663
28d1deca302e Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents: 63643
diff changeset
     8
  "~~/src/HOL/Library/Log_Nat"
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     9
begin
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    10
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    11
fun bal :: "'a list \<Rightarrow> nat \<Rightarrow> 'a tree * 'a list" where
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    12
"bal xs n = (if n=0 then (Leaf,xs) else
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    13
 (let m = n div 2;
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    14
      (l, ys) = bal xs m;
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    15
      (r, zs) = bal (tl ys) (n-1-m)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    16
  in (Node l (hd ys) r, zs)))"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    17
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    18
declare bal.simps[simp del]
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    19
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    20
definition balance_list :: "'a list \<Rightarrow> 'a tree" where
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    21
"balance_list xs = fst (bal xs (length xs))"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    22
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    23
definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    24
"balance_tree = balance_list o inorder"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    25
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    26
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    27
lemma bal_inorder:
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63663
diff changeset
    28
  "\<lbrakk> bal xs n = (t,ys); n \<le> length xs \<rbrakk>
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63663
diff changeset
    29
  \<Longrightarrow> inorder t = take n xs \<and> ys = drop n xs"
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    30
proof(induction xs n arbitrary: t ys rule: bal.induct)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    31
  case (1 xs n) show ?case
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    32
  proof cases
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    33
    assume "n = 0" thus ?thesis using 1 by (simp add: bal.simps)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    34
  next
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    35
    assume [arith]: "n \<noteq> 0"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    36
    let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    37
    from "1.prems" obtain l r xs' where
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    38
      b1: "bal xs ?n1 = (l,xs')" and
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    39
      b2: "bal (tl xs') ?n2 = (r,ys)" and
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    40
      t: "t = \<langle>l, hd xs', r\<rangle>"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    41
      using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    42
    have IH1: "inorder l = take ?n1 xs \<and> xs' = drop ?n1 xs"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    43
      using b1 "1.prems" by(intro "1.IH"(1)) auto
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    44
    have IH2: "inorder r = take ?n2 (tl xs') \<and> ys = drop ?n2 (tl xs')"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    45
      using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    46
    have "drop (n div 2) xs \<noteq> []" using "1.prems"(2) by simp
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    47
    hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    48
      by (metis Suc_eq_plus1 take_Suc)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    49
    hence *: "inorder t = take n xs" using t IH1 IH2
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    50
      using take_add[of ?n1 "?n2+1" xs] by(simp)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    51
    have "n - n div 2 + n div 2 = n" by simp
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    52
    hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric])
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    53
    thus ?thesis using * by blast
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    54
  qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    55
qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    56
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    57
corollary inorder_balance_list: "inorder(balance_list xs) = xs"
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    58
using bal_inorder[of xs "length xs"]
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    59
by (metis balance_list_def order_refl prod.collapse take_all)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    60
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    61
lemma bal_height: "bal xs n = (t,ys) \<Longrightarrow> height t = floorlog 2 n"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    62
proof(induction xs n arbitrary: t ys rule: bal.induct)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    63
  case (1 xs n) show ?case
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    64
  proof cases
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    65
    assume "n = 0" thus ?thesis
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    66
      using "1.prems" by (simp add: floorlog_def bal.simps)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    67
  next
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    68
    assume [arith]: "n \<noteq> 0"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    69
    from "1.prems" obtain l r xs' where
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    70
      b1: "bal xs (n div 2) = (l,xs')" and
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    71
      b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    72
      t: "t = \<langle>l, hd xs', r\<rangle>"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    73
      using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    74
    let ?log1 = "floorlog 2 (n div 2)"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    75
    let ?log2 = "floorlog 2 (n - 1 - n div 2)"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    76
    have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    77
    have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    78
    have "n div 2 \<ge> n - 1 - n div 2" by arith
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    79
    hence le: "?log2 \<le> ?log1" by(simp add:floorlog_mono)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    80
    have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    81
    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: 63643
diff changeset
    82
    also have "\<dots> = floorlog 2 n" by (simp add: compute_floorlog)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    83
    finally show ?thesis .
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    84
  qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    85
qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    86
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    87
lemma bal_min_height:
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    88
  "bal xs n = (t,ys) \<Longrightarrow> min_height t = floorlog 2 (n + 1) - 1"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    89
proof(induction xs n arbitrary: t ys rule: bal.induct)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    90
  case (1 xs n) show ?case
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    91
  proof cases
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    92
    assume "n = 0" thus ?thesis
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    93
      using "1.prems" by (simp add: floorlog_def bal.simps)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    94
  next
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    95
    assume [arith]: "n \<noteq> 0"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    96
    from "1.prems" obtain l r xs' where
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    97
      b1: "bal xs (n div 2) = (l,xs')" and
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    98
      b2: "bal (tl xs') (n - 1 - n div 2) = (r,ys)" and
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    99
      t: "t = \<langle>l, hd xs', r\<rangle>"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   100
      using bal.simps[of xs n] by(auto simp: Let_def split: prod.splits)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   101
    let ?log1 = "floorlog 2 (n div 2 + 1) - 1"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   102
    let ?log2 = "floorlog 2 (n - 1 - n div 2 + 1) - 1"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   103
    let ?log2' = "floorlog 2 (n - n div 2) - 1"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   104
    have "n - 1 - n div 2 + 1 = n - n div 2" by arith
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   105
    hence IH2: "min_height r = ?log2'" using "1.IH"(2) b1 b2 by simp
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   106
    have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   107
    have *: "floorlog 2 (n - n div 2) \<ge> 1" by (simp add: floorlog_def)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   108
    have "n div 2 + 1 \<ge> n - n div 2" by arith
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   109
    with * have le: "?log2' \<le> ?log1" by(simp add: floorlog_mono diff_le_mono)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   110
    have "min_height t = min ?log1 ?log2' + 1" by (simp add: t IH1 IH2)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   111
    also have "\<dots> = ?log2' + 1" using le by (simp add: min_absorb2)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   112
    also have "\<dots> = floorlog 2 (n - n div 2)" by(simp add: floorlog_def)
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   113
    also have "n - n div 2 = (n+1) div 2" by arith
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   114
    also have "floorlog 2 \<dots> = floorlog 2 (n+1) - 1"
63663
28d1deca302e Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents: 63643
diff changeset
   115
      by (simp add: compute_floorlog)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   116
    finally show ?thesis .
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   117
  qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   118
qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   119
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   120
lemma balanced_bal:
63755
182c111190e5 Renamed balanced to complete; added balanced; more about both
nipkow
parents: 63663
diff changeset
   121
  assumes "bal xs n = (t,ys)" shows "balanced t"
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   122
proof -
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   123
  have "floorlog 2 n \<le> floorlog 2 (n+1)" by (rule floorlog_mono) auto
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   124
  thus ?thesis unfolding balanced_def
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   125
    using bal_height[OF assms] bal_min_height[OF assms] by linarith
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   126
qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   127
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   128
corollary size_balance_list[simp]: "size(balance_list xs) = length xs"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   129
by (metis inorder_balance_list length_inorder)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   130
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   131
corollary balanced_balance_list[simp]: "balanced (balance_list xs)"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   132
by (metis balance_list_def balanced_bal prod.collapse)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   133
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   134
lemma height_balance_list: "height(balance_list xs) = floorlog 2 (length xs)"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   135
by (metis bal_height balance_list_def prod.collapse)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   136
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   137
lemma inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   138
by(simp add: balance_tree_def inorder_balance_list)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   139
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   140
lemma size_balance_tree[simp]: "size(balance_tree t) = size t"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   141
by(simp add: balance_tree_def inorder_balance_list)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   142
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   143
corollary balanced_balance_tree[simp]: "balanced (balance_tree t)"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   144
by (simp add: balance_tree_def)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   145
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   146
lemma height_balance_tree: "height(balance_tree t) = floorlog 2 (size t)"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   147
by(simp add: balance_tree_def height_balance_list)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   148
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   149
hide_const (open) bal
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   150
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   151
end