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