diff --git a/src/HOL/Data_Structures/Binomial_Heap.thy b/src/HOL/Data_Structures/Binomial_Heap.thy --- a/src/HOL/Data_Structures/Binomial_Heap.thy +++ b/src/HOL/Data_Structures/Binomial_Heap.thy @@ -1,720 +1,700 @@ (* Author: Peter Lammich Tobias Nipkow (tuning) *) section \Binomial Heap\ theory Binomial_Heap imports "HOL-Library.Pattern_Aliases" Complex_Main Priority_Queue_Specs begin text \ We formalize the binomial heap presentation from Okasaki's book. We show the functional correctness and complexity of all operations. The presentation is engineered for simplicity, and most proofs are straightforward and automatic. \ subsection \Binomial Tree and Heap Datatype\ datatype 'a tree = Node (rank: nat) (root: 'a) (children: "'a tree list") type_synonym 'a heap = "'a tree list" subsubsection \Multiset of elements\ fun mset_tree :: "'a::linorder tree \ 'a multiset" where "mset_tree (Node _ a ts) = {#a#} + (\t\#mset ts. mset_tree t)" definition mset_heap :: "'a::linorder heap \ 'a multiset" where "mset_heap ts = (\t\#mset ts. mset_tree t)" lemma mset_tree_simp_alt[simp]: "mset_tree (Node r a ts) = {#a#} + mset_heap ts" unfolding mset_heap_def by auto declare mset_tree.simps[simp del] lemma mset_tree_nonempty[simp]: "mset_tree t \ {#}" by (cases t) auto lemma mset_heap_Nil[simp]: "mset_heap [] = {#}" by (auto simp: mset_heap_def) lemma mset_heap_Cons[simp]: "mset_heap (t#ts) = mset_tree t + mset_heap ts" by (auto simp: mset_heap_def) lemma mset_heap_empty_iff[simp]: "mset_heap ts = {#} \ ts=[]" by (auto simp: mset_heap_def) lemma root_in_mset[simp]: "root t \# mset_tree t" by (cases t) auto lemma mset_heap_rev_eq[simp]: "mset_heap (rev ts) = mset_heap ts" by (auto simp: mset_heap_def) subsubsection \Invariants\ text \Binomial tree\ fun invar_btree :: "'a::linorder tree \ bool" where "invar_btree (Node r x ts) \ (\t\set ts. invar_btree t) \ map rank ts = rev [0..Ordering (heap) invariant\ fun invar_otree :: "'a::linorder tree \ bool" where "invar_otree (Node _ x ts) \ (\t\set ts. invar_otree t \ x \ root t)" definition "invar_tree t \ invar_btree t \ invar_otree t" text \Binomial Heap invariant\ definition "invar ts \ (\t\set ts. invar_tree t) \ (sorted_wrt (<) (map rank ts))" text \The children of a node are a valid heap\ lemma invar_children: "invar_tree (Node r v ts) \ invar (rev ts)" by (auto simp: invar_tree_def invar_def rev_map[symmetric]) subsection \Operations and Their Functional Correctness\ subsubsection \\link\\ context includes pattern_aliases begin fun link :: "('a::linorder) tree \ 'a tree \ 'a tree" where "link (Node r x\<^sub>1 ts\<^sub>1 =: t\<^sub>1) (Node r' x\<^sub>2 ts\<^sub>2 =: t\<^sub>2) = (if x\<^sub>1\x\<^sub>2 then Node (r+1) x\<^sub>1 (t\<^sub>2#ts\<^sub>1) else Node (r+1) x\<^sub>2 (t\<^sub>1#ts\<^sub>2))" end lemma invar_link: assumes "invar_tree t\<^sub>1" assumes "invar_tree t\<^sub>2" assumes "rank t\<^sub>1 = rank t\<^sub>2" shows "invar_tree (link t\<^sub>1 t\<^sub>2)" using assms unfolding invar_tree_def by (cases "(t\<^sub>1, t\<^sub>2)" rule: link.cases) auto lemma rank_link[simp]: "rank (link t\<^sub>1 t\<^sub>2) = rank t\<^sub>1 + 1" by (cases "(t\<^sub>1, t\<^sub>2)" rule: link.cases) simp lemma mset_link[simp]: "mset_tree (link t\<^sub>1 t\<^sub>2) = mset_tree t\<^sub>1 + mset_tree t\<^sub>2" by (cases "(t\<^sub>1, t\<^sub>2)" rule: link.cases) simp subsubsection \\ins_tree\\ fun ins_tree :: "'a::linorder tree \ 'a heap \ 'a heap" where "ins_tree t [] = [t]" | "ins_tree t\<^sub>1 (t\<^sub>2#ts) = (if rank t\<^sub>1 < rank t\<^sub>2 then t\<^sub>1#t\<^sub>2#ts else ins_tree (link t\<^sub>1 t\<^sub>2) ts)" lemma invar_tree0[simp]: "invar_tree (Node 0 x [])" unfolding invar_tree_def by auto lemma invar_Cons[simp]: "invar (t#ts) \ invar_tree t \ invar ts \ (\t'\set ts. rank t < rank t')" by (auto simp: invar_def) lemma invar_ins_tree: assumes "invar_tree t" assumes "invar ts" assumes "\t'\set ts. rank t \ rank t'" shows "invar (ins_tree t ts)" using assms by (induction t ts rule: ins_tree.induct) (auto simp: invar_link less_eq_Suc_le[symmetric]) lemma mset_heap_ins_tree[simp]: "mset_heap (ins_tree t ts) = mset_tree t + mset_heap ts" by (induction t ts rule: ins_tree.induct) auto lemma ins_tree_rank_bound: assumes "t' \ set (ins_tree t ts)" assumes "\t'\set ts. rank t\<^sub>0 < rank t'" assumes "rank t\<^sub>0 < rank t" shows "rank t\<^sub>0 < rank t'" using assms by (induction t ts rule: ins_tree.induct) (auto split: if_splits) subsubsection \\insert\\ hide_const (open) insert definition insert :: "'a::linorder \ 'a heap \ 'a heap" where "insert x ts = ins_tree (Node 0 x []) ts" lemma invar_insert[simp]: "invar t \ invar (insert x t)" by (auto intro!: invar_ins_tree simp: insert_def) lemma mset_heap_insert[simp]: "mset_heap (insert x t) = {#x#} + mset_heap t" by(auto simp: insert_def) subsubsection \\merge\\ context includes pattern_aliases begin fun merge :: "'a::linorder heap \ 'a heap \ 'a heap" where "merge ts\<^sub>1 [] = ts\<^sub>1" | "merge [] ts\<^sub>2 = ts\<^sub>2" | "merge (t\<^sub>1#ts\<^sub>1 =: h\<^sub>1) (t\<^sub>2#ts\<^sub>2 =: h\<^sub>2) = ( if rank t\<^sub>1 < rank t\<^sub>2 then t\<^sub>1 # merge ts\<^sub>1 h\<^sub>2 else if rank t\<^sub>2 < rank t\<^sub>1 then t\<^sub>2 # merge h\<^sub>1 ts\<^sub>2 else ins_tree (link t\<^sub>1 t\<^sub>2) (merge ts\<^sub>1 ts\<^sub>2) )" end lemma merge_simp2[simp]: "merge [] ts\<^sub>2 = ts\<^sub>2" by (cases ts\<^sub>2) auto lemma merge_rank_bound: assumes "t' \ set (merge ts\<^sub>1 ts\<^sub>2)" assumes "\t\<^sub>1\set ts\<^sub>1. rank t < rank t\<^sub>1" assumes "\t\<^sub>2\set ts\<^sub>2. rank t < rank t\<^sub>2" shows "rank t < rank t'" using assms by (induction ts\<^sub>1 ts\<^sub>2 arbitrary: t' rule: merge.induct) (auto split: if_splits simp: ins_tree_rank_bound) lemma invar_merge[simp]: assumes "invar ts\<^sub>1" assumes "invar ts\<^sub>2" shows "invar (merge ts\<^sub>1 ts\<^sub>2)" using assms by (induction ts\<^sub>1 ts\<^sub>2 rule: merge.induct) (auto 0 3 simp: Suc_le_eq intro!: invar_ins_tree invar_link elim!: merge_rank_bound) text \Longer, more explicit proof of @{thm [source] invar_merge}, to illustrate the application of the @{thm [source] merge_rank_bound} lemma.\ lemma assumes "invar ts\<^sub>1" assumes "invar ts\<^sub>2" shows "invar (merge ts\<^sub>1 ts\<^sub>2)" using assms proof (induction ts\<^sub>1 ts\<^sub>2 rule: merge.induct) case (3 t\<^sub>1 ts\<^sub>1 t\<^sub>2 ts\<^sub>2) \ \Invariants of the parts can be shown automatically\ from "3.prems" have [simp]: "invar_tree t\<^sub>1" "invar_tree t\<^sub>2" (*"invar (merge (t\<^sub>1#ts\<^sub>1) ts\<^sub>2)" "invar (merge ts\<^sub>1 (t\<^sub>2#ts\<^sub>2))" "invar (merge ts\<^sub>1 ts\<^sub>2)"*) by auto \ \These are the three cases of the @{const merge} function\ consider (LT) "rank t\<^sub>1 < rank t\<^sub>2" | (GT) "rank t\<^sub>1 > rank t\<^sub>2" | (EQ) "rank t\<^sub>1 = rank t\<^sub>2" using antisym_conv3 by blast then show ?case proof cases case LT \ \@{const merge} takes the first tree from the left heap\ then have "merge (t\<^sub>1 # ts\<^sub>1) (t\<^sub>2 # ts\<^sub>2) = t\<^sub>1 # merge ts\<^sub>1 (t\<^sub>2 # ts\<^sub>2)" by simp also have "invar \" proof (simp, intro conjI) \ \Invariant follows from induction hypothesis\ show "invar (merge ts\<^sub>1 (t\<^sub>2 # ts\<^sub>2))" using LT "3.IH" "3.prems" by simp \ \It remains to show that \t\<^sub>1\ has smallest rank.\ show "\t'\set (merge ts\<^sub>1 (t\<^sub>2 # ts\<^sub>2)). rank t\<^sub>1 < rank t'" \ \Which is done by auxiliary lemma @{thm [source] merge_rank_bound}\ using LT "3.prems" by (force elim!: merge_rank_bound) qed finally show ?thesis . next \ \@{const merge} takes the first tree from the right heap\ case GT \ \The proof is anaologous to the \LT\ case\ then show ?thesis using "3.prems" "3.IH" by (force elim!: merge_rank_bound) next case [simp]: EQ \ \@{const merge} links both first trees, and inserts them into the merged remaining heaps\ have "merge (t\<^sub>1 # ts\<^sub>1) (t\<^sub>2 # ts\<^sub>2) = ins_tree (link t\<^sub>1 t\<^sub>2) (merge ts\<^sub>1 ts\<^sub>2)" by simp also have "invar \" proof (intro invar_ins_tree invar_link) \ \Invariant of merged remaining heaps follows by IH\ show "invar (merge ts\<^sub>1 ts\<^sub>2)" using EQ "3.prems" "3.IH" by auto \ \For insertion, we have to show that the rank of the linked tree is \\\ the ranks in the merged remaining heaps\ show "\t'\set (merge ts\<^sub>1 ts\<^sub>2). rank (link t\<^sub>1 t\<^sub>2) \ rank t'" proof - \ \Which is, again, done with the help of @{thm [source] merge_rank_bound}\ have "rank (link t\<^sub>1 t\<^sub>2) = Suc (rank t\<^sub>2)" by simp thus ?thesis using "3.prems" by (auto simp: Suc_le_eq elim!: merge_rank_bound) qed qed simp_all finally show ?thesis . qed qed auto lemma mset_heap_merge[simp]: "mset_heap (merge ts\<^sub>1 ts\<^sub>2) = mset_heap ts\<^sub>1 + mset_heap ts\<^sub>2" by (induction ts\<^sub>1 ts\<^sub>2 rule: merge.induct) auto subsubsection \\get_min\\ fun get_min :: "'a::linorder heap \ 'a" where "get_min [t] = root t" | "get_min (t#ts) = min (root t) (get_min ts)" lemma invar_tree_root_min: assumes "invar_tree t" assumes "x \# mset_tree t" shows "root t \ x" using assms unfolding invar_tree_def by (induction t arbitrary: x rule: mset_tree.induct) (fastforce simp: mset_heap_def) lemma get_min_mset: assumes "ts\[]" assumes "invar ts" assumes "x \# mset_heap ts" shows "get_min ts \ x" using assms apply (induction ts arbitrary: x rule: get_min.induct) apply (auto simp: invar_tree_root_min min_def intro: order_trans; meson linear order_trans invar_tree_root_min )+ done lemma get_min_member: "ts\[] \ get_min ts \# mset_heap ts" by (induction ts rule: get_min.induct) (auto simp: min_def) lemma get_min: assumes "mset_heap ts \ {#}" assumes "invar ts" shows "get_min ts = Min_mset (mset_heap ts)" using assms get_min_member get_min_mset by (auto simp: eq_Min_iff) subsubsection \\get_min_rest\\ fun get_min_rest :: "'a::linorder heap \ 'a tree \ 'a heap" where "get_min_rest [t] = (t,[])" | "get_min_rest (t#ts) = (let (t',ts') = get_min_rest ts in if root t \ root t' then (t,ts) else (t',t#ts'))" lemma get_min_rest_get_min_same_root: assumes "ts\[]" assumes "get_min_rest ts = (t',ts')" shows "root t' = get_min ts" using assms by (induction ts arbitrary: t' ts' rule: get_min.induct) (auto simp: min_def split: prod.splits) lemma mset_get_min_rest: assumes "get_min_rest ts = (t',ts')" assumes "ts\[]" shows "mset ts = {#t'#} + mset ts'" using assms by (induction ts arbitrary: t' ts' rule: get_min.induct) (auto split: prod.splits if_splits) lemma set_get_min_rest: assumes "get_min_rest ts = (t', ts')" assumes "ts\[]" shows "set ts = Set.insert t' (set ts')" using mset_get_min_rest[OF assms, THEN arg_cong[where f=set_mset]] by auto lemma invar_get_min_rest: assumes "get_min_rest ts = (t',ts')" assumes "ts\[]" assumes "invar ts" shows "invar_tree t'" and "invar ts'" proof - have "invar_tree t' \ invar ts'" using assms proof (induction ts arbitrary: t' ts' rule: get_min.induct) case (2 t v va) then show ?case apply (clarsimp split: prod.splits if_splits) apply (drule set_get_min_rest; fastforce) done qed auto thus "invar_tree t'" and "invar ts'" by auto qed subsubsection \\del_min\\ definition del_min :: "'a::linorder heap \ 'a::linorder heap" where "del_min ts = (case get_min_rest ts of (Node r x ts\<^sub>1, ts\<^sub>2) \ merge (rev ts\<^sub>1) ts\<^sub>2)" lemma invar_del_min[simp]: assumes "ts \ []" assumes "invar ts" shows "invar (del_min ts)" using assms unfolding del_min_def by (auto split: prod.split tree.split intro!: invar_merge invar_children dest: invar_get_min_rest ) lemma mset_heap_del_min: assumes "ts \ []" shows "mset_heap ts = mset_heap (del_min ts) + {# get_min ts #}" using assms unfolding del_min_def apply (clarsimp split: tree.split prod.split) apply (frule (1) get_min_rest_get_min_same_root) apply (frule (1) mset_get_min_rest) apply (auto simp: mset_heap_def) done subsubsection \Instantiating the Priority Queue Locale\ text \Last step of functional correctness proof: combine all the above lemmas to show that binomial heaps satisfy the specification of priority queues with merge.\ interpretation binheap: Priority_Queue_Merge where empty = "[]" and is_empty = "(=) []" and insert = insert and get_min = get_min and del_min = del_min and merge = merge and invar = invar and mset = mset_heap proof (unfold_locales, goal_cases) case 1 thus ?case by simp next case 2 thus ?case by auto next case 3 thus ?case by auto next case (4 q) thus ?case using mset_heap_del_min[of q] get_min[OF _ \invar q\] by (auto simp: union_single_eq_diff) next case (5 q) thus ?case using get_min[of q] by auto next case 6 thus ?case by (auto simp add: invar_def) next case 7 thus ?case by simp next case 8 thus ?case by simp next case 9 thus ?case by simp next case 10 thus ?case by simp qed subsection \Complexity\ text \The size of a binomial tree is determined by its rank\ lemma size_mset_btree: assumes "invar_btree t" shows "size (mset_tree t) = 2^rank t" using assms proof (induction t) case (Node r v ts) hence IH: "size (mset_tree t) = 2^rank t" if "t \ set ts" for t using that by auto from Node have COMPL: "map rank ts = rev [0..t\ts. size (mset_tree t))" by (induction ts) auto also have "\ = (\t\ts. 2^rank t)" using IH by (auto cong: map_cong) also have "\ = (\r\map rank ts. 2^r)" by (induction ts) auto also have "\ = (\i\{0.. = 2^r - 1" by (induction r) auto finally show ?case by (simp) qed lemma size_mset_tree: assumes "invar_tree t" shows "size (mset_tree t) = 2^rank t" using assms unfolding invar_tree_def by (simp add: size_mset_btree) text \The length of a binomial heap is bounded by the number of its elements\ lemma size_mset_heap: assumes "invar ts" shows "2^length ts \ size (mset_heap ts) + 1" proof - from \invar ts\ have ASC: "sorted_wrt (<) (map rank ts)" and TINV: "\t\set ts. invar_tree t" unfolding invar_def by auto have "(2::nat)^length ts = (\i\{0.. \ (\t\ts. 2^rank t) + 1" using sorted_wrt_less_sum_mono_lowerbound[OF _ ASC, of "(^) (2::nat)"] using power_increasing[where a="2::nat"] by (auto simp: o_def) also have "\ = (\t\ts. size (mset_tree t)) + 1" using TINV by (auto cong: map_cong simp: size_mset_tree) also have "\ = size (mset_heap ts) + 1" unfolding mset_heap_def by (induction ts) auto finally show ?thesis . qed subsubsection \Timing Functions\ text \ We define timing functions for each operation, and provide estimations of their complexity. \ definition T_link :: "'a::linorder tree \ 'a tree \ nat" where [simp]: "T_link _ _ = 1" fun T_ins_tree :: "'a::linorder tree \ 'a heap \ nat" where "T_ins_tree t [] = 1" | "T_ins_tree t\<^sub>1 (t\<^sub>2 # rest) = ( (if rank t\<^sub>1 < rank t\<^sub>2 then 1 else T_link t\<^sub>1 t\<^sub>2 + T_ins_tree (link t\<^sub>1 t\<^sub>2) rest) )" definition T_insert :: "'a::linorder \ 'a heap \ nat" where "T_insert x ts = T_ins_tree (Node 0 x []) ts" lemma T_ins_tree_simple_bound: "T_ins_tree t ts \ length ts + 1" by (induction t ts rule: T_ins_tree.induct) auto subsubsection \\T_insert\\ lemma T_insert_bound: assumes "invar ts" shows "T_insert x ts \ log 2 (size (mset_heap ts) + 1) + 1" proof - have 1: "T_insert x ts \ length ts + 1" unfolding T_insert_def by (rule T_ins_tree_simple_bound) also have "\ \ log 2 (2 * (size (mset_heap ts) + 1))" proof - from size_mset_heap[of ts] assms have "2 ^ length ts \ size (mset_heap ts) + 1" unfolding invar_def by auto hence "2 ^ (length ts + 1) \ 2 * (size (mset_heap ts) + 1)" by auto thus ?thesis using le_log2_of_power by blast qed finally show ?thesis by (simp only: log_mult of_nat_mult) auto qed subsubsection \\T_merge\\ context includes pattern_aliases begin fun T_merge :: "'a::linorder heap \ 'a heap \ nat" where "T_merge ts\<^sub>1 [] = 1" | "T_merge [] ts\<^sub>2 = 1" | "T_merge (t\<^sub>1#ts\<^sub>1 =: h\<^sub>1) (t\<^sub>2#ts\<^sub>2 =: h\<^sub>2) = 1 + ( if rank t\<^sub>1 < rank t\<^sub>2 then T_merge ts\<^sub>1 h\<^sub>2 else if rank t\<^sub>2 < rank t\<^sub>1 then T_merge h\<^sub>1 ts\<^sub>2 else T_ins_tree (link t\<^sub>1 t\<^sub>2) (merge ts\<^sub>1 ts\<^sub>2) + T_merge ts\<^sub>1 ts\<^sub>2 )" end text \A crucial idea is to estimate the time in correlation with the result length, as each carry reduces the length of the result.\ lemma T_ins_tree_length: "T_ins_tree t ts + length (ins_tree t ts) = 2 + length ts" by (induction t ts rule: ins_tree.induct) auto lemma T_merge_length: "length (merge ts\<^sub>1 ts\<^sub>2) + T_merge ts\<^sub>1 ts\<^sub>2 \ 2 * (length ts\<^sub>1 + length ts\<^sub>2) + 1" by (induction ts\<^sub>1 ts\<^sub>2 rule: T_merge.induct) (auto simp: T_ins_tree_length algebra_simps) text \Finally, we get the desired logarithmic bound\ -lemma T_merge_bound_aux: +lemma T_merge_bound: fixes ts\<^sub>1 ts\<^sub>2 defines "n\<^sub>1 \ size (mset_heap ts\<^sub>1)" defines "n\<^sub>2 \ size (mset_heap ts\<^sub>2)" - assumes BINVARS: "invar ts\<^sub>1" "invar ts\<^sub>2" + assumes "invar ts\<^sub>1" "invar ts\<^sub>2" shows "T_merge ts\<^sub>1 ts\<^sub>2 \ 4*log 2 (n\<^sub>1 + n\<^sub>2 + 1) + 2" proof - define n where "n = n\<^sub>1 + n\<^sub>2" + note BINVARS = \invar ts\<^sub>1\ \invar ts\<^sub>2\ + from T_merge_length[of ts\<^sub>1 ts\<^sub>2] have "T_merge ts\<^sub>1 ts\<^sub>2 \ 2 * (length ts\<^sub>1 + length ts\<^sub>2) + 1" by auto hence "(2::nat)^T_merge ts\<^sub>1 ts\<^sub>2 \ 2^(2 * (length ts\<^sub>1 + length ts\<^sub>2) + 1)" by (rule power_increasing) auto also have "\ = 2*(2^length ts\<^sub>1)\<^sup>2*(2^length ts\<^sub>2)\<^sup>2" by (auto simp: algebra_simps power_add power_mult) also note BINVARS(1)[THEN size_mset_heap] also note BINVARS(2)[THEN size_mset_heap] finally have "2 ^ T_merge ts\<^sub>1 ts\<^sub>2 \ 2 * (n\<^sub>1 + 1)\<^sup>2 * (n\<^sub>2 + 1)\<^sup>2" by (auto simp: power2_nat_le_eq_le n\<^sub>1_def n\<^sub>2_def) from le_log2_of_power[OF this] have "T_merge ts\<^sub>1 ts\<^sub>2 \ log 2 \" by simp also have "\ = log 2 2 + 2*log 2 (n\<^sub>1 + 1) + 2*log 2 (n\<^sub>2 + 1)" by (simp add: log_mult log_nat_power) also have "n\<^sub>2 \ n" by (auto simp: n_def) finally have "T_merge ts\<^sub>1 ts\<^sub>2 \ log 2 2 + 2*log 2 (n\<^sub>1 + 1) + 2*log 2 (n + 1)" by auto also have "n\<^sub>1 \ n" by (auto simp: n_def) finally have "T_merge ts\<^sub>1 ts\<^sub>2 \ log 2 2 + 4*log 2 (n + 1)" by auto also have "log 2 2 \ 2" by auto finally have "T_merge ts\<^sub>1 ts\<^sub>2 \ 4*log 2 (n + 1) + 2" by auto thus ?thesis unfolding n_def by (auto simp: algebra_simps) qed -lemma T_merge_bound: - fixes ts\<^sub>1 ts\<^sub>2 - defines "n\<^sub>1 \ size (mset_heap ts\<^sub>1)" - defines "n\<^sub>2 \ size (mset_heap ts\<^sub>2)" - assumes "invar ts\<^sub>1" "invar ts\<^sub>2" - shows "T_merge ts\<^sub>1 ts\<^sub>2 \ 4*log 2 (n\<^sub>1 + n\<^sub>2 + 1) + 2" -using assms T_merge_bound_aux unfolding invar_def by blast - subsubsection \\T_get_min\\ fun T_get_min :: "'a::linorder heap \ nat" where "T_get_min [t] = 1" | "T_get_min (t#ts) = 1 + T_get_min ts" lemma T_get_min_estimate: "ts\[] \ T_get_min ts = length ts" by (induction ts rule: T_get_min.induct) auto lemma T_get_min_bound: assumes "invar ts" assumes "ts\[]" shows "T_get_min ts \ log 2 (size (mset_heap ts) + 1)" proof - have 1: "T_get_min ts = length ts" using assms T_get_min_estimate by auto also have "\ \ log 2 (size (mset_heap ts) + 1)" proof - from size_mset_heap[of ts] assms have "2 ^ length ts \ size (mset_heap ts) + 1" unfolding invar_def by auto thus ?thesis using le_log2_of_power by blast qed finally show ?thesis by auto qed subsubsection \\T_del_min\\ fun T_get_min_rest :: "'a::linorder heap \ nat" where "T_get_min_rest [t] = 1" | "T_get_min_rest (t#ts) = 1 + T_get_min_rest ts" lemma T_get_min_rest_estimate: "ts\[] \ T_get_min_rest ts = length ts" by (induction ts rule: T_get_min_rest.induct) auto -lemma T_get_min_rest_bound_aux: +lemma T_get_min_rest_bound: assumes "invar ts" assumes "ts\[]" shows "T_get_min_rest ts \ log 2 (size (mset_heap ts) + 1)" proof - have 1: "T_get_min_rest ts = length ts" using assms T_get_min_rest_estimate by auto also have "\ \ log 2 (size (mset_heap ts) + 1)" proof - from size_mset_heap[of ts] assms have "2 ^ length ts \ size (mset_heap ts) + 1" by auto thus ?thesis using le_log2_of_power by blast qed finally show ?thesis by auto qed -lemma T_get_min_rest_bound: - assumes "invar ts" - assumes "ts\[]" - shows "T_get_min_rest ts \ log 2 (size (mset_heap ts) + 1)" -using assms T_get_min_rest_bound_aux unfolding invar_def by blast - text\Note that although the definition of function \<^const>\rev\ has quadratic complexity, it can and is implemented (via suitable code lemmas) as a linear time function. Thus the following definition is justified:\ definition "T_rev xs = length xs + 1" definition T_del_min :: "'a::linorder heap \ nat" where "T_del_min ts = T_get_min_rest ts + (case get_min_rest ts of (Node _ x ts\<^sub>1, ts\<^sub>2) \ T_rev ts\<^sub>1 + T_merge (rev ts\<^sub>1) ts\<^sub>2 )" lemma T_rev_ts1_bound_aux: fixes ts defines "n \ size (mset_heap ts)" assumes BINVAR: "invar (rev ts)" shows "T_rev ts \ 1 + log 2 (n+1)" proof - have "T_rev ts = length ts + 1" by (auto simp: T_rev_def) hence "2^T_rev ts = 2*2^length ts" by auto also have "\ \ 2*n+2" using size_mset_heap[OF BINVAR] by (auto simp: n_def) finally have "2 ^ T_rev ts \ 2 * n + 2" . from le_log2_of_power[OF this] have "T_rev ts \ log 2 (2 * (n + 1))" by auto also have "\ = 1 + log 2 (n+1)" by (simp only: of_nat_mult log_mult) auto finally show ?thesis by (auto simp: algebra_simps) qed -lemma T_del_min_bound_aux: +lemma T_del_min_bound: fixes ts defines "n \ size (mset_heap ts)" - assumes BINVAR: "invar ts" + assumes "invar ts" assumes "ts\[]" shows "T_del_min ts \ 6 * log 2 (n+1) + 3" proof - obtain r x ts\<^sub>1 ts\<^sub>2 where GM: "get_min_rest ts = (Node r x ts\<^sub>1, ts\<^sub>2)" by (metis surj_pair tree.exhaust_sel) - note BINVAR' = invar_get_min_rest[OF GM \ts\[]\ BINVAR] + note BINVAR' = invar_get_min_rest[OF GM \ts\[]\ \invar ts\] hence BINVAR1: "invar (rev ts\<^sub>1)" by (blast intro: invar_children) define n\<^sub>1 where "n\<^sub>1 = size (mset_heap ts\<^sub>1)" define n\<^sub>2 where "n\<^sub>2 = size (mset_heap ts\<^sub>2)" have T_rev_ts1_bound: "T_rev ts\<^sub>1 \ 1 + log 2 (n+1)" proof - note T_rev_ts1_bound_aux[OF BINVAR1, simplified, folded n\<^sub>1_def] also have "n\<^sub>1 \ n" unfolding n\<^sub>1_def n_def using mset_get_min_rest[OF GM \ts\[]\] by (auto simp: mset_heap_def) finally show ?thesis by (auto simp: algebra_simps) qed have "T_del_min ts = T_get_min_rest ts + T_rev ts\<^sub>1 + T_merge (rev ts\<^sub>1) ts\<^sub>2" unfolding T_del_min_def by (simp add: GM) also have "\ \ log 2 (n+1) + T_rev ts\<^sub>1 + T_merge (rev ts\<^sub>1) ts\<^sub>2" - using T_get_min_rest_bound_aux[OF assms(2-)] by (auto simp: n_def) + using T_get_min_rest_bound[OF assms(2-)] by (auto simp: n_def) also have "\ \ 2*log 2 (n+1) + T_merge (rev ts\<^sub>1) ts\<^sub>2 + 1" using T_rev_ts1_bound by auto also have "\ \ 2*log 2 (n+1) + 4 * log 2 (n\<^sub>1 + n\<^sub>2 + 1) + 3" - using T_merge_bound_aux[OF \invar (rev ts\<^sub>1)\ \invar ts\<^sub>2\] + using T_merge_bound[OF \invar (rev ts\<^sub>1)\ \invar ts\<^sub>2\] by (auto simp: n\<^sub>1_def n\<^sub>2_def algebra_simps) also have "n\<^sub>1 + n\<^sub>2 \ n" unfolding n\<^sub>1_def n\<^sub>2_def n_def using mset_get_min_rest[OF GM \ts\[]\] by (auto simp: mset_heap_def) finally have "T_del_min ts \ 6 * log 2 (n+1) + 3" by auto thus ?thesis by (simp add: algebra_simps) qed -lemma T_del_min_bound: - fixes ts - defines "n \ size (mset_heap ts)" - assumes "invar ts" - assumes "ts\[]" - shows "T_del_min ts \ 6 * log 2 (n+1) + 3" -using assms T_del_min_bound_aux unfolding invar_def by blast - end