author | huffman |
Tue, 29 May 2012 17:53:43 +0200 | |
changeset 48040 | 4caf6cd063be |
parent 48039 | daab96c3e2a9 |
child 48042 | 918a92d4079f |
--- a/src/HOL/Library/Multiset.thy Wed May 30 10:04:05 2012 +0200 +++ b/src/HOL/Library/Multiset.thy Tue May 29 17:53:43 2012 +0200 @@ -827,6 +827,9 @@ apply (simp add: add_assoc [symmetric] image_mset_insert) done +lemma set_of_image_mset [simp]: "set_of (image_mset f M) = image f (set_of M)" +by (induct M) simp_all + lemma size_image_mset [simp]: "size (image_mset f M) = size M" by (induct M) simp_all