diff --git a/src/HOL/Combinatorics/Permutations.thy b/src/HOL/Combinatorics/Permutations.thy --- a/src/HOL/Combinatorics/Permutations.thy +++ b/src/HOL/Combinatorics/Permutations.thy @@ -1,1670 +1,1779 @@ (* Author: Amine Chaieb, University of Cambridge *) section \Permutations, both general and specifically on finite sets.\ theory Permutations imports "HOL-Library.Multiset" "HOL-Library.Disjoint_Sets" begin subsection \Auxiliary\ abbreviation (input) fixpoints :: \('a \ 'a) \ 'a set\ where \fixpoints f \ {x. f x = x}\ lemma inj_on_fixpoints: \inj_on f (fixpoints f)\ by (rule inj_onI) simp lemma bij_betw_fixpoints: \bij_betw f (fixpoints f) (fixpoints f)\ using inj_on_fixpoints by (auto simp add: bij_betw_def) subsection \Basic definition and consequences\ definition permutes :: \('a \ 'a) \ 'a set \ bool\ (infixr \permutes\ 41) where \p permutes S \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)\ lemma bij_imp_permutes: \p permutes S\ if \bij_betw p S S\ and stable: \\x. x \ S \ p x = x\ proof - note \bij_betw p S S\ moreover have \bij_betw p (- S) (- S)\ by (auto simp add: stable intro!: bij_betw_imageI inj_onI) ultimately have \bij_betw p (S \ - S) (S \ - S)\ by (rule bij_betw_combine) simp then have \\!x. p x = y\ for y by (simp add: bij_iff) with stable show ?thesis by (simp add: permutes_def) qed context fixes p :: \'a \ 'a\ and S :: \'a set\ assumes perm: \p permutes S\ begin lemma permutes_inj: \inj p\ using perm by (auto simp: permutes_def inj_on_def) lemma permutes_image: \p ` S = S\ proof (rule set_eqI) fix x show \x \ p ` S \ x \ S\ proof assume \x \ p ` S\ then obtain y where \y \ S\ \p y = x\ by blast with perm show \x \ S\ by (cases \y = x\) (auto simp add: permutes_def) next assume \x \ S\ with perm obtain y where \y \ S\ \p y = x\ by (metis permutes_def) then show \x \ p ` S\ by blast qed qed lemma permutes_not_in: \x \ S \ p x = x\ using perm by (auto simp: permutes_def) lemma permutes_image_complement: \p ` (- S) = - S\ by (auto simp add: permutes_not_in) lemma permutes_in_image: \p x \ S \ x \ S\ using permutes_image permutes_inj by (auto dest: inj_image_mem_iff) lemma permutes_surj: \surj p\ proof - have \p ` (S \ - S) = p ` S \ p ` (- S)\ by (rule image_Un) then show ?thesis by (simp add: permutes_image permutes_image_complement) qed lemma permutes_inv_o: shows "p \ inv p = id" and "inv p \ p = id" using permutes_inj permutes_surj unfolding inj_iff [symmetric] surj_iff [symmetric] by auto lemma permutes_inverses: shows "p (inv p x) = x" and "inv p (p x) = x" using permutes_inv_o [unfolded fun_eq_iff o_def] by auto lemma permutes_inv_eq: \inv p y = x \ p x = y\ by (auto simp add: permutes_inverses) lemma permutes_inj_on: \inj_on p A\ by (rule inj_on_subset [of _ UNIV]) (auto intro: permutes_inj) lemma permutes_bij: \bij p\ unfolding bij_def by (metis permutes_inj permutes_surj) lemma permutes_imp_bij: \bij_betw p S S\ by (simp add: bij_betw_def permutes_image permutes_inj_on) lemma permutes_subset: \p permutes T\ if \S \ T\ proof (rule bij_imp_permutes) define R where \R = T - S\ with that have \T = R \ S\ \R \ S = {}\ by auto then have \p x = x\ if \x \ R\ for x using that by (auto intro: permutes_not_in) then have \p ` R = R\ by simp with \T = R \ S\ show \bij_betw p T T\ by (simp add: bij_betw_def permutes_inj_on image_Un permutes_image) fix x assume \x \ T\ with \T = R \ S\ show \p x = x\ by (simp add: permutes_not_in) qed lemma permutes_imp_permutes_insert: \p permutes insert x S\ by (rule permutes_subset) auto end lemma permutes_id [simp]: \id permutes S\ by (auto intro: bij_imp_permutes) lemma permutes_empty [simp]: \p permutes {} \ p = id\ proof assume \p permutes {}\ then show \p = id\ by (auto simp add: fun_eq_iff permutes_not_in) next assume \p = id\ then show \p permutes {}\ by simp qed lemma permutes_sing [simp]: \p permutes {a} \ p = id\ proof assume perm: \p permutes {a}\ show \p = id\ proof fix x from perm have \p ` {a} = {a}\ by (rule permutes_image) with perm show \p x = id x\ by (cases \x = a\) (auto simp add: permutes_not_in) qed next assume \p = id\ then show \p permutes {a}\ by simp qed lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)" by (simp add: permutes_def) lemma permutes_swap_id: "a \ S \ b \ S \ Fun.swap a b id permutes S" by (rule bij_imp_permutes) (auto simp add: swap_id_eq) lemma permutes_superset: \p permutes T\ if \p permutes S\ \\x. x \ S - T \ p x = x\ proof - define R U where \R = T \ S\ and \U = S - T\ then have \T = R \ (T - S)\ \S = R \ U\ \R \ U = {}\ by auto from that \U = S - T\ have \p ` U = U\ by simp from \p permutes S\ have \bij_betw p (R \ U) (R \ U)\ by (simp add: permutes_imp_bij \S = R \ U\) moreover have \bij_betw p U U\ using that \U = S - T\ by (simp add: bij_betw_def permutes_inj_on) ultimately have \bij_betw p R R\ using \R \ U = {}\ \R \ U = {}\ by (rule bij_betw_partition) then have \p permutes R\ proof (rule bij_imp_permutes) fix x assume \x \ R\ with \R = T \ S\ \p permutes S\ show \p x = x\ by (cases \x \ S\) (auto simp add: permutes_not_in that(2)) qed then have \p permutes R \ (T - S)\ by (rule permutes_subset) simp with \T = R \ (T - S)\ show ?thesis by simp qed lemma permutes_bij_inv_into: \<^marker>\contributor \Lukas Bulwahn\\ fixes A :: "'a set" and B :: "'b set" assumes "p permutes A" and "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) from assms have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A" by (auto simp add: permutes_imp_bij bij_betw_inv_into) then have "bij_betw (f \ p \ inv_into A f) B B" by (simp add: bij_betw_trans) then 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 \ p \ inv_into A f"]) auto next fix x assume "x \ B" then show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto qed lemma permutes_image_mset: \<^marker>\contributor \Lukas Bulwahn\\ 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: \<^marker>\contributor \Lukas Bulwahn\\ 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 \# mset_set A" for x using assms(2)[of x] that by (cases "finite A") auto with assms have "image_mset f (mset_set A) = image_mset (f' \ p) (mset_set A)" 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 permutes_image_mset have "image_mset p (mset_set A) = mset_set A" by blast then show ?thesis by simp qed finally show ?thesis .. qed subsection \Group properties\ lemma permutes_compose: "p permutes S \ q permutes S \ q \ p permutes S" unfolding permutes_def o_def by metis lemma permutes_inv: assumes "p permutes S" shows "inv p permutes S" using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis lemma permutes_inv_inv: assumes "p permutes S" shows "inv (inv p) = p" unfolding fun_eq_iff permutes_inv_eq[OF assms] permutes_inv_eq[OF permutes_inv[OF assms]] by blast lemma permutes_invI: assumes perm: "p permutes S" and inv: "\x. x \ S \ p' (p x) = x" and outside: "\x. x \ S \ p' x = x" shows "inv p = p'" proof show "inv p x = p' x" for x proof (cases "x \ S") case True from assms have "p' x = p' (p (inv p x))" by (simp add: permutes_inverses) also from permutes_inv[OF perm] True have "\ = inv p x" by (subst inv) (simp_all add: permutes_in_image) finally show ?thesis .. next case False with permutes_inv[OF perm] show ?thesis by (simp_all add: outside permutes_not_in) qed qed lemma permutes_vimage: "f permutes A \ f -` A = A" by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv]) subsection \Mapping permutations with bijections\ lemma bij_betw_permutations: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A} {\. \ permutes B}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 show ?case using permutes_bij_inv_into[OF _ assms] by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) { fix \ assume "\ permutes B" from permutes_bij_inv_into[OF this bij_inv] and assms have "(\x. if x \ A then inv_into A f (\ (f x)) else x) permutes A" by (simp add: inv_into_inv_into_eq cong: if_cong) } from this show ?case by (auto simp: permutes_inv) next case 1 thus ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed lemma bij_betw_derangements: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A \ (\x\A. \ x \ x)} {\. \ permutes B \ (\x\B. \ x \ x)}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 have "?f \ x \ x" if "\ permutes A" "\x. x \ A \ \ x \ x" "x \ B" for \ x using that and assms by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on inv_into_f_f inv_into_into permutes_imp_bij) with permutes_bij_inv_into[OF _ assms] show ?case by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) have "?g \ permutes A" if "\ permutes B" for \ using permutes_bij_inv_into[OF that bij_inv] and assms by (simp add: inv_into_inv_into_eq cong: if_cong) moreover have "?g \ x \ x" if "\ permutes B" "\x. x \ B \ \ x \ x" "x \ A" for \ x using that and assms by (metis bij_betwE bij_betw_imp_surj_on f_inv_into_f permutes_imp_bij) ultimately show ?case by auto next case 1 thus ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed subsection \The number of permutations on a finite set\ lemma permutes_insert_lemma: assumes "p permutes (insert a S)" shows "Fun.swap a (p a) id \ p permutes S" apply (rule permutes_superset[where S = "insert a S"]) apply (rule permutes_compose[OF assms]) apply (rule permutes_swap_id, simp) using permutes_in_image[OF assms, of a] apply simp apply (auto simp add: Ball_def Fun.swap_def) done lemma permutes_insert: "{p. p permutes (insert a S)} = (\(b, p). Fun.swap a b id \ p) ` {(b, p). b \ insert a S \ p \ {p. p permutes S}}" proof - have "p permutes insert a S \ (\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S)" for p proof - have "\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S" if p: "p permutes insert a S" proof - let ?b = "p a" let ?q = "Fun.swap a (p a) id \ p" have *: "p = Fun.swap a ?b id \ ?q" by (simp add: fun_eq_iff o_assoc) have **: "?b \ insert a S" unfolding permutes_in_image[OF p] by simp from permutes_insert_lemma[OF p] * ** show ?thesis by blast qed moreover have "p permutes insert a S" if bq: "p = Fun.swap a b id \ q" "b \ insert a S" "q permutes S" for b q proof - from permutes_subset[OF bq(3), of "insert a S"] have q: "q permutes insert a S" by auto have a: "a \ insert a S" by simp from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis by simp qed ultimately show ?thesis by blast qed then show ?thesis by auto qed lemma card_permutations: assumes "card S = n" and "finite S" shows "card {p. p permutes S} = fact n" using assms(2,1) proof (induct arbitrary: n) case empty then show ?case by simp next case (insert x F) { fix n assume card_insert: "card (insert x F) = n" let ?xF = "{p. p permutes insert x F}" let ?pF = "{p. p permutes F}" let ?pF' = "{(b, p). b \ insert x F \ p \ ?pF}" let ?g = "(\(b, p). Fun.swap x b id \ p)" have xfgpF': "?xF = ?g ` ?pF'" by (rule permutes_insert[of x F]) from \x \ F\ \finite F\ card_insert have Fs: "card F = n - 1" by auto from \finite F\ insert.hyps Fs have pFs: "card ?pF = fact (n - 1)" by auto then have "finite ?pF" by (auto intro: card_ge_0_finite) with \finite F\ card.insert_remove have pF'f: "finite ?pF'" apply (simp only: Collect_case_prod Collect_mem_eq) apply (rule finite_cartesian_product) apply simp_all done have ginj: "inj_on ?g ?pF'" proof - { fix b p c q assume bp: "(b, p) \ ?pF'" assume cq: "(c, q) \ ?pF'" assume eq: "?g (b, p) = ?g (c, q)" from bp cq have pF: "p permutes F" and qF: "q permutes F" by auto from pF \x \ F\ eq have "b = ?g (b, p) x" by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff) also from qF \x \ F\ eq have "\ = ?g (c, q) x" by (auto simp: fun_upd_def fun_eq_iff) also from qF \x \ F\ have "\ = c" by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff) finally have "b = c" . then have "Fun.swap x b id = Fun.swap x c id" by simp with eq have "Fun.swap x b id \ p = Fun.swap x b id \ q" by simp then have "Fun.swap x b id \ (Fun.swap x b id \ p) = Fun.swap x b id \ (Fun.swap x b id \ q)" by simp then have "p = q" by (simp add: o_assoc) with \b = c\ have "(b, p) = (c, q)" by simp } then show ?thesis unfolding inj_on_def by blast qed from \x \ F\ \finite F\ card_insert have "n \ 0" by auto then have "\m. n = Suc m" by presburger then obtain m where n: "n = Suc m" by blast from pFs card_insert have *: "card ?xF = fact n" unfolding xfgpF' card_image[OF ginj] using \finite F\ \finite ?pF\ by (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) (simp add: n) from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" by (simp add: xfgpF' n) from * have "card ?xF = fact n" unfolding xFf by blast } with insert show ?case by simp qed lemma finite_permutations: assumes "finite S" shows "finite {p. p permutes S}" using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite) subsection \Hence a sort of induction principle composing by swaps\ lemma permutes_induct [consumes 2, case_names id swap]: \P p\ if \p permutes S\ \finite S\ and id: \P id\ and swap: \\a b p. a \ S \ b \ S \ p permutes S \ P p \ P (Fun.swap a b id \ p)\ using \finite S\ \p permutes S\ swap proof (induction S arbitrary: p) case empty with id show ?case by (simp only: permutes_empty) next case (insert x S p) define q where \q = Fun.swap x (p x) id \ p\ then have swap_q: \Fun.swap x (p x) id \ q = p\ by (simp add: o_assoc) from \p permutes insert x S\ have \q permutes S\ by (simp add: q_def permutes_insert_lemma) then have \q permutes insert x S\ by (simp add: permutes_imp_permutes_insert) from \q permutes S\ have \P q\ by (auto intro: insert.IH insert.prems(2) permutes_imp_permutes_insert) have \x \ insert x S\ by simp moreover from \p permutes insert x S\ have \p x \ insert x S\ using permutes_in_image [of p \insert x S\ x] by simp ultimately have \P (Fun.swap x (p x) id \ q)\ using \q permutes insert x S\ \P q\ by (rule insert.prems(2)) then show ?case by (simp add: swap_q) qed lemma permutes_rev_induct [consumes 2, case_names id swap]: \P p\ if \p permutes S\ \finite S\ and id': \P id\ and swap': \\a b p. a \ S \ b \ S \ p permutes S \ P p \ P (Fun.swap a b p)\ using \p permutes S\ \finite S\ proof (induction rule: permutes_induct) case id from id' show ?case . next case (swap a b p) have \P (Fun.swap (inv p a) (inv p b) p)\ by (rule swap') (auto simp add: swap permutes_in_image permutes_inv) also have \Fun.swap (inv p a) (inv p b) p = Fun.swap a b id \ p\ by (rule bij_swap_comp [symmetric]) (rule permutes_bij, rule swap) finally show ?case . qed subsection \Permutations of index set for iterated operations\ lemma (in comm_monoid_set) permute: assumes "p permutes S" shows "F g S = F (g \ p) S" proof - from \p permutes S\ have "inj p" by (rule permutes_inj) then have "inj_on p S" by (auto intro: subset_inj_on) then have "F g (p ` S) = F (g \ p) S" by (rule reindex) moreover from \p permutes S\ have "p ` S = S" by (rule permutes_image) ultimately show ?thesis by simp qed subsection \Permutations as transposition sequences\ inductive swapidseq :: "nat \ ('a \ 'a) \ bool" where id[simp]: "swapidseq 0 id" | comp_Suc: "swapidseq n p \ a \ b \ swapidseq (Suc n) (Fun.swap a b id \ p)" declare id[unfolded id_def, simp] definition "permutation p \ (\n. swapidseq n p)" subsection \Some closure properties of the set of permutations, with lengths\ lemma permutation_id[simp]: "permutation id" unfolding permutation_def by (rule exI[where x=0]) simp declare permutation_id[unfolded id_def, simp] lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)" apply clarsimp using comp_Suc[of 0 id a b] apply simp done lemma permutation_swap_id: "permutation (Fun.swap a b id)" proof (cases "a = b") case True then show ?thesis by simp next case False then show ?thesis unfolding permutation_def using swapidseq_swap[of a b] by blast qed lemma swapidseq_comp_add: "swapidseq n p \ swapidseq m q \ swapidseq (n + m) (p \ q)" proof (induct n p arbitrary: m q rule: swapidseq.induct) case (id m q) then show ?case by simp next case (comp_Suc n p a b m q) have eq: "Suc n + m = Suc (n + m)" by arith show ?case apply (simp only: eq comp_assoc) apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3) apply blast+ done qed lemma permutation_compose: "permutation p \ permutation q \ permutation (p \ q)" unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis lemma swapidseq_endswap: "swapidseq n p \ a \ b \ swapidseq (Suc n) (p \ Fun.swap a b id)" by (induct n p rule: swapidseq.induct) (use swapidseq_swap[of a b] in \auto simp add: comp_assoc intro: swapidseq.comp_Suc\) lemma swapidseq_inverse_exists: "swapidseq n p \ \q. swapidseq n q \ p \ q = id \ q \ p = id" proof (induct n p rule: swapidseq.induct) case id then show ?case by (rule exI[where x=id]) simp next case (comp_Suc n p a b) from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast let ?q = "q \ Fun.swap a b id" note H = comp_Suc.hyps from swapidseq_swap[of a b] H(3) have *: "swapidseq 1 (Fun.swap a b id)" by simp from swapidseq_comp_add[OF q(1) *] have **: "swapidseq (Suc n) ?q" by simp have "Fun.swap a b id \ p \ ?q = Fun.swap a b id \ (p \ q) \ Fun.swap a b id" by (simp add: o_assoc) also have "\ = id" by (simp add: q(2)) finally have ***: "Fun.swap a b id \ p \ ?q = id" . have "?q \ (Fun.swap a b id \ p) = q \ (Fun.swap a b id \ Fun.swap a b id) \ p" by (simp only: o_assoc) then have "?q \ (Fun.swap a b id \ p) = id" by (simp add: q(3)) with ** *** show ?case by blast qed lemma swapidseq_inverse: assumes "swapidseq n p" shows "swapidseq n (inv p)" using swapidseq_inverse_exists[OF assms] inv_unique_comp[of p] by auto lemma permutation_inverse: "permutation p \ permutation (inv p)" using permutation_def swapidseq_inverse by blast subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ lemma swap_id_common:" a \ c \ b \ c \ Fun.swap a b id \ Fun.swap a c id = Fun.swap b c id \ Fun.swap a b id" by (simp add: fun_eq_iff Fun.swap_def) lemma swap_id_common': "a \ b \ a \ c \ Fun.swap a c id \ Fun.swap b c id = Fun.swap b c id \ Fun.swap a b id" by (simp add: fun_eq_iff Fun.swap_def) lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ Fun.swap a b id \ Fun.swap c d id = Fun.swap c d id \ Fun.swap a b id" by (simp add: fun_eq_iff Fun.swap_def) subsection \The identity map only has even transposition sequences\ lemma symmetry_lemma: assumes "\a b c d. P a b c d \ P a b d c" and "\a b c d. a \ b \ c \ d \ a = c \ b = d \ a = c \ b \ d \ a \ c \ b = d \ a \ c \ a \ d \ b \ c \ b \ d \ P a b c d" shows "\a b c d. a \ b \ c \ d \ P a b c d" using assms by metis lemma swap_general: "a \ b \ c \ d \ Fun.swap a b id \ Fun.swap c d id = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id)" proof - assume neq: "a \ b" "c \ d" have "a \ b \ c \ d \ (Fun.swap a b id \ Fun.swap c d id = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id))" apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d]) apply (simp_all only: swap_commute) apply (case_tac "a = c \ b = d") apply (clarsimp simp only: swap_commute swap_id_idempotent) apply (case_tac "a = c \ b \ d") apply (rule disjI2) apply (rule_tac x="b" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="b" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) apply (case_tac "a \ c \ b = d") apply (rule disjI2) apply (rule_tac x="c" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="c" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) apply (rule disjI2) apply (rule_tac x="c" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="b" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) done with neq show ?thesis by metis qed lemma swapidseq_id_iff[simp]: "swapidseq 0 p \ p = id" using swapidseq.cases[of 0 p "p = id"] by auto lemma swapidseq_cases: "swapidseq n p \ n = 0 \ p = id \ (\a b q m. n = Suc m \ p = Fun.swap a b id \ q \ swapidseq m q \ a \ b)" apply (rule iffI) apply (erule swapidseq.cases[of n p]) apply simp apply (rule disjI2) apply (rule_tac x= "a" in exI) apply (rule_tac x= "b" in exI) apply (rule_tac x= "pa" in exI) apply (rule_tac x= "na" in exI) apply simp apply auto apply (rule comp_Suc, simp_all) done lemma fixing_swapidseq_decrease: assumes "swapidseq n p" and "a \ b" and "(Fun.swap a b id \ p) a = a" shows "n \ 0 \ swapidseq (n - 1) (Fun.swap a b id \ p)" using assms proof (induct n arbitrary: p a b) case 0 then show ?case by (auto simp add: Fun.swap_def fun_upd_def) next case (Suc n p a b) from Suc.prems(1) swapidseq_cases[of "Suc n" p] obtain c d q m where cdqm: "Suc n = Suc m" "p = Fun.swap c d id \ q" "swapidseq m q" "c \ d" "n = m" by auto consider "Fun.swap a b id \ Fun.swap c d id = id" | x y z where "x \ a" "y \ a" "z \ a" "x \ y" "Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id" using swap_general[OF Suc.prems(2) cdqm(4)] by metis then show ?case proof cases case 1 then show ?thesis by (simp only: cdqm o_assoc) (simp add: cdqm) next case prems: 2 then have az: "a \ z" by simp from prems have *: "(Fun.swap x y id \ h) a = a \ h a = a" for h by (simp add: Fun.swap_def) from cdqm(2) have "Fun.swap a b id \ p = Fun.swap a b id \ (Fun.swap c d id \ q)" by simp then have "Fun.swap a b id \ p = Fun.swap x y id \ (Fun.swap a z id \ q)" by (simp add: o_assoc prems) then have "(Fun.swap a b id \ p) a = (Fun.swap x y id \ (Fun.swap a z id \ q)) a" by simp then have "(Fun.swap x y id \ (Fun.swap a z id \ q)) a = a" unfolding Suc by metis then have "(Fun.swap a z id \ q) a = a" by (simp only: *) from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az this] have **: "swapidseq (n - 1) (Fun.swap a z id \ q)" "n \ 0" by blast+ from \n \ 0\ have ***: "Suc n - 1 = Suc (n - 1)" by auto show ?thesis apply (simp only: cdqm(2) prems o_assoc ***) apply (simp only: Suc_not_Zero simp_thms comp_assoc) apply (rule comp_Suc) using ** prems apply blast+ done qed qed lemma swapidseq_identity_even: assumes "swapidseq n (id :: 'a \ 'a)" shows "even n" using \swapidseq n id\ proof (induct n rule: nat_less_induct) case H: (1 n) consider "n = 0" | a b :: 'a and q m where "n = Suc m" "id = Fun.swap a b id \ q" "swapidseq m q" "a \ b" using H(2)[unfolded swapidseq_cases[of n id]] by auto then show ?case proof cases case 1 then show ?thesis by presburger next case h: 2 from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]] have m: "m \ 0" "swapidseq (m - 1) (id :: 'a \ 'a)" by auto from h m have mn: "m - 1 < n" by arith from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis by presburger qed qed subsection \Therefore we have a welldefined notion of parity\ definition "evenperm p = even (SOME n. swapidseq n p)" lemma swapidseq_even_even: assumes m: "swapidseq m p" and n: "swapidseq n p" shows "even m \ even n" proof - from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis by arith qed lemma evenperm_unique: assumes p: "swapidseq n p" and n:"even n = b" shows "evenperm p = b" unfolding n[symmetric] evenperm_def apply (rule swapidseq_even_even[where p = p]) apply (rule someI[where x = n]) using p apply blast+ done subsection \And it has the expected composition properties\ lemma evenperm_id[simp]: "evenperm id = True" by (rule evenperm_unique[where n = 0]) simp_all +lemma evenperm_identity [simp]: + \evenperm (\x. x)\ + using evenperm_id by (simp add: id_def [abs_def]) + lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)" by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap) lemma evenperm_comp: assumes "permutation p" "permutation q" shows "evenperm (p \ q) \ evenperm p = evenperm q" proof - from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q" unfolding permutation_def by blast have "even (n + m) \ (even n \ even m)" by arith from evenperm_unique[OF n refl] evenperm_unique[OF m refl] and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis by blast qed lemma evenperm_inv: assumes "permutation p" shows "evenperm (inv p) = evenperm p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast show ?thesis by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]) qed subsection \A more abstract characterization of permutations\ lemma permutation_bijective: assumes "permutation p" shows "bij p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast then show ?thesis unfolding bij_iff apply (auto simp add: fun_eq_iff) apply metis done qed lemma permutation_finite_support: assumes "permutation p" shows "finite {x. p x \ x}" proof - from assms obtain n where "swapidseq n p" unfolding permutation_def by blast then show ?thesis proof (induct n p rule: swapidseq.induct) case id then show ?case by simp next case (comp_Suc n p a b) let ?S = "insert a (insert b {x. p x \ x})" from comp_Suc.hyps(2) have *: "finite ?S" by simp from \a \ b\ have **: "{x. (Fun.swap a b id \ p) x \ x} \ ?S" by (auto simp: Fun.swap_def) show ?case by (rule finite_subset[OF ** *]) qed qed lemma permutation_lemma: assumes "finite S" and "bij p" and "\x. x \ S \ p x = x" shows "permutation p" using assms proof (induct S arbitrary: p rule: finite_induct) case empty then show ?case by simp next case (insert a F p) let ?r = "Fun.swap a (p a) id \ p" let ?q = "Fun.swap a (p a) id \ ?r" have *: "?r a = a" by (simp add: Fun.swap_def) from insert * have **: "\x. x \ F \ ?r x = x" by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3)) have "bij ?r" by (rule bij_swap_compose_bij[OF insert(4)]) have "permutation ?r" by (rule insert(3)[OF \bij ?r\ **]) then have "permutation ?q" by (simp add: permutation_compose permutation_swap_id) then show ?case by (simp add: o_assoc) qed lemma permutation: "permutation p \ bij p \ finite {x. p x \ x}" (is "?lhs \ ?b \ ?f") proof assume ?lhs with permutation_bijective permutation_finite_support show "?b \ ?f" by auto next assume "?b \ ?f" then have "?f" "?b" by blast+ from permutation_lemma[OF this] show ?lhs by blast qed lemma permutation_inverse_works: assumes "permutation p" shows "inv p \ p = id" and "p \ inv p = id" using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff) lemma permutation_inverse_compose: assumes p: "permutation p" and q: "permutation q" shows "inv (p \ q) = inv q \ inv p" proof - note ps = permutation_inverse_works[OF p] note qs = permutation_inverse_works[OF q] have "p \ q \ (inv q \ inv p) = p \ (q \ inv q) \ inv p" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have *: "p \ q \ (inv q \ inv p) = id" . have "inv q \ inv p \ (p \ q) = inv q \ (inv p \ p) \ q" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have **: "inv q \ inv p \ (p \ q) = id" . show ?thesis by (rule inv_unique_comp[OF * **]) qed subsection \Relation to \permutes\\ lemma permutes_imp_permutation: \permutation p\ if \finite S\ \p permutes S\ proof - from \p permutes S\ have \{x. p x \ x} \ S\ by (auto dest: permutes_not_in) then have \finite {x. p x \ x}\ using \finite S\ by (rule finite_subset) moreover from \p permutes S\ have \bij p\ by (auto dest: permutes_bij) ultimately show ?thesis by (simp add: permutation) qed lemma permutation_permutesE: assumes \permutation p\ obtains S where \finite S\ \p permutes S\ proof - from assms have fin: \finite {x. p x \ x}\ by (simp add: permutation) from assms have \bij p\ by (simp add: permutation) also have \UNIV = {x. p x \ x} \ {x. p x = x}\ by auto finally have \bij_betw p {x. p x \ x} {x. p x \ x}\ by (rule bij_betw_partition) (auto simp add: bij_betw_fixpoints) then have \p permutes {x. p x \ x}\ by (auto intro: bij_imp_permutes) with fin show thesis .. qed lemma permutation_permutes: "permutation p \ (\S. finite S \ p permutes S)" by (auto elim: permutation_permutesE intro: permutes_imp_permutation) subsection \Sign of a permutation as a real number\ definition sign :: \('a \ 'a) \ int\ \ \TODO: prefer less generic name\ - where \sign p = (if evenperm p then (1::int) else -1)\ + where \sign p = (if evenperm p then 1 else - 1)\ -lemma sign_nz: "sign p \ 0" +lemma sign_cases [case_names even odd]: + obtains \sign p = 1\ | \sign p = - 1\ + by (cases \evenperm p\) (simp_all add: sign_def) + +lemma sign_nz [simp]: "sign p \ 0" + by (cases p rule: sign_cases) simp_all + +lemma sign_id [simp]: "sign id = 1" by (simp add: sign_def) -lemma sign_id: "sign id = 1" +lemma sign_identity [simp]: + \sign (\x. x) = 1\ by (simp add: sign_def) lemma sign_inverse: "permutation p \ sign (inv p) = sign p" by (simp add: sign_def evenperm_inv) lemma sign_compose: "permutation p \ permutation q \ sign (p \ q) = sign p * sign q" by (simp add: sign_def evenperm_comp) -lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)" +lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else - 1)" by (simp add: sign_def evenperm_swap) -lemma sign_idempotent: "sign p * sign p = 1" +lemma sign_idempotent [simp]: "sign p * sign p = 1" by (simp add: sign_def) +lemma sign_left_idempotent [simp]: + \sign p * (sign p * sign q) = sign q\ + by (simp add: sign_def) + +term "(bij, bij_betw, permutation)" + subsection \Permuting a list\ text \This function permutes a list by applying a permutation to the indices.\ definition permute_list :: "(nat \ nat) \ 'a list \ 'a list" where "permute_list f xs = map (\i. xs ! (f i)) [0.. g) xs = permute_list g (permute_list f xs)" using assms[THEN permutes_in_image] by (auto simp add: permute_list_def) lemma permute_list_ident [simp]: "permute_list (\x. x) xs = xs" by (simp add: permute_list_def map_nth) lemma permute_list_id [simp]: "permute_list id xs = xs" by (simp add: id_def) lemma mset_permute_list [simp]: fixes xs :: "'a list" assumes "f permutes {.. x < length xs" for x using permutes_in_image[OF assms] by auto have "count (mset (permute_list f xs)) y = card ((\i. xs ! f i) -` {y} \ {..i. xs ! f i) -` {y} \ {.. y = xs ! i}" by auto also from assms have "card \ = card {i. i < length xs \ y = xs ! i}" by (intro card_vimage_inj) (auto simp: permutes_inj permutes_surj) also have "\ = count (mset xs) y" by (simp add: count_mset length_filter_conv_card) finally show "count (mset (permute_list f xs)) y = count (mset xs) y" by simp qed lemma set_permute_list [simp]: assumes "f permutes {.. i < length ys" for i by simp have "permute_list f (zip xs ys) = map (\i. zip xs ys ! f i) [0.. = map (\(x, y). (xs ! f x, ys ! f y)) (zip [0.. = zip (permute_list f xs) (permute_list f ys)" by (simp_all add: permute_list_def zip_map_map) finally show ?thesis . qed lemma map_of_permute: assumes "\ permutes fst ` set xs" shows "map_of xs \ \ = map_of (map (\(x,y). (inv \ x, y)) xs)" (is "_ = map_of (map ?f _)") proof from assms have "inj \" "surj \" by (simp_all add: permutes_inj permutes_surj) then show "(map_of xs \ \) x = map_of (map ?f xs) x" for x by (induct xs) (auto simp: inv_f_f surj_f_inv_f) qed subsection \More lemmas about permutations\ lemma permutes_in_funpow_image: \<^marker>\contributor \Lars Noschinski\\ assumes "f permutes S" "x \ S" shows "(f ^^ n) x \ S" using assms by (induction n) (auto simp: permutes_in_image) lemma permutation_self: \<^marker>\contributor \Lars Noschinski\\ assumes \permutation p\ obtains n where \n > 0\ \(p ^^ n) x = x\ proof (cases \p x = x\) case True with that [of 1] show thesis by simp next case False from \permutation p\ have \inj p\ by (intro permutation_bijective bij_is_inj) moreover from \p x \ x\ have \(p ^^ Suc n) x \ (p ^^ n) x\ for n proof (induction n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have "p (p x) \ p x" proof (rule notI) assume "p (p x) = p x" then show False using \p x \ x\ \inj p\ by (simp add: inj_eq) qed have "(p ^^ Suc (Suc n)) x = (p ^^ Suc n) (p x)" by (simp add: funpow_swap1) also have "\ \ (p ^^ n) (p x)" by (rule Suc) fact also have "(p ^^ n) (p x) = (p ^^ Suc n) x" by (simp add: funpow_swap1) finally show ?case by simp qed then have "{y. \n. y = (p ^^ n) x} \ {x. p x \ x}" by auto then have "finite {y. \n. y = (p ^^ n) x}" using permutation_finite_support[OF assms] by (rule finite_subset) ultimately obtain n where \n > 0\ \(p ^^ n) x = x\ by (rule funpow_inj_finite) with that [of n] show thesis by blast qed text \The following few lemmas were contributed by Lukas Bulwahn.\ lemma count_image_mset_eq_card_vimage: assumes "finite A" shows "count (image_mset f (mset_set A)) b = card {a \ A. f a = b}" using assms proof (induct A) case empty show ?case by simp next case (insert x F) show ?case proof (cases "f x = b") case True with insert.hyps have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \ F. f a = f x})" by auto also from insert.hyps(1,2) have "\ = card (insert x {a \ F. f a = f x})" by simp also from \f x = b\ have "card (insert x {a \ F. f a = f x}) = card {a \ insert x F. f a = b}" by (auto intro: arg_cong[where f="card"]) finally show ?thesis using insert by auto next case False then have "{a \ F. f a = b} = {a \ insert x F. f a = b}" by auto with insert False 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" and 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 \finite A\ have [simp]: "finite {a \ A. f a = (b::'b)}" for f b by auto have "f ` A = f' ` A" proof - from \finite A\ have "f ` A = f ` (set_mset (mset_set A))" by simp also have "\ = f' ` set_mset (mset_set A)" by (metis mset_eq multiset.set_map) also from \finite A\ have "\ = f' ` 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 b from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b" by simp with \finite A\ have "card {a \ A. f a = b} = card {a \ A. f' a = b}" by (simp add: count_image_mset_eq_card_vimage) then show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}" by (intro finite_same_card_bij) simp_all qed then have "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" by (rule bchoice) then obtain p where p: "\b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" .. 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)" by (auto simp: disjoint_family_on_def) moreover have "bij_betw (\a. p (f a) a) {a \ A. f a = b} {a \ A. f' a = b}" if "b \ f ` A" for b using p that 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 from \f ` A = f' ` A\ have "(\b\f ` A. {a \ A. f' a = b}) = 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 show "\x. x \ A \ p' x = x" by (simp add: p'_def) qed moreover from p have "\x\A. f x = f' (p' x)" unfolding p'_def using bij_betwE by fastforce ultimately show ?thesis .. qed \ \... and derive the existing property:\ lemma mset_eq_permutation: fixes xs ys :: "'a list" assumes mset_eq: "mset xs = mset ys" obtains p where "p permutes {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i\{..i \ S. p i \ i" shows "p = id" proof - have "p n = n" for n using assms proof (induct n arbitrary: S rule: less_induct) case (less n) show ?case proof (cases "n \ S") case False with less(2) show ?thesis unfolding permutes_def by metis next case True with less(3) have "p n < n \ p n = n" by auto then show ?thesis proof assume "p n < n" with less have "p (p n) = p n" by metis with permutes_inj[OF less(2)] have "p n = n" unfolding inj_def by blast with \p n < n\ have False by simp then show ?thesis .. qed qed qed then show ?thesis by (auto simp: fun_eq_iff) qed lemma permutes_natset_ge: fixes S :: "'a::wellorder set" assumes p: "p permutes S" and le: "\i \ S. p i \ i" shows "p = id" proof - have "i \ inv p i" if "i \ S" for i proof - from that permutes_in_image[OF permutes_inv[OF p]] have "inv p i \ S" by simp with le have "p (inv p i) \ inv p i" by blast with permutes_inverses[OF p] show ?thesis by simp qed then have "\i\S. inv p i \ i" by blast from permutes_natset_le[OF permutes_inv[OF p] this] have "inv p = inv id" by simp then show ?thesis apply (subst permutes_inv_inv[OF p, symmetric]) apply (rule inv_unique_comp) apply simp_all done qed lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto using permutes_inv_inv permutes_inv apply auto apply (rule_tac x="inv x" in exI) apply auto done lemma image_compose_permutations_left: assumes "q permutes S" shows "{q \ p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "inv q \ x" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o) done lemma image_compose_permutations_right: assumes "q permutes S" shows "{p \ q | p. p permutes S} = {p . p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "x \ inv q" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc) done lemma permutes_in_seg: "p permutes {1 ..n} \ i \ {1..n} \ 1 \ p i \ p i \ n" by (simp add: permutes_def) metis lemma sum_permutations_inverse: "sum f {p. p permutes S} = sum (\p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p . p permutes S}" have *: "inj_on inv ?S" proof (auto simp add: inj_on_def) fix q r assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r" then have "inv (inv q) = inv (inv r)" by simp with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r" by metis qed have **: "inv ` ?S = ?S" using image_inverse_permutations by blast have ***: "?rhs = sum (f \ inv) ?S" by (simp add: o_def) from sum.reindex[OF *, of f] show ?thesis by (simp only: ** ***) qed lemma setum_permutations_compose_left: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(q \ p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ ((\) q)) ?S" by (simp add: o_def) have **: "inj_on ((\) q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "q \ p = q \ r" then have "inv q \ q \ p = inv q \ q \ r" by (simp add: comp_assoc) with permutes_inj[OF q, unfolded inj_iff] show "p = r" by simp qed have "((\) q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed lemma sum_permutations_compose_right: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(p \ q)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ (\p. p \ q)) ?S" by (simp add: o_def) have **: "inj_on (\p. p \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "p \ q = r \ q" then have "p \ (q \ inv q) = r \ (q \ inv q)" by (simp add: o_assoc) with permutes_surj[OF q, unfolded surj_iff] show "p = r" by simp qed from image_compose_permutations_right[OF q] have "(\p. p \ q) ` ?S = ?S" by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed +lemma inv_inj_on_permutes: + \inj_on inv {p. p permutes S}\ +proof (intro inj_onI, unfold mem_Collect_eq) + fix p q + assume p: "p permutes S" and q: "q permutes S" and eq: "inv p = inv q" + have "inv (inv p) = inv (inv q)" using eq by simp + thus "p = q" + using inv_inv_eq[OF permutes_bij] p q by metis +qed + +lemma permutes_pair_eq: + \{(p s, s) |s. s \ S} = {(s, inv p s) |s. s \ S}\ (is \?L = ?R\) if \p permutes S\ +proof + show "?L \ ?R" + proof + fix x assume "x \ ?L" + then obtain s where x: "x = (p s, s)" and s: "s \ S" by auto + note x + also have "(p s, s) = (p s, Hilbert_Choice.inv p (p s))" + using permutes_inj [OF that] inv_f_f by auto + also have "... \ ?R" using s permutes_in_image[OF that] by auto + finally show "x \ ?R". + qed + show "?R \ ?L" + proof + fix x assume "x \ ?R" + then obtain s + where x: "x = (s, Hilbert_Choice.inv p s)" (is "_ = (s, ?ips)") + and s: "s \ S" by auto + note x + also have "(s, ?ips) = (p ?ips, ?ips)" + using inv_f_f[OF permutes_inj[OF permutes_inv[OF that]]] + using inv_inv_eq[OF permutes_bij[OF that]] by auto + also have "... \ ?L" + using s permutes_in_image[OF permutes_inv[OF that]] by auto + finally show "x \ ?L". + qed +qed + +context + fixes p and n i :: nat + assumes p: \p permutes {0.. and i: \i < n\ +begin + +lemma permutes_nat_less: + \p i < n\ +proof - + have \?thesis \ p i \ {0.. + by simp + also from p have \p i \ {0.. i \ {0.. + by (rule permutes_in_image) + finally show ?thesis + using i by simp +qed + +lemma permutes_nat_inv_less: + \inv p i < n\ +proof - + from p have \inv p permutes {0.. + by (rule permutes_inv) + then show ?thesis + using i by (rule Permutations.permutes_nat_less) +qed + +end + +context comm_monoid_set +begin + +lemma permutes_inv: + \F (\s. g (p s) s) S = F (\s. g s (inv p s)) S\ (is \?l = ?r\) + if \p permutes S\ +proof - + let ?g = "\(x, y). g x y" + let ?ps = "\s. (p s, s)" + let ?ips = "\s. (s, inv p s)" + have inj1: "inj_on ?ps S" by (rule inj_onI) auto + have inj2: "inj_on ?ips S" by (rule inj_onI) auto + have "?l = F ?g (?ps ` S)" + using reindex [OF inj1, of ?g] by simp + also have "?ps ` S = {(p s, s) |s. s \ S}" by auto + also have "... = {(s, inv p s) |s. s \ S}" + unfolding permutes_pair_eq [OF that] by simp + also have "... = ?ips ` S" by auto + also have "F ?g ... = ?r" + using reindex [OF inj2, of ?g] by simp + finally show ?thesis. +qed + +end + subsection \Sum over a set of permutations (could generalize to iteration)\ lemma sum_over_permutations_insert: assumes fS: "finite S" and aS: "a \ S" shows "sum f {p. p permutes (insert a S)} = sum (\b. sum (\q. f (Fun.swap a b id \ q)) {p. p permutes S}) (insert a S)" proof - have *: "\f a b. (\(b, p). f (Fun.swap a b id \ p)) = f \ (\(b,p). Fun.swap a b id \ p)" by (simp add: fun_eq_iff) have **: "\P Q. {(a, b). a \ P \ b \ Q} = P \ Q" by blast show ?thesis unfolding * ** sum.cartesian_product permutes_insert proof (rule sum.reindex) let ?f = "(\(b, y). Fun.swap a b id \ y)" let ?P = "{p. p permutes S}" { fix b c p q assume b: "b \ insert a S" assume c: "c \ insert a S" assume p: "p permutes S" assume q: "q permutes S" assume eq: "Fun.swap a b id \ p = Fun.swap a c id \ q" from p q aS have pa: "p a = a" and qa: "q a = a" unfolding permutes_def by metis+ from eq have "(Fun.swap a b id \ p) a = (Fun.swap a c id \ q) a" by simp then have bc: "b = c" by (simp add: permutes_def pa qa o_def fun_upd_def Fun.swap_def id_def cong del: if_weak_cong split: if_split_asm) from eq[unfolded bc] have "(\p. Fun.swap a c id \ p) (Fun.swap a c id \ p) = (\p. Fun.swap a c id \ p) (Fun.swap a c id \ q)" by simp then have "p = q" unfolding o_assoc swap_id_idempotent by simp with bc have "b = c \ p = q" by blast } then show "inj_on ?f (insert a S \ ?P)" unfolding inj_on_def by clarify metis qed qed subsection \Constructing permutations from association lists\ definition list_permutes :: "('a \ 'a) list \ 'a set \ bool" where "list_permutes xs A \ set (map fst xs) \ A \ set (map snd xs) = set (map fst xs) \ distinct (map fst xs) \ distinct (map snd xs)" lemma list_permutesI [simp]: assumes "set (map fst xs) \ A" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)" shows "list_permutes xs A" proof - from assms(2,3) have "distinct (map snd xs)" by (intro card_distinct) (simp_all add: distinct_card del: set_map) with assms show ?thesis by (simp add: list_permutes_def) qed definition permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "permutation_of_list xs x = (case map_of xs x of None \ x | Some y \ y)" lemma permutation_of_list_Cons: "permutation_of_list ((x, y) # xs) x' = (if x = x' then y else permutation_of_list xs x')" by (simp add: permutation_of_list_def) fun inverse_permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "inverse_permutation_of_list [] x = x" | "inverse_permutation_of_list ((y, x') # xs) x = (if x = x' then y else inverse_permutation_of_list xs x)" declare inverse_permutation_of_list.simps [simp del] lemma inj_on_map_of: assumes "distinct (map snd xs)" shows "inj_on (map_of xs) (set (map fst xs))" proof (rule inj_onI) fix x y assume xy: "x \ set (map fst xs)" "y \ set (map fst xs)" assume eq: "map_of xs x = map_of xs y" from xy obtain x' y' where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'" by (cases "map_of xs x"; cases "map_of xs y") (simp_all add: map_of_eq_None_iff) moreover from x'y' have *: "(x, x') \ set xs" "(y, y') \ set xs" by (force dest: map_of_SomeD)+ moreover from * eq x'y' have "x' = y'" by simp ultimately show "x = y" using assms by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"]) qed lemma inj_on_the: "None \ A \ inj_on the A" by (auto simp: inj_on_def option.the_def split: option.splits) lemma inj_on_map_of': assumes "distinct (map snd xs)" shows "inj_on (the \ map_of xs) (set (map fst xs))" by (intro comp_inj_on inj_on_map_of assms inj_on_the) (force simp: eq_commute[of None] map_of_eq_None_iff) lemma image_map_of: assumes "distinct (map fst xs)" shows "map_of xs ` set (map fst xs) = Some ` set (map snd xs)" using assms by (auto simp: rev_image_eqI) lemma the_Some_image [simp]: "the ` Some ` A = A" by (subst image_image) simp lemma image_map_of': assumes "distinct (map fst xs)" shows "(the \ map_of xs) ` set (map fst xs) = set (map snd xs)" by (simp only: image_comp [symmetric] image_map_of assms the_Some_image) lemma permutation_of_list_permutes [simp]: assumes "list_permutes xs A" shows "permutation_of_list xs permutes A" (is "?f permutes _") proof (rule permutes_subset[OF bij_imp_permutes]) from assms show "set (map fst xs) \ A" by (simp add: list_permutes_def) from assms have "inj_on (the \ map_of xs) (set (map fst xs))" (is ?P) by (intro inj_on_map_of') (simp_all add: list_permutes_def) also have "?P \ inj_on ?f (set (map fst xs))" by (intro inj_on_cong) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) finally have "bij_betw ?f (set (map fst xs)) (?f ` set (map fst xs))" by (rule inj_on_imp_bij_betw) also from assms have "?f ` set (map fst xs) = (the \ map_of xs) ` set (map fst xs)" by (intro image_cong refl) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) also from assms have "\ = set (map fst xs)" by (subst image_map_of') (simp_all add: list_permutes_def) finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" . qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+ lemma eval_permutation_of_list [simp]: "permutation_of_list [] x = x" "x = x' \ permutation_of_list ((x',y)#xs) x = y" "x \ x' \ permutation_of_list ((x',y')#xs) x = permutation_of_list xs x" by (simp_all add: permutation_of_list_def) lemma eval_inverse_permutation_of_list [simp]: "inverse_permutation_of_list [] x = x" "x = x' \ inverse_permutation_of_list ((y,x')#xs) x = y" "x \ x' \ inverse_permutation_of_list ((y',x')#xs) x = inverse_permutation_of_list xs x" by (simp_all add: inverse_permutation_of_list.simps) lemma permutation_of_list_id: "x \ set (map fst xs) \ permutation_of_list xs x = x" by (induct xs) (auto simp: permutation_of_list_Cons) lemma permutation_of_list_unique': "distinct (map fst xs) \ (x, y) \ set xs \ permutation_of_list xs x = y" by (induct xs) (force simp: permutation_of_list_Cons)+ lemma permutation_of_list_unique: "list_permutes xs A \ (x, y) \ set xs \ permutation_of_list xs x = y" by (intro permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_id: "x \ set (map snd xs) \ inverse_permutation_of_list xs x = x" by (induct xs) auto lemma inverse_permutation_of_list_unique': "distinct (map snd xs) \ (x, y) \ set xs \ inverse_permutation_of_list xs y = x" by (induct xs) (force simp: inverse_permutation_of_list.simps(2))+ lemma inverse_permutation_of_list_unique: "list_permutes xs A \ (x,y) \ set xs \ inverse_permutation_of_list xs y = x" by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_correct: fixes A :: "'a set" assumes "list_permutes xs A" shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)" proof (rule ext, rule sym, subst permutes_inv_eq) from assms show "permutation_of_list xs permutes A" by simp show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" for x proof (cases "x \ set (map snd xs)") case True then obtain y where "(y, x) \ set xs" by auto with assms show ?thesis by (simp add: inverse_permutation_of_list_unique permutation_of_list_unique) next case False with assms show ?thesis by (auto simp: list_permutes_def inverse_permutation_of_list_id permutation_of_list_id) qed qed end