src/HOL/Data_Structures/Leftist_Heap_List.thy
author haftmann
Mon, 16 Jun 2025 15:25:38 +0200
changeset 82730 3b98b1b57435
parent 82010 dfde9a8296f5
permissions -rw-r--r--
more explicit theorem names for list quantifiers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
78231
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     2
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     3
theory Leftist_Heap_List
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     4
imports
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     5
  Leftist_Heap
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     6
  Complex_Main
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     7
begin
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     8
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
     9
subsection "Converting a list into a leftist heap"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    10
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    11
fun merge_adj :: "('a::ord) lheap list \<Rightarrow> 'a lheap list" where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    12
"merge_adj [] = []" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    13
"merge_adj [t] = [t]" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    14
"merge_adj (t1 # t2 # ts) = merge t1 t2 # merge_adj ts"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    15
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    16
text \<open>For the termination proof of \<open>merge_all\<close> below.\<close>
82010
nipkow
parents: 80776
diff changeset
    17
lemma length_merge_adjacent[termination_simp]: "length (merge_adj ts) = (length ts + 1) div 2"
78231
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    18
by (induction ts rule: merge_adj.induct) auto
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    19
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    20
fun merge_all :: "('a::ord) lheap list \<Rightarrow> 'a lheap" where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    21
"merge_all [] = Leaf" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    22
"merge_all [t] = t" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    23
"merge_all ts = merge_all (merge_adj ts)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    24
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    25
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    26
subsubsection \<open>Functional correctness\<close>
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    27
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    28
lemma heap_merge_adj: "\<forall>t \<in> set ts. heap t \<Longrightarrow> \<forall>t \<in> set (merge_adj ts). heap t"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    29
by(induction ts rule: merge_adj.induct) (auto simp: heap_merge)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    30
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    31
lemma ltree_merge_adj: "\<forall>t \<in> set ts. ltree t \<Longrightarrow> \<forall>t \<in> set (merge_adj ts). ltree t"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    32
by(induction ts rule: merge_adj.induct) (auto simp: ltree_merge)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    33
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    34
lemma heap_merge_all: "\<forall>t \<in> set ts. heap t \<Longrightarrow> heap (merge_all ts)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    35
apply(induction ts rule: merge_all.induct)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    36
using [[simp_depth_limit=3]] by (auto simp add: heap_merge_adj)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    37
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    38
lemma ltree_merge_all: "\<forall>t \<in> set ts. ltree t \<Longrightarrow> ltree (merge_all ts)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    39
apply(induction ts rule: merge_all.induct)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    40
using [[simp_depth_limit=3]] by (auto simp add: ltree_merge_adj)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    41
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    42
lemma mset_merge_adj:
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    43
  "\<Sum>\<^sub># (image_mset mset_tree (mset (merge_adj ts))) =
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    44
   \<Sum>\<^sub># (image_mset mset_tree (mset ts))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    45
by(induction ts rule: merge_adj.induct) (auto simp: mset_merge)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    46
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    47
lemma mset_merge_all:
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    48
  "mset_tree (merge_all ts) = \<Sum>\<^sub># (mset (map mset_tree ts))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    49
by(induction ts rule: merge_all.induct) (auto simp: mset_merge mset_merge_adj)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    50
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    51
fun lheap_list :: "'a::ord list \<Rightarrow> 'a lheap" where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    52
"lheap_list xs = merge_all (map (\<lambda>x. Node Leaf (x,1) Leaf) xs)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    53
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    54
lemma mset_lheap_list: "mset_tree (lheap_list xs) = mset xs"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    55
by (simp add: mset_merge_all o_def)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    56
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    57
lemma ltree_lheap_list: "ltree (lheap_list ts)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    58
by(simp add: ltree_merge_all)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    59
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    60
lemma heap_lheap_list: "heap (lheap_list ts)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    61
by(simp add: heap_merge_all)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    62
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    63
lemma size_merge: "size(merge t1 t2) = size t1 + size t2"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    64
by(induction t1 t2 rule: merge.induct) (auto simp: node_def)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    65
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    66
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    67
subsubsection \<open>Running time\<close>
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    68
79984
nipkow
parents: 78231
diff changeset
    69
text \<open>Not defined automatically because we only count the time for @{const merge}.\<close>
nipkow
parents: 78231
diff changeset
    70
78231
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    71
fun T_merge_adj :: "('a::ord) lheap list \<Rightarrow> nat" where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    72
"T_merge_adj [] = 0" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    73
"T_merge_adj [t] = 0" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    74
"T_merge_adj (t1 # t2 # ts) = T_merge t1 t2 + T_merge_adj ts"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    75
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    76
fun T_merge_all :: "('a::ord) lheap list  \<Rightarrow> nat" where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    77
"T_merge_all [] = 0" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    78
"T_merge_all [t] = 0" |
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    79
"T_merge_all ts = T_merge_adj ts + T_merge_all (merge_adj ts)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    80
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    81
fun T_lheap_list :: "'a::ord list \<Rightarrow> nat" where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    82
"T_lheap_list xs = T_merge_all (map (\<lambda>x. Node Leaf (x,1) Leaf) xs)"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    83
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    84
abbreviation Tm where
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    85
"Tm n == 2 * log 2 (n+1) + 1"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    86
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    87
lemma T_merge_adj: "\<lbrakk> \<forall>t \<in> set ts. ltree t; \<forall>t \<in> set ts. size t = n \<rbrakk>
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    88
  \<Longrightarrow> T_merge_adj ts \<le> (length ts div 2) * Tm n"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    89
proof(induction ts rule: T_merge_adj.induct)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    90
  case 1 thus ?case by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    91
next
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    92
  case 2 thus ?case by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    93
next
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    94
  case (3 t1 t2) thus ?case using T_merge_log[of t1 t2] by (simp add: algebra_simps size1_size)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    95
qed
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    96
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    97
lemma size_merge_adj:
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    98
  "\<lbrakk> even(length ts); \<forall>t \<in> set ts. ltree t; \<forall>t \<in> set ts. size t = n \<rbrakk>
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
    99
   \<Longrightarrow> \<forall>t \<in> set (merge_adj ts). size t = 2*n"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   100
by(induction ts rule: merge_adj.induct) (auto simp: size_merge)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   101
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   102
lemma T_merge_all:
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   103
 "\<lbrakk> \<forall>t \<in> set ts. ltree t; \<forall>t \<in> set ts. size t = n; length ts = 2^k \<rbrakk>
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   104
  \<Longrightarrow> T_merge_all ts \<le> (\<Sum>i=1..k. 2^(k-i) * Tm(2 ^ (i-1) * n))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   105
proof (induction ts arbitrary: k n rule: merge_all.induct)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   106
  case 1 thus ?case by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   107
next
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   108
  case 2 thus ?case by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   109
next
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   110
  case (3 t1 t2 ts)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   111
  let ?ts = "t1 # t2 # ts"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   112
  let ?ts2 = "merge_adj ?ts"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   113
  obtain k' where k': "k = Suc k'" using "3.prems"(3)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   114
    by (metis length_Cons nat.inject nat_power_eq_Suc_0_iff nat.exhaust)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   115
  have 1: "\<forall>x \<in> set(merge_adj ?ts). ltree x"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   116
    by(rule ltree_merge_adj[OF "3.prems"(1)])
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   117
  have "even (length ts)" using "3.prems"(3) even_Suc_Suc_iff by fastforce
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   118
  from "3.prems"(2) size_merge_adj[OF this] "3.prems"(1)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   119
  have 2: "\<forall>x \<in> set(merge_adj ?ts). size x = 2*n" by(auto simp: size_merge)
82010
nipkow
parents: 80776
diff changeset
   120
  have 3: "length ?ts2 = 2 ^ k'" using "3.prems"(3) k' by (simp add: length_merge_adjacent)
78231
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   121
  have 4: "length ?ts div 2 = 2 ^ k'"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   122
    using "3.prems"(3) k' by(simp add: power_eq_if[of 2 k] split: if_splits)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   123
  have "T_merge_all ?ts = T_merge_adj ?ts + T_merge_all ?ts2" by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   124
  also have "\<dots> \<le> 2^k' * Tm n + T_merge_all ?ts2"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   125
    using 4 T_merge_adj[OF "3.prems"(1,2)] by auto
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   126
  also have "\<dots> \<le> 2^k' * Tm n + (\<Sum>i=1..k'. 2^(k'-i) * Tm(2 ^ (i-1) * (2*n)))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   127
    using "3.IH"[OF 1 2 3] by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   128
  also have "\<dots> = 2^k' * Tm n + (\<Sum>i=1..k'. 2^(k'-i) * Tm(2 ^ (Suc(i-1)) * n))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   129
    by (simp add: mult_ac cong del: sum.cong)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   130
  also have "\<dots> = 2^k' * Tm n + (\<Sum>i=1..k'. 2^(k'-i) * Tm(2 ^ i * n))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   131
     by (simp)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   132
  also have "\<dots> = (\<Sum>i=1..k. 2^(k-i) * Tm(2 ^ (i-1) * real n))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   133
    by(simp add: sum.atLeast_Suc_atMost[of "Suc 0" "Suc k'"] sum.atLeast_Suc_atMost_Suc_shift[of _ "Suc 0"] k'
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   134
        del: sum.cl_ivl_Suc)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   135
  finally show ?case .
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   136
qed
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   137
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   138
lemma summation: "(\<Sum>i=1..k. 2^(k-i) * ((2::real)*i+1)) = 5*2^k - (2::real)*k - 5"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   139
proof (induction k)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   140
  case 0 thus ?case by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   141
next
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   142
  case (Suc k)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   143
  have "(\<Sum>i=1..Suc k. 2^(Suc k - i) * ((2::real)*i+1))
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   144
    = (\<Sum>i=1..k. 2^(k+1-i) * ((2::real)*i+1)) + 2*k+3"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   145
    by(simp)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   146
  also have "\<dots> = (\<Sum>i=1..k. (2::real)*(2^(k-i) * ((2::real)*i+1))) + 2*k+3"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   147
    by (simp add: Suc_diff_le mult.assoc)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   148
  also have "\<dots> = 2*(\<Sum>i=1..k. 2^(k-i) * ((2::real)*i+1)) + 2*k+3"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   149
    by(simp add: sum_distrib_left)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   150
  also have "\<dots> = (2::real)*(5*2^k - (2::real)*k - 5) + 2*k+3"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   151
    using Suc.IH by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   152
  also have "\<dots> = 5*2^(Suc k) - (2::real)*(Suc k) - 5"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   153
    by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   154
  finally show ?case .
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   155
qed
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   156
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   157
lemma T_lheap_list: assumes "length xs = 2 ^ k"
80776
3a9e570c916d more precise bound
nipkow
parents: 79984
diff changeset
   158
shows "T_lheap_list xs \<le> 5 * length xs - 2 * log 2 (length xs)"
78231
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   159
proof -
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   160
  let ?ts = "map (\<lambda>x. Node Leaf (x,1) Leaf) xs"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   161
  have "T_lheap_list xs = T_merge_all ?ts" by simp
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   162
  also have "\<dots> \<le> (\<Sum>i = 1..k. 2^(k-i) * (2 * log 2 (2^(i-1) + 1) + 1))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   163
    using T_merge_all[of ?ts 1 k] assms by (simp)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   164
  also have "\<dots> \<le> (\<Sum>i = 1..k. 2^(k-i) * (2 * log 2 (2*2^(i-1)) + 1))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   165
    apply(rule sum_mono)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   166
    using zero_le_power[of "2::real"] by (simp add: add_pos_nonneg)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   167
  also have "\<dots> = (\<Sum>i = 1..k. 2^(k-i) * (2 * log 2 (2^(1+(i-1))) + 1))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   168
    by (simp del: Suc_pred)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   169
  also have "\<dots> = (\<Sum>i = 1..k. 2^(k-i) * (2 * log 2 (2^i) + 1))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   170
    by (simp)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   171
  also have "\<dots> = (\<Sum>i = 1..k. 2^(k-i) * ((2::real)*i+1))"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   172
    by (simp add:log_nat_power algebra_simps)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   173
  also have "\<dots> = 5*(2::real)^k - (2::real)*k - 5"
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   174
    using summation by (simp)
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   175
  finally show ?thesis
80776
3a9e570c916d more precise bound
nipkow
parents: 79984
diff changeset
   176
    using assms of_nat_le_iff by simp
78231
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   177
qed
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   178
3e8d443b9512 New theory Leftist_Heap_List
nipkow
parents:
diff changeset
   179
end