src/HOL/Data_Structures/Balance.thy
author wenzelm
Tue, 03 Dec 2019 16:40:04 +0100
changeset 71222 2bc39c80a95d
parent 70751 fd9614c98dd6
child 71846 1a884605a08b
permissions -rw-r--r--
clarified export of consts: recursion is accessible via spec_rules;
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
66510
ca7a369301f6 reorganization of tree lemmas; new lemmas
nipkow
parents: 66453
diff changeset
     7
  "HOL-Library.Tree_Real"
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     8
begin
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
     9
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    10
fun bal :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree * 'a list" where
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    11
"bal n xs = (if n=0 then (Leaf,xs) else
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    12
 (let m = n div 2;
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    13
      (l, ys) = bal m xs;
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    14
      (r, zs) = bal (n-1-m) (tl ys)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    15
  in (Node l (hd ys) r, zs)))"
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    16
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    17
declare bal.simps[simp del]
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    18
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    19
definition bal_list :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree" where
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    20
"bal_list n xs = fst (bal n xs)"
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    21
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    22
definition balance_list :: "'a list \<Rightarrow> 'a tree" where
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    23
"balance_list xs = bal_list (length xs) xs"
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    24
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    25
definition bal_tree :: "nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    26
"bal_tree n t = bal_list n (inorder t)"
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    27
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    28
definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    29
"balance_tree t = bal_tree (size t) t"
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
    30
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    31
lemma bal_simps:
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    32
  "bal 0 xs = (Leaf, xs)"
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    33
  "n > 0 \<Longrightarrow>
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    34
   bal n xs =
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    35
  (let m = n div 2;
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    36
      (l, ys) = bal m xs;
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    37
      (r, zs) = bal (n-1-m) (tl ys)
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    38
  in (Node l (hd ys) r, zs))"
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    39
by(simp_all add: bal.simps)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    40
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    41
lemma bal_inorder:
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    42
  "\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk>
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    43
  \<Longrightarrow> xs = inorder t @ zs \<and> size t = n"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    44
proof(induction n xs arbitrary: t zs rule: bal.induct)
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    45
  case (1 n xs) show ?case
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    46
  proof cases
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    47
    assume "n = 0" thus ?thesis using 1 by (simp add: bal_simps)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    48
  next
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    49
    assume [arith]: "n \<noteq> 0"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    50
    let ?m = "n div 2" let ?m' = "n - 1 - ?m"
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    51
    from "1.prems"(2) obtain l r ys where
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    52
      b1: "bal ?m xs = (l,ys)" and
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    53
      b2: "bal ?m' (tl ys) = (r,zs)" and
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    54
      t: "t = \<langle>l, hd ys, r\<rangle>"
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
    55
      by(auto simp: Let_def bal_simps split: prod.splits)
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    56
    have IH1: "xs = inorder l @ ys \<and> size l = ?m"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    57
      using b1 "1.prems"(1) by(intro "1.IH"(1)) auto
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    58
    have IH2: "tl ys = inorder r @ zs \<and> size r = ?m'"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    59
      using b1 b2 IH1 "1.prems"(1) by(intro "1.IH"(2)) auto
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    60
    show ?thesis using t IH1 IH2  "1.prems"(1) hd_Cons_tl[of ys] by fastforce
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    61
  qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    62
qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    63
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    64
corollary inorder_bal_list[simp]:
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    65
  "n \<le> length xs \<Longrightarrow> inorder(bal_list n xs) = take n xs"
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    66
unfolding bal_list_def
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    67
by (metis (mono_tags) prod.collapse[of "bal n xs"] append_eq_conv_conj bal_inorder length_inorder)
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    68
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    69
corollary inorder_balance_list[simp]: "inorder(balance_list xs) = xs"
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    70
by(simp add: balance_list_def)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    71
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    72
corollary inorder_bal_tree:
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    73
  "n \<le> size t \<Longrightarrow> inorder(bal_tree n t) = take n (inorder t)"
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    74
by(simp add: bal_tree_def)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
    75
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
    76
corollary inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    77
by(simp add: balance_tree_def inorder_bal_tree)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    78
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    79
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    80
text\<open>The length/size lemmas below do not require the precondition @{prop"n \<le> length xs"}
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    81
(or  @{prop"n \<le> size t"}) that they come with. They could take advantage of the fact
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    82
that @{term "bal xs n"} yields a result even if @{prop "n > length xs"}.
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    83
In that case the result will contain one or more occurrences of @{term "hd []"}.
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    84
However, this is counter-intuitive and does not reflect the execution
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    85
in an eager functional language.\<close>
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    86
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    87
lemma bal_length: "\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk> \<Longrightarrow> length zs = length xs - n"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    88
using bal_inorder by fastforce
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    89
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    90
corollary size_bal_list[simp]: "n \<le> length xs \<Longrightarrow> size(bal_list n xs) = n"
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
    91
unfolding bal_list_def using bal_inorder prod.exhaust_sel by blast
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
    92
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
    93
corollary size_balance_list[simp]: "size(balance_list xs) = length xs"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    94
by (simp add: balance_list_def)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    95
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
    96
corollary size_bal_tree[simp]: "n \<le> size t \<Longrightarrow> size(bal_tree n t) = n"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
    97
by(simp add: bal_tree_def)
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
    98
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
    99
corollary size_balance_tree[simp]: "size(balance_tree t) = size t"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   100
by(simp add: balance_tree_def)
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   101
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   102
lemma min_height_bal:
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   103
  "\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk> \<Longrightarrow> min_height t = nat(\<lfloor>log 2 (n + 1)\<rfloor>)"
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   104
proof(induction n xs arbitrary: t zs rule: bal.induct)
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   105
  case (1 n xs)
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   106
  show ?case
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   107
  proof cases
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   108
    assume "n = 0" thus ?thesis using "1.prems"(2) by (simp add: bal_simps)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   109
  next
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   110
    assume [arith]: "n \<noteq> 0"
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   111
    let ?m = "n div 2" let ?m' = "n - 1 - ?m"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   112
    from "1.prems" obtain l r ys where
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   113
      b1: "bal ?m xs = (l,ys)" and
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   114
      b2: "bal ?m' (tl ys) = (r,zs)" and
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   115
      t: "t = \<langle>l, hd ys, r\<rangle>"
63843
ade7c3a20917 more simp rules
nipkow
parents: 63829
diff changeset
   116
      by(auto simp: bal_simps Let_def split: prod.splits)
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   117
    let ?hl = "nat (floor(log 2 (?m + 1)))"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   118
    let ?hr = "nat (floor(log 2 (?m' + 1)))"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   119
    have IH1: "min_height l = ?hl" using "1.IH"(1) b1 "1.prems"(1) by simp
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   120
    have IH2: "min_height r = ?hr"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   121
      using "1.prems"(1) bal_length[OF _ b1] b1 b2 by(intro "1.IH"(2)) auto
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   122
    have "(n+1) div 2 \<ge> 1" by arith
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   123
    hence 0: "log 2 ((n+1) div 2) \<ge> 0" by simp
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   124
    have "?m' \<le> ?m" by arith
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   125
    hence le: "?hr \<le> ?hl" by(simp add: nat_mono floor_mono)
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   126
    have "min_height t = min ?hl ?hr + 1" by (simp add: t IH1 IH2)
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   127
    also have "\<dots> = ?hr + 1" using le by (simp add: min_absorb2)
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   128
    also have "?m' + 1 = (n+1) div 2" by linarith
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   129
    also have "nat (floor(log 2 ((n+1) div 2))) + 1
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   130
       = nat (floor(log 2 ((n+1) div 2) + 1))"
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   131
      using 0 by linarith
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   132
    also have "\<dots> = nat (floor(log 2 (n + 1)))"
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   133
      using floor_log2_div2[of "n+1"] by (simp add: log_mult)
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   134
    finally show ?thesis .
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   135
  qed
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   136
qed
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   137
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   138
lemma height_bal:
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   139
  "\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk> \<Longrightarrow> height t = nat \<lceil>log 2 (n + 1)\<rceil>"
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   140
proof(induction n xs arbitrary: t zs rule: bal.induct)
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   141
  case (1 n xs) show ?case
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   142
  proof cases
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   143
    assume "n = 0" thus ?thesis
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   144
      using "1.prems" by (simp add: bal_simps)
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   145
  next
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   146
    assume [arith]: "n \<noteq> 0"
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   147
    let ?m = "n div 2" let ?m' = "n - 1 - ?m"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   148
    from "1.prems" obtain l r ys where
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   149
      b1: "bal ?m xs = (l,ys)" and
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   150
      b2: "bal ?m' (tl ys) = (r,zs)" and
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   151
      t: "t = \<langle>l, hd ys, r\<rangle>"
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   152
      by(auto simp: bal_simps Let_def split: prod.splits)
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   153
    let ?hl = "nat \<lceil>log 2 (?m + 1)\<rceil>"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   154
    let ?hr = "nat \<lceil>log 2 (?m' + 1)\<rceil>"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   155
    have IH1: "height l = ?hl" using "1.IH"(1) b1 "1.prems"(1) by simp
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   156
    have IH2: "height r = ?hr"
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   157
      using b1 b2 bal_length[OF _ b1] "1.prems"(1) by(intro "1.IH"(2)) auto
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   158
    have 0: "log 2 (?m + 1) \<ge> 0" by simp
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   159
    have "?m' \<le> ?m" by arith
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   160
    hence le: "?hr \<le> ?hl"
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   161
      by(simp add: nat_mono ceiling_mono del: nat_ceiling_le_eq)
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   162
    have "height t = max ?hl ?hr + 1" by (simp add: t IH1 IH2)
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   163
    also have "\<dots> = ?hl + 1" using le by (simp add: max_absorb1)
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   164
    also have "\<dots> = nat \<lceil>log 2 (?m + 1) + 1\<rceil>" using 0 by linarith
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   165
    also have "\<dots> = nat \<lceil>log 2 (n + 1)\<rceil>"
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   166
      using ceiling_log2_div2[of "n+1"] by (simp)
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   167
    finally show ?thesis .
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   168
  qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   169
qed
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   170
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   171
lemma balanced_bal:
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   172
  assumes "n \<le> length xs" "bal n xs = (t,ys)" shows "balanced t"
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   173
unfolding balanced_def
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   174
using height_bal[OF assms] min_height_bal[OF assms]
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   175
by linarith
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   176
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   177
lemma height_bal_list:
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   178
  "n \<le> length xs \<Longrightarrow> height (bal_list n xs) = nat \<lceil>log 2 (n + 1)\<rceil>"
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   179
unfolding bal_list_def by (metis height_bal prod.collapse)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   180
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   181
lemma height_balance_list:
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   182
  "height (balance_list xs) = nat \<lceil>log 2 (length xs + 1)\<rceil>"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   183
by (simp add: balance_list_def height_bal_list)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   184
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   185
corollary height_bal_tree:
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   186
  "n \<le> size t \<Longrightarrow> height (bal_tree n t) = nat\<lceil>log 2 (n + 1)\<rceil>"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   187
unfolding bal_list_def bal_tree_def
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   188
by (metis bal_list_def height_bal_list length_inorder)
64018
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   189
c6eb691770d8 replaced floorlog by floor/ceiling(log .)
nipkow
parents: 63861
diff changeset
   190
corollary height_balance_tree:
66515
85c505c98332 reorganized and added log-related lemmas
nipkow
parents: 66510
diff changeset
   191
  "height (balance_tree t) = nat\<lceil>log 2 (size t + 1)\<rceil>"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   192
by (simp add: bal_tree_def balance_tree_def height_bal_list)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   193
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   194
corollary balanced_bal_list[simp]: "n \<le> length xs \<Longrightarrow> balanced (bal_list n xs)"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   195
unfolding bal_list_def by (metis  balanced_bal prod.collapse)
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   196
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   197
corollary balanced_balance_list[simp]: "balanced (balance_list xs)"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   198
by (simp add: balance_list_def)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   199
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   200
corollary balanced_bal_tree[simp]: "n \<le> size t \<Longrightarrow> balanced (bal_tree n t)"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   201
by (simp add: bal_tree_def)
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   202
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   203
corollary balanced_balance_tree[simp]: "balanced (balance_tree t)"
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   204
by (simp add: balance_tree_def)
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   205
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   206
lemma wbalanced_bal: "\<lbrakk> n \<le> length xs; bal n xs = (t,ys) \<rbrakk> \<Longrightarrow> wbalanced t"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   207
proof(induction n xs arbitrary: t ys rule: bal.induct)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   208
  case (1 n xs)
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   209
  show ?case
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   210
  proof cases
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   211
    assume "n = 0"
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   212
    thus ?thesis using "1.prems"(2) by(simp add: bal_simps)
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   213
  next
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   214
    assume [arith]: "n \<noteq> 0"
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   215
    with "1.prems" obtain l ys r zs where
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   216
      rec1: "bal (n div 2) xs = (l, ys)" and
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   217
      rec2: "bal (n - 1 - n div 2) (tl ys) = (r, zs)" and
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   218
      t: "t = \<langle>l, hd ys, r\<rangle>"
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   219
      by(auto simp add: bal_simps Let_def split: prod.splits)
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   220
    have l: "wbalanced l" using "1.IH"(1)[OF \<open>n\<noteq>0\<close> refl _ rec1] "1.prems"(1) by linarith
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   221
    have "wbalanced r"
70751
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   222
      using rec1 rec2 bal_length[OF _ rec1] "1.prems"(1) by(intro "1.IH"(2)) auto
fd9614c98dd6 simplified proofs
nipkow
parents: 70747
diff changeset
   223
    with l t bal_length[OF _ rec1] "1.prems"(1) bal_inorder[OF _ rec1] bal_inorder[OF _ rec2]
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   224
    show ?thesis by auto
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   225
  qed
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   226
qed
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   227
64541
3d4331b65861 more lemmas
nipkow
parents: 64540
diff changeset
   228
text\<open>An alternative proof via @{thm balanced_if_wbalanced}:\<close>
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   229
lemma "\<lbrakk> n \<le> length xs; bal n xs = (t,ys) \<rbrakk> \<Longrightarrow> balanced t"
64541
3d4331b65861 more lemmas
nipkow
parents: 64540
diff changeset
   230
by(rule balanced_if_wbalanced[OF wbalanced_bal])
3d4331b65861 more lemmas
nipkow
parents: 64540
diff changeset
   231
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   232
lemma wbalanced_bal_list[simp]: "n \<le> length xs \<Longrightarrow> wbalanced (bal_list n xs)"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   233
by(simp add: bal_list_def) (metis prod.collapse wbalanced_bal)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   234
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   235
lemma wbalanced_balance_list[simp]: "wbalanced (balance_list xs)"
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   236
by(simp add: balance_list_def)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   237
70747
548420d389ea Enforced precodition "n <= length xs" to avoid relying on "hd []".
nipkow
parents: 66516
diff changeset
   238
lemma wbalanced_bal_tree[simp]: "n \<le> size t \<Longrightarrow> wbalanced (bal_tree n t)"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   239
by(simp add: bal_tree_def)
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   240
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   241
lemma wbalanced_balance_tree: "wbalanced (balance_tree t)"
64444
daae191c9344 provided more efficient interface
nipkow
parents: 64065
diff changeset
   242
by (simp add: balance_tree_def)
63861
90360390a916 reorganization, more funs and lemmas
nipkow
parents: 63843
diff changeset
   243
63829
6a05c8cbf7de More on balancing; renamed theory to Balance
nipkow
parents: 63755
diff changeset
   244
hide_const (open) bal
63643
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   245
f9ad2e591957 New theory Balance_List
nipkow
parents:
diff changeset
   246
end