# HG changeset patch # User wenzelm # Date 1491065426 -7200 # Node ID e32eb488c3a3077d36956606ef190262ec3429de # Parent c82a1620b274c1a96c0b67619e094714167dc14f misc tuning and modernization; diff -r c82a1620b274 -r e32eb488c3a3 src/HOL/Library/Permutations.thy --- a/src/HOL/Library/Permutations.thy Sat Apr 01 15:35:32 2017 +0200 +++ b/src/HOL/Library/Permutations.thy Sat Apr 01 18:50:26 2017 +0200 @@ -5,39 +5,33 @@ section \Permutations, both general and specifically on finite sets.\ theory Permutations -imports Binomial Multiset Disjoint_Sets + imports Binomial Multiset Disjoint_Sets begin subsection \Transpositions\ -lemma swap_id_idempotent [simp]: - "Fun.swap a b id \ Fun.swap a b id = id" - by (rule ext, auto simp add: Fun.swap_def) +lemma swap_id_idempotent [simp]: "Fun.swap a b id \ Fun.swap a b id = id" + by (rule ext) (auto simp add: Fun.swap_def) -lemma inv_swap_id: - "inv (Fun.swap a b id) = Fun.swap a b id" +lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id" by (rule inv_unique_comp) simp_all -lemma swap_id_eq: - "Fun.swap a b id x = (if x = a then b else if x = b then a else x)" +lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)" by (simp add: Fun.swap_def) lemma bij_inv_eq_iff: "bij p \ x = inv p y \ p x = y" using surj_f_inv_f[of p] by (auto simp add: bij_def) lemma bij_swap_comp: - assumes bp: "bij p" + assumes "bij p" shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p" - using surj_f_inv_f[OF bij_is_surj[OF bp]] - by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF bp]) + using surj_f_inv_f[OF bij_is_surj[OF \bij p\]] + by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \bij p\]) -lemma bij_swap_compose_bij: "bij p \ bij (Fun.swap a b id \ p)" -proof - - assume H: "bij p" - show ?thesis - unfolding bij_swap_comp[OF H] bij_swap_iff - using H . -qed +lemma bij_swap_compose_bij: + assumes "bij p" + shows "bij (Fun.swap a b id \ p)" + by (simp only: bij_swap_comp[OF \bij p\] bij_swap_iff \bij p\) subsection \Basic consequences of the definition\ @@ -48,9 +42,8 @@ lemma permutes_in_image: "p permutes S \ p x \ S \ x \ S" unfolding permutes_def by metis -lemma permutes_not_in: - assumes "f permutes S" "x \ S" shows "f x = x" - using assms by (auto simp: permutes_def) +lemma permutes_not_in: "f permutes S \ x \ S \ f x = x" + by (auto simp: permutes_def) lemma permutes_image: "p permutes S \ p ` S = S" unfolding permutes_def @@ -63,46 +56,46 @@ unfolding permutes_def inj_def by blast lemma permutes_inj_on: "f permutes S \ inj_on f A" - unfolding permutes_def inj_on_def by auto + by (auto simp: permutes_def inj_on_def) lemma permutes_surj: "p permutes s \ surj p" unfolding permutes_def surj_def by metis lemma permutes_bij: "p permutes s \ bij p" -unfolding bij_def by (metis permutes_inj permutes_surj) + unfolding bij_def by (metis permutes_inj permutes_surj) lemma permutes_imp_bij: "p permutes S \ bij_betw p S S" -by (metis UNIV_I bij_betw_subset permutes_bij permutes_image subsetI) + by (metis UNIV_I bij_betw_subset permutes_bij permutes_image subsetI) lemma bij_imp_permutes: "bij_betw p S S \ (\x. x \ S \ p x = x) \ p permutes S" unfolding permutes_def bij_betw_def inj_on_def by auto (metis image_iff)+ lemma permutes_inv_o: - assumes pS: "p permutes S" + assumes permutes: "p permutes S" shows "p \ inv p = id" and "inv p \ p = id" - using permutes_inj[OF pS] permutes_surj[OF pS] + using permutes_inj[OF permutes] permutes_surj[OF permutes] unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+ lemma permutes_inverses: fixes p :: "'a \ 'a" - assumes pS: "p permutes S" + assumes permutes: "p permutes S" shows "p (inv p x) = x" and "inv p (p x) = x" - using permutes_inv_o[OF pS, unfolded fun_eq_iff o_def] by auto + using permutes_inv_o[OF permutes, unfolded fun_eq_iff o_def] by auto lemma permutes_subset: "p permutes S \ S \ T \ p permutes T" unfolding permutes_def by blast lemma permutes_empty[simp]: "p permutes {} \ p = id" - unfolding fun_eq_iff permutes_def by simp metis + by (auto simp add: fun_eq_iff permutes_def) lemma permutes_sing[simp]: "p permutes {a} \ p = id" - unfolding fun_eq_iff permutes_def by simp metis + by (simp add: fun_eq_iff permutes_def) metis (*somewhat slow*) lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)" - unfolding permutes_def by simp + by (simp add: permutes_def) lemma permutes_inv_eq: "p permutes S \ inv p y = x \ p x = y" unfolding permutes_def inv_def @@ -124,42 +117,44 @@ (* Next three lemmas contributed by Lukas Bulwahn *) lemma permutes_bij_inv_into: - fixes A :: "'a set" and B :: "'b set" + fixes A :: "'a set" + and B :: "'b set" assumes "p permutes A" - assumes "bij_betw f A B" + 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) - 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 + 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" - from this show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto + then 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) + 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) + 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 have "image_mset p (mset_set A) = mset_set A" - using permutes_image_mset by blast - from this show ?thesis by simp + 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 @@ -168,36 +163,41 @@ subsection \Group properties\ lemma permutes_id: "id permutes S" - unfolding permutes_def by simp + by (simp add: permutes_def) lemma permutes_compose: "p permutes S \ q permutes S \ q \ p permutes S" unfolding permutes_def o_def by metis lemma permutes_inv: - assumes pS: "p permutes S" + assumes "p permutes S" shows "inv p permutes S" - using pS unfolding permutes_def permutes_inv_eq[OF pS] by metis + using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis lemma permutes_inv_inv: - assumes pS: "p permutes S" + assumes "p permutes S" shows "inv (inv p) = p" - unfolding fun_eq_iff permutes_inv_eq[OF pS] permutes_inv_eq[OF permutes_inv[OF pS]] + 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'" + and inv: "\x. x \ S \ p' (p x) = x" + and outside: "\x. x \ S \ p' x = x" + shows "inv p = p'" proof - fix x show "inv p x = p' x" + show "inv p x = p' x" for x proof (cases "x \ S") - assume [simp]: "x \ S" - from assms have "p' x = p' (p (inv p x))" by (simp add: permutes_inverses) - also from permutes_inv[OF perm] - have "\ = inv p x" by (subst inv) (simp_all add: permutes_in_image) - finally show "inv p x = p' x" .. - qed (insert permutes_inv[OF perm], simp_all add: outside permutes_not_in) + 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" @@ -207,58 +207,54 @@ subsection \The number of permutations on a finite set\ lemma permutes_insert_lemma: - assumes pS: "p permutes (insert a S)" + 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 pS]) + apply (rule permutes_compose[OF assms]) apply (rule permutes_swap_id, simp) - using permutes_in_image[OF pS, of a] + 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}}" + (\(b, p). Fun.swap a b id \ p) ` {(b, p). b \ insert a S \ p \ {p. p permutes S}}" proof - - { - fix p - { - assume pS: "p permutes insert a S" + 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 th0: "p = Fun.swap a ?b id \ ?q" - unfolding fun_eq_iff o_assoc by simp - have th1: "?b \ insert a S" - unfolding permutes_in_image[OF pS] by simp - from permutes_insert_lemma[OF pS] th0 th1 - have "\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S" by blast - } - moreover - { - fix b q - assume bq: "p = Fun.swap a b id \ q" "b \ insert a S" "q permutes S" - from permutes_subset[OF bq(3), of "insert a S"] - have qS: "q permutes insert a S" + 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 aS: "a \ insert a S" + have a: "a \ insert a S" by simp - from bq(1) permutes_compose[OF qS permutes_swap_id[OF aS bq(2)]] - have "p permutes insert a S" + from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis by simp - } - ultimately have "p permutes insert a S \ - (\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S)" - by blast - } - then show ?thesis - by auto + qed + ultimately show ?thesis by blast + qed + then show ?thesis by auto qed lemma card_permutations: - assumes Sn: "card S = n" - and fS: "finite S" + assumes "card S = n" + and "finite S" shows "card {p. p permutes S} = fact n" - using fS Sn + using assms(2,1) proof (induct arbitrary: n) case empty then show ?case by simp @@ -266,21 +262,20 @@ case (insert x F) { fix n - assume H0: "card (insert x F) = 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)" - from permutes_insert[of x F] - have xfgpF': "?xF = ?g ` ?pF'" . - have Fs: "card F = n - 1" - using \x \ F\ H0 \finite F\ by auto - from insert.hyps Fs have pFs: "card ?pF = fact (n - 1)" - using \finite F\ by auto + 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) - then have pF'f: "finite ?pF'" - using H0 \finite F\ + with \finite F\ card_insert have pF'f: "finite ?pF'" apply (simp only: Collect_case_prod Collect_mem_eq) apply (rule finite_cartesian_product) apply simp_all @@ -290,64 +285,54 @@ 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 ths: "b \ insert x F" "c \ insert x F" "x \ insert x F" - "p permutes F" "q permutes F" + 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 ths(4) \x \ F\ eq have "b = ?g (b,p) x" - unfolding permutes_def - by (auto simp add: Fun.swap_def fun_upd_def fun_eq_iff) - also have "\ = ?g (c,q) x" - using ths(5) \x \ F\ eq - by (auto simp add: swap_def fun_upd_def fun_eq_iff) - also have "\ = c" - using ths(5) \x \ F\ - unfolding permutes_def - by (auto simp add: Fun.swap_def fun_upd_def fun_eq_iff) - finally have bc: "b = c" . + 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: swap_def 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)" + 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 bc have "(b, p) = (c, q)" + with \b = c\ have "(b, p) = (c, q)" by simp } then show ?thesis unfolding inj_on_def by blast qed - from \x \ F\ H0 have n0: "n \ 0" - using \finite F\ by auto + 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[simp]: "n = Suc m" + then obtain m where n: "n = Suc m" by blast - from pFs H0 have xFc: "card ?xF = fact n" + from pFs card_insert have *: "card ?xF = fact n" unfolding xfgpF' card_image[OF ginj] using \finite F\ \finite ?pF\ - apply (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) - apply simp - done + 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" - unfolding xfgpF' by simp - have "card ?xF = fact n" - using xFf xFc unfolding xFf by blast + by (simp add: xfgpF' n) + from * have "card ?xF = fact n" + unfolding xFf by blast } - then show ?case - using insert by simp + with insert show ?case by simp qed lemma finite_permutations: - assumes fS: "finite S" + assumes "finite S" shows "finite {p. p permutes S}" - using card_permutations[OF refl fS] - by (auto intro: card_ge_0_finite) + using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite) subsection \Permutations of index set for iterated operations\ @@ -387,9 +372,9 @@ 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)" + 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] @@ -410,12 +395,16 @@ done lemma permutation_swap_id: "permutation (Fun.swap a b id)" - apply (cases "a = b") - apply simp_all - unfolding permutation_def - using swapidseq_swap[of a b] - apply blast - done +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) @@ -423,13 +412,13 @@ then show ?case by simp next case (comp_Suc n p a b m q) - have th: "Suc n + m = Suc (n + m)" + have eq: "Suc n + m = Suc (n + m)" by arith show ?case - unfolding th comp_assoc + 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+ + apply blast+ done qed @@ -437,10 +426,8 @@ 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)" - apply (induct n p rule: swapidseq.induct) - using swapidseq_swap[of a b] - apply (auto simp add: comp_assoc intro: swapidseq.comp_Suc) - done + 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) @@ -453,27 +440,27 @@ by blast let ?q = "q \ Fun.swap a b id" note H = comp_Suc.hyps - from swapidseq_swap[of a b] H(3) have th0: "swapidseq 1 (Fun.swap a b id)" + 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) th0] have th1: "swapidseq (Suc n) ?q" + 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 th2: "Fun.swap a b id \ p \ ?q = id" . + 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 th1 th2 show ?case + with ** *** show ?case by blast qed lemma swapidseq_inverse: - assumes H: "swapidseq n p" + assumes "swapidseq n p" shows "swapidseq n (inv p)" - using swapidseq_inverse_exists[OF H] inv_unique_comp[of p] by auto + 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 @@ -494,61 +481,60 @@ (\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 H: "a \ b" "c \ d" + 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 (simp_all only: swap_commute) apply (case_tac "a = c \ b = d") - apply (clarsimp simp only: swap_commute swap_id_idempotent) + 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 (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="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 H show ?thesis by metis + 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 + 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)" + 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 (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 spn: "swapidseq n p" - and ab: "a \ b" - and pa: "(Fun.swap a b id \ p) a = a" + 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 spn ab pa + using assms proof (induct n arbitrary: p a b) case 0 then show ?case @@ -559,49 +545,44 @@ 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 - { - assume H: "Fun.swap a b id \ Fun.swap c d id = id" - have ?case by (simp only: cdqm o_assoc H) (simp add: cdqm) - } - moreover - { - fix x y z - assume H: "x \ a" "y \ a" "z \ a" "x \ y" + 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" - from H have az: "a \ z" + 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 - - { - fix h - have "(Fun.swap x y id \ h) a = a \ h a = a" - using H by (simp add: Fun.swap_def) - } - note th3 = this + 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 H) + 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 th1: "(Fun.swap a z id \ q) a = a" - unfolding th3 . - from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az th1] - have th2: "swapidseq (n - 1) (Fun.swap a z id \ q)" "n \ 0" + 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+ - have th: "Suc n - 1 = Suc (n - 1)" - using th2(2) by auto - have ?case - unfolding cdqm(2) H o_assoc th + 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 th2 H - apply blast+ + using ** prems + apply blast+ done - } - ultimately show ?case - using swap_general[OF Suc.prems(2) cdqm(4)] by metis + qed qed lemma swapidseq_identity_even: @@ -609,26 +590,24 @@ shows "even n" using \swapidseq n id\ proof (induct n rule: nat_less_induct) - fix n - assume H: "\m 'a) \ even m" "swapidseq n (id :: 'a \ 'a)" - { - assume "n = 0" - then have "even n" by presburger - } - moreover - { - fix a b :: 'a and q m - assume h: "n = Suc m" "(id :: 'a \ 'a) = Fun.swap a b id \ q" "swapidseq m q" "a \ b" + 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) have "even n" + from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis by presburger - } - ultimately show "even n" - using H(2)[unfolded swapidseq_cases[of n id]] by auto + qed qed @@ -641,11 +620,9 @@ 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" + 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 + from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis by arith qed @@ -655,9 +632,9 @@ shows "evenperm p = b" unfolding n[symmetric] evenperm_def apply (rule swapidseq_even_even[where p = p]) - apply (rule someI[where x = n]) + apply (rule someI[where x = n]) using p - apply blast+ + apply blast+ done @@ -670,29 +647,26 @@ by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap) lemma evenperm_comp: - assumes p: "permutation p" - and q:"permutation q" - shows "evenperm (p \ q) = (evenperm p = evenperm q)" + assumes "permutation p" "permutation q" + shows "evenperm (p \ q) \ evenperm p = evenperm q" proof - - from p q obtain n m where n: "swapidseq n p" and m: "swapidseq m q" + from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q" unfolding permutation_def by blast - note nm = swapidseq_comp_add[OF n m] - have th: "even (n + m) = (even n \ even m)" + have "even (n + m) \ (even n \ even m)" by arith from evenperm_unique[OF n refl] evenperm_unique[OF m refl] - evenperm_unique[OF nm th] - show ?thesis + and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis by blast qed lemma evenperm_inv: - assumes p: "permutation p" + assumes "permutation p" shows "evenperm (inv p) = evenperm p" proof - - from p obtain n where n: "swapidseq n p" + from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast - from evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]] - show ?thesis . + show ?thesis + by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]) qed @@ -701,67 +675,71 @@ lemma bij_iff: "bij f \ (\x. \!y. f y = x)" unfolding bij_def inj_def surj_def apply auto - apply metis + apply metis apply metis done lemma permutation_bijective: - assumes p: "permutation p" + assumes "permutation p" shows "bij p" proof - - from p obtain n where n: "swapidseq n p" + 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" + 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 + then show ?thesis + unfolding bij_iff apply (auto simp add: fun_eq_iff) apply metis done qed lemma permutation_finite_support: - assumes p: "permutation p" + assumes "permutation p" shows "finite {x. p x \ x}" proof - - from p obtain n where n: "swapidseq n p" + from assms obtain n where "swapidseq n p" unfolding permutation_def by blast - from n show ?thesis + 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 fS: "finite ?S" + from comp_Suc.hyps(2) have *: "finite ?S" by simp - from \a \ b\ have th: "{x. (Fun.swap a b id \ p) x \ x} \ ?S" - by (auto simp add: Fun.swap_def) - from finite_subset[OF th fS] show ?case . + 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 fS: "finite S" - and p: "bij p" - and pS: "\x. x\ S \ p x = x" + assumes "finite S" + and "bij p" + and "\x. x\ S \ p x = x" shows "permutation p" - using fS p pS + using assms proof (induct S arbitrary: p rule: finite_induct) - case (empty p) - then show ?case by simp + 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 raa: "?r a = a" + have *: "?r a = a" by (simp add: Fun.swap_def) - from bij_swap_compose_bij[OF insert(4)] have br: "bij ?r" . - from insert raa have th: "\x. x \ F \ ?r x = x" + from insert * have **: "\x. x \ F \ ?r x = x" by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3)) - from insert(3)[OF br th] have rp: "permutation ?r" . - have "permutation ?q" - by (simp add: permutation_compose permutation_swap_id rp) + 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 @@ -769,8 +747,8 @@ lemma permutation: "permutation p \ bij p \ finite {x. p x \ x}" (is "?lhs \ ?b \ ?f") proof - assume p: ?lhs - from p permutation_bijective permutation_finite_support show "?b \ ?f" + assume ?lhs + with permutation_bijective permutation_finite_support show "?b \ ?f" by auto next assume "?b \ ?f" @@ -780,11 +758,10 @@ qed lemma permutation_inverse_works: - assumes p: "permutation p" + assumes "permutation p" shows "inv p \ p = id" and "p \ inv p = id" - using permutation_bijective [OF p] - unfolding bij_def inj_iff surj_iff by auto + using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff) lemma permutation_inverse_compose: assumes p: "permutation p" @@ -797,33 +774,34 @@ by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) - finally have th0: "p \ q \ (inv q \ inv p) = id" . + 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 th1: "inv q \ inv p \ (p \ q) = id" . - from inv_unique_comp[OF th0 th1] show ?thesis . + finally have **: "inv q \ inv p \ (p \ q) = id" . + show ?thesis + by (rule inv_unique_comp[OF * **]) qed -subsection \Relation to "permutes"\ +subsection \Relation to \permutes\\ lemma permutation_permutes: "permutation p \ (\S. finite S \ p permutes S)" unfolding permutation permutes_def bij_iff[symmetric] apply (rule iffI, clarify) - apply (rule exI[where x="{x. p x \ x}"]) - apply simp + apply (rule exI[where x="{x. p x \ x}"]) + apply simp apply clarsimp apply (rule_tac B="S" in finite_subset) - apply auto + apply auto done subsection \Hence a sort of induction principle composing by swaps\ lemma permutes_induct: "finite S \ P id \ - (\ a b p. a \ S \ b \ S \ P p \ P p \ permutation p \ P (Fun.swap a b id \ p)) \ + (\a b p. a \ S \ b \ S \ P p \ P p \ permutation p \ P (Fun.swap a b id \ p)) \ (\p. p permutes S \ P p)" proof (induct S rule: finite_induct) case empty @@ -842,12 +820,11 @@ have xF: "x \ insert x F" by simp have rp: "permutation ?r" - unfolding permutation_permutes using insert.hyps(1) - permutes_insert_lemma[OF insert.prems(3)] + unfolding permutation_permutes + using insert.hyps(1) permutes_insert_lemma[OF insert.prems(3)] by blast - from insert.prems(2)[OF xF pxF Pr Pr rp] - show ?case - unfolding qp . + from insert.prems(2)[OF xF pxF Pr Pr rp] qp show ?case + by (simp only:) qed @@ -878,17 +855,17 @@ 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.. 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)" + shows "permute_list (f \ 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" @@ -910,44 +887,46 @@ by (simp add: id_def) lemma mset_permute_list [simp]: - 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} \ {..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 + 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 + from permutes_in_image[OF assms(1)] assms(2) have *: "f i < length ys \ 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 . @@ -955,20 +934,19 @@ 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 _)") + shows "map_of xs \ \ = map_of (map (\(x,y). (inv \ x, y)) xs)" + (is "_ = map_of (map ?f _)") proof - fix x - from assms have "inj \" "surj \" by (simp_all add: permutes_inj permutes_surj) - thus "(map_of xs \ \) x = map_of (map ?f xs) x" - by (induction xs) (auto simp: inv_f_f surj_f_inv_f) + 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\ -text \ - The following few lemmas were contributed by Lukas Bulwahn. -\ +text \The following few lemmas were contributed by Lukas Bulwahn.\ lemma count_image_mset_eq_card_vimage: assumes "finite A" @@ -980,19 +958,23 @@ 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 + 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 - 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 + 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 @@ -1000,123 +982,116 @@ 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)" + 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 - - have "f ` A = f ` (set_mset (mset_set A))" using \finite A\ by simp - also have "\ = f' ` (set_mset (mset_set A))" + 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 have "\ = f' ` A" using \finite A\ by simp + 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 - from this have "card {a \ A. f a = b} = card {a \ A. f' a = b}" - using \finite A\ + 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) - from this show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}" + 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 - hence "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" + then have "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" by (rule bchoice) - then guess p .. note p = this + 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)" - 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 (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 have "(\b\f ` A. {a \ A. f' a = b}) = A" using \f ` A = f' ` A\ by auto + 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 - fix x - assume "x \ A" - from this show "p' x = x" - unfolding p'_def by simp + 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 by (rule that) + ultimately show ?thesis .. qed -lemma mset_set_upto_eq_mset_upto: - "mset_set {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i\{..i\{..i \ S. p i \ i" + assumes "p permutes S" + and "\i \ S. p i \ i" shows "p = id" proof - - { - fix n - have "p n = n" - using p le - proof (induct n arbitrary: S rule: less_induct) - fix n S - assume H: - "\m S. m < n \ p permutes S \ \i\S. p i \ i \ p m = m" - "p permutes S" "\i \S. p i \ i" - { - assume "n \ S" - with H(2) have "p n = n" - unfolding permutes_def by metis - } - moreover - { - assume ns: "n \ S" - from H(3) ns have "p n < n \ p n = n" - by auto - moreover { - assume h: "p n < n" - from H h have "p (p n) = p n" - by metis - with permutes_inj[OF H(2)] have "p n = n" - unfolding inj_def by blast - with h have False - by simp - } - ultimately have "p n = n" - by blast - } - ultimately show "p n = n" - by blast + 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 - } - then show ?thesis - by (auto simp add: fun_eq_iff) + qed + then show ?thesis by (auto simp: fun_eq_iff) qed lemma permutes_natset_ge: @@ -1125,25 +1100,23 @@ and le: "\i \ S. p i \ i" shows "p = id" proof - - { - fix i - assume i: "i \ S" - from i permutes_in_image[OF permutes_inv[OF p]] have "inv p i \ S" + 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] have "i \ inv p i" + with permutes_inverses[OF p] show ?thesis by simp - } - then have th: "\i\S. inv p i \ i" + qed + then have "\i\S. inv p i \ i" by blast - from permutes_natset_le[OF permutes_inv[OF p] th] - have "inv p = inv id" + 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 + apply simp_all done qed @@ -1151,31 +1124,31 @@ apply (rule set_eqI) apply auto using permutes_inv_inv permutes_inv - apply auto + apply auto apply (rule_tac x="inv x" in exI) apply auto done lemma image_compose_permutations_left: - assumes q: "q permutes S" - shows "{q \ p | p. p permutes S} = {p . p permutes S}" + 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 q - 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: "q permutes S" + 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 q - 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 @@ -1183,12 +1156,11 @@ 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}" +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 th0: "inj_on inv ?S" + have *: "inj_on inv ?S" proof (auto simp add: inj_on_def) fix q r assume q: "q permutes S" @@ -1199,11 +1171,12 @@ with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r" by metis qed - have th1: "inv ` ?S = ?S" + have **: "inv ` ?S = ?S" using image_inverse_permutations by blast - have th2: "?rhs = sum (f \ inv) ?S" + have ***: "?rhs = sum (f \ inv) ?S" by (simp add: o_def) - from sum.reindex[OF th0, of f] show ?thesis unfolding th1 th2 . + from sum.reindex[OF *, of f] show ?thesis + by (simp only: ** ***) qed lemma setum_permutations_compose_left: @@ -1212,9 +1185,9 @@ (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" - have th0: "?rhs = sum (f \ (op \ q)) ?S" + have *: "?rhs = sum (f \ (op \ q)) ?S" by (simp add: o_def) - have th1: "inj_on (op \ q) ?S" + have **: "inj_on (op \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" @@ -1225,9 +1198,10 @@ with permutes_inj[OF q, unfolded inj_iff] show "p = r" by simp qed - have th3: "(op \ q) ` ?S = ?S" + have "(op \ q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto - from sum.reindex[OF th1, of f] show ?thesis unfolding th0 th1 th3 . + with * sum.reindex[OF **, of f] show ?thesis + by (simp only:) qed lemma sum_permutations_compose_right: @@ -1236,9 +1210,9 @@ (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" - have th0: "?rhs = sum (f \ (\p. p \ q)) ?S" + have *: "?rhs = sum (f \ (\p. p \ q)) ?S" by (simp add: o_def) - have th1: "inj_on (\p. p \ q) ?S" + have **: "inj_on (\p. p \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" @@ -1249,10 +1223,10 @@ with permutes_surj[OF q, unfolded surj_iff] show "p = r" by simp qed - have th3: "(\p. p \ q) ` ?S = ?S" - using image_compose_permutations_right[OF q] by auto - from sum.reindex[OF th1, of f] - show ?thesis unfolding th0 th1 th3 . + 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 @@ -1264,17 +1238,12 @@ 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 th0: "\f a b. (\(b,p). f (Fun.swap a b id \ p)) = f \ (\(b,p). Fun.swap a b id \ p)" + 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 th1: "\P Q. P \ Q = {(a,b). a \ P \ b \ Q}" - by blast - have th2: "\P Q. P \ (P \ Q) \ P \ Q" + have **: "\P Q. {(a, b). a \ P \ b \ Q} = P \ Q" by blast show ?thesis - unfolding permutes_insert - unfolding sum.cartesian_product - unfolding th1[symmetric] - unfolding th0 + 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}" @@ -1295,8 +1264,7 @@ 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 add: o_def) + unfolding o_assoc swap_id_idempotent by simp with bc have "b = c \ p = q" by blast } @@ -1308,48 +1276,53 @@ subsection \Constructing permutations from association lists\ -definition list_permutes 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)" +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" + 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) + with assms show ?thesis + by (simp add: list_permutes_def) qed -definition permutation_of_list where - "permutation_of_list xs x = (case map_of xs x of None \ x | Some y \ y)" +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')" + "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 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)" +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))" + 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)" + 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" + 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')"]) + 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" @@ -1357,13 +1330,13 @@ lemma inj_on_map_of': assumes "distinct (map snd xs)" - shows "inj_on (the \ map_of xs) (set (map fst 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) + (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)" + 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" @@ -1371,12 +1344,13 @@ lemma image_map_of': assumes "distinct (map fst xs)" - shows "(the \ map_of xs) ` set (map fst xs) = set (map snd 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 _") + 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) @@ -1384,12 +1358,12 @@ 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) + (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) + (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))" . @@ -1407,52 +1381,47 @@ "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: - assumes "x \ set (map fst xs)" - shows "permutation_of_list xs x = x" - using assms by (induction xs) (auto simp: permutation_of_list_Cons) +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': - assumes "distinct (map fst xs)" "(x, y) \ set xs" - shows "permutation_of_list xs x = y" - using assms by (induction xs) (force simp: permutation_of_list_Cons)+ + "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: - assumes "list_permutes xs A" "(x,y) \ set xs" - shows "permutation_of_list xs x = y" - using assms by (intro permutation_of_list_unique') (simp_all add: list_permutes_def) + "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: - assumes "x \ set (map snd xs)" - shows "inverse_permutation_of_list xs x = x" - using assms by (induction xs) auto + "x \ set (map snd xs) \ inverse_permutation_of_list xs x = x" + by (induct xs) auto lemma inverse_permutation_of_list_unique': - assumes "distinct (map snd xs)" "(x, y) \ set xs" - shows "inverse_permutation_of_list xs y = x" - using assms by (induction xs) (force simp: inverse_permutation_of_list.simps)+ + "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)+ lemma inverse_permutation_of_list_unique: - assumes "list_permutes xs A" "(x,y) \ set xs" - shows "inverse_permutation_of_list xs y = x" - using assms by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def) + "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: - assumes "list_permutes xs (A :: 'a set)" - shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)" + 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 -next - fix x - show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" + 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 force + 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) - qed (insert assms, auto simp: list_permutes_def - inverse_permutation_of_list_id permutation_of_list_id) + 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 -