author | wenzelm |
Thu, 08 Dec 2022 11:24:43 +0100 | |
changeset 76598 | 9f97eda3fcf1 |
parent 75501 | 426afab39a55 |
child 77922 | d28dcd57d2f3 |
permissions | -rw-r--r-- |
66543 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
70250 | 3 |
section "Sorting" |
4 |
||
66543 | 5 |
theory Sorting |
6 |
imports |
|
7 |
Complex_Main |
|
8 |
"HOL-Library.Multiset" |
|
9 |
begin |
|
10 |
||
68160 | 11 |
hide_const List.insort |
66543 | 12 |
|
13 |
declare Let_def [simp] |
|
14 |
||
15 |
||
16 |
subsection "Insertion Sort" |
|
17 |
||
75501 | 18 |
fun insort1 :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
19 |
"insort1 x [] = [x]" | |
|
20 |
"insort1 x (y#ys) = |
|
21 |
(if x \<le> y then x#y#ys else y#(insort1 x ys))" |
|
66543 | 22 |
|
75501 | 23 |
fun insort :: "'a::linorder list \<Rightarrow> 'a list" where |
24 |
"insort [] = []" | |
|
25 |
"insort (x#xs) = insort1 x (insort xs)" |
|
66543 | 26 |
|
68078 | 27 |
|
66543 | 28 |
subsubsection "Functional Correctness" |
29 |
||
75501 | 30 |
lemma mset_insort1: "mset (insort1 x xs) = {#x#} + mset xs" |
66543 | 31 |
apply(induction xs) |
32 |
apply auto |
|
33 |
done |
|
34 |
||
75501 | 35 |
lemma mset_insort: "mset (insort xs) = mset xs" |
66543 | 36 |
apply(induction xs) |
37 |
apply simp |
|
75501 | 38 |
apply (simp add: mset_insort1) |
66543 | 39 |
done |
40 |
||
75501 | 41 |
lemma set_insort1: "set (insort1 x xs) = {x} \<union> set xs" |
42 |
by(simp add: mset_insort1 flip: set_mset_mset) |
|
66543 | 43 |
|
75501 | 44 |
lemma sorted_insort1: "sorted (insort1 a xs) = sorted xs" |
66543 | 45 |
apply(induction xs) |
75501 | 46 |
apply(auto simp add: set_insort1) |
66543 | 47 |
done |
48 |
||
75501 | 49 |
lemma sorted_insort: "sorted (insort xs)" |
66543 | 50 |
apply(induction xs) |
75501 | 51 |
apply(auto simp: sorted_insort1) |
66543 | 52 |
done |
53 |
||
68078 | 54 |
|
55 |
subsubsection "Time Complexity" |
|
66543 | 56 |
|
57 |
text \<open>We count the number of function calls.\<close> |
|
58 |
||
59 |
text\<open> |
|
75501 | 60 |
\<open>insort1 x [] = [x]\<close> |
61 |
\<open>insort1 x (y#ys) = |
|
62 |
(if x \<le> y then x#y#ys else y#(insort1 x ys))\<close> |
|
66543 | 63 |
\<close> |
75501 | 64 |
fun T_insort1 :: "'a::linorder \<Rightarrow> 'a list \<Rightarrow> nat" where |
65 |
"T_insort1 x [] = 1" | |
|
66 |
"T_insort1 x (y#ys) = |
|
67 |
(if x \<le> y then 0 else T_insort1 x ys) + 1" |
|
66543 | 68 |
|
69 |
text\<open> |
|
75501 | 70 |
\<open>insort [] = []\<close> |
71 |
\<open>insort (x#xs) = insort1 x (insort xs)\<close> |
|
66543 | 72 |
\<close> |
75501 | 73 |
fun T_insort :: "'a::linorder list \<Rightarrow> nat" where |
74 |
"T_insort [] = 1" | |
|
75 |
"T_insort (x#xs) = T_insort xs + T_insort1 x (insort xs) + 1" |
|
66543 | 76 |
|
77 |
||
75501 | 78 |
lemma T_insort1_length: "T_insort1 x xs \<le> length xs + 1" |
66543 | 79 |
apply(induction xs) |
80 |
apply auto |
|
81 |
done |
|
82 |
||
75501 | 83 |
lemma length_insort1: "length (insort1 x xs) = length xs + 1" |
66543 | 84 |
apply(induction xs) |
85 |
apply auto |
|
86 |
done |
|
87 |
||
75501 | 88 |
lemma length_insort: "length (insort xs) = length xs" |
66543 | 89 |
apply(induction xs) |
75501 | 90 |
apply (auto simp: length_insort1) |
66543 | 91 |
done |
92 |
||
75501 | 93 |
lemma T_insort_length: "T_insort xs \<le> (length xs + 1) ^ 2" |
66543 | 94 |
proof(induction xs) |
95 |
case Nil show ?case by simp |
|
96 |
next |
|
97 |
case (Cons x xs) |
|
75501 | 98 |
have "T_insort (x#xs) = T_insort xs + T_insort1 x (insort xs) + 1" by simp |
99 |
also have "\<dots> \<le> (length xs + 1) ^ 2 + T_insort1 x (insort xs) + 1" |
|
66543 | 100 |
using Cons.IH by simp |
101 |
also have "\<dots> \<le> (length xs + 1) ^ 2 + length xs + 1 + 1" |
|
75501 | 102 |
using T_insort1_length[of x "insort xs"] by (simp add: length_insort) |
66543 | 103 |
also have "\<dots> \<le> (length(x#xs) + 1) ^ 2" |
104 |
by (simp add: power2_eq_square) |
|
105 |
finally show ?case . |
|
106 |
qed |
|
107 |
||
108 |
||
109 |
subsection "Merge Sort" |
|
110 |
||
111 |
fun merge :: "'a::linorder list \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
112 |
"merge [] ys = ys" | |
|
113 |
"merge xs [] = xs" | |
|
114 |
"merge (x#xs) (y#ys) = (if x \<le> y then x # merge xs (y#ys) else y # merge (x#xs) ys)" |
|
115 |
||
116 |
fun msort :: "'a::linorder list \<Rightarrow> 'a list" where |
|
117 |
"msort xs = (let n = length xs in |
|
118 |
if n \<le> 1 then xs |
|
119 |
else merge (msort (take (n div 2) xs)) (msort (drop (n div 2) xs)))" |
|
120 |
||
121 |
declare msort.simps [simp del] |
|
122 |
||
68078 | 123 |
|
67983 | 124 |
subsubsection "Functional Correctness" |
125 |
||
126 |
lemma mset_merge: "mset(merge xs ys) = mset xs + mset ys" |
|
127 |
by(induction xs ys rule: merge.induct) auto |
|
128 |
||
68963 | 129 |
lemma mset_msort: "mset (msort xs) = mset xs" |
67983 | 130 |
proof(induction xs rule: msort.induct) |
131 |
case (1 xs) |
|
132 |
let ?n = "length xs" |
|
68966 | 133 |
let ?ys = "take (?n div 2) xs" |
134 |
let ?zs = "drop (?n div 2) xs" |
|
67983 | 135 |
show ?case |
136 |
proof cases |
|
137 |
assume "?n \<le> 1" |
|
138 |
thus ?thesis by(simp add: msort.simps[of xs]) |
|
139 |
next |
|
140 |
assume "\<not> ?n \<le> 1" |
|
68966 | 141 |
hence "mset (msort xs) = mset (msort ?ys) + mset (msort ?zs)" |
67983 | 142 |
by(simp add: msort.simps[of xs] mset_merge) |
68966 | 143 |
also have "\<dots> = mset ?ys + mset ?zs" |
67983 | 144 |
using \<open>\<not> ?n \<le> 1\<close> by(simp add: "1.IH") |
68966 | 145 |
also have "\<dots> = mset (?ys @ ?zs)" by (simp del: append_take_drop_id) |
67983 | 146 |
also have "\<dots> = mset xs" by simp |
147 |
finally show ?thesis . |
|
148 |
qed |
|
149 |
qed |
|
150 |
||
68966 | 151 |
text \<open>Via the previous lemma or directly:\<close> |
68963 | 152 |
|
67983 | 153 |
lemma set_merge: "set(merge xs ys) = set xs \<union> set ys" |
68963 | 154 |
by (metis mset_merge set_mset_mset set_mset_union) |
155 |
||
156 |
lemma "set(merge xs ys) = set xs \<union> set ys" |
|
67983 | 157 |
by(induction xs ys rule: merge.induct) (auto) |
158 |
||
159 |
lemma sorted_merge: "sorted (merge xs ys) \<longleftrightarrow> (sorted xs \<and> sorted ys)" |
|
160 |
by(induction xs ys rule: merge.induct) (auto simp: set_merge) |
|
161 |
||
68966 | 162 |
lemma sorted_msort: "sorted (msort xs)" |
67983 | 163 |
proof(induction xs rule: msort.induct) |
164 |
case (1 xs) |
|
165 |
let ?n = "length xs" |
|
166 |
show ?case |
|
167 |
proof cases |
|
168 |
assume "?n \<le> 1" |
|
169 |
thus ?thesis by(simp add: msort.simps[of xs] sorted01) |
|
170 |
next |
|
171 |
assume "\<not> ?n \<le> 1" |
|
172 |
thus ?thesis using "1.IH" |
|
68966 | 173 |
by(simp add: sorted_merge msort.simps[of xs]) |
67983 | 174 |
qed |
175 |
qed |
|
176 |
||
68078 | 177 |
|
178 |
subsubsection "Time Complexity" |
|
67983 | 179 |
|
180 |
text \<open>We only count the number of comparisons between list elements.\<close> |
|
66543 | 181 |
|
72501 | 182 |
fun C_merge :: "'a::linorder list \<Rightarrow> 'a list \<Rightarrow> nat" where |
183 |
"C_merge [] ys = 0" | |
|
184 |
"C_merge xs [] = 0" | |
|
185 |
"C_merge (x#xs) (y#ys) = 1 + (if x \<le> y then C_merge xs (y#ys) else C_merge (x#xs) ys)" |
|
66543 | 186 |
|
72501 | 187 |
lemma C_merge_ub: "C_merge xs ys \<le> length xs + length ys" |
188 |
by (induction xs ys rule: C_merge.induct) auto |
|
66543 | 189 |
|
72501 | 190 |
fun C_msort :: "'a::linorder list \<Rightarrow> nat" where |
191 |
"C_msort xs = |
|
66543 | 192 |
(let n = length xs; |
193 |
ys = take (n div 2) xs; |
|
194 |
zs = drop (n div 2) xs |
|
195 |
in if n \<le> 1 then 0 |
|
72501 | 196 |
else C_msort ys + C_msort zs + C_merge (msort ys) (msort zs))" |
66543 | 197 |
|
72501 | 198 |
declare C_msort.simps [simp del] |
66543 | 199 |
|
200 |
lemma length_merge: "length(merge xs ys) = length xs + length ys" |
|
201 |
apply (induction xs ys rule: merge.induct) |
|
202 |
apply auto |
|
203 |
done |
|
204 |
||
205 |
lemma length_msort: "length(msort xs) = length xs" |
|
206 |
proof (induction xs rule: msort.induct) |
|
207 |
case (1 xs) |
|
71918 | 208 |
show ?case |
209 |
by (auto simp: msort.simps [of xs] 1 length_merge) |
|
66543 | 210 |
qed |
211 |
text \<open>Why structured proof? |
|
212 |
To have the name "xs" to specialize msort.simps with xs |
|
213 |
to ensure that msort.simps cannot be used recursively. |
|
214 |
Also works without this precaution, but that is just luck.\<close> |
|
215 |
||
72501 | 216 |
lemma C_msort_le: "length xs = 2^k \<Longrightarrow> C_msort xs \<le> k * 2^k" |
66543 | 217 |
proof(induction k arbitrary: xs) |
72501 | 218 |
case 0 thus ?case by (simp add: C_msort.simps) |
66543 | 219 |
next |
220 |
case (Suc k) |
|
221 |
let ?n = "length xs" |
|
222 |
let ?ys = "take (?n div 2) xs" |
|
223 |
let ?zs = "drop (?n div 2) xs" |
|
224 |
show ?case |
|
225 |
proof (cases "?n \<le> 1") |
|
226 |
case True |
|
72501 | 227 |
thus ?thesis by(simp add: C_msort.simps) |
66543 | 228 |
next |
229 |
case False |
|
72501 | 230 |
have "C_msort(xs) = |
231 |
C_msort ?ys + C_msort ?zs + C_merge (msort ?ys) (msort ?zs)" |
|
232 |
by (simp add: C_msort.simps msort.simps) |
|
233 |
also have "\<dots> \<le> C_msort ?ys + C_msort ?zs + length ?ys + length ?zs" |
|
234 |
using C_merge_ub[of "msort ?ys" "msort ?zs"] length_msort[of ?ys] length_msort[of ?zs] |
|
66543 | 235 |
by arith |
72501 | 236 |
also have "\<dots> \<le> k * 2^k + C_msort ?zs + length ?ys + length ?zs" |
66543 | 237 |
using Suc.IH[of ?ys] Suc.prems by simp |
238 |
also have "\<dots> \<le> k * 2^k + k * 2^k + length ?ys + length ?zs" |
|
239 |
using Suc.IH[of ?zs] Suc.prems by simp |
|
240 |
also have "\<dots> = 2 * k * 2^k + 2 * 2 ^ k" |
|
241 |
using Suc.prems by simp |
|
242 |
finally show ?thesis by simp |
|
243 |
qed |
|
244 |
qed |
|
245 |
||
70295 | 246 |
(* Beware of implicit conversions: *) |
72501 | 247 |
lemma C_msort_log: "length xs = 2^k \<Longrightarrow> C_msort xs \<le> length xs * log 2 (length xs)" |
248 |
using C_msort_le[of xs k] apply (simp add: log_nat_power algebra_simps) |
|
66912
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents:
66543
diff
changeset
|
249 |
by (metis (mono_tags) numeral_power_eq_of_nat_cancel_iff of_nat_le_iff of_nat_mult) |
66543 | 250 |
|
67983 | 251 |
|
252 |
subsection "Bottom-Up Merge Sort" |
|
253 |
||
254 |
fun merge_adj :: "('a::linorder) list list \<Rightarrow> 'a list list" where |
|
255 |
"merge_adj [] = []" | |
|
256 |
"merge_adj [xs] = [xs]" | |
|
257 |
"merge_adj (xs # ys # zss) = merge xs ys # merge_adj zss" |
|
258 |
||
259 |
text \<open>For the termination proof of \<open>merge_all\<close> below.\<close> |
|
260 |
lemma length_merge_adjacent[simp]: "length (merge_adj xs) = (length xs + 1) div 2" |
|
261 |
by (induction xs rule: merge_adj.induct) auto |
|
262 |
||
263 |
fun merge_all :: "('a::linorder) list list \<Rightarrow> 'a list" where |
|
68970 | 264 |
"merge_all [] = []" | |
67983 | 265 |
"merge_all [xs] = xs" | |
266 |
"merge_all xss = merge_all (merge_adj xss)" |
|
267 |
||
268 |
definition msort_bu :: "('a::linorder) list \<Rightarrow> 'a list" where |
|
68971 | 269 |
"msort_bu xs = merge_all (map (\<lambda>x. [x]) xs)" |
67983 | 270 |
|
68078 | 271 |
|
67983 | 272 |
subsubsection "Functional Correctness" |
273 |
||
72802 | 274 |
abbreviation mset_mset :: "'a list list \<Rightarrow> 'a multiset" where |
73047
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72802
diff
changeset
|
275 |
"mset_mset xss \<equiv> \<Sum>\<^sub># (image_mset mset (mset xss))" |
72802 | 276 |
|
67983 | 277 |
lemma mset_merge_adj: |
72802 | 278 |
"mset_mset (merge_adj xss) = mset_mset xss" |
67983 | 279 |
by(induction xss rule: merge_adj.induct) (auto simp: mset_merge) |
280 |
||
68967 | 281 |
lemma mset_merge_all: |
72802 | 282 |
"mset (merge_all xss) = mset_mset xss" |
67983 | 283 |
by(induction xss rule: merge_all.induct) (auto simp: mset_merge mset_merge_adj) |
284 |
||
68968 | 285 |
lemma mset_msort_bu: "mset (msort_bu xs) = mset xs" |
72802 | 286 |
by(simp add: msort_bu_def mset_merge_all multiset.map_comp comp_def) |
68968 | 287 |
|
67983 | 288 |
lemma sorted_merge_adj: |
289 |
"\<forall>xs \<in> set xss. sorted xs \<Longrightarrow> \<forall>xs \<in> set (merge_adj xss). sorted xs" |
|
290 |
by(induction xss rule: merge_adj.induct) (auto simp: sorted_merge) |
|
291 |
||
292 |
lemma sorted_merge_all: |
|
68971 | 293 |
"\<forall>xs \<in> set xss. sorted xs \<Longrightarrow> sorted (merge_all xss)" |
67983 | 294 |
apply(induction xss rule: merge_all.induct) |
295 |
using [[simp_depth_limit=3]] by (auto simp add: sorted_merge_adj) |
|
296 |
||
297 |
lemma sorted_msort_bu: "sorted (msort_bu xs)" |
|
298 |
by(simp add: msort_bu_def sorted_merge_all) |
|
299 |
||
68078 | 300 |
|
301 |
subsubsection "Time Complexity" |
|
67983 | 302 |
|
72501 | 303 |
fun C_merge_adj :: "('a::linorder) list list \<Rightarrow> nat" where |
304 |
"C_merge_adj [] = 0" | |
|
305 |
"C_merge_adj [xs] = 0" | |
|
306 |
"C_merge_adj (xs # ys # zss) = C_merge xs ys + C_merge_adj zss" |
|
67983 | 307 |
|
72501 | 308 |
fun C_merge_all :: "('a::linorder) list list \<Rightarrow> nat" where |
309 |
"C_merge_all [] = 0" | |
|
310 |
"C_merge_all [xs] = 0" | |
|
311 |
"C_merge_all xss = C_merge_adj xss + C_merge_all (merge_adj xss)" |
|
67983 | 312 |
|
72501 | 313 |
definition C_msort_bu :: "('a::linorder) list \<Rightarrow> nat" where |
314 |
"C_msort_bu xs = C_merge_all (map (\<lambda>x. [x]) xs)" |
|
67983 | 315 |
|
316 |
lemma length_merge_adj: |
|
68974 | 317 |
"\<lbrakk> even(length xss); \<forall>xs \<in> set xss. length xs = m \<rbrakk> |
318 |
\<Longrightarrow> \<forall>xs \<in> set (merge_adj xss). length xs = 2*m" |
|
68161 | 319 |
by(induction xss rule: merge_adj.induct) (auto simp: length_merge) |
67983 | 320 |
|
72501 | 321 |
lemma C_merge_adj: "\<forall>xs \<in> set xss. length xs = m \<Longrightarrow> C_merge_adj xss \<le> m * length xss" |
322 |
proof(induction xss rule: C_merge_adj.induct) |
|
67983 | 323 |
case 1 thus ?case by simp |
324 |
next |
|
325 |
case 2 thus ?case by simp |
|
326 |
next |
|
72501 | 327 |
case (3 x y) thus ?case using C_merge_ub[of x y] by (simp add: algebra_simps) |
67983 | 328 |
qed |
329 |
||
72501 | 330 |
lemma C_merge_all: "\<lbrakk> \<forall>xs \<in> set xss. length xs = m; length xss = 2^k \<rbrakk> |
331 |
\<Longrightarrow> C_merge_all xss \<le> m * k * 2^k" |
|
332 |
proof (induction xss arbitrary: k m rule: C_merge_all.induct) |
|
67983 | 333 |
case 1 thus ?case by simp |
334 |
next |
|
68158 | 335 |
case 2 thus ?case by simp |
67983 | 336 |
next |
68162 | 337 |
case (3 xs ys xss) |
338 |
let ?xss = "xs # ys # xss" |
|
339 |
let ?xss2 = "merge_adj ?xss" |
|
67983 | 340 |
obtain k' where k': "k = Suc k'" using "3.prems"(2) |
341 |
by (metis length_Cons nat.inject nat_power_eq_Suc_0_iff nat.exhaust) |
|
68972 | 342 |
have "even (length ?xss)" using "3.prems"(2) k' by auto |
343 |
from length_merge_adj[OF this "3.prems"(1)] |
|
344 |
have *: "\<forall>x \<in> set(merge_adj ?xss). length x = 2*m" . |
|
68162 | 345 |
have **: "length ?xss2 = 2 ^ k'" using "3.prems"(2) k' by auto |
72501 | 346 |
have "C_merge_all ?xss = C_merge_adj ?xss + C_merge_all ?xss2" by simp |
347 |
also have "\<dots> \<le> m * 2^k + C_merge_all ?xss2" |
|
348 |
using "3.prems"(2) C_merge_adj[OF "3.prems"(1)] by (auto simp: algebra_simps) |
|
67983 | 349 |
also have "\<dots> \<le> m * 2^k + (2*m) * k' * 2^k'" |
68079 | 350 |
using "3.IH"[OF * **] by simp |
67983 | 351 |
also have "\<dots> = m * k * 2^k" |
352 |
using k' by (simp add: algebra_simps) |
|
353 |
finally show ?case . |
|
354 |
qed |
|
355 |
||
72501 | 356 |
corollary C_msort_bu: "length xs = 2 ^ k \<Longrightarrow> C_msort_bu xs \<le> k * 2 ^ k" |
357 |
using C_merge_all[of "map (\<lambda>x. [x]) xs" 1] by (simp add: C_msort_bu_def) |
|
67983 | 358 |
|
68993 | 359 |
|
360 |
subsection "Quicksort" |
|
361 |
||
362 |
fun quicksort :: "('a::linorder) list \<Rightarrow> 'a list" where |
|
363 |
"quicksort [] = []" | |
|
364 |
"quicksort (x#xs) = quicksort (filter (\<lambda>y. y < x) xs) @ [x] @ quicksort (filter (\<lambda>y. x \<le> y) xs)" |
|
365 |
||
366 |
lemma mset_quicksort: "mset (quicksort xs) = mset xs" |
|
367 |
apply (induction xs rule: quicksort.induct) |
|
368 |
apply (auto simp: not_le) |
|
369 |
done |
|
370 |
||
371 |
lemma set_quicksort: "set (quicksort xs) = set xs" |
|
372 |
by(rule mset_eq_setD[OF mset_quicksort]) |
|
373 |
||
374 |
lemma sorted_quicksort: "sorted (quicksort xs)" |
|
375 |
apply (induction xs rule: quicksort.induct) |
|
376 |
apply (auto simp add: sorted_append set_quicksort) |
|
377 |
done |
|
378 |
||
69005 | 379 |
|
380 |
subsection "Insertion Sort w.r.t. Keys and Stability" |
|
381 |
||
75501 | 382 |
hide_const List.insort_key |
69005 | 383 |
|
75501 | 384 |
fun insort1_key :: "('a \<Rightarrow> 'k::linorder) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
385 |
"insort1_key f x [] = [x]" | |
|
386 |
"insort1_key f x (y # ys) = (if f x \<le> f y then x # y # ys else y # insort1_key f x ys)" |
|
387 |
||
388 |
fun insort_key :: "('a \<Rightarrow> 'k::linorder) \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
389 |
"insort_key f [] = []" | |
|
390 |
"insort_key f (x # xs) = insort1_key f x (insort_key f xs)" |
|
69005 | 391 |
|
392 |
||
393 |
subsubsection "Standard functional correctness" |
|
394 |
||
75501 | 395 |
lemma mset_insort1_key: "mset (insort1_key f x xs) = {#x#} + mset xs" |
69005 | 396 |
by(induction xs) simp_all |
397 |
||
75501 | 398 |
lemma mset_insort_key: "mset (insort_key f xs) = mset xs" |
399 |
by(induction xs) (simp_all add: mset_insort1_key) |
|
69005 | 400 |
|
75501 | 401 |
(* Inductive proof simpler than derivation from mset lemma: *) |
402 |
lemma set_insort1_key: "set (insort1_key f x xs) = {x} \<union> set xs" |
|
403 |
by (induction xs) auto |
|
69005 | 404 |
|
75501 | 405 |
lemma sorted_insort1_key: "sorted (map f (insort1_key f a xs)) = sorted (map f xs)" |
406 |
by(induction xs)(auto simp: set_insort1_key) |
|
407 |
||
408 |
lemma sorted_insort_key: "sorted (map f (insort_key f xs))" |
|
409 |
by(induction xs)(simp_all add: sorted_insort1_key) |
|
69005 | 410 |
|
411 |
||
412 |
subsubsection "Stability" |
|
413 |
||
75501 | 414 |
lemma insort1_is_Cons: "\<forall>x\<in>set xs. f a \<le> f x \<Longrightarrow> insort1_key f a xs = a # xs" |
69005 | 415 |
by (cases xs) auto |
416 |
||
75501 | 417 |
lemma filter_insort1_key_neg: |
418 |
"\<not> P x \<Longrightarrow> filter P (insort1_key f x xs) = filter P xs" |
|
69005 | 419 |
by (induction xs) simp_all |
420 |
||
75501 | 421 |
lemma filter_insort1_key_pos: |
422 |
"sorted (map f xs) \<Longrightarrow> P x \<Longrightarrow> filter P (insort1_key f x xs) = insort1_key f x (filter P xs)" |
|
423 |
by (induction xs) (auto, subst insort1_is_Cons, auto) |
|
69005 | 424 |
|
75501 | 425 |
lemma sort_key_stable: "filter (\<lambda>y. f y = k) (insort_key f xs) = filter (\<lambda>y. f y = k) xs" |
69005 | 426 |
proof (induction xs) |
427 |
case Nil thus ?case by simp |
|
428 |
next |
|
429 |
case (Cons a xs) |
|
430 |
thus ?case |
|
431 |
proof (cases "f a = k") |
|
75501 | 432 |
case False thus ?thesis by (simp add: Cons.IH filter_insort1_key_neg) |
69005 | 433 |
next |
434 |
case True |
|
75501 | 435 |
have "filter (\<lambda>y. f y = k) (insort_key f (a # xs)) |
436 |
= filter (\<lambda>y. f y = k) (insort1_key f a (insort_key f xs))" by simp |
|
437 |
also have "\<dots> = insort1_key f a (filter (\<lambda>y. f y = k) (insort_key f xs))" |
|
438 |
by (simp add: True filter_insort1_key_pos sorted_insort_key) |
|
439 |
also have "\<dots> = insort1_key f a (filter (\<lambda>y. f y = k) xs)" by (simp add: Cons.IH) |
|
440 |
also have "\<dots> = a # (filter (\<lambda>y. f y = k) xs)" by(simp add: True insort1_is_Cons) |
|
69005 | 441 |
also have "\<dots> = filter (\<lambda>y. f y = k) (a # xs)" by (simp add: True) |
442 |
finally show ?thesis . |
|
443 |
qed |
|
444 |
qed |
|
445 |
||
66543 | 446 |
end |