# HG changeset patch # User haftmann # Date 1288711916 -3600 # Node ID 2d507370e8797aab0883fb5657f2fab57a275774 # Parent ac4d75f86d97d6aea40878d47fe194d4245cf01b lemmas multiset_of_filter, sort_key_by_quicksort diff -r ac4d75f86d97 -r 2d507370e879 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Sun Oct 31 13:26:37 2010 +0100 +++ b/src/HOL/Library/Multiset.thy Tue Nov 02 16:31:56 2010 +0100 @@ -735,6 +735,10 @@ "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" by (induct xs arbitrary: ys) (auto simp: add_ac) +lemma multiset_of_filter: + "multiset_of (filter P xs) = {#x :# multiset_of xs. P x #}" + by (induct xs) simp_all + lemma surj_multiset_of: "surj multiset_of" apply (unfold surj_def) apply (rule allI) @@ -884,6 +888,59 @@ by (simp add: replicate_length_filter) qed +lemma sort_key_by_quicksort: + "sort_key f xs = sort_key f [x\xs. f x < f (xs ! (length xs div 2))] + @ [x\xs. f x = f (xs ! (length xs div 2))] + @ sort_key f [x\xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") +proof (rule properties_for_sort_key) + show "multiset_of ?rhs = multiset_of ?lhs" + by (rule multiset_eqI) (auto simp add: multiset_of_filter) +next + show "sorted (map f ?rhs)" + by (auto simp add: sorted_append intro: sorted_map_same) +next + show "\k\ f ` set ?rhs. [x \ ?rhs. k = f x] = [x \ ?lhs. k = f x]" + proof + let ?pivot = "f (xs ! (length xs div 2))" + fix k + assume k: "k \ f ` set ?rhs" + then obtain l where f_l: "k = f l" by blast + have *: "\x P. P (f x) \ f x = f l \ P (f l) \ f x = f l" by auto + have **: "\x. k = f x \ f x = k" by auto + have [simp]: "[x \ sort_key f xs . f x = f l] = [x \ xs. f x = f l]" + unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) + have "[x \ ?rhs. f x = k] = [x \ ?lhs. f x = k]" + proof (cases "f l" ?pivot rule: linorder_cases) + case less then show ?thesis + apply (auto simp add: filter_sort [symmetric] f_l) + apply (subst *) apply simp + apply (frule less_imp_neq) apply simp + apply (subst *) apply (frule less_not_sym) apply simp + done + next + case equal from this [symmetric] show ?thesis + apply (auto simp add: filter_sort [symmetric] f_l) + apply (subst *) apply simp + apply (subst *) apply simp + done + next + case greater then show ?thesis + apply (auto simp add: filter_sort [symmetric] f_l) + apply (subst *) apply (frule less_not_sym) apply simp + apply (frule less_imp_neq) apply simp + apply (subst *) apply simp + done + qed + then show "[x \ ?rhs. k = f x] = [x \ ?lhs. k = f x]" by (simp add: **) + qed +qed + +lemma sort_by_quicksort: + "sort xs = sort [x\xs. x < xs ! (length xs div 2)] + @ [x\xs. x = xs ! (length xs div 2)] + @ sort [x\xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") + using sort_key_by_quicksort [of "\x. x", symmetric] by simp + end lemma multiset_of_remdups_le: "multiset_of (remdups xs) \ multiset_of xs"