src/HOL/Data_Structures/Leftist_Heap.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 80484 0ca36dcdcbd3
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
     2
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
     3
section \<open>Leftist Heap\<close>
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
     4
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
     5
theory Leftist_Heap
66419
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
     6
imports
70450
7c2724cefcb8 reduced dependencies
nipkow
parents: 70363
diff changeset
     7
  "HOL-Library.Pattern_Aliases"
66419
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
     8
  Tree2
68492
c7e0a7bcacaf added lemmas; uniform names
nipkow
parents: 68413
diff changeset
     9
  Priority_Queue_Specs
66419
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
    10
  Complex_Main
79490
b287510a4202 Added time function automation
nipkow
parents: 77937
diff changeset
    11
  Define_Time_Function
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    12
begin
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    13
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    14
fun mset_tree :: "('a*'b) tree \<Rightarrow> 'a multiset" where
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    15
"mset_tree Leaf = {#}" |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    16
"mset_tree (Node l (a, _) r) = {#a#} + mset_tree l + mset_tree r"
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    17
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    18
type_synonym 'a lheap = "('a*nat)tree"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    19
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    20
fun mht :: "'a lheap \<Rightarrow> nat" where
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    21
"mht Leaf = 0" |
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    22
"mht (Node _ (_, n) _) = n"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    23
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 66565
diff changeset
    24
text\<open>The invariants:\<close>
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    25
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    26
fun (in linorder) heap :: "('a*'b) tree \<Rightarrow> bool" where
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    27
"heap Leaf = True" |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    28
"heap (Node l (m, _) r) =
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    29
  ((\<forall>x \<in> set_tree l \<union> set_tree r. m \<le> x) \<and> heap l \<and> heap r)"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    30
64973
nipkow
parents: 64971
diff changeset
    31
fun ltree :: "'a lheap \<Rightarrow> bool" where
nipkow
parents: 64971
diff changeset
    32
"ltree Leaf = True" |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    33
"ltree (Node l (a, n) r) =
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    34
 (min_height l \<ge> min_height r \<and> n = min_height r + 1 \<and> ltree l & ltree r)"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    35
70585
nipkow
parents: 70450
diff changeset
    36
definition empty :: "'a lheap" where
nipkow
parents: 70450
diff changeset
    37
"empty = Leaf"
nipkow
parents: 70450
diff changeset
    38
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    39
definition node :: "'a lheap \<Rightarrow> 'a \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    40
"node l a r =
73619
nipkow
parents: 72540
diff changeset
    41
 (let mhl = mht l; mhr = mht r
nipkow
parents: 72540
diff changeset
    42
  in if mhl \<ge> mhr then Node l (a,mhr+1) r else Node r (a,mhl+1) l)"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    43
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    44
fun get_min :: "'a lheap \<Rightarrow> 'a" where
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    45
"get_min(Node l (a, n) r) = a"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    46
66499
nipkow
parents: 66491
diff changeset
    47
text \<open>For function \<open>merge\<close>:\<close>
nipkow
parents: 66491
diff changeset
    48
unbundle pattern_aliases
66491
nipkow
parents: 66425
diff changeset
    49
66499
nipkow
parents: 66491
diff changeset
    50
fun merge :: "'a::ord lheap \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
70585
nipkow
parents: 70450
diff changeset
    51
"merge Leaf t = t" |
nipkow
parents: 70450
diff changeset
    52
"merge t Leaf = t" |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    53
"merge (Node l1 (a1, n1) r1 =: t1) (Node l2 (a2, n2) r2 =: t2) =
66491
nipkow
parents: 66425
diff changeset
    54
   (if a1 \<le> a2 then node l1 a1 (merge r1 t2)
68600
bdd6536bd57c more symmetric
nipkow
parents: 68492
diff changeset
    55
    else node l2 a2 (merge t1 r2))"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    56
70585
nipkow
parents: 70450
diff changeset
    57
text \<open>Termination of @{const merge}: by sum or lexicographic product of the sizes
nipkow
parents: 70450
diff changeset
    58
of the two arguments. Isabelle uses a lexicographic product.\<close>
nipkow
parents: 70450
diff changeset
    59
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
    60
lemma merge_code: "merge t1 t2 = (case (t1,t2) of
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    61
  (Leaf, _) \<Rightarrow> t2 |
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    62
  (_, Leaf) \<Rightarrow> t1 |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    63
  (Node l1 (a1, n1) r1, Node l2 (a2, n2) r2) \<Rightarrow>
68600
bdd6536bd57c more symmetric
nipkow
parents: 68492
diff changeset
    64
    if a1 \<le> a2 then node l1 a1 (merge r1 t2) else node l2 a2 (merge t1 r2))"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
    65
by(induction t1 t2 rule: merge.induct) (simp_all split: tree.split)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    66
66522
nipkow
parents: 66499
diff changeset
    67
hide_const (open) insert
nipkow
parents: 66499
diff changeset
    68
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    69
definition insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    70
"insert x t = merge (Node Leaf (x,1) Leaf) t"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    71
68021
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
    72
fun del_min :: "'a::ord lheap \<Rightarrow> 'a lheap" where
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
    73
"del_min Leaf = Leaf" |
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
    74
"del_min (Node l _ r) = merge l r"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    75
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    76
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    77
subsection "Lemmas"
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    78
66419
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
    79
lemma mset_tree_empty: "mset_tree t = {#} \<longleftrightarrow> t = Leaf"
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
    80
by(cases t) auto
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
    81
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    82
lemma mht_eq_min_height: "ltree t \<Longrightarrow> mht t = min_height t"
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    83
by(cases t) auto
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    84
64973
nipkow
parents: 64971
diff changeset
    85
lemma ltree_node: "ltree (node l a r) \<longleftrightarrow> ltree l \<and> ltree r"
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
    86
by(auto simp add: node_def mht_eq_min_height)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    87
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    88
lemma heap_node: "heap (node l a r) \<longleftrightarrow>
70585
nipkow
parents: 70450
diff changeset
    89
  heap l \<and> heap r \<and> (\<forall>x \<in> set_tree l \<union> set_tree r. a \<le> x)"
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    90
by(auto simp add: node_def)
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
    91
70585
nipkow
parents: 70450
diff changeset
    92
lemma set_tree_mset: "set_tree t = set_mset(mset_tree t)"
nipkow
parents: 70450
diff changeset
    93
by(induction t) auto
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    94
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    95
subsection "Functional Correctness"
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    96
72282
nipkow
parents: 70755
diff changeset
    97
lemma mset_merge: "mset_tree (merge t1 t2) = mset_tree t1 + mset_tree t2"
nipkow
parents: 70755
diff changeset
    98
by (induction t1 t2 rule: merge.induct) (auto simp add: node_def ac_simps)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
    99
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   100
lemma mset_insert: "mset_tree (insert x t) = mset_tree t + {#x#}"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   101
by (auto simp add: insert_def mset_merge)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   102
72282
nipkow
parents: 70755
diff changeset
   103
lemma get_min: "\<lbrakk> heap t;  t \<noteq> Leaf \<rbrakk> \<Longrightarrow> get_min t = Min(set_tree t)"
nipkow
parents: 70755
diff changeset
   104
by (cases t) (auto simp add: eq_Min_iff)
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   105
72282
nipkow
parents: 70755
diff changeset
   106
lemma mset_del_min: "mset_tree (del_min t) = mset_tree t - {# get_min t #}"
nipkow
parents: 70755
diff changeset
   107
by (cases t) (auto simp: mset_merge)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   108
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   109
lemma ltree_merge: "\<lbrakk> ltree l; ltree r \<rbrakk> \<Longrightarrow> ltree (merge l r)"
72282
nipkow
parents: 70755
diff changeset
   110
by(induction l r rule: merge.induct)(auto simp: ltree_node)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   111
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   112
lemma heap_merge: "\<lbrakk> heap l; heap r \<rbrakk> \<Longrightarrow> heap (merge l r)"
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   113
proof(induction l r rule: merge.induct)
70585
nipkow
parents: 70450
diff changeset
   114
  case 3 thus ?case by(auto simp: heap_node mset_merge ball_Un set_tree_mset)
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   115
qed simp_all
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   116
64973
nipkow
parents: 64971
diff changeset
   117
lemma ltree_insert: "ltree t \<Longrightarrow> ltree(insert x t)"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   118
by(simp add: insert_def ltree_merge del: merge.simps split: tree.split)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   119
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   120
lemma heap_insert: "heap t \<Longrightarrow> heap(insert x t)"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   121
by(simp add: insert_def heap_merge del: merge.simps split: tree.split)
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   122
68021
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
   123
lemma ltree_del_min: "ltree t \<Longrightarrow> ltree(del_min t)"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   124
by(cases t)(auto simp add: ltree_merge simp del: merge.simps)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   125
68021
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
   126
lemma heap_del_min: "heap t \<Longrightarrow> heap(del_min t)"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   127
by(cases t)(auto simp add: heap_merge simp del: merge.simps)
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   128
66565
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   129
text \<open>Last step of functional correctness proof: combine all the above lemmas
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   130
to show that leftist heaps satisfy the specification of priority queues with merge.\<close>
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   131
66565
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   132
interpretation lheap: Priority_Queue_Merge
72282
nipkow
parents: 70755
diff changeset
   133
where empty = empty and is_empty = "\<lambda>t. t = Leaf"
68021
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
   134
and insert = insert and del_min = del_min
66565
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   135
and get_min = get_min and merge = merge
72282
nipkow
parents: 70755
diff changeset
   136
and invar = "\<lambda>t. heap t \<and> ltree t" and mset = mset_tree
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   137
proof(standard, goal_cases)
70585
nipkow
parents: 70450
diff changeset
   138
  case 1 show ?case by (simp add: empty_def)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   139
next
64975
96b66d5c0fc1 added is_empty
nipkow
parents: 64973
diff changeset
   140
  case (2 q) show ?case by (cases q) auto
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   141
next
64975
96b66d5c0fc1 added is_empty
nipkow
parents: 64973
diff changeset
   142
  case 3 show ?case by(rule mset_insert)
96b66d5c0fc1 added is_empty
nipkow
parents: 64973
diff changeset
   143
next
68021
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
   144
  case 4 show ?case by(rule mset_del_min)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   145
next
70585
nipkow
parents: 70450
diff changeset
   146
  case 5 thus ?case by(simp add: get_min mset_tree_empty set_tree_mset)
66419
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
   147
next
70585
nipkow
parents: 70450
diff changeset
   148
  case 6 thus ?case by(simp add: empty_def)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   149
next
66419
8194ed7cf2cb separate file for priority queue interface; extended Leftist_Heap.
nipkow
parents: 64977
diff changeset
   150
  case 7 thus ?case by(simp add: heap_insert ltree_insert)
64968
a7ea55c1be52 proper priority queue spec
nipkow
parents: 62706
diff changeset
   151
next
68021
b91a043c0dcb dont rename PQ.del_min
nipkow
parents: 68020
diff changeset
   152
  case 8 thus ?case by(simp add: heap_del_min ltree_del_min)
66565
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   153
next
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   154
  case 9 thus ?case by (simp add: mset_merge)
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   155
next
ff561d9cb661 added PQ with merge
nipkow
parents: 66522
diff changeset
   156
  case 10 thus ?case by (simp add: heap_merge ltree_merge)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   157
qed
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   158
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   159
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   160
subsection "Complexity"
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   161
79490
b287510a4202 Added time function automation
nipkow
parents: 77937
diff changeset
   162
text \<open>Auxiliary time functions (which are both 0):\<close>
79969
4aeb25ba90f3 more uniform command names
nipkow
parents: 79490
diff changeset
   163
time_fun mht
4aeb25ba90f3 more uniform command names
nipkow
parents: 79490
diff changeset
   164
time_fun node
66491
nipkow
parents: 66425
diff changeset
   165
79490
b287510a4202 Added time function automation
nipkow
parents: 77937
diff changeset
   166
lemma T_mht_0[simp]: "T_mht t = 0"
b287510a4202 Added time function automation
nipkow
parents: 77937
diff changeset
   167
by(cases t)auto
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   168
79490
b287510a4202 Added time function automation
nipkow
parents: 77937
diff changeset
   169
text \<open>Define timing function\<close>
79969
4aeb25ba90f3 more uniform command names
nipkow
parents: 79490
diff changeset
   170
time_fun merge
4aeb25ba90f3 more uniform command names
nipkow
parents: 79490
diff changeset
   171
time_fun insert
4aeb25ba90f3 more uniform command names
nipkow
parents: 79490
diff changeset
   172
time_fun del_min
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   173
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   174
lemma T_merge_min_height: "ltree l \<Longrightarrow> ltree r \<Longrightarrow> T_merge l r \<le> min_height l + min_height r + 1"
64976
1a4cb9403a10 renaming
nipkow
parents: 64975
diff changeset
   175
proof(induction l r rule: merge.induct)
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   176
  case 3 thus ?case by(auto)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   177
qed simp_all
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   178
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   179
corollary T_merge_log: assumes "ltree l" "ltree r"
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   180
  shows "T_merge l r \<le> log 2 (size1 l) + log 2 (size1 r) + 1"
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   181
using le_log2_of_power[OF min_height_size1[of l]]
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   182
  le_log2_of_power[OF min_height_size1[of r]] T_merge_min_height[of l r] assms
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   183
by linarith
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   184
77937
8fa4e4fd852e streamlined
nipkow
parents: 73619
diff changeset
   185
corollary T_insert_log: "ltree t \<Longrightarrow> T_insert x t \<le> log 2 (size1 t) + 2"
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   186
using T_merge_log[of "Node Leaf (x, 1) Leaf" t]
79490
b287510a4202 Added time function automation
nipkow
parents: 77937
diff changeset
   187
by(simp split: tree.split)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   188
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   189
corollary T_del_min_log: assumes "ltree t"
80484
0ca36dcdcbd3 simpler theorem
nipkow
parents: 79969
diff changeset
   190
  shows "T_del_min t \<le> 2 * log 2 (size1 t) + 1"
70755
3fb16bed5d6c replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents: 70585
diff changeset
   191
proof(cases t rule: tree2_cases)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   192
  case Leaf thus ?thesis using assms by simp
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   193
next
72540
8eabaf951e6b use min_height as in (much of?) the literature
nipkow
parents: 72282
diff changeset
   194
  case [simp]: (Node l _ _ r)
80484
0ca36dcdcbd3 simpler theorem
nipkow
parents: 79969
diff changeset
   195
  show ?thesis
0ca36dcdcbd3 simpler theorem
nipkow
parents: 79969
diff changeset
   196
    using \<open>ltree t\<close> T_merge_log[of l r]
0ca36dcdcbd3 simpler theorem
nipkow
parents: 79969
diff changeset
   197
      log_mono[of 2 "size1 l" "size1 t"] log_mono[of 2 "size1 r" "size1 t"]
0ca36dcdcbd3 simpler theorem
nipkow
parents: 79969
diff changeset
   198
    by (simp del: T_merge.simps)
62706
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   199
qed
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   200
49c6a54ceab6 added Leftist_Heap
nipkow
parents:
diff changeset
   201
end