72688
|
1 |
(* Author: Tobias Nipkow *)
|
|
2 |
|
|
3 |
section \<open>Heaps\<close>
|
|
4 |
|
|
5 |
theory Heaps
|
|
6 |
imports
|
|
7 |
"HOL-Library.Tree_Multiset"
|
|
8 |
Priority_Queue_Specs
|
|
9 |
begin
|
|
10 |
|
|
11 |
text \<open>Heap = priority queue on trees:\<close>
|
|
12 |
|
|
13 |
locale Heap =
|
|
14 |
fixes insert :: "('a::linorder) \<Rightarrow> 'a tree \<Rightarrow> 'a tree"
|
|
15 |
and del_min :: "'a tree \<Rightarrow> 'a tree"
|
|
16 |
assumes mset_insert: "heap q \<Longrightarrow> mset_tree (insert x q) = {#x#} + mset_tree q"
|
|
17 |
and mset_del_min: "\<lbrakk> heap q; q \<noteq> Leaf \<rbrakk> \<Longrightarrow> mset_tree (del_min q) = mset_tree q - {#value q#}"
|
|
18 |
and heap_insert: "heap q \<Longrightarrow> heap(insert x q)"
|
|
19 |
and heap_del_min: "heap q \<Longrightarrow> heap(del_min q)"
|
|
20 |
begin
|
|
21 |
|
|
22 |
definition empty :: "'a tree" where
|
|
23 |
"empty = Leaf"
|
|
24 |
|
|
25 |
fun is_empty :: "'a tree \<Rightarrow> bool" where
|
|
26 |
"is_empty t = (t = Leaf)"
|
|
27 |
|
81680
|
28 |
fun get_min :: "'a tree \<Rightarrow> 'a" where
|
|
29 |
"get_min (Node l a r) = a"
|
|
30 |
|
72688
|
31 |
sublocale Priority_Queue where empty = empty and is_empty = is_empty and insert = insert
|
81680
|
32 |
and get_min = get_min and del_min = del_min and invar = heap and mset = mset_tree
|
72688
|
33 |
proof (standard, goal_cases)
|
|
34 |
case 1 thus ?case by (simp add: empty_def)
|
|
35 |
next
|
|
36 |
case 2 thus ?case by(auto)
|
|
37 |
next
|
|
38 |
case 3 thus ?case by(simp add: mset_insert)
|
|
39 |
next
|
81680
|
40 |
case 4 thus ?case by(auto simp add: mset_del_min neq_Leaf_iff)
|
72688
|
41 |
next
|
|
42 |
case 5 thus ?case by(auto simp: neq_Leaf_iff Min_insert2 simp del: Un_iff)
|
|
43 |
next
|
|
44 |
case 6 thus ?case by(simp add: empty_def)
|
|
45 |
next
|
|
46 |
case 7 thus ?case by(simp add: heap_insert)
|
|
47 |
next
|
|
48 |
case 8 thus ?case by(simp add: heap_del_min)
|
|
49 |
qed
|
|
50 |
|
|
51 |
end
|
|
52 |
|
|
53 |
|
|
54 |
text \<open>Once you have \<open>merge\<close>, \<open>insert\<close> and \<open>del_min\<close> are easy:\<close>
|
|
55 |
|
|
56 |
locale Heap_Merge =
|
|
57 |
fixes merge :: "'a::linorder tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree"
|
|
58 |
assumes mset_merge: "\<lbrakk> heap q1; heap q2 \<rbrakk> \<Longrightarrow> mset_tree (merge q1 q2) = mset_tree q1 + mset_tree q2"
|
|
59 |
and invar_merge: "\<lbrakk> heap q1; heap q2 \<rbrakk> \<Longrightarrow> heap (merge q1 q2)"
|
|
60 |
begin
|
|
61 |
|
|
62 |
fun insert :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
|
|
63 |
"insert x t = merge (Node Leaf x Leaf) t"
|
|
64 |
|
|
65 |
fun del_min :: "'a tree \<Rightarrow> 'a tree" where
|
|
66 |
"del_min Leaf = Leaf" |
|
|
67 |
"del_min (Node l x r) = merge l r"
|
|
68 |
|
|
69 |
interpretation Heap insert del_min
|
|
70 |
proof(standard, goal_cases)
|
|
71 |
case 1 thus ?case by(simp add:mset_merge)
|
|
72 |
next
|
|
73 |
case (2 q) thus ?case by(cases q)(auto simp: mset_merge)
|
|
74 |
next
|
|
75 |
case 3 thus ?case by (simp add: invar_merge)
|
|
76 |
next
|
|
77 |
case (4 q) thus ?case by (cases q)(auto simp: invar_merge)
|
|
78 |
qed
|
|
79 |
|
81680
|
80 |
lemmas local_empty_def = local.empty_def
|
|
81 |
lemmas local_get_min_def = local.get_min.simps
|
|
82 |
|
72688
|
83 |
sublocale PQM: Priority_Queue_Merge where empty = empty and is_empty = is_empty and insert = insert
|
81680
|
84 |
and get_min = get_min and del_min = del_min and invar = heap and mset = mset_tree and merge = merge
|
72688
|
85 |
proof(standard, goal_cases)
|
|
86 |
case 1 thus ?case by (simp add: mset_merge)
|
|
87 |
next
|
|
88 |
case 2 thus ?case by (simp add: invar_merge)
|
|
89 |
qed
|
|
90 |
|
|
91 |
end
|
|
92 |
|
|
93 |
end
|