author | haftmann |
Mon, 27 Sep 2010 14:13:22 +0200 | |
changeset 39727 | 5dab9549c80d |
parent 39726 | ba01ecc2252a |
child 39728 | 832c42be723e |
--- a/src/HOL/Library/Dlist.thy Mon Sep 27 13:28:54 2010 +0200 +++ b/src/HOL/Library/Dlist.thy Mon Sep 27 14:13:22 2010 +0200 @@ -35,6 +35,10 @@ "list_of_dlist (Dlist xs) = remdups xs" by (simp add: Dlist_def Abs_dlist_inverse) +lemma remdups_list_of_dlist [simp]: + "remdups (list_of_dlist dxs) = list_of_dlist dxs" + by simp + lemma Dlist_list_of_dlist [simp, code abstype]: "Dlist (list_of_dlist dxs) = dxs" by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)