# HG changeset patch # User eberlm # Date 1474299442 -7200 # Node ID 0a5184877cb7068f9a6a4c22ef3c76f9b81eb7b0 # Parent 40d1c5e7f407a63d2666d1967b13bfc0e4c89ac1 Additions to permutations (contributed by Lukas Bulwahn) diff -r 40d1c5e7f407 -r 0a5184877cb7 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Mon Sep 19 12:53:30 2016 +0200 +++ b/src/HOL/Library/Multiset.thy Mon Sep 19 17:37:22 2016 +0200 @@ -1989,6 +1989,21 @@ finally show ?case by simp qed simp_all +(* Contributed by Lukas Bulwahn *) +lemma image_mset_mset_set: + assumes "inj_on f A" + shows "image_mset f (mset_set A) = mset_set (f ` A)" +proof cases + assume "finite A" + from this \inj_on f A\ show ?thesis + by (induct A) auto +next + assume "infinite A" + from this \inj_on f A\ have "infinite (f ` A)" + using finite_imageD by blast + from \infinite A\ \infinite (f ` A)\ show ?thesis by simp +qed + subsection \More properties of the replicate and repeat operations\ diff -r 40d1c5e7f407 -r 0a5184877cb7 src/HOL/Library/Permutations.thy --- a/src/HOL/Library/Permutations.thy Mon Sep 19 12:53:30 2016 +0200 +++ b/src/HOL/Library/Permutations.thy Mon Sep 19 17:37:22 2016 +0200 @@ -105,6 +105,48 @@ lemma permutes_superset: "p permutes S \ (\x \ S - T. p x = x) \ p permutes T" by (simp add: Ball_def permutes_def) metis +(* Next three lemmas contributed by Lukas Bulwahn *) +lemma permutes_bij_inv_into: + fixes A :: "'a set" and B :: "'b set" + assumes "p permutes A" + assumes "bij_betw f A B" + shows "(\x. if x \ B then f (p (inv_into A f x)) else x) permutes B" +proof (rule bij_imp_permutes) + have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A" + using assms by (auto simp add: permutes_imp_bij bij_betw_inv_into) + from this have "bij_betw (f o p o inv_into A f) B B" by (simp add: bij_betw_trans) + from this show "bij_betw (\x. if x \ B then f (p (inv_into A f x)) else x) B B" + by (subst bij_betw_cong[where g="f o p o inv_into A f"]) auto +next + fix x + assume "x \ B" + from this show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto +qed + +lemma permutes_image_mset: + assumes "p permutes A" + shows "image_mset p (mset_set A) = mset_set A" +using assms by (metis image_mset_mset_set bij_betw_imp_inj_on permutes_imp_bij permutes_image) + +lemma permutes_implies_image_mset_eq: + assumes "p permutes A" "\x. x \ A \ f x = f' (p x)" + shows "image_mset f' (mset_set A) = image_mset f (mset_set A)" +proof - + have "f x = f' (p x)" if x: "x \# mset_set A" for x + using assms(2)[of x] x by (cases "finite A") auto + from this have "image_mset f (mset_set A) = image_mset (f' o p) (mset_set A)" + using assms by (auto intro!: image_mset_cong) + also have "\ = image_mset f' (image_mset p (mset_set A))" + by (simp add: image_mset.compositionality) + also have "\ = image_mset f' (mset_set A)" + proof - + from assms have "image_mset p (mset_set A) = mset_set A" + using permutes_image_mset by blast + from this show ?thesis by simp + qed + finally show ?thesis .. +qed + subsection \Group properties\ @@ -893,7 +935,7 @@ finally show "count (mset (permute_list f xs)) y = count (mset xs) y" by simp qed -lemma set_permute_list [simp]: +lemma set_permute_list [simp]: assumes "f permutes {..More lemmas about permutations\ text \ - If two lists correspond to the same multiset, there exists a permutation - on the list indices that maps one to the other. + The following few lemmas were contributed by Lukas Bulwahn. \ -lemma mset_eq_permutation: - assumes mset_eq: "mset (xs::'a list) = mset ys" - defines [simp]: "n \ length xs" - obtains f where "f permutes {.. A. f a = b}" + using assms +proof (induct A) + case empty + show ?case by simp +next + case (insert x F) + show ?case + proof cases + assume "f x = b" + from this have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \ F. f a = f x})" + using insert.hyps by auto + also have "\ = card (insert x {a \ F. f a = f x})" + using insert.hyps(1,2) by simp + also have "card (insert x {a \ F. f a = f x}) = card {a \ insert x F. f a = b}" + using \f x = b\ by (auto intro: arg_cong[where f="card"]) + finally show ?thesis using insert by auto + next + assume A: "f x \ b" + hence "{a \ F. f a = b} = {a \ insert x F. f a = b}" by auto + with insert A show ?thesis by simp + qed +qed + +(* Prove image_mset_eq_implies_permutes *) +lemma image_mset_eq_implies_permutes: + fixes f :: "'a \ 'b" + assumes "finite A" + assumes mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)" + obtains p where "p permutes A" and "\x\A. f x = f' (p x)" proof - - from mset_eq have [simp]: "length xs = length ys" - by (rule mset_eq_length) - define indices_of :: "'a \ 'a list \ nat set" - where "indices_of x xs = {i. i < length xs \ x = xs ! i}" for x xs - have indices_of_subset: "indices_of x xs \ {..x\set xs. \f. bij_betw f (indices_of x xs) (indices_of x ys)" + from \finite A\ have [simp]: "finite {a \ A. f a = (b::'b)}" for f b by auto + have "f ` A = f' ` A" + proof - + have "f ` A = f ` (set_mset (mset_set A))" using \finite A\ by simp + also have "\ = f' ` (set_mset (mset_set A))" + by (metis mset_eq multiset.set_map) + also have "\ = f' ` A" using \finite A\ by simp + finally show ?thesis . + qed + have "\b\(f ` A). \p. bij_betw p {a \ A. f a = b} {a \ A. f' a = b}" proof - fix x - from mset_eq have "count (mset xs) x = count (mset ys) x" by simp - hence "card (indices_of x xs) = card (indices_of x ys)" - by (simp add: count_mset length_filter_conv_card indices_of_def) - thus "\f. bij_betw f (indices_of x xs) (indices_of x ys)" + fix b + from mset_eq have + "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b" by simp + from this have "card {a \ A. f a = b} = card {a \ A. f' a = b}" + using \finite A\ + by (simp add: count_image_mset_eq_card_vimage) + from this show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}" by (intro finite_same_card_bij) simp_all qed - hence "\f. \x\set xs. bij_betw (f x) (indices_of x xs) (indices_of x ys)" + hence "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" by (rule bchoice) - then guess f .. note f = this - define g where "g i = (if i < n then f (xs ! i) i else i)" for i - - have bij_f: "bij_betw (\i. f (xs ! i) i) (indices_of x xs) (indices_of x ys)" - if x: "x \ set xs" for x - proof (subst bij_betw_cong) - from f x show "bij_betw (f x) (indices_of x xs) (indices_of x ys)" by blast - fix i assume "i \ indices_of x xs" - thus "f (xs ! i) i = f x i" by (simp add: indices_of_def) + then guess p .. note p = this + define p' where "p' = (\a. if a \ A then p (f a) a else a)" + have "p' permutes A" + proof (rule bij_imp_permutes) + have "disjoint_family_on (\i. {a \ A. f' a = i}) (f ` A)" + unfolding disjoint_family_on_def by auto + moreover have "bij_betw (\a. p (f a) a) {a \ A. f a = b} {a \ A. f' a = b}" if b: "b \ f ` A" for b + using p b by (subst bij_betw_cong[where g="p b"]) auto + ultimately have "bij_betw (\a. p (f a) a) (\b\f ` A. {a \ A. f a = b}) (\b\f ` A. {a \ A. f' a = b})" + by (rule bij_betw_UNION_disjoint) + moreover have "(\b\f ` A. {a \ A. f a = b}) = A" by auto + moreover have "(\b\f ` A. {a \ A. f' a = b}) = A" using \f ` A = f' ` A\ by auto + ultimately show "bij_betw p' A A" + unfolding p'_def by (subst bij_betw_cong[where g="(\a. p (f a) a)"]) auto + next + fix x + assume "x \ A" + from this show "p' x = x" + unfolding p'_def by simp qed + moreover from p have "\x\A. f x = f' (p' x)" + unfolding p'_def using bij_betwE by fastforce + ultimately show ?thesis by (rule that) +qed - hence "bij_betw (\i. f (xs ! i) i) (\x\set xs. indices_of x xs) (\x\set xs. indices_of x ys)" - by (intro bij_betw_UNION_disjoint) (auto simp add: disjoint_family_on_def indices_of_def) - also have "(\x\set xs. indices_of x xs) = {..x\set ys. indices_of x ys) = {..i. f (xs ! i) i) {.. bij_betw g {.. indices_of (xs ! i) xs" by (simp add: indices_of_def) - with bij_f[of "xs ! i"] i have "f (xs ! i) i \ indices_of (xs ! i) ys" - by (auto simp: bij_betw_def) - hence "ys ! f (xs ! i) i = xs ! i" by (simp add: indices_of_def) - finally show "xs ! i = permute_list g ys ! i" .. - qed simp_all - - ultimately show ?thesis by (rule that) +(* and derive the existing property: *) +lemma mset_eq_permutation: + assumes mset_eq: "mset (xs::'a list) = mset ys" + obtains p where "p permutes {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i\{..