# HG changeset patch # User haftmann # Date 1291412060 -3600 # Node ID 1d46d893d682ed499a7bbbcd21f63540f0dad76c # Parent ad8535384c340c6204dc9dd4644a0eb53ca38ddd lemmas fold_remove1_split and fold_multiset_equiv diff -r ad8535384c34 -r 1d46d893d682 src/HOL/Library/More_List.thy --- a/src/HOL/Library/More_List.thy Fri Dec 03 22:08:14 2010 +0100 +++ b/src/HOL/Library/More_List.thy Fri Dec 03 22:34:20 2010 +0100 @@ -3,7 +3,7 @@ header {* Operations on lists beyond the standard List theory *} theory More_List -imports Main +imports Main Multiset begin hide_const (open) Finite_Set.fold @@ -78,10 +78,32 @@ "fold g (map f xs) = fold (g o f) xs" by (induct xs) simp_all +lemma fold_remove1_split: + assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" + and x: "x \ set xs" + shows "fold f xs = fold f (remove1 x xs) \ f x" + using assms by (induct xs) (auto simp add: o_assoc [symmetric]) + +lemma fold_multiset_equiv: + assumes f: "\x y. x \# multiset_of xs \ y \# multiset_of xs \ f x \ f y = f y \ f x" + and equiv: "multiset_of xs = multiset_of ys" + shows "fold f xs = fold f ys" +using f equiv [symmetric] proof (induct xs arbitrary: ys) + case Nil then show ?case by simp +next + case (Cons x xs) + have "\x y. x \ set ys \ y \ set ys \ f x \ f y = f y \ f x" + by (rule Cons.prems(1)) (simp_all add: mem_set_multiset_eq Cons.prems(2)) + moreover from Cons.prems have "x \ set ys" by (simp add: mem_set_multiset_eq) + ultimately have "fold f ys = fold f (remove1 x ys) \ f x" by (fact fold_remove1_split) + moreover from Cons.prems have "fold f xs = fold f (remove1 x ys)" by (auto intro: Cons.hyps) + ultimately show ?case by simp +qed + lemma fold_rev: assumes "\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y" shows "fold f (rev xs) = fold f xs" - using assms by (induct xs) (simp_all del: o_apply add: fold_commute) + by (rule fold_multiset_equiv, rule assms) (simp_all add: in_multiset_in_set) lemma foldr_fold: assumes "\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y"