# HG changeset patch # User chaieb # Date 1234197735 0 # Node ID cfab6a76aa13fbfb346d215eaa97e4e93f89e2e6 # Parent 018ac1fa1ed31914fc46d15a9b071cb201326d0d Permutations, both general and specifically on finite sets. diff -r 018ac1fa1ed3 -r cfab6a76aa13 src/HOL/Library/Permutations.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Permutations.thy Mon Feb 09 16:42:15 2009 +0000 @@ -0,0 +1,862 @@ +(* Title: Library/Permutations + ID: $Id: + Author: Amine Chaieb, University of Cambridge +*) + +header {* Permutations, both general and specifically on finite sets.*} + +theory Permutations +imports Main Finite_Cartesian_Product Parity +begin + + (* Why should I import Main just to solve the Typerep problem! *) + +definition permutes (infixr "permutes" 41) where + "(p permutes S) \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)" + +(* ------------------------------------------------------------------------- *) +(* Transpositions. *) +(* ------------------------------------------------------------------------- *) + +declare swap_self[simp] +lemma swapid_sym: "Fun.swap a b id = Fun.swap b a id" + by (auto simp add: expand_fun_eq swap_def fun_upd_def) +lemma swap_id_refl: "Fun.swap a a id = id" by simp +lemma swap_id_sym: "Fun.swap a b id = Fun.swap b a id" + by (rule ext, simp add: swap_def) +lemma swap_id_idempotent[simp]: "Fun.swap a b id o Fun.swap a b id = id" + by (rule ext, auto simp add: swap_def) + +lemma inv_unique_comp: assumes fg: "f o g = id" and gf: "g o f = id" + shows "inv f = g" + using fg gf inv_equality[of g f] by (auto simp add: expand_fun_eq) + +lemma inverse_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)" + by (simp add: swap_def) + +(* ------------------------------------------------------------------------- *) +(* Basic consequences of the definition. *) +(* ------------------------------------------------------------------------- *) + +lemma permutes_in_image: "p permutes S \ p x \ S \ x \ S" + unfolding permutes_def by metis + +lemma permutes_image: assumes pS: "p permutes S" shows "p ` S = S" + using pS + unfolding permutes_def + apply - + apply (rule set_ext) + apply (simp add: image_iff) + apply metis + done + +lemma permutes_inj: "p permutes S ==> inj p " + unfolding permutes_def inj_on_def by blast + +lemma permutes_surj: "p permutes s ==> surj p" + unfolding permutes_def surj_def by metis + +lemma permutes_inv_o: assumes pS: "p permutes S" + shows " p o inv p = id" + and "inv p o p = id" + using permutes_inj[OF pS] permutes_surj[OF pS] + unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+ + + +lemma permutes_inverses: + fixes p :: "'a \ 'a" + assumes pS: "p permutes S" + shows "p (inv p x) = x" + and "inv p (p x) = x" + using permutes_inv_o[OF pS, unfolded expand_fun_eq 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 expand_fun_eq permutes_def apply simp by metis + +lemma permutes_sing[simp]: "p permutes {a} \ p = id" + unfolding expand_fun_eq permutes_def apply simp by metis + +lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)" + unfolding permutes_def by simp + +lemma permutes_inv_eq: "p permutes S ==> inv p y = x \ p x = y" + unfolding permutes_def inv_def apply auto + apply (erule allE[where x=y]) + apply (erule allE[where x=y]) + apply (rule someI_ex) apply blast + apply (rule some1_equality) + apply blast + apply blast + done + +lemma permutes_swap_id: "a \ S \ b \ S ==> Fun.swap a b id permutes S" + unfolding permutes_def swap_def fun_upd_def apply auto apply metis done + +lemma permutes_superset: "p permutes S \ (\x \ S - T. p x = x) \ p permutes T" +apply (simp add: Ball_def permutes_def Diff_iff) by metis + +(* ------------------------------------------------------------------------- *) +(* Group properties. *) +(* ------------------------------------------------------------------------- *) + +lemma permutes_id: "id permutes S" unfolding permutes_def by simp + +lemma permutes_compose: "p permutes S \ q permutes S ==> q o p permutes S" + unfolding permutes_def o_def by metis + +lemma permutes_inv: assumes pS: "p permutes S" shows "inv p permutes S" + using pS unfolding permutes_def permutes_inv_eq[OF pS] by metis + +lemma permutes_inv_inv: assumes pS: "p permutes S" shows "inv (inv p) = p" + unfolding expand_fun_eq permutes_inv_eq[OF pS] permutes_inv_eq[OF permutes_inv[OF pS]] + by blast + +(* ------------------------------------------------------------------------- *) +(* The number of permutations on a finite set. *) +(* ------------------------------------------------------------------------- *) + +lemma permutes_insert_lemma: + assumes pS: "p permutes (insert a S)" + shows "Fun.swap a (p a) id o p permutes S" + apply (rule permutes_superset[where S = "insert a S"]) + apply (rule permutes_compose[OF pS]) + apply (rule permutes_swap_id, simp) + using permutes_in_image[OF pS, of a] apply simp + apply (auto simp add: Ball_def Diff_iff swap_def) + done + +lemma permutes_insert: "{p. p permutes (insert a S)} = + (\(b,p). Fun.swap a b id o p) ` {(b,p). b \ insert a S \ p \ {p. p permutes S}}" +proof- + + {fix p + {assume pS: "p permutes insert a S" + let ?b = "p a" + let ?q = "Fun.swap a (p a) id o p" + have th0: "p = Fun.swap a ?b id o ?q" unfolding expand_fun_eq 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 o q \ b \ insert a S \ q permutes S" by blast} + moreover + {fix b q assume bq: "p = Fun.swap a b id o 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" by auto + have aS: "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" by simp } + ultimately have "p permutes insert a S \ (\ b q. p = Fun.swap a b id o q \ b \ insert a S \ q permutes S)" by blast} + thus ?thesis by auto +qed + +lemma hassize_insert: "a \ F \ insert a F hassize n \ F hassize (n - 1)" + by (auto simp add: hassize_def) + +lemma hassize_permutations: assumes Sn: "S hassize n" + shows "{p. p permutes S} hassize (fact n)" +proof- + from Sn have fS:"finite S" by (simp add: hassize_def) + + have "\n. (S hassize n) \ ({p. p permutes S} hassize (fact n))" + proof(rule finite_induct[where F = S]) + from fS show "finite S" . + next + show "\n. ({} hassize n) \ ({p. p permutes {}} hassize fact n)" + by (simp add: hassize_def permutes_empty) + next + fix x F + assume fF: "finite F" and xF: "x \ F" + and H: "\n. (F hassize n) \ ({p. p permutes F} hassize fact n)" + {fix n assume H0: "insert x F hassize 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'" . + from hassize_insert[OF xF H0] have Fs: "F hassize (n - 1)" . + from H Fs have pFs: "?pF hassize fact (n - 1)" by blast + hence pF'f: "finite ?pF'" using H0 unfolding hassize_def + apply (simp only: Collect_split 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'" and cq: "(c,q) \ ?pF'" + and 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" by auto + from ths(4) xF eq have "b = ?g (b,p) x" unfolding permutes_def + by (auto simp add: swap_def fun_upd_def expand_fun_eq) + also have "\ = ?g (c,q) x" using ths(5) xF eq + by (auto simp add: swap_def fun_upd_def expand_fun_eq) + also have "\ = c"using ths(5) xF unfolding permutes_def + by (auto simp add: swap_def fun_upd_def expand_fun_eq) + finally have bc: "b = c" . + hence "Fun.swap x b id = Fun.swap x c id" by simp + with eq have "Fun.swap x b id o p = Fun.swap x b id o q" by simp + hence "Fun.swap x b id o (Fun.swap x b id o p) = Fun.swap x b id o (Fun.swap x b id o q)" by simp + hence "p = q" by (simp add: o_assoc) + with bc have "(b,p) = (c,q)" by simp } + thus ?thesis unfolding inj_on_def by blast + qed + from xF H0 have n0: "n \ 0 " by (auto simp add: hassize_def) + hence "\m. n = Suc m" by arith + then obtain m where n[simp]: "n = Suc m" by blast + from pFs H0 have xFc: "card ?xF = fact n" + unfolding xfgpF' card_image[OF ginj] hassize_def + apply (simp only: Collect_split Collect_mem_eq card_cartesian_product) + by simp + from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" unfolding xfgpF' by simp + have "?xF hassize fact n" + using xFf xFc + unfolding hassize_def xFf by blast } + thus "\n. (insert x F hassize n) \ ({p. p permutes insert x F} hassize fact n)" + by blast + qed + with Sn show ?thesis by blast +qed + +lemma finite_permutations: "finite S ==> finite {p. p permutes S}" + using hassize_permutations[of S] unfolding hassize_def by blast + +(* ------------------------------------------------------------------------- *) +(* Permutations of index set for iterated operations. *) +(* ------------------------------------------------------------------------- *) + +lemma (in ab_semigroup_mult) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S" + shows "fold_image times f z S = fold_image times (f o p) z S" + using fold_image_reindex[OF fS subset_inj_on[OF permutes_inj[OF pS], of S, simplified], of f z] + unfolding permutes_image[OF pS] . +lemma (in ab_semigroup_add) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S" + shows "fold_image plus f z S = fold_image plus (f o p) z S" +proof- + interpret ab_semigroup_mult plus apply unfold_locales apply (simp add: add_assoc) + apply (simp add: add_commute) done + from fold_image_reindex[OF fS subset_inj_on[OF permutes_inj[OF pS], of S, simplified], of f z] + show ?thesis + unfolding permutes_image[OF pS] . +qed + +lemma setsum_permute: assumes pS: "p permutes S" + shows "setsum f S = setsum (f o p) S" + unfolding setsum_def using fold_image_permute[of S p f 0] pS by clarsimp + +lemma setsum_permute_natseg:assumes pS: "p permutes {m .. n}" + shows "setsum f {m .. n} = setsum (f o p) {m .. n}" + using setsum_permute[OF pS, of f ] pS by blast + +lemma setprod_permute: assumes pS: "p permutes S" + shows "setprod f S = setprod (f o p) S" + unfolding setprod_def + using ab_semigroup_mult_class.fold_image_permute[of S p f 1] pS by clarsimp + +lemma setprod_permute_natseg:assumes pS: "p permutes {m .. n}" + shows "setprod f {m .. n} = setprod (f o p) {m .. n}" + using setprod_permute[OF pS, of f ] pS by blast + +(* ------------------------------------------------------------------------- *) +(* Various combinations of transpositions with 2, 1 and 0 common elements. *) +(* ------------------------------------------------------------------------- *) + +lemma swap_id_common:" a \ c \ b \ c \ Fun.swap a b id o Fun.swap a c id = Fun.swap b c id o Fun.swap a b id" by (simp add: expand_fun_eq swap_def) + +lemma swap_id_common': "~(a = b) \ ~(a = c) \ Fun.swap a c id o Fun.swap b c id = Fun.swap b c id o Fun.swap a b id" by (simp add: expand_fun_eq swap_def) + +lemma swap_id_independent: "~(a = c) \ ~(a = d) \ ~(b = c) \ ~(b = d) ==> Fun.swap a b id o Fun.swap c d id = Fun.swap c d id o Fun.swap a b id" + by (simp add: swap_def expand_fun_eq) + +(* ------------------------------------------------------------------------- *) +(* 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 o p)" + +declare id[unfolded id_def, simp] +definition "permutation p \ (\n. swapidseq n p)" + +(* ------------------------------------------------------------------------- *) +(* 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] by simp + +lemma permutation_swap_id: "permutation (Fun.swap a b id)" + apply (cases "a=b", simp_all) + unfolding permutation_def using swapidseq_swap[of a b] by blast + +lemma swapidseq_comp_add: "swapidseq n p \ swapidseq m q ==> swapidseq (n + m) (p o q)" + proof (induct n p arbitrary: m q rule: swapidseq.induct) + case (id m q) thus ?case by simp + next + case (comp_Suc n p a b m q) + have th: "Suc n + m = Suc (n + m)" by arith + show ?case unfolding th o_assoc[symmetric] + apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3) by blast+ +qed + +lemma permutation_compose: "permutation p \ permutation q ==> permutation(p o 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 o Fun.swap a b id)" + apply (induct n p rule: swapidseq.induct) + using swapidseq_swap[of a b] + by (auto simp add: o_assoc[symmetric] intro: swapidseq.comp_Suc) + +lemma swapidseq_inverse_exists: "swapidseq n p ==> \q. swapidseq n q \ p o q = id \ q o p = id" +proof(induct n p rule: swapidseq.induct) + case id thus ?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 o 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)" by simp + from swapidseq_comp_add[OF q(1) th0] have th1:"swapidseq (Suc n) ?q" by simp + have "Fun.swap a b id o p o ?q = Fun.swap a b id o (p o q) o 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 o p o ?q = id" . + have "?q \ (Fun.swap a b id \ p) = q \ (Fun.swap a b id o Fun.swap a b id) \ p" by (simp only: o_assoc) + hence "?q \ (Fun.swap a b id \ p) = id" by (simp add: q(3)) + with th1 th2 show ?case by blast +qed + + +lemma swapidseq_inverse: assumes H: "swapidseq n p" shows "swapidseq n (inv p)" + using swapidseq_inverse_exists[OF H] inv_unique_comp[of p] by auto + +lemma permutation_inverse: "permutation p ==> permutation (inv p)" + using permutation_def swapidseq_inverse by blast + +(* ------------------------------------------------------------------------- *) +(* The identity map only has even transposition sequences. *) +(* ------------------------------------------------------------------------- *) + +lemma symmetry_lemma:"(\a b c d. P a b c d ==> P a b d c) \ + (\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) + ==> (\a b c d. a \ b --> c \ d \ P a b c d)" by metis + +lemma swap_general: "a \ b \ c \ d \ Fun.swap a b id o Fun.swap c d id = id \ + (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id)" +proof- + assume H: "a\b" "c\d" +have "a \ b \ c \ d \ +( Fun.swap a b id o Fun.swap c d id = id \ + (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o 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: swapid_sym) + apply (case_tac "a = c \ b = d", clarsimp simp only: swapid_sym 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: expand_fun_eq 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: expand_fun_eq 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: expand_fun_eq swap_def) + done +with H 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 o 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 spn: "swapidseq n p" and ab: "a\b" and pa: "(Fun.swap a b id o p) a = a" + shows "n \ 0 \ swapidseq (n - 1) (Fun.swap a b id o p)" + using spn ab pa +proof(induct n arbitrary: p a b) + case 0 thus ?case by (auto simp add: 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 o q" "swapidseq m q" "c \ d" "n = m" + by auto + {assume H: "Fun.swap a b id o Fun.swap c d id = id" + + have ?case apply (simp only: cdqm o_assoc H) + by (simp add: cdqm)} + moreover + { fix x y z + assume H: "x\a" "y\a" "z \a" "x \y" + "Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id" + from H have az: "a \ z" by simp + + {fix h have "(Fun.swap x y id o h) a = a \ h a = a" + using H by (simp add: swap_def)} + note th3 = this + from cdqm(2) have "Fun.swap a b id o p = Fun.swap a b id o (Fun.swap c d id o q)" by simp + hence "Fun.swap a b id o p = Fun.swap x y id o (Fun.swap a z id o q)" by (simp add: o_assoc H) + hence "(Fun.swap a b id o p) a = (Fun.swap x y id o (Fun.swap a z id o q)) a" by simp + hence "(Fun.swap x y id o (Fun.swap a z id o q)) a = a" unfolding Suc by metis + hence th1: "(Fun.swap a z id o 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 o 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 + apply (simp only: Suc_not_Zero simp_thms o_assoc[symmetric]) + apply (rule comp_Suc) + using th2 H apply blast+ + done} + ultimately show ?case using swap_general[OF Suc.prems(2) cdqm(4)] by metis +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) + fix n + assume H: "\m 'a) \ even m" "swapidseq n (id :: 'a \ 'a)" + {assume "n = 0" hence "even n" by arith} + 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" + 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" apply arith done} + ultimately show "even n" using H(2)[unfolded swapidseq_cases[of n id]] by auto +qed + +(* ------------------------------------------------------------------------- *) +(* 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 by blast+ + +(* ------------------------------------------------------------------------- *) +(* And it has the expected composition properties. *) +(* ------------------------------------------------------------------------- *) + +lemma evenperm_id[simp]: "evenperm id = True" + apply (rule evenperm_unique[where n = 0]) by simp_all + +lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)" +apply (rule evenperm_unique[where n="if a = b then 0 else 1"]) +by (simp_all add: swapidseq_swap) + +lemma evenperm_comp: + assumes p: "permutation p" and q:"permutation q" + shows "evenperm (p o q) = (evenperm p = evenperm q)" +proof- + from p q 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)" by arith + from evenperm_unique[OF n refl] evenperm_unique[OF m refl] + evenperm_unique[OF nm th] + show ?thesis by blast +qed + +lemma evenperm_inv: assumes p: "permutation p" + shows "evenperm (inv p) = evenperm p" +proof- + from p 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 . +qed + +(* ------------------------------------------------------------------------- *) +(* A more abstract characterization of permutations. *) +(* ------------------------------------------------------------------------- *) + + +lemma bij_iff: "bij f \ (\x. \!y. f y = x)" + unfolding bij_def inj_on_def surj_def + apply auto + apply metis + apply metis + done + +lemma permutation_bijective: + assumes p: "permutation p" + shows "bij p" +proof- + from p 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 + thus ?thesis unfolding bij_iff apply (auto simp add: expand_fun_eq) apply metis done +qed + +lemma permutation_finite_support: assumes p: "permutation p" + shows "finite {x. p x \ x}" +proof- + from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast + from n show ?thesis + proof(induct n p rule: swapidseq.induct) + case id thus ?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" by simp + from `a \ b` have th: "{x. (Fun.swap a b id o p) x \ x} \ ?S" + by (auto simp add: swap_def) + from finite_subset[OF th fS] show ?case . +qed +qed + +lemma bij_inv_eq_iff: "bij p ==> x = inv p y \ p x = y" + using surj_f_inv_f[of p] inv_f_f[of f] by (auto simp add: bij_def) + +lemma bij_swap_comp: + assumes bp: "bij p" shows "Fun.swap a b id o p = Fun.swap (inv p a) (inv p b) p" + using surj_f_inv_f[OF bij_is_surj[OF bp]] + by (simp add: expand_fun_eq swap_def bij_inv_eq_iff[OF bp]) + +lemma bij_swap_ompose_bij: "bij p \ bij (Fun.swap a b id o p)" +proof- + assume H: "bij p" + show ?thesis + unfolding bij_swap_comp[OF H] bij_swap_iff + using H . +qed + +lemma permutation_lemma: + assumes fS: "finite S" and p: "bij p" and pS: "\x. x\ S \ p x = x" + shows "permutation p" +using fS p pS +proof(induct S arbitrary: p rule: finite_induct) + case (empty p) thus ?case by simp +next + case (insert a F p) + let ?r = "Fun.swap a (p a) id o p" + let ?q = "Fun.swap a (p a) id o ?r " + have raa: "?r a = a" by (simp add: swap_def) + from bij_swap_ompose_bij[OF insert(4)] + have br: "bij ?r" . + + from insert raa have th: "\x. x \ F \ ?r x = x" + apply (clarsimp simp add: swap_def) + apply (erule_tac x="x" in allE) + apply auto + unfolding bij_iff apply metis + done + from insert(3)[OF br th] + have rp: "permutation ?r" . + have "permutation ?q" by (simp add: permutation_compose permutation_swap_id rp) + thus ?case by (simp add: o_assoc) +qed + +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" by auto +next + assume bf: "?b \ ?f" + hence bf: "?f" "?b" by blast+ + from permutation_lemma[OF bf] show ?lhs by blast +qed + +lemma permutation_inverse_works: assumes p: "permutation p" + shows "inv p o p = id" "p o inv p = id" +using permutation_bijective[OF p] surj_iff bij_def inj_iff by auto + +lemma permutation_inverse_compose: + assumes p: "permutation p" and q: "permutation q" + shows "inv (p o q) = inv q o inv p" +proof- + note ps = permutation_inverse_works[OF p] + note qs = permutation_inverse_works[OF q] + have "p o q o (inv q o inv p) = p o (q o inv q) o inv p" by (simp add: o_assoc) + also have "\ = id" by (simp add: ps qs) + finally have th0: "p o q o (inv q o inv p) = id" . + have "inv q o inv p o (p o q) = inv q o (inv p o p) o q" by (simp add: o_assoc) + also have "\ = id" by (simp add: ps qs) + finally have th1: "inv q o inv p o (p o q) = id" . + from inv_unique_comp[OF th0 th1] show ?thesis . +qed + +(* ------------------------------------------------------------------------- *) +(* 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 clarsimp +apply (rule_tac B="S" in finite_subset) +apply auto +done + +(* ------------------------------------------------------------------------- *) +(* 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 o p)) + ==> (\p. p permutes S ==> P p)" +proof(induct S rule: finite_induct) + case empty thus ?case by auto +next + case (insert x F p) + let ?r = "Fun.swap x (p x) id o p" + let ?q = "Fun.swap x (p x) id o ?r" + have qp: "?q = p" by (simp add: o_assoc) + from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r" by blast + from permutes_in_image[OF insert.prems(3), of x] + have pxF: "p x \ insert x F" by simp + 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)] by blast + from insert.prems(2)[OF xF pxF Pr Pr rp] + show ?case unfolding qp . +qed + +(* ------------------------------------------------------------------------- *) +(* Sign of a permutation as a real number. *) +(* ------------------------------------------------------------------------- *) + +definition "sign p = (if evenperm p then (1::int) else -1)" + +lemma sign_nz: "sign p \ 0" by (simp add: sign_def) +lemma sign_id: "sign id = 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 o 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)" + by (simp add: sign_def evenperm_swap) +lemma sign_idempotent: "sign p * sign p = 1" by (simp add: sign_def) + +(* ------------------------------------------------------------------------- *) +(* More lemmas about permutations. *) +(* ------------------------------------------------------------------------- *) + +lemma permutes_natset_le: + assumes p: "p permutes (S:: nat set)" and le: "\i \ S. p i <= i" shows "p = id" +proof- + {fix n + have "p n = n" + using p le + proof(induct n arbitrary: S rule: nat_less_induct) + fix n S assume H: "\ m< n. \S. 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_on_def by blast + with h have False by arith} + ultimately have "p n = n" by blast } + ultimately show "p n = n" by blast + qed} + thus ?thesis by (auto simp add: expand_fun_eq) +qed + +lemma permutes_natset_ge: + assumes p: "p permutes (S:: nat set)" 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" by simp + with le have "p (inv p i) \ inv p i" by blast + with permutes_inverses[OF p] have "i \ inv p i" by simp} + then have th: "\i\S. inv p i \ i" by blast + from permutes_natset_le[OF permutes_inv[OF p] th] + 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_ext) +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: "q permutes S" shows "{q o p | p. p permutes S} = {p . p permutes S}" +apply (rule set_ext) +apply auto +apply (rule permutes_compose) +using q apply auto +apply (rule_tac x = "inv q o x" in exI) +by (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o) + +lemma image_compose_permutations_right: + assumes q: "q permutes S" + shows "{p o q | p. p permutes S} = {p . p permutes S}" +apply (rule set_ext) +apply auto +apply (rule permutes_compose) +using q apply auto +apply (rule_tac x = "x o inv q" in exI) +by (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o o_assoc[symmetric]) + +lemma permutes_in_seg: "p permutes {1 ..n} \ i \ {1..n} ==> 1 <= p i \ p i <= n" + +apply (simp add: permutes_def) +apply metis +done + +term setsum +lemma setsum_permutations_inverse: "setsum f {p. p permutes {m..n}} = setsum (\p. f(inv p)) {p. p permutes {m..n}}" (is "?lhs = ?rhs") +proof- + let ?S = "{p . p permutes {m .. n}}" +have th0: "inj_on inv ?S" +proof(auto simp add: inj_on_def) + fix q r + assume q: "q permutes {m .. n}" and r: "r permutes {m .. n}" and qr: "inv q = inv r" + hence "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 th1: "inv ` ?S = ?S" using image_inverse_permutations by blast + have th2: "?rhs = setsum (f o inv) ?S" by (simp add: o_def) + from setsum_reindex[OF th0, of f] show ?thesis unfolding th1 th2 . +qed + +lemma setum_permutations_compose_left: + assumes q: "q permutes {m..n}" + shows "setsum f {p. p permutes {m..n}} = + setsum (\p. f(q o p)) {p. p permutes {m..n}}" (is "?lhs = ?rhs") +proof- + let ?S = "{p. p permutes {m..n}}" + have th0: "?rhs = setsum (f o (op o q)) ?S" by (simp add: o_def) + have th1: "inj_on (op o q) ?S" + apply (auto simp add: inj_on_def) + proof- + fix p r + assume "p permutes {m..n}" and r:"r permutes {m..n}" and rp: "q \ p = q \ r" + hence "inv q o q o p = inv q o q o r" by (simp add: o_assoc[symmetric]) + with permutes_inj[OF q, unfolded inj_iff] + + show "p = r" by simp + qed + have th3: "(op o q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto + from setsum_reindex[OF th1, of f] + show ?thesis unfolding th0 th1 th3 . +qed + +lemma sum_permutations_compose_right: + assumes q: "q permutes {m..n}" + shows "setsum f {p. p permutes {m..n}} = + setsum (\p. f(p o q)) {p. p permutes {m..n}}" (is "?lhs = ?rhs") +proof- + let ?S = "{p. p permutes {m..n}}" + have th0: "?rhs = setsum (f o (\p. p o q)) ?S" by (simp add: o_def) + have th1: "inj_on (\p. p o q) ?S" + apply (auto simp add: inj_on_def) + proof- + fix p r + assume "p permutes {m..n}" and r:"r permutes {m..n}" and rp: "p o q = r o q" + hence "p o (q o inv q) = r o (q o inv q)" by (simp add: o_assoc) + with permutes_surj[OF q, unfolded surj_iff] + + show "p = r" by simp + qed + have th3: "(\p. p o q) ` ?S = ?S" using image_compose_permutations_right[OF q] by auto + from setsum_reindex[OF th1, of f] + show ?thesis unfolding th0 th1 th3 . +qed + +(* ------------------------------------------------------------------------- *) +(* Sum over a set of permutations (could generalize to iteration). *) +(* ------------------------------------------------------------------------- *) + +lemma setsum_over_permutations_insert: + assumes fS: "finite S" and aS: "a \ S" + shows "setsum f {p. p permutes (insert a S)} = setsum (\b. setsum (\q. f (Fun.swap a b id o q)) {p. p permutes S}) (insert a S)" +proof- + have th0: "\f a b. (\(b,p). f (Fun.swap a b id o p)) = f o (\(b,p). Fun.swap a b id o p)" + by (simp add: expand_fun_eq) + have th1: "\P Q. P \ Q = {(a,b). a \ P \ b \ Q}" by blast + have th2: "\P Q. P \ (P \ Q) \ P \ Q" by blast + show ?thesis + unfolding permutes_insert + unfolding setsum_cartesian_product + unfolding th1[symmetric] + unfolding th0 + proof(rule setsum_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" and c: "c \ insert a S" + and p: "p permutes S" and q: "q permutes S" + and eq: "Fun.swap a b id o p = Fun.swap a c id o 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 o p) a = (Fun.swap a c id o q) a" by simp + hence bc: "b = c" + apply (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong) + apply (cases "a = b", auto) + by (cases "b = c", auto) + from eq[unfolded bc] have "(\p. Fun.swap a c id o p) (Fun.swap a c id o p) = (\p. Fun.swap a c id o p) (Fun.swap a c id o q)" by simp + hence "p = q" unfolding o_assoc swap_id_idempotent + by (simp add: o_def) + with bc have "b = c \ p = q" by blast + } + + then show "inj_on ?f (insert a S \ ?P)" + unfolding inj_on_def + apply clarify by metis + qed +qed + +end