diff -r 1b73c3e17d9f -r c484e9da2546 src/HOL/List.thy --- a/src/HOL/List.thy Tue Mar 04 19:34:12 2025 +0100 +++ b/src/HOL/List.thy Wed Mar 05 15:34:01 2025 +0100 @@ -2488,6 +2488,11 @@ "(\x. x \ set xs \ P(x)) \ dropWhile P (xs @ ys) = dropWhile P ys" by (induct xs) auto +lemma dropWhile_id[simp]: + "(\x. x \ set xs \ \ P x) \ dropWhile P xs = xs" +using takeWhile_dropWhile_id[of P xs] takeWhile_eq_Nil_iff[of P xs] +by fastforce + lemma dropWhile_append3: "\ P y \dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys" by (induct xs) auto @@ -5644,14 +5649,20 @@ lemma assumes "sorted_wrt f xs" - shows sorted_wrt_take: "sorted_wrt f (take n xs)" - and sorted_wrt_drop: "sorted_wrt f (drop n xs)" + shows sorted_wrt_take[simp]: "sorted_wrt f (take n xs)" + and sorted_wrt_drop[simp]: "sorted_wrt f (drop n xs)" proof - from assms have "sorted_wrt f (take n xs @ drop n xs)" by simp thus "sorted_wrt f (take n xs)" and "sorted_wrt f (drop n xs)" unfolding sorted_wrt_append by simp_all qed +lemma sorted_wrt_dropWhile[simp]: "sorted_wrt R xs \ sorted_wrt R (dropWhile P xs)" +by (auto dest: sorted_wrt_drop simp: dropWhile_eq_drop) + +lemma sorted_wrt_takeWhile[simp]: "sorted_wrt R xs \ sorted_wrt R (takeWhile P xs)" +by (subst takeWhile_eq_take) (auto dest: sorted_wrt_take) + lemma sorted_wrt_filter: "sorted_wrt f xs \ sorted_wrt f (filter P xs)" by (induction xs) auto @@ -6500,6 +6511,10 @@ by standard simp qed (simp_all add: sorted_list_of_set_def) +lemma ex1_sorted_list_for_set_if_finite: + "finite X \ \!xs. sorted_wrt (<) xs \ set xs = X" + by (metis sorted_list_of_set.finite_set_strict_sorted strict_sorted_equal) + text \Alias theorems for backwards compatibility and ease of use.\ lemmas sorted_list_of_set = sorted_list_of_set.sorted_key_list_of_set and sorted_list_of_set_empty = sorted_list_of_set.sorted_key_list_of_set_empty and