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