# HG changeset patch # User haftmann # Date 1273655922 -7200 # Node ID ef6ba914e2091cbc4af90ba1a07b00e111ab46dd # Parent 33e5b40ec4bb9b26a9dc9eb51ec36652a7ddcb8a# Parent 5135adb331577626af5be1fc34e73038cf8e333a merged diff -r 33e5b40ec4bb -r ef6ba914e209 src/HOL/List.thy --- a/src/HOL/List.thy Wed May 12 11:13:33 2010 +0200 +++ b/src/HOL/List.thy Wed May 12 11:18:42 2010 +0200 @@ -1529,6 +1529,14 @@ apply(simp add:list_update_append split:nat.splits) done +lemma last_map: + "xs \ [] \ last (map f xs) = f (last xs)" + by (cases xs rule: rev_cases) simp_all + +lemma map_butlast: + "map f (butlast xs) = butlast (map f xs)" + by (induct xs) simp_all + subsubsection {* @{text take} and @{text drop} *} @@ -2910,6 +2918,14 @@ "remdups (remdups xs) = remdups xs" by (induct xs) simp_all +lemma distinct_butlast: + assumes "xs \ []" and "distinct xs" + shows "distinct (butlast xs)" +proof - + from `xs \ []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto + with `distinct xs` show ?thesis by simp +qed + subsubsection {* @{const insert} *} @@ -3605,6 +3621,18 @@ theorem sorted_sort[simp]: "sorted (sort xs)" by(induct xs)(auto simp:sorted_insort) +lemma sorted_butlast: + assumes "xs \ []" and "sorted xs" + shows "sorted (butlast xs)" +proof - + from `xs \ []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto + with `sorted xs` show ?thesis by (simp add: sorted_append) +qed + +lemma insort_not_Nil [simp]: + "insort_key f a xs \ []" + by (induct xs) simp_all + lemma insort_is_Cons: "\x\set xs. f a \ f x \ insort_key f a xs = a # xs" by (cases xs) auto