# HG changeset patch # User wenzelm # Date 1536917661 -7200 # Node ID 10da16970d82ed4f2b36df9ae811dba154ded7db # Parent e6678381151899a8e67b6ec6cded77a2176aca4e# Parent d961e11e0e87f903da5f00c071fad579ff2fe83c merged diff -r d961e11e0e87 -r 10da16970d82 src/HOL/Data_Structures/Sorting.thy --- a/src/HOL/Data_Structures/Sorting.thy Thu Sep 13 21:18:43 2018 +0200 +++ b/src/HOL/Data_Structures/Sorting.thy Fri Sep 14 11:34:21 2018 +0200 @@ -351,4 +351,24 @@ corollary c_msort_bu: "length xs = 2 ^ k \ c_msort_bu xs \ k * 2 ^ k" using c_merge_all[of "map (\x. [x]) xs" 1] by (simp add: c_msort_bu_def) + +subsection "Quicksort" + +fun quicksort :: "('a::linorder) list \ 'a list" where +"quicksort [] = []" | +"quicksort (x#xs) = quicksort (filter (\y. y < x) xs) @ [x] @ quicksort (filter (\y. x \ y) xs)" + +lemma mset_quicksort: "mset (quicksort xs) = mset xs" +apply (induction xs rule: quicksort.induct) +apply (auto simp: not_le) +done + +lemma set_quicksort: "set (quicksort xs) = set xs" +by(rule mset_eq_setD[OF mset_quicksort]) + +lemma sorted_quicksort: "sorted (quicksort xs)" +apply (induction xs rule: quicksort.induct) +apply (auto simp add: sorted_append set_quicksort) +done + end diff -r d961e11e0e87 -r 10da16970d82 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Thu Sep 13 21:18:43 2018 +0200 +++ b/src/HOL/Library/Multiset.thy Fri Sep 14 11:34:21 2018 +0200 @@ -1520,8 +1520,12 @@ lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" by (cases "B = {#}") (auto dest: multi_member_split) +lemma union_filter_mset_complement[simp]: + "\x. P x = (\ Q x) \ filter_mset P M + filter_mset Q M = M" +by (subst multiset_eq_iff) auto + lemma multiset_partition: "M = {#x \# M. P x#} + {#x \# M. \ P x#}" - by (subst multiset_eq_iff) auto +by simp lemma mset_subset_size: "A \# B \ size A < size B" proof (induct A arbitrary: B) @@ -1887,9 +1891,6 @@ apply simp done -lemma union_mset_compl[simp]: "\x. P x = (\ Q x) \ filter_mset P M + filter_mset Q M = M" -by (induction M) simp_all - lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" proof (induct ls arbitrary: i) case Nil @@ -2924,7 +2925,7 @@ next case [simp]: False have K: "K = {#x \# K. (x, a) \ r#} + {#x \# K. (x, a) \ r#}" - by (rule multiset_partition) + by simp have "(I + K, (I + {# x \# K. (x, a) \ r #}) + J') \ mult r" using IH[of J' "{# x \# K. (x, a) \ r#}" "I + {# x \# K. (x, a) \ r#}"] J Suc.prems K size_J by (auto simp: ac_simps)