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