| author | haftmann | 
| Sun, 01 Nov 2020 16:54:49 +0100 | |
| changeset 72536 | 589645894305 | 
| parent 72304 | 6fdeef6d6335 | 
| child 73328 | ff24fe85ee57 | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Library/Permutations.thy  | 
2  | 
Author: Amine Chaieb, University of Cambridge  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
4  | 
|
| 60500 | 5  | 
section \<open>Permutations, both general and specifically on finite sets.\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
7  | 
theory Permutations  | 
| 
65552
 
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
 
wenzelm 
parents: 
65342 
diff
changeset
 | 
8  | 
imports Multiset Disjoint_Sets  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
10  | 
|
| 60500 | 11  | 
subsection \<open>Transpositions\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
12  | 
|
| 65342 | 13  | 
lemma swap_id_idempotent [simp]: "Fun.swap a b id \<circ> Fun.swap a b id = id"  | 
14  | 
by (rule ext) (auto simp add: Fun.swap_def)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
15  | 
|
| 65342 | 16  | 
lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id"  | 
| 54681 | 17  | 
by (rule inv_unique_comp) simp_all  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
18  | 
|
| 65342 | 19  | 
lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)"  | 
| 56545 | 20  | 
by (simp add: Fun.swap_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
21  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
22  | 
lemma bij_swap_comp:  | 
| 65342 | 23  | 
assumes "bij p"  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
24  | 
shows "Fun.swap a b id \<circ> p = Fun.swap (inv p a) (inv p b) p"  | 
| 65342 | 25  | 
using surj_f_inv_f[OF bij_is_surj[OF \<open>bij p\<close>]]  | 
26  | 
by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \<open>bij p\<close>])  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
27  | 
|
| 65342 | 28  | 
lemma bij_swap_compose_bij:  | 
29  | 
assumes "bij p"  | 
|
30  | 
shows "bij (Fun.swap a b id \<circ> p)"  | 
|
31  | 
by (simp only: bij_swap_comp[OF \<open>bij p\<close>] bij_swap_iff \<open>bij p\<close>)  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
32  | 
|
| 54681 | 33  | 
|
| 60500 | 34  | 
subsection \<open>Basic consequences of the definition\<close>  | 
| 54681 | 35  | 
|
36  | 
definition permutes (infixr "permutes" 41)  | 
|
37  | 
where "(p permutes S) \<longleftrightarrow> (\<forall>x. x \<notin> S \<longrightarrow> p x = x) \<and> (\<forall>y. \<exists>!x. p x = y)"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
39  | 
lemma permutes_in_image: "p permutes S \<Longrightarrow> p x \<in> S \<longleftrightarrow> x \<in> S"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
40  | 
unfolding permutes_def by metis  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
41  | 
|
| 65342 | 42  | 
lemma permutes_not_in: "f permutes S \<Longrightarrow> x \<notin> S \<Longrightarrow> f x = x"  | 
43  | 
by (auto simp: permutes_def)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
44  | 
|
| 54681 | 45  | 
lemma permutes_image: "p permutes S \<Longrightarrow> p ` S = S"  | 
| 30488 | 46  | 
unfolding permutes_def  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
47  | 
apply (rule set_eqI)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
48  | 
apply (simp add: image_iff)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
49  | 
apply metis  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
50  | 
done  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
51  | 
|
| 54681 | 52  | 
lemma permutes_inj: "p permutes S \<Longrightarrow> inj p"  | 
| 
64966
 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 
wenzelm 
parents: 
64543 
diff
changeset
 | 
53  | 
unfolding permutes_def inj_def by blast  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
54  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
55  | 
lemma permutes_inj_on: "f permutes S \<Longrightarrow> inj_on f A"  | 
| 65342 | 56  | 
by (auto simp: permutes_def inj_on_def)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
57  | 
|
| 54681 | 58  | 
lemma permutes_surj: "p permutes s \<Longrightarrow> surj p"  | 
| 30488 | 59  | 
unfolding permutes_def surj_def by metis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
60  | 
|
| 60601 | 61  | 
lemma permutes_bij: "p permutes s \<Longrightarrow> bij p"  | 
| 65342 | 62  | 
unfolding bij_def by (metis permutes_inj permutes_surj)  | 
| 60601 | 63  | 
|
| 59474 | 64  | 
lemma permutes_imp_bij: "p permutes S \<Longrightarrow> bij_betw p S S"  | 
| 65342 | 65  | 
by (metis UNIV_I bij_betw_subset permutes_bij permutes_image subsetI)  | 
| 
59669
 
de7792ea4090
renaming HOL/Fact.thy -> Binomial.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
59474 
diff
changeset
 | 
66  | 
|
| 59474 | 67  | 
lemma bij_imp_permutes: "bij_betw p S S \<Longrightarrow> (\<And>x. x \<notin> S \<Longrightarrow> p x = x) \<Longrightarrow> p permutes S"  | 
68  | 
unfolding permutes_def bij_betw_def inj_on_def  | 
|
69  | 
by auto (metis image_iff)+  | 
|
70  | 
||
| 54681 | 71  | 
lemma permutes_inv_o:  | 
| 65342 | 72  | 
assumes permutes: "p permutes S"  | 
| 54681 | 73  | 
shows "p \<circ> inv p = id"  | 
74  | 
and "inv p \<circ> p = id"  | 
|
| 65342 | 75  | 
using permutes_inj[OF permutes] permutes_surj[OF permutes]  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
76  | 
unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
77  | 
|
| 30488 | 78  | 
lemma permutes_inverses:  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
79  | 
fixes p :: "'a \<Rightarrow> 'a"  | 
| 65342 | 80  | 
assumes permutes: "p permutes S"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
81  | 
shows "p (inv p x) = x"  | 
| 54681 | 82  | 
and "inv p (p x) = x"  | 
| 65342 | 83  | 
using permutes_inv_o[OF permutes, unfolded fun_eq_iff o_def] by auto  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
84  | 
|
| 54681 | 85  | 
lemma permutes_subset: "p permutes S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> p permutes T"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
86  | 
unfolding permutes_def by blast  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
88  | 
lemma permutes_empty[simp]: "p permutes {} \<longleftrightarrow> p = id"
 | 
| 65342 | 89  | 
by (auto simp add: fun_eq_iff permutes_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
91  | 
lemma permutes_sing[simp]: "p permutes {a} \<longleftrightarrow> p = id"
 | 
| 65342 | 92  | 
by (simp add: fun_eq_iff permutes_def) metis (*somewhat slow*)  | 
| 30488 | 93  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
94  | 
lemma permutes_univ: "p permutes UNIV \<longleftrightarrow> (\<forall>y. \<exists>!x. p x = y)"  | 
| 65342 | 95  | 
by (simp add: permutes_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
96  | 
|
| 54681 | 97  | 
lemma permutes_inv_eq: "p permutes S \<Longrightarrow> inv p y = x \<longleftrightarrow> p x = y"  | 
98  | 
unfolding permutes_def inv_def  | 
|
99  | 
apply auto  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
100  | 
apply (erule allE[where x=y])  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
101  | 
apply (erule allE[where x=y])  | 
| 54681 | 102  | 
apply (rule someI_ex)  | 
103  | 
apply blast  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
104  | 
apply (rule some1_equality)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
105  | 
apply blast  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
106  | 
apply blast  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
107  | 
done  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
108  | 
|
| 54681 | 109  | 
lemma permutes_swap_id: "a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> Fun.swap a b id permutes S"  | 
| 56545 | 110  | 
unfolding permutes_def Fun.swap_def fun_upd_def by auto metis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
111  | 
|
| 54681 | 112  | 
lemma permutes_superset: "p permutes S \<Longrightarrow> (\<forall>x \<in> S - T. p x = x) \<Longrightarrow> p permutes T"  | 
113  | 
by (simp add: Ball_def permutes_def) metis  | 
|
114  | 
||
| 
69895
 
6b03a8cf092d
more formal contributors (with the help of the history);
 
wenzelm 
parents: 
67673 
diff
changeset
 | 
115  | 
lemma permutes_bij_inv_into: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>  | 
| 65342 | 116  | 
fixes A :: "'a set"  | 
117  | 
and B :: "'b set"  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
118  | 
assumes "p permutes A"  | 
| 65342 | 119  | 
and "bij_betw f A B"  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
120  | 
shows "(\<lambda>x. if x \<in> B then f (p (inv_into A f x)) else x) permutes B"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
121  | 
proof (rule bij_imp_permutes)  | 
| 65342 | 122  | 
from assms have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A"  | 
123  | 
by (auto simp add: permutes_imp_bij bij_betw_inv_into)  | 
|
124  | 
then have "bij_betw (f \<circ> p \<circ> inv_into A f) B B"  | 
|
125  | 
by (simp add: bij_betw_trans)  | 
|
126  | 
then show "bij_betw (\<lambda>x. if x \<in> B then f (p (inv_into A f x)) else x) B B"  | 
|
127  | 
by (subst bij_betw_cong[where g="f \<circ> p \<circ> inv_into A f"]) auto  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
128  | 
next  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
129  | 
fix x  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
130  | 
assume "x \<notin> B"  | 
| 65342 | 131  | 
then show "(if x \<in> B then f (p (inv_into A f x)) else x) = x" by auto  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
132  | 
qed  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
133  | 
|
| 
69895
 
6b03a8cf092d
more formal contributors (with the help of the history);
 
wenzelm 
parents: 
67673 
diff
changeset
 | 
134  | 
lemma permutes_image_mset: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
135  | 
assumes "p permutes A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
136  | 
shows "image_mset p (mset_set A) = mset_set A"  | 
| 65342 | 137  | 
using assms by (metis image_mset_mset_set bij_betw_imp_inj_on permutes_imp_bij permutes_image)  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
138  | 
|
| 
69895
 
6b03a8cf092d
more formal contributors (with the help of the history);
 
wenzelm 
parents: 
67673 
diff
changeset
 | 
139  | 
lemma permutes_implies_image_mset_eq: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
140  | 
assumes "p permutes A" "\<And>x. x \<in> A \<Longrightarrow> f x = f' (p x)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
141  | 
shows "image_mset f' (mset_set A) = image_mset f (mset_set A)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
142  | 
proof -  | 
| 65342 | 143  | 
have "f x = f' (p x)" if "x \<in># mset_set A" for x  | 
144  | 
using assms(2)[of x] that by (cases "finite A") auto  | 
|
145  | 
with assms have "image_mset f (mset_set A) = image_mset (f' \<circ> p) (mset_set A)"  | 
|
146  | 
by (auto intro!: image_mset_cong)  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
147  | 
also have "\<dots> = image_mset f' (image_mset p (mset_set A))"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
148  | 
by (simp add: image_mset.compositionality)  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
149  | 
also have "\<dots> = image_mset f' (mset_set A)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
150  | 
proof -  | 
| 65342 | 151  | 
from assms permutes_image_mset have "image_mset p (mset_set A) = mset_set A"  | 
152  | 
by blast  | 
|
153  | 
then show ?thesis by simp  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
154  | 
qed  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
155  | 
finally show ?thesis ..  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
156  | 
qed  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
157  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
158  | 
|
| 60500 | 159  | 
subsection \<open>Group properties\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
160  | 
|
| 54681 | 161  | 
lemma permutes_id: "id permutes S"  | 
| 65342 | 162  | 
by (simp add: permutes_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
163  | 
|
| 54681 | 164  | 
lemma permutes_compose: "p permutes S \<Longrightarrow> q permutes S \<Longrightarrow> q \<circ> p permutes S"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
165  | 
unfolding permutes_def o_def by metis  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
166  | 
|
| 54681 | 167  | 
lemma permutes_inv:  | 
| 65342 | 168  | 
assumes "p permutes S"  | 
| 54681 | 169  | 
shows "inv p permutes S"  | 
| 65342 | 170  | 
using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
171  | 
|
| 54681 | 172  | 
lemma permutes_inv_inv:  | 
| 65342 | 173  | 
assumes "p permutes S"  | 
| 54681 | 174  | 
shows "inv (inv p) = p"  | 
| 65342 | 175  | 
unfolding fun_eq_iff permutes_inv_eq[OF assms] permutes_inv_eq[OF permutes_inv[OF assms]]  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
176  | 
by blast  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
177  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
178  | 
lemma permutes_invI:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
179  | 
assumes perm: "p permutes S"  | 
| 65342 | 180  | 
and inv: "\<And>x. x \<in> S \<Longrightarrow> p' (p x) = x"  | 
181  | 
and outside: "\<And>x. x \<notin> S \<Longrightarrow> p' x = x"  | 
|
182  | 
shows "inv p = p'"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
183  | 
proof  | 
| 65342 | 184  | 
show "inv p x = p' x" for x  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
185  | 
proof (cases "x \<in> S")  | 
| 65342 | 186  | 
case True  | 
187  | 
from assms have "p' x = p' (p (inv p x))"  | 
|
188  | 
by (simp add: permutes_inverses)  | 
|
189  | 
also from permutes_inv[OF perm] True have "\<dots> = inv p x"  | 
|
190  | 
by (subst inv) (simp_all add: permutes_in_image)  | 
|
191  | 
finally show ?thesis ..  | 
|
192  | 
next  | 
|
193  | 
case False  | 
|
194  | 
with permutes_inv[OF perm] show ?thesis  | 
|
195  | 
by (simp_all add: outside permutes_not_in)  | 
|
196  | 
qed  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
197  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
198  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
199  | 
lemma permutes_vimage: "f permutes A \<Longrightarrow> f -` A = A"  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
200  | 
by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv])  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
201  | 
|
| 54681 | 202  | 
|
| 
66486
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
203  | 
subsection \<open>Mapping permutations with bijections\<close>  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
204  | 
|
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
205  | 
lemma bij_betw_permutations:  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
206  | 
assumes "bij_betw f A B"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
207  | 
shows "bij_betw (\<lambda>\<pi> x. if x \<in> B then f (\<pi> (inv_into A f x)) else x)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
208  | 
             {\<pi>. \<pi> permutes A} {\<pi>. \<pi> permutes B}" (is "bij_betw ?f _ _")
 | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
209  | 
proof -  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
210  | 
let ?g = "(\<lambda>\<pi> x. if x \<in> A then inv_into A f (\<pi> (f x)) else x)"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
211  | 
show ?thesis  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
212  | 
proof (rule bij_betw_byWitness [of _ ?g], goal_cases)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
213  | 
case 3  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
214  | 
show ?case using permutes_bij_inv_into[OF _ assms] by auto  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
215  | 
next  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
216  | 
case 4  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
217  | 
have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
218  | 
    {
 | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
219  | 
fix \<pi> assume "\<pi> permutes B"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
220  | 
from permutes_bij_inv_into[OF this bij_inv] and assms  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
221  | 
have "(\<lambda>x. if x \<in> A then inv_into A f (\<pi> (f x)) else x) permutes A"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
222  | 
by (simp add: inv_into_inv_into_eq cong: if_cong)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
223  | 
}  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
224  | 
from this show ?case by (auto simp: permutes_inv)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
225  | 
next  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
226  | 
case 1  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
227  | 
thus ?case using assms  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
228  | 
by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
229  | 
dest: bij_betwE)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
230  | 
next  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
231  | 
case 2  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
232  | 
moreover have "bij_betw (inv_into A f) B A"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
233  | 
by (intro bij_betw_inv_into assms)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
234  | 
ultimately show ?case using assms  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
235  | 
by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
236  | 
dest: bij_betwE)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
237  | 
qed  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
238  | 
qed  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
239  | 
|
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
240  | 
lemma bij_betw_derangements:  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
241  | 
assumes "bij_betw f A B"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
242  | 
shows "bij_betw (\<lambda>\<pi> x. if x \<in> B then f (\<pi> (inv_into A f x)) else x)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
243  | 
             {\<pi>. \<pi> permutes A \<and> (\<forall>x\<in>A. \<pi> x \<noteq> x)} {\<pi>. \<pi> permutes B \<and> (\<forall>x\<in>B. \<pi> x \<noteq> x)}" 
 | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
244  | 
(is "bij_betw ?f _ _")  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
245  | 
proof -  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
246  | 
let ?g = "(\<lambda>\<pi> x. if x \<in> A then inv_into A f (\<pi> (f x)) else x)"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
247  | 
show ?thesis  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
248  | 
proof (rule bij_betw_byWitness [of _ ?g], goal_cases)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
249  | 
case 3  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
250  | 
have "?f \<pi> x \<noteq> x" if "\<pi> permutes A" "\<And>x. x \<in> A \<Longrightarrow> \<pi> x \<noteq> x" "x \<in> B" for \<pi> x  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
251  | 
using that and assms by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
252  | 
inv_into_f_f inv_into_into permutes_imp_bij)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
253  | 
with permutes_bij_inv_into[OF _ assms] show ?case by auto  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
254  | 
next  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
255  | 
case 4  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
256  | 
have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
257  | 
have "?g \<pi> permutes A" if "\<pi> permutes B" for \<pi>  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
258  | 
using permutes_bij_inv_into[OF that bij_inv] and assms  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
259  | 
by (simp add: inv_into_inv_into_eq cong: if_cong)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
260  | 
moreover have "?g \<pi> x \<noteq> x" if "\<pi> permutes B" "\<And>x. x \<in> B \<Longrightarrow> \<pi> x \<noteq> x" "x \<in> A" for \<pi> x  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
261  | 
using that and assms by (metis bij_betwE bij_betw_imp_surj_on f_inv_into_f permutes_imp_bij)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
262  | 
ultimately show ?case by auto  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
263  | 
next  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
264  | 
case 1  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
265  | 
thus ?case using assms  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
266  | 
by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
267  | 
dest: bij_betwE)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
268  | 
next  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
269  | 
case 2  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
270  | 
moreover have "bij_betw (inv_into A f) B A"  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
271  | 
by (intro bij_betw_inv_into assms)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
272  | 
ultimately show ?case using assms  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
273  | 
by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
274  | 
dest: bij_betwE)  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
275  | 
qed  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
276  | 
qed  | 
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
277  | 
|
| 
 
ffaaa83543b2
Lemmas about analysis and permutations
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
65552 
diff
changeset
 | 
278  | 
|
| 60500 | 279  | 
subsection \<open>The number of permutations on a finite set\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
280  | 
|
| 30488 | 281  | 
lemma permutes_insert_lemma:  | 
| 65342 | 282  | 
assumes "p permutes (insert a S)"  | 
| 54681 | 283  | 
shows "Fun.swap a (p a) id \<circ> p permutes S"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
284  | 
apply (rule permutes_superset[where S = "insert a S"])  | 
| 65342 | 285  | 
apply (rule permutes_compose[OF assms])  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
286  | 
apply (rule permutes_swap_id, simp)  | 
| 65342 | 287  | 
using permutes_in_image[OF assms, of a]  | 
| 54681 | 288  | 
apply simp  | 
| 56545 | 289  | 
apply (auto simp add: Ball_def Fun.swap_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
290  | 
done  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
291  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
292  | 
lemma permutes_insert: "{p. p permutes (insert a S)} =
 | 
| 65342 | 293  | 
  (\<lambda>(b, p). Fun.swap a b id \<circ> p) ` {(b, p). b \<in> insert a S \<and> p \<in> {p. p permutes S}}"
 | 
| 54681 | 294  | 
proof -  | 
| 65342 | 295  | 
have "p permutes insert a S \<longleftrightarrow>  | 
296  | 
(\<exists>b q. p = Fun.swap a b id \<circ> q \<and> b \<in> insert a S \<and> q permutes S)" for p  | 
|
297  | 
proof -  | 
|
298  | 
have "\<exists>b q. p = Fun.swap a b id \<circ> q \<and> b \<in> insert a S \<and> q permutes S"  | 
|
299  | 
if p: "p permutes insert a S"  | 
|
300  | 
proof -  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
301  | 
let ?b = "p a"  | 
| 54681 | 302  | 
let ?q = "Fun.swap a (p a) id \<circ> p"  | 
| 65342 | 303  | 
have *: "p = Fun.swap a ?b id \<circ> ?q"  | 
304  | 
by (simp add: fun_eq_iff o_assoc)  | 
|
305  | 
have **: "?b \<in> insert a S"  | 
|
306  | 
unfolding permutes_in_image[OF p] by simp  | 
|
307  | 
from permutes_insert_lemma[OF p] * ** show ?thesis  | 
|
308  | 
by blast  | 
|
309  | 
qed  | 
|
310  | 
moreover have "p permutes insert a S"  | 
|
311  | 
if bq: "p = Fun.swap a b id \<circ> q" "b \<in> insert a S" "q permutes S" for b q  | 
|
312  | 
proof -  | 
|
313  | 
from permutes_subset[OF bq(3), of "insert a S"] have q: "q permutes insert a S"  | 
|
| 54681 | 314  | 
by auto  | 
| 65342 | 315  | 
have a: "a \<in> insert a S"  | 
| 54681 | 316  | 
by simp  | 
| 65342 | 317  | 
from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis  | 
| 54681 | 318  | 
by simp  | 
| 65342 | 319  | 
qed  | 
320  | 
ultimately show ?thesis by blast  | 
|
321  | 
qed  | 
|
322  | 
then show ?thesis by auto  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
323  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
324  | 
|
| 54681 | 325  | 
lemma card_permutations:  | 
| 65342 | 326  | 
assumes "card S = n"  | 
327  | 
and "finite S"  | 
|
| 33715 | 328  | 
  shows "card {p. p permutes S} = fact n"
 | 
| 65342 | 329  | 
using assms(2,1)  | 
| 54681 | 330  | 
proof (induct arbitrary: n)  | 
331  | 
case empty  | 
|
332  | 
then show ?case by simp  | 
|
| 33715 | 333  | 
next  | 
334  | 
case (insert x F)  | 
|
| 54681 | 335  | 
  {
 | 
336  | 
fix n  | 
|
| 72304 | 337  | 
assume card_insert: "card (insert x F) = n"  | 
| 33715 | 338  | 
    let ?xF = "{p. p permutes insert x F}"
 | 
339  | 
    let ?pF = "{p. p permutes F}"
 | 
|
340  | 
    let ?pF' = "{(b, p). b \<in> insert x F \<and> p \<in> ?pF}"
 | 
|
341  | 
let ?g = "(\<lambda>(b, p). Fun.swap x b id \<circ> p)"  | 
|
| 65342 | 342  | 
have xfgpF': "?xF = ?g ` ?pF'"  | 
343  | 
by (rule permutes_insert[of x F])  | 
|
| 72304 | 344  | 
from \<open>x \<notin> F\<close> \<open>finite F\<close> card_insert have Fs: "card F = n - 1"  | 
| 65342 | 345  | 
by auto  | 
346  | 
from \<open>finite F\<close> insert.hyps Fs have pFs: "card ?pF = fact (n - 1)"  | 
|
347  | 
by auto  | 
|
| 54681 | 348  | 
then have "finite ?pF"  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59669 
diff
changeset
 | 
349  | 
by (auto intro: card_ge_0_finite)  | 
| 
72302
 
d7d90ed4c74e
fixed some remarkably ugly proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
69895 
diff
changeset
 | 
350  | 
with \<open>finite F\<close> card.insert_remove have pF'f: "finite ?pF'"  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60601 
diff
changeset
 | 
351  | 
apply (simp only: Collect_case_prod Collect_mem_eq)  | 
| 33715 | 352  | 
apply (rule finite_cartesian_product)  | 
353  | 
apply simp_all  | 
|
354  | 
done  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
355  | 
|
| 33715 | 356  | 
have ginj: "inj_on ?g ?pF'"  | 
| 54681 | 357  | 
proof -  | 
| 33715 | 358  | 
      {
 | 
| 54681 | 359  | 
fix b p c q  | 
| 65342 | 360  | 
assume bp: "(b, p) \<in> ?pF'"  | 
361  | 
assume cq: "(c, q) \<in> ?pF'"  | 
|
362  | 
assume eq: "?g (b, p) = ?g (c, q)"  | 
|
363  | 
from bp cq have pF: "p permutes F" and qF: "q permutes F"  | 
|
| 54681 | 364  | 
by auto  | 
| 65342 | 365  | 
from pF \<open>x \<notin> F\<close> eq have "b = ?g (b, p) x"  | 
366  | 
by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff)  | 
|
367  | 
also from qF \<open>x \<notin> F\<close> eq have "\<dots> = ?g (c, q) x"  | 
|
368  | 
by (auto simp: swap_def fun_upd_def fun_eq_iff)  | 
|
369  | 
also from qF \<open>x \<notin> F\<close> have "\<dots> = c"  | 
|
370  | 
by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff)  | 
|
371  | 
finally have "b = c" .  | 
|
| 54681 | 372  | 
then have "Fun.swap x b id = Fun.swap x c id"  | 
373  | 
by simp  | 
|
374  | 
with eq have "Fun.swap x b id \<circ> p = Fun.swap x b id \<circ> q"  | 
|
375  | 
by simp  | 
|
| 65342 | 376  | 
then have "Fun.swap x b id \<circ> (Fun.swap x b id \<circ> p) = Fun.swap x b id \<circ> (Fun.swap x b id \<circ> q)"  | 
| 54681 | 377  | 
by simp  | 
378  | 
then have "p = q"  | 
|
379  | 
by (simp add: o_assoc)  | 
|
| 65342 | 380  | 
with \<open>b = c\<close> have "(b, p) = (c, q)"  | 
| 54681 | 381  | 
by simp  | 
| 33715 | 382  | 
}  | 
| 54681 | 383  | 
then show ?thesis  | 
384  | 
unfolding inj_on_def by blast  | 
|
| 33715 | 385  | 
qed  | 
| 72304 | 386  | 
from \<open>x \<notin> F\<close> \<open>finite F\<close> card_insert have "n \<noteq> 0"  | 
| 65342 | 387  | 
by auto  | 
| 54681 | 388  | 
then have "\<exists>m. n = Suc m"  | 
389  | 
by presburger  | 
|
| 65342 | 390  | 
then obtain m where n: "n = Suc m"  | 
| 54681 | 391  | 
by blast  | 
| 72304 | 392  | 
from pFs card_insert have *: "card ?xF = fact n"  | 
| 54681 | 393  | 
unfolding xfgpF' card_image[OF ginj]  | 
| 60500 | 394  | 
using \<open>finite F\<close> \<open>finite ?pF\<close>  | 
| 65342 | 395  | 
by (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) (simp add: n)  | 
| 54681 | 396  | 
from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF"  | 
| 65342 | 397  | 
by (simp add: xfgpF' n)  | 
398  | 
from * have "card ?xF = fact n"  | 
|
399  | 
unfolding xFf by blast  | 
|
| 33715 | 400  | 
}  | 
| 65342 | 401  | 
with insert show ?case by simp  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
402  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
403  | 
|
| 54681 | 404  | 
lemma finite_permutations:  | 
| 65342 | 405  | 
assumes "finite S"  | 
| 54681 | 406  | 
  shows "finite {p. p permutes S}"
 | 
| 65342 | 407  | 
using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
408  | 
|
| 54681 | 409  | 
|
| 60500 | 410  | 
subsection \<open>Permutations of index set for iterated operations\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
411  | 
|
| 51489 | 412  | 
lemma (in comm_monoid_set) permute:  | 
413  | 
assumes "p permutes S"  | 
|
| 54681 | 414  | 
shows "F g S = F (g \<circ> p) S"  | 
| 51489 | 415  | 
proof -  | 
| 60500 | 416  | 
from \<open>p permutes S\<close> have "inj p"  | 
| 54681 | 417  | 
by (rule permutes_inj)  | 
418  | 
then have "inj_on p S"  | 
|
419  | 
by (auto intro: subset_inj_on)  | 
|
420  | 
then have "F g (p ` S) = F (g \<circ> p) S"  | 
|
421  | 
by (rule reindex)  | 
|
| 60500 | 422  | 
moreover from \<open>p permutes S\<close> have "p ` S = S"  | 
| 54681 | 423  | 
by (rule permutes_image)  | 
424  | 
ultimately show ?thesis  | 
|
425  | 
by simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
426  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
427  | 
|
| 54681 | 428  | 
|
| 60500 | 429  | 
subsection \<open>Various combinations of transpositions with 2, 1 and 0 common elements\<close>  | 
| 54681 | 430  | 
|
431  | 
lemma swap_id_common:" a \<noteq> c \<Longrightarrow> b \<noteq> c \<Longrightarrow>  | 
|
432  | 
Fun.swap a b id \<circ> Fun.swap a c id = Fun.swap b c id \<circ> Fun.swap a b id"  | 
|
| 56545 | 433  | 
by (simp add: fun_eq_iff Fun.swap_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
434  | 
|
| 54681 | 435  | 
lemma swap_id_common': "a \<noteq> b \<Longrightarrow> a \<noteq> c \<Longrightarrow>  | 
436  | 
Fun.swap a c id \<circ> Fun.swap b c id = Fun.swap b c id \<circ> Fun.swap a b id"  | 
|
| 56545 | 437  | 
by (simp add: fun_eq_iff Fun.swap_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
438  | 
|
| 54681 | 439  | 
lemma swap_id_independent: "a \<noteq> c \<Longrightarrow> a \<noteq> d \<Longrightarrow> b \<noteq> c \<Longrightarrow> b \<noteq> d \<Longrightarrow>  | 
440  | 
Fun.swap a b id \<circ> Fun.swap c d id = Fun.swap c d id \<circ> Fun.swap a b id"  | 
|
| 56545 | 441  | 
by (simp add: fun_eq_iff Fun.swap_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
442  | 
|
| 54681 | 443  | 
|
| 60500 | 444  | 
subsection \<open>Permutations as transposition sequences\<close>  | 
| 54681 | 445  | 
|
446  | 
inductive swapidseq :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|
| 65342 | 447  | 
where  | 
448  | 
id[simp]: "swapidseq 0 id"  | 
|
449  | 
| comp_Suc: "swapidseq n p \<Longrightarrow> a \<noteq> b \<Longrightarrow> swapidseq (Suc n) (Fun.swap a b id \<circ> p)"  | 
|
| 54681 | 450  | 
|
451  | 
declare id[unfolded id_def, simp]  | 
|
452  | 
||
453  | 
definition "permutation p \<longleftrightarrow> (\<exists>n. swapidseq n p)"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
454  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
455  | 
|
| 60500 | 456  | 
subsection \<open>Some closure properties of the set of permutations, with lengths\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
457  | 
|
| 54681 | 458  | 
lemma permutation_id[simp]: "permutation id"  | 
459  | 
unfolding permutation_def by (rule exI[where x=0]) simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
461  | 
declare permutation_id[unfolded id_def, simp]  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
462  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
463  | 
lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
464  | 
apply clarsimp  | 
| 54681 | 465  | 
using comp_Suc[of 0 id a b]  | 
466  | 
apply simp  | 
|
467  | 
done  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
468  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
469  | 
lemma permutation_swap_id: "permutation (Fun.swap a b id)"  | 
| 65342 | 470  | 
proof (cases "a = b")  | 
471  | 
case True  | 
|
472  | 
then show ?thesis by simp  | 
|
473  | 
next  | 
|
474  | 
case False  | 
|
475  | 
then show ?thesis  | 
|
476  | 
unfolding permutation_def  | 
|
477  | 
using swapidseq_swap[of a b] by blast  | 
|
478  | 
qed  | 
|
479  | 
||
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
480  | 
|
| 54681 | 481  | 
lemma swapidseq_comp_add: "swapidseq n p \<Longrightarrow> swapidseq m q \<Longrightarrow> swapidseq (n + m) (p \<circ> q)"  | 
482  | 
proof (induct n p arbitrary: m q rule: swapidseq.induct)  | 
|
483  | 
case (id m q)  | 
|
484  | 
then show ?case by simp  | 
|
485  | 
next  | 
|
486  | 
case (comp_Suc n p a b m q)  | 
|
| 65342 | 487  | 
have eq: "Suc n + m = Suc (n + m)"  | 
| 54681 | 488  | 
by arith  | 
489  | 
show ?case  | 
|
| 65342 | 490  | 
apply (simp only: eq comp_assoc)  | 
| 54681 | 491  | 
apply (rule swapidseq.comp_Suc)  | 
492  | 
using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3)  | 
|
| 65342 | 493  | 
apply blast+  | 
| 54681 | 494  | 
done  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
495  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
496  | 
|
| 54681 | 497  | 
lemma permutation_compose: "permutation p \<Longrightarrow> permutation q \<Longrightarrow> permutation (p \<circ> q)"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
498  | 
unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
499  | 
|
| 54681 | 500  | 
lemma swapidseq_endswap: "swapidseq n p \<Longrightarrow> a \<noteq> b \<Longrightarrow> swapidseq (Suc n) (p \<circ> Fun.swap a b id)"  | 
| 65342 | 501  | 
by (induct n p rule: swapidseq.induct)  | 
502  | 
(use swapidseq_swap[of a b] in \<open>auto simp add: comp_assoc intro: swapidseq.comp_Suc\<close>)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
503  | 
|
| 54681 | 504  | 
lemma swapidseq_inverse_exists: "swapidseq n p \<Longrightarrow> \<exists>q. swapidseq n q \<and> p \<circ> q = id \<and> q \<circ> p = id"  | 
505  | 
proof (induct n p rule: swapidseq.induct)  | 
|
506  | 
case id  | 
|
507  | 
then show ?case  | 
|
508  | 
by (rule exI[where x=id]) simp  | 
|
| 30488 | 509  | 
next  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
510  | 
case (comp_Suc n p a b)  | 
| 54681 | 511  | 
from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id"  | 
512  | 
by blast  | 
|
513  | 
let ?q = "q \<circ> Fun.swap a b id"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
514  | 
note H = comp_Suc.hyps  | 
| 65342 | 515  | 
from swapidseq_swap[of a b] H(3) have *: "swapidseq 1 (Fun.swap a b id)"  | 
| 54681 | 516  | 
by simp  | 
| 65342 | 517  | 
from swapidseq_comp_add[OF q(1) *] have **: "swapidseq (Suc n) ?q"  | 
| 54681 | 518  | 
by simp  | 
519  | 
have "Fun.swap a b id \<circ> p \<circ> ?q = Fun.swap a b id \<circ> (p \<circ> q) \<circ> Fun.swap a b id"  | 
|
520  | 
by (simp add: o_assoc)  | 
|
521  | 
also have "\<dots> = id"  | 
|
522  | 
by (simp add: q(2))  | 
|
| 65342 | 523  | 
finally have ***: "Fun.swap a b id \<circ> p \<circ> ?q = id" .  | 
| 54681 | 524  | 
have "?q \<circ> (Fun.swap a b id \<circ> p) = q \<circ> (Fun.swap a b id \<circ> Fun.swap a b id) \<circ> p"  | 
525  | 
by (simp only: o_assoc)  | 
|
526  | 
then have "?q \<circ> (Fun.swap a b id \<circ> p) = id"  | 
|
527  | 
by (simp add: q(3))  | 
|
| 65342 | 528  | 
with ** *** show ?case  | 
| 54681 | 529  | 
by blast  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
530  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
531  | 
|
| 54681 | 532  | 
lemma swapidseq_inverse:  | 
| 65342 | 533  | 
assumes "swapidseq n p"  | 
| 54681 | 534  | 
shows "swapidseq n (inv p)"  | 
| 65342 | 535  | 
using swapidseq_inverse_exists[OF assms] inv_unique_comp[of p] by auto  | 
| 54681 | 536  | 
|
537  | 
lemma permutation_inverse: "permutation p \<Longrightarrow> permutation (inv p)"  | 
|
538  | 
using permutation_def swapidseq_inverse by blast  | 
|
539  | 
||
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
540  | 
|
| 60500 | 541  | 
subsection \<open>The identity map only has even transposition sequences\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
542  | 
|
| 54681 | 543  | 
lemma symmetry_lemma:  | 
544  | 
assumes "\<And>a b c d. P a b c d \<Longrightarrow> P a b d c"  | 
|
545  | 
and "\<And>a b c d. a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow>  | 
|
546  | 
a = c \<and> b = d \<or> a = c \<and> b \<noteq> d \<or> a \<noteq> c \<and> b = d \<or> a \<noteq> c \<and> a \<noteq> d \<and> b \<noteq> c \<and> b \<noteq> d \<Longrightarrow>  | 
|
547  | 
P a b c d"  | 
|
548  | 
shows "\<And>a b c d. a \<noteq> b \<longrightarrow> c \<noteq> d \<longrightarrow> P a b c d"  | 
|
549  | 
using assms by metis  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
550  | 
|
| 54681 | 551  | 
lemma swap_general: "a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow>  | 
552  | 
Fun.swap a b id \<circ> Fun.swap c d id = id \<or>  | 
|
553  | 
(\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and>  | 
|
554  | 
Fun.swap a b id \<circ> Fun.swap c d id = Fun.swap x y id \<circ> Fun.swap a z id)"  | 
|
555  | 
proof -  | 
|
| 65342 | 556  | 
assume neq: "a \<noteq> b" "c \<noteq> d"  | 
| 54681 | 557  | 
have "a \<noteq> b \<longrightarrow> c \<noteq> d \<longrightarrow>  | 
558  | 
(Fun.swap a b id \<circ> Fun.swap c d id = id \<or>  | 
|
559  | 
(\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and>  | 
|
560  | 
Fun.swap a b id \<circ> Fun.swap c d id = Fun.swap x y id \<circ> Fun.swap a z id))"  | 
|
561  | 
apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d])  | 
|
| 65342 | 562  | 
apply (simp_all only: swap_commute)  | 
| 54681 | 563  | 
apply (case_tac "a = c \<and> b = d")  | 
| 65342 | 564  | 
apply (clarsimp simp only: swap_commute swap_id_idempotent)  | 
| 54681 | 565  | 
apply (case_tac "a = c \<and> b \<noteq> d")  | 
| 65342 | 566  | 
apply (rule disjI2)  | 
567  | 
apply (rule_tac x="b" in exI)  | 
|
568  | 
apply (rule_tac x="d" in exI)  | 
|
569  | 
apply (rule_tac x="b" in exI)  | 
|
570  | 
apply (clarsimp simp add: fun_eq_iff Fun.swap_def)  | 
|
| 54681 | 571  | 
apply (case_tac "a \<noteq> c \<and> b = d")  | 
| 65342 | 572  | 
apply (rule disjI2)  | 
573  | 
apply (rule_tac x="c" in exI)  | 
|
574  | 
apply (rule_tac x="d" in exI)  | 
|
575  | 
apply (rule_tac x="c" in exI)  | 
|
576  | 
apply (clarsimp simp add: fun_eq_iff Fun.swap_def)  | 
|
| 54681 | 577  | 
apply (rule disjI2)  | 
578  | 
apply (rule_tac x="c" in exI)  | 
|
579  | 
apply (rule_tac x="d" in exI)  | 
|
580  | 
apply (rule_tac x="b" in exI)  | 
|
| 56545 | 581  | 
apply (clarsimp simp add: fun_eq_iff Fun.swap_def)  | 
| 54681 | 582  | 
done  | 
| 65342 | 583  | 
with neq show ?thesis by metis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
584  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
585  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
586  | 
lemma swapidseq_id_iff[simp]: "swapidseq 0 p \<longleftrightarrow> p = id"  | 
| 65342 | 587  | 
using swapidseq.cases[of 0 p "p = id"] by auto  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
588  | 
|
| 54681 | 589  | 
lemma swapidseq_cases: "swapidseq n p \<longleftrightarrow>  | 
| 65342 | 590  | 
n = 0 \<and> p = id \<or> (\<exists>a b q m. n = Suc m \<and> p = Fun.swap a b id \<circ> q \<and> swapidseq m q \<and> a \<noteq> b)"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
591  | 
apply (rule iffI)  | 
| 65342 | 592  | 
apply (erule swapidseq.cases[of n p])  | 
593  | 
apply simp  | 
|
594  | 
apply (rule disjI2)  | 
|
595  | 
apply (rule_tac x= "a" in exI)  | 
|
596  | 
apply (rule_tac x= "b" in exI)  | 
|
597  | 
apply (rule_tac x= "pa" in exI)  | 
|
598  | 
apply (rule_tac x= "na" in exI)  | 
|
599  | 
apply simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
600  | 
apply auto  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
601  | 
apply (rule comp_Suc, simp_all)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
602  | 
done  | 
| 54681 | 603  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
604  | 
lemma fixing_swapidseq_decrease:  | 
| 65342 | 605  | 
assumes "swapidseq n p"  | 
606  | 
and "a \<noteq> b"  | 
|
607  | 
and "(Fun.swap a b id \<circ> p) a = a"  | 
|
| 54681 | 608  | 
shows "n \<noteq> 0 \<and> swapidseq (n - 1) (Fun.swap a b id \<circ> p)"  | 
| 65342 | 609  | 
using assms  | 
| 54681 | 610  | 
proof (induct n arbitrary: p a b)  | 
611  | 
case 0  | 
|
612  | 
then show ?case  | 
|
| 56545 | 613  | 
by (auto simp add: Fun.swap_def fun_upd_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
614  | 
next  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
615  | 
case (Suc n p a b)  | 
| 54681 | 616  | 
from Suc.prems(1) swapidseq_cases[of "Suc n" p]  | 
617  | 
obtain c d q m where  | 
|
618  | 
cdqm: "Suc n = Suc m" "p = Fun.swap c d id \<circ> q" "swapidseq m q" "c \<noteq> d" "n = m"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
619  | 
by auto  | 
| 65342 | 620  | 
consider "Fun.swap a b id \<circ> Fun.swap c d id = id"  | 
621  | 
| x y z where "x \<noteq> a" "y \<noteq> a" "z \<noteq> a" "x \<noteq> y"  | 
|
| 54681 | 622  | 
"Fun.swap a b id \<circ> Fun.swap c d id = Fun.swap x y id \<circ> Fun.swap a z id"  | 
| 65342 | 623  | 
using swap_general[OF Suc.prems(2) cdqm(4)] by metis  | 
624  | 
then show ?case  | 
|
625  | 
proof cases  | 
|
626  | 
case 1  | 
|
627  | 
then show ?thesis  | 
|
628  | 
by (simp only: cdqm o_assoc) (simp add: cdqm)  | 
|
629  | 
next  | 
|
630  | 
case prems: 2  | 
|
631  | 
then have az: "a \<noteq> z"  | 
|
| 54681 | 632  | 
by simp  | 
| 65342 | 633  | 
from prems have *: "(Fun.swap x y id \<circ> h) a = a \<longleftrightarrow> h a = a" for h  | 
634  | 
by (simp add: Fun.swap_def)  | 
|
| 54681 | 635  | 
from cdqm(2) have "Fun.swap a b id \<circ> p = Fun.swap a b id \<circ> (Fun.swap c d id \<circ> q)"  | 
636  | 
by simp  | 
|
637  | 
then have "Fun.swap a b id \<circ> p = Fun.swap x y id \<circ> (Fun.swap a z id \<circ> q)"  | 
|
| 65342 | 638  | 
by (simp add: o_assoc prems)  | 
| 54681 | 639  | 
then have "(Fun.swap a b id \<circ> p) a = (Fun.swap x y id \<circ> (Fun.swap a z id \<circ> q)) a"  | 
640  | 
by simp  | 
|
641  | 
then have "(Fun.swap x y id \<circ> (Fun.swap a z id \<circ> q)) a = a"  | 
|
642  | 
unfolding Suc by metis  | 
|
| 65342 | 643  | 
then have "(Fun.swap a z id \<circ> q) a = a"  | 
644  | 
by (simp only: *)  | 
|
645  | 
from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az this]  | 
|
646  | 
have **: "swapidseq (n - 1) (Fun.swap a z id \<circ> q)" "n \<noteq> 0"  | 
|
| 54681 | 647  | 
by blast+  | 
| 65342 | 648  | 
from \<open>n \<noteq> 0\<close> have ***: "Suc n - 1 = Suc (n - 1)"  | 
649  | 
by auto  | 
|
650  | 
show ?thesis  | 
|
651  | 
apply (simp only: cdqm(2) prems o_assoc ***)  | 
|
| 49739 | 652  | 
apply (simp only: Suc_not_Zero simp_thms comp_assoc)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
653  | 
apply (rule comp_Suc)  | 
| 65342 | 654  | 
using ** prems  | 
655  | 
apply blast+  | 
|
| 54681 | 656  | 
done  | 
| 65342 | 657  | 
qed  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
658  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
659  | 
|
| 30488 | 660  | 
lemma swapidseq_identity_even:  | 
| 54681 | 661  | 
assumes "swapidseq n (id :: 'a \<Rightarrow> 'a)"  | 
662  | 
shows "even n"  | 
|
| 60500 | 663  | 
using \<open>swapidseq n id\<close>  | 
| 54681 | 664  | 
proof (induct n rule: nat_less_induct)  | 
| 65342 | 665  | 
case H: (1 n)  | 
666  | 
consider "n = 0"  | 
|
667  | 
| a b :: 'a and q m where "n = Suc m" "id = Fun.swap a b id \<circ> q" "swapidseq m q" "a \<noteq> b"  | 
|
668  | 
using H(2)[unfolded swapidseq_cases[of n id]] by auto  | 
|
669  | 
then show ?case  | 
|
670  | 
proof cases  | 
|
671  | 
case 1  | 
|
672  | 
then show ?thesis by presburger  | 
|
673  | 
next  | 
|
674  | 
case h: 2  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
675  | 
from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]]  | 
| 54681 | 676  | 
have m: "m \<noteq> 0" "swapidseq (m - 1) (id :: 'a \<Rightarrow> 'a)"  | 
677  | 
by auto  | 
|
678  | 
from h m have mn: "m - 1 < n"  | 
|
679  | 
by arith  | 
|
| 65342 | 680  | 
from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis  | 
| 54681 | 681  | 
by presburger  | 
| 65342 | 682  | 
qed  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
683  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
684  | 
|
| 54681 | 685  | 
|
| 60500 | 686  | 
subsection \<open>Therefore we have a welldefined notion of parity\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
687  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
688  | 
definition "evenperm p = even (SOME n. swapidseq n p)"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
689  | 
|
| 54681 | 690  | 
lemma swapidseq_even_even:  | 
691  | 
assumes m: "swapidseq m p"  | 
|
692  | 
and n: "swapidseq n p"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
693  | 
shows "even m \<longleftrightarrow> even n"  | 
| 54681 | 694  | 
proof -  | 
| 65342 | 695  | 
from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id"  | 
| 54681 | 696  | 
by blast  | 
| 65342 | 697  | 
from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis  | 
| 54681 | 698  | 
by arith  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
699  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
700  | 
|
| 54681 | 701  | 
lemma evenperm_unique:  | 
702  | 
assumes p: "swapidseq n p"  | 
|
703  | 
and n:"even n = b"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
704  | 
shows "evenperm p = b"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
705  | 
unfolding n[symmetric] evenperm_def  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
706  | 
apply (rule swapidseq_even_even[where p = p])  | 
| 65342 | 707  | 
apply (rule someI[where x = n])  | 
| 54681 | 708  | 
using p  | 
| 65342 | 709  | 
apply blast+  | 
| 54681 | 710  | 
done  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
711  | 
|
| 54681 | 712  | 
|
| 60500 | 713  | 
subsection \<open>And it has the expected composition properties\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
714  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
715  | 
lemma evenperm_id[simp]: "evenperm id = True"  | 
| 54681 | 716  | 
by (rule evenperm_unique[where n = 0]) simp_all  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
717  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
718  | 
lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)"  | 
| 54681 | 719  | 
by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
720  | 
|
| 30488 | 721  | 
lemma evenperm_comp:  | 
| 65342 | 722  | 
assumes "permutation p" "permutation q"  | 
723  | 
shows "evenperm (p \<circ> q) \<longleftrightarrow> evenperm p = evenperm q"  | 
|
| 54681 | 724  | 
proof -  | 
| 65342 | 725  | 
from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
726  | 
unfolding permutation_def by blast  | 
| 65342 | 727  | 
have "even (n + m) \<longleftrightarrow> (even n \<longleftrightarrow> even m)"  | 
| 54681 | 728  | 
by arith  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
729  | 
from evenperm_unique[OF n refl] evenperm_unique[OF m refl]  | 
| 65342 | 730  | 
and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis  | 
| 54681 | 731  | 
by blast  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
732  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
733  | 
|
| 54681 | 734  | 
lemma evenperm_inv:  | 
| 65342 | 735  | 
assumes "permutation p"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
736  | 
shows "evenperm (inv p) = evenperm p"  | 
| 54681 | 737  | 
proof -  | 
| 65342 | 738  | 
from assms obtain n where n: "swapidseq n p"  | 
| 54681 | 739  | 
unfolding permutation_def by blast  | 
| 65342 | 740  | 
show ?thesis  | 
741  | 
by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]])  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
742  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
743  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
744  | 
|
| 60500 | 745  | 
subsection \<open>A more abstract characterization of permutations\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
746  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
747  | 
lemma bij_iff: "bij f \<longleftrightarrow> (\<forall>x. \<exists>!y. f y = x)"  | 
| 
64966
 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 
wenzelm 
parents: 
64543 
diff
changeset
 | 
748  | 
unfolding bij_def inj_def surj_def  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
749  | 
apply auto  | 
| 65342 | 750  | 
apply metis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
751  | 
apply metis  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
752  | 
done  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
753  | 
|
| 30488 | 754  | 
lemma permutation_bijective:  | 
| 65342 | 755  | 
assumes "permutation p"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
756  | 
shows "bij p"  | 
| 54681 | 757  | 
proof -  | 
| 65342 | 758  | 
from assms obtain n where n: "swapidseq n p"  | 
| 54681 | 759  | 
unfolding permutation_def by blast  | 
| 65342 | 760  | 
from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id"  | 
| 54681 | 761  | 
by blast  | 
| 65342 | 762  | 
then show ?thesis  | 
763  | 
unfolding bij_iff  | 
|
| 54681 | 764  | 
apply (auto simp add: fun_eq_iff)  | 
765  | 
apply metis  | 
|
766  | 
done  | 
|
| 30488 | 767  | 
qed  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
768  | 
|
| 54681 | 769  | 
lemma permutation_finite_support:  | 
| 65342 | 770  | 
assumes "permutation p"  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
771  | 
  shows "finite {x. p x \<noteq> x}"
 | 
| 54681 | 772  | 
proof -  | 
| 65342 | 773  | 
from assms obtain n where "swapidseq n p"  | 
| 54681 | 774  | 
unfolding permutation_def by blast  | 
| 65342 | 775  | 
then show ?thesis  | 
| 54681 | 776  | 
proof (induct n p rule: swapidseq.induct)  | 
777  | 
case id  | 
|
778  | 
then show ?case by simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
779  | 
next  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
780  | 
case (comp_Suc n p a b)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
781  | 
    let ?S = "insert a (insert b {x. p x \<noteq> x})"
 | 
| 65342 | 782  | 
from comp_Suc.hyps(2) have *: "finite ?S"  | 
| 54681 | 783  | 
by simp  | 
| 65342 | 784  | 
    from \<open>a \<noteq> b\<close> have **: "{x. (Fun.swap a b id \<circ> p) x \<noteq> x} \<subseteq> ?S"
 | 
785  | 
by (auto simp: Fun.swap_def)  | 
|
786  | 
show ?case  | 
|
787  | 
by (rule finite_subset[OF ** *])  | 
|
| 54681 | 788  | 
qed  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
789  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
790  | 
|
| 30488 | 791  | 
lemma permutation_lemma:  | 
| 65342 | 792  | 
assumes "finite S"  | 
793  | 
and "bij p"  | 
|
794  | 
and "\<forall>x. x\<notin> S \<longrightarrow> p x = x"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
795  | 
shows "permutation p"  | 
| 65342 | 796  | 
using assms  | 
| 54681 | 797  | 
proof (induct S arbitrary: p rule: finite_induct)  | 
| 65342 | 798  | 
case empty  | 
799  | 
then show ?case  | 
|
800  | 
by simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
801  | 
next  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
802  | 
case (insert a F p)  | 
| 54681 | 803  | 
let ?r = "Fun.swap a (p a) id \<circ> p"  | 
804  | 
let ?q = "Fun.swap a (p a) id \<circ> ?r"  | 
|
| 65342 | 805  | 
have *: "?r a = a"  | 
| 56545 | 806  | 
by (simp add: Fun.swap_def)  | 
| 65342 | 807  | 
from insert * have **: "\<forall>x. x \<notin> F \<longrightarrow> ?r x = x"  | 
| 
64966
 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 
wenzelm 
parents: 
64543 
diff
changeset
 | 
808  | 
by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3))  | 
| 65342 | 809  | 
have "bij ?r"  | 
810  | 
by (rule bij_swap_compose_bij[OF insert(4)])  | 
|
811  | 
have "permutation ?r"  | 
|
812  | 
by (rule insert(3)[OF \<open>bij ?r\<close> **])  | 
|
813  | 
then have "permutation ?q"  | 
|
814  | 
by (simp add: permutation_compose permutation_swap_id)  | 
|
| 54681 | 815  | 
then show ?case  | 
816  | 
by (simp add: o_assoc)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
817  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
818  | 
|
| 30488 | 819  | 
lemma permutation: "permutation p \<longleftrightarrow> bij p \<and> finite {x. p x \<noteq> x}"
 | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
820  | 
(is "?lhs \<longleftrightarrow> ?b \<and> ?f")  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
821  | 
proof  | 
| 65342 | 822  | 
assume ?lhs  | 
823  | 
with permutation_bijective permutation_finite_support show "?b \<and> ?f"  | 
|
| 54681 | 824  | 
by auto  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
825  | 
next  | 
| 54681 | 826  | 
assume "?b \<and> ?f"  | 
827  | 
then have "?f" "?b" by blast+  | 
|
828  | 
from permutation_lemma[OF this] show ?lhs  | 
|
829  | 
by blast  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
830  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
831  | 
|
| 54681 | 832  | 
lemma permutation_inverse_works:  | 
| 65342 | 833  | 
assumes "permutation p"  | 
| 54681 | 834  | 
shows "inv p \<circ> p = id"  | 
835  | 
and "p \<circ> inv p = id"  | 
|
| 65342 | 836  | 
using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
837  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
838  | 
lemma permutation_inverse_compose:  | 
| 54681 | 839  | 
assumes p: "permutation p"  | 
840  | 
and q: "permutation q"  | 
|
841  | 
shows "inv (p \<circ> q) = inv q \<circ> inv p"  | 
|
842  | 
proof -  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
843  | 
note ps = permutation_inverse_works[OF p]  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
844  | 
note qs = permutation_inverse_works[OF q]  | 
| 54681 | 845  | 
have "p \<circ> q \<circ> (inv q \<circ> inv p) = p \<circ> (q \<circ> inv q) \<circ> inv p"  | 
846  | 
by (simp add: o_assoc)  | 
|
847  | 
also have "\<dots> = id"  | 
|
848  | 
by (simp add: ps qs)  | 
|
| 65342 | 849  | 
finally have *: "p \<circ> q \<circ> (inv q \<circ> inv p) = id" .  | 
| 54681 | 850  | 
have "inv q \<circ> inv p \<circ> (p \<circ> q) = inv q \<circ> (inv p \<circ> p) \<circ> q"  | 
851  | 
by (simp add: o_assoc)  | 
|
852  | 
also have "\<dots> = id"  | 
|
853  | 
by (simp add: ps qs)  | 
|
| 65342 | 854  | 
finally have **: "inv q \<circ> inv p \<circ> (p \<circ> q) = id" .  | 
855  | 
show ?thesis  | 
|
856  | 
by (rule inv_unique_comp[OF * **])  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
857  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
858  | 
|
| 54681 | 859  | 
|
| 65342 | 860  | 
subsection \<open>Relation to \<open>permutes\<close>\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
861  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
862  | 
lemma permutation_permutes: "permutation p \<longleftrightarrow> (\<exists>S. finite S \<and> p permutes S)"  | 
| 54681 | 863  | 
unfolding permutation permutes_def bij_iff[symmetric]  | 
864  | 
apply (rule iffI, clarify)  | 
|
| 65342 | 865  | 
   apply (rule exI[where x="{x. p x \<noteq> x}"])
 | 
866  | 
apply simp  | 
|
| 54681 | 867  | 
apply clarsimp  | 
868  | 
apply (rule_tac B="S" in finite_subset)  | 
|
| 65342 | 869  | 
apply auto  | 
| 54681 | 870  | 
done  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
871  | 
|
| 54681 | 872  | 
|
| 60500 | 873  | 
subsection \<open>Hence a sort of induction principle composing by swaps\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
874  | 
|
| 54681 | 875  | 
lemma permutes_induct: "finite S \<Longrightarrow> P id \<Longrightarrow>  | 
| 65342 | 876  | 
(\<And>a b p. a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> P p \<Longrightarrow> P p \<Longrightarrow> permutation p \<Longrightarrow> P (Fun.swap a b id \<circ> p)) \<Longrightarrow>  | 
| 54681 | 877  | 
(\<And>p. p permutes S \<Longrightarrow> P p)"  | 
878  | 
proof (induct S rule: finite_induct)  | 
|
879  | 
case empty  | 
|
880  | 
then show ?case by auto  | 
|
| 30488 | 881  | 
next  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
882  | 
case (insert x F p)  | 
| 54681 | 883  | 
let ?r = "Fun.swap x (p x) id \<circ> p"  | 
884  | 
let ?q = "Fun.swap x (p x) id \<circ> ?r"  | 
|
885  | 
have qp: "?q = p"  | 
|
886  | 
by (simp add: o_assoc)  | 
|
887  | 
from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r"  | 
|
888  | 
by blast  | 
|
| 30488 | 889  | 
from permutes_in_image[OF insert.prems(3), of x]  | 
| 54681 | 890  | 
have pxF: "p x \<in> insert x F"  | 
891  | 
by simp  | 
|
892  | 
have xF: "x \<in> insert x F"  | 
|
893  | 
by simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
894  | 
have rp: "permutation ?r"  | 
| 65342 | 895  | 
unfolding permutation_permutes  | 
896  | 
using insert.hyps(1) permutes_insert_lemma[OF insert.prems(3)]  | 
|
| 54681 | 897  | 
by blast  | 
| 65342 | 898  | 
from insert.prems(2)[OF xF pxF Pr Pr rp] qp show ?case  | 
899  | 
by (simp only:)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
900  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
901  | 
|
| 54681 | 902  | 
|
| 60500 | 903  | 
subsection \<open>Sign of a permutation as a real number\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
904  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
905  | 
definition "sign p = (if evenperm p then (1::int) else -1)"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
906  | 
|
| 54681 | 907  | 
lemma sign_nz: "sign p \<noteq> 0"  | 
908  | 
by (simp add: sign_def)  | 
|
909  | 
||
910  | 
lemma sign_id: "sign id = 1"  | 
|
911  | 
by (simp add: sign_def)  | 
|
912  | 
||
913  | 
lemma sign_inverse: "permutation p \<Longrightarrow> sign (inv p) = sign p"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
914  | 
by (simp add: sign_def evenperm_inv)  | 
| 54681 | 915  | 
|
916  | 
lemma sign_compose: "permutation p \<Longrightarrow> permutation q \<Longrightarrow> sign (p \<circ> q) = sign p * sign q"  | 
|
917  | 
by (simp add: sign_def evenperm_comp)  | 
|
918  | 
||
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
919  | 
lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
920  | 
by (simp add: sign_def evenperm_swap)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
921  | 
|
| 54681 | 922  | 
lemma sign_idempotent: "sign p * sign p = 1"  | 
923  | 
by (simp add: sign_def)  | 
|
924  | 
||
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
925  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
926  | 
subsection \<open>Permuting a list\<close>  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
927  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
928  | 
text \<open>This function permutes a list by applying a permutation to the indices.\<close>  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
929  | 
|
| 65342 | 930  | 
definition permute_list :: "(nat \<Rightarrow> nat) \<Rightarrow> 'a list \<Rightarrow> 'a list"  | 
931  | 
where "permute_list f xs = map (\<lambda>i. xs ! (f i)) [0..<length xs]"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
932  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
933  | 
lemma permute_list_map:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
934  | 
  assumes "f permutes {..<length xs}"
 | 
| 65342 | 935  | 
shows "permute_list f (map g xs) = map g (permute_list f xs)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
936  | 
using permutes_in_image[OF assms] by (auto simp: permute_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
937  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
938  | 
lemma permute_list_nth:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
939  | 
  assumes "f permutes {..<length xs}" "i < length xs"
 | 
| 65342 | 940  | 
shows "permute_list f xs ! i = xs ! f i"  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
941  | 
using permutes_in_image[OF assms(1)] assms(2)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
942  | 
by (simp add: permute_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
943  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
944  | 
lemma permute_list_Nil [simp]: "permute_list f [] = []"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
945  | 
by (simp add: permute_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
946  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
947  | 
lemma length_permute_list [simp]: "length (permute_list f xs) = length xs"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
948  | 
by (simp add: permute_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
949  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
950  | 
lemma permute_list_compose:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
951  | 
  assumes "g permutes {..<length xs}"
 | 
| 65342 | 952  | 
shows "permute_list (f \<circ> g) xs = permute_list g (permute_list f xs)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
953  | 
using assms[THEN permutes_in_image] by (auto simp add: permute_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
954  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
955  | 
lemma permute_list_ident [simp]: "permute_list (\<lambda>x. x) xs = xs"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
956  | 
by (simp add: permute_list_def map_nth)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
957  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
958  | 
lemma permute_list_id [simp]: "permute_list id xs = xs"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
959  | 
by (simp add: id_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
960  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
961  | 
lemma mset_permute_list [simp]:  | 
| 65342 | 962  | 
fixes xs :: "'a list"  | 
963  | 
  assumes "f permutes {..<length xs}"
 | 
|
964  | 
shows "mset (permute_list f xs) = mset xs"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
965  | 
proof (rule multiset_eqI)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
966  | 
fix y :: 'a  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
967  | 
from assms have [simp]: "f x < length xs \<longleftrightarrow> x < length xs" for x  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
968  | 
using permutes_in_image[OF assms] by auto  | 
| 65342 | 969  | 
  have "count (mset (permute_list f xs)) y = card ((\<lambda>i. xs ! f i) -` {y} \<inter> {..<length xs})"
 | 
| 
64543
 
6b13586ef1a2
remove typo in bij_swap_compose_bij theorem name; tune proof
 
bulwahn 
parents: 
64284 
diff
changeset
 | 
970  | 
by (simp add: permute_list_def count_image_mset atLeast0LessThan)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
971  | 
  also have "(\<lambda>i. xs ! f i) -` {y} \<inter> {..<length xs} = f -` {i. i < length xs \<and> y = xs ! i}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
972  | 
by auto  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
973  | 
  also from assms have "card \<dots> = card {i. i < length xs \<and> y = xs ! i}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
974  | 
by (intro card_vimage_inj) (auto simp: permutes_inj permutes_surj)  | 
| 65342 | 975  | 
also have "\<dots> = count (mset xs) y"  | 
976  | 
by (simp add: count_mset length_filter_conv_card)  | 
|
977  | 
finally show "count (mset (permute_list f xs)) y = count (mset xs) y"  | 
|
978  | 
by simp  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
979  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
980  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
981  | 
lemma set_permute_list [simp]:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
982  | 
  assumes "f permutes {..<length xs}"
 | 
| 65342 | 983  | 
shows "set (permute_list f xs) = set xs"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
984  | 
by (rule mset_eq_setD[OF mset_permute_list]) fact  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
985  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
986  | 
lemma distinct_permute_list [simp]:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
987  | 
  assumes "f permutes {..<length xs}"
 | 
| 65342 | 988  | 
shows "distinct (permute_list f xs) = distinct xs"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
989  | 
by (simp add: distinct_count_atmost_1 assms)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
990  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
991  | 
lemma permute_list_zip:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
992  | 
  assumes "f permutes A" "A = {..<length xs}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
993  | 
assumes [simp]: "length xs = length ys"  | 
| 65342 | 994  | 
shows "permute_list f (zip xs ys) = zip (permute_list f xs) (permute_list f ys)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
995  | 
proof -  | 
| 65342 | 996  | 
from permutes_in_image[OF assms(1)] assms(2) have *: "f i < length ys \<longleftrightarrow> i < length ys" for i  | 
997  | 
by simp  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
998  | 
have "permute_list f (zip xs ys) = map (\<lambda>i. zip xs ys ! f i) [0..<length ys]"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
999  | 
by (simp_all add: permute_list_def zip_map_map)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1000  | 
also have "\<dots> = map (\<lambda>(x, y). (xs ! f x, ys ! f y)) (zip [0..<length ys] [0..<length ys])"  | 
| 65342 | 1001  | 
by (intro nth_equalityI) (simp_all add: *)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1002  | 
also have "\<dots> = zip (permute_list f xs) (permute_list f ys)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1003  | 
by (simp_all add: permute_list_def zip_map_map)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1004  | 
finally show ?thesis .  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1005  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1006  | 
|
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
1007  | 
lemma map_of_permute:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1008  | 
assumes "\<sigma> permutes fst ` set xs"  | 
| 65342 | 1009  | 
shows "map_of xs \<circ> \<sigma> = map_of (map (\<lambda>(x,y). (inv \<sigma> x, y)) xs)"  | 
1010  | 
(is "_ = map_of (map ?f _)")  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1011  | 
proof  | 
| 65342 | 1012  | 
from assms have "inj \<sigma>" "surj \<sigma>"  | 
1013  | 
by (simp_all add: permutes_inj permutes_surj)  | 
|
1014  | 
then show "(map_of xs \<circ> \<sigma>) x = map_of (map ?f xs) x" for x  | 
|
1015  | 
by (induct xs) (auto simp: inv_f_f surj_f_inv_f)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1016  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1017  | 
|
| 54681 | 1018  | 
|
| 60500 | 1019  | 
subsection \<open>More lemmas about permutations\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1020  | 
|
| 65342 | 1021  | 
text \<open>The following few lemmas were contributed by Lukas Bulwahn.\<close>  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1022  | 
|
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1023  | 
lemma count_image_mset_eq_card_vimage:  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1024  | 
assumes "finite A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1025  | 
  shows "count (image_mset f (mset_set A)) b = card {a \<in> A. f a = b}"
 | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1026  | 
using assms  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1027  | 
proof (induct A)  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1028  | 
case empty  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1029  | 
show ?case by simp  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1030  | 
next  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1031  | 
case (insert x F)  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1032  | 
show ?case  | 
| 65342 | 1033  | 
proof (cases "f x = b")  | 
1034  | 
case True  | 
|
1035  | 
with insert.hyps  | 
|
1036  | 
    have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \<in> F. f a = f x})"
 | 
|
1037  | 
by auto  | 
|
1038  | 
    also from insert.hyps(1,2) have "\<dots> = card (insert x {a \<in> F. f a = f x})"
 | 
|
1039  | 
by simp  | 
|
1040  | 
    also from \<open>f x = b\<close> have "card (insert x {a \<in> F. f a = f x}) = card {a \<in> insert x F. f a = b}"
 | 
|
1041  | 
by (auto intro: arg_cong[where f="card"])  | 
|
1042  | 
finally show ?thesis  | 
|
1043  | 
using insert by auto  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1044  | 
next  | 
| 65342 | 1045  | 
case False  | 
1046  | 
    then have "{a \<in> F. f a = b} = {a \<in> insert x F. f a = b}"
 | 
|
1047  | 
by auto  | 
|
1048  | 
with insert False show ?thesis  | 
|
1049  | 
by simp  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1050  | 
qed  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1051  | 
qed  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
1052  | 
|
| 67408 | 1053  | 
\<comment> \<open>Prove \<open>image_mset_eq_implies_permutes\<close> ...\<close>  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1054  | 
lemma image_mset_eq_implies_permutes:  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1055  | 
fixes f :: "'a \<Rightarrow> 'b"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1056  | 
assumes "finite A"  | 
| 65342 | 1057  | 
and mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)"  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1058  | 
obtains p where "p permutes A" and "\<forall>x\<in>A. f x = f' (p x)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1059  | 
proof -  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1060  | 
  from \<open>finite A\<close> have [simp]: "finite {a \<in> A. f a = (b::'b)}" for f b by auto
 | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1061  | 
have "f ` A = f' ` A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1062  | 
proof -  | 
| 65342 | 1063  | 
from \<open>finite A\<close> have "f ` A = f ` (set_mset (mset_set A))"  | 
1064  | 
by simp  | 
|
1065  | 
also have "\<dots> = f' ` set_mset (mset_set A)"  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1066  | 
by (metis mset_eq multiset.set_map)  | 
| 65342 | 1067  | 
also from \<open>finite A\<close> have "\<dots> = f' ` A"  | 
1068  | 
by simp  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1069  | 
finally show ?thesis .  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1070  | 
qed  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1071  | 
  have "\<forall>b\<in>(f ` A). \<exists>p. bij_betw p {a \<in> A. f a = b} {a \<in> A. f' a = b}"
 | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1072  | 
proof  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1073  | 
fix b  | 
| 65342 | 1074  | 
from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b"  | 
1075  | 
by simp  | 
|
1076  | 
    with \<open>finite A\<close> have "card {a \<in> A. f a = b} = card {a \<in> A. f' a = b}"
 | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1077  | 
by (simp add: count_image_mset_eq_card_vimage)  | 
| 65342 | 1078  | 
    then show "\<exists>p. bij_betw p {a\<in>A. f a = b} {a \<in> A. f' a = b}"
 | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1079  | 
by (intro finite_same_card_bij) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1080  | 
qed  | 
| 65342 | 1081  | 
  then have "\<exists>p. \<forall>b\<in>f ` A. bij_betw (p b) {a \<in> A. f a = b} {a \<in> A. f' a = b}"
 | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1082  | 
by (rule bchoice)  | 
| 65342 | 1083  | 
  then obtain p where p: "\<forall>b\<in>f ` A. bij_betw (p b) {a \<in> A. f a = b} {a \<in> A. f' a = b}" ..
 | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1084  | 
define p' where "p' = (\<lambda>a. if a \<in> A then p (f a) a else a)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1085  | 
have "p' permutes A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1086  | 
proof (rule bij_imp_permutes)  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1087  | 
    have "disjoint_family_on (\<lambda>i. {a \<in> A. f' a = i}) (f ` A)"
 | 
| 65342 | 1088  | 
by (auto simp: disjoint_family_on_def)  | 
1089  | 
moreover  | 
|
1090  | 
    have "bij_betw (\<lambda>a. p (f a) a) {a \<in> A. f a = b} {a \<in> A. f' a = b}" if "b \<in> f ` A" for b
 | 
|
1091  | 
using p that by (subst bij_betw_cong[where g="p b"]) auto  | 
|
1092  | 
ultimately  | 
|
1093  | 
    have "bij_betw (\<lambda>a. p (f a) a) (\<Union>b\<in>f ` A. {a \<in> A. f a = b}) (\<Union>b\<in>f ` A. {a \<in> A. f' a = b})"
 | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1094  | 
by (rule bij_betw_UNION_disjoint)  | 
| 65342 | 1095  | 
    moreover have "(\<Union>b\<in>f ` A. {a \<in> A. f a = b}) = A"
 | 
1096  | 
by auto  | 
|
1097  | 
    moreover from \<open>f ` A = f' ` A\<close> have "(\<Union>b\<in>f ` A. {a \<in> A. f' a = b}) = A"
 | 
|
1098  | 
by auto  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1099  | 
ultimately show "bij_betw p' A A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1100  | 
unfolding p'_def by (subst bij_betw_cong[where g="(\<lambda>a. p (f a) a)"]) auto  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1101  | 
next  | 
| 65342 | 1102  | 
show "\<And>x. x \<notin> A \<Longrightarrow> p' x = x"  | 
1103  | 
by (simp add: p'_def)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1104  | 
qed  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1105  | 
moreover from p have "\<forall>x\<in>A. f x = f' (p' x)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1106  | 
unfolding p'_def using bij_betwE by fastforce  | 
| 65342 | 1107  | 
ultimately show ?thesis ..  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1108  | 
qed  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1109  | 
|
| 65342 | 1110  | 
lemma mset_set_upto_eq_mset_upto: "mset_set {..<n} = mset [0..<n]"
 | 
1111  | 
by (induct n) (auto simp: add.commute lessThan_Suc)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1112  | 
|
| 67408 | 1113  | 
\<comment> \<open>... and derive the existing property:\<close>  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1114  | 
lemma mset_eq_permutation:  | 
| 65342 | 1115  | 
fixes xs ys :: "'a list"  | 
1116  | 
assumes mset_eq: "mset xs = mset ys"  | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1117  | 
  obtains p where "p permutes {..<length ys}" "permute_list p ys = xs"
 | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1118  | 
proof -  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1119  | 
from mset_eq have length_eq: "length xs = length ys"  | 
| 65342 | 1120  | 
by (rule mset_eq_length)  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1121  | 
  have "mset_set {..<length ys} = mset [0..<length ys]"
 | 
| 65342 | 1122  | 
by (rule mset_set_upto_eq_mset_upto)  | 
1123  | 
  with mset_eq length_eq have "image_mset (\<lambda>i. xs ! i) (mset_set {..<length ys}) =
 | 
|
1124  | 
    image_mset (\<lambda>i. ys ! i) (mset_set {..<length ys})"
 | 
|
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1125  | 
by (metis map_nth mset_map)  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63539 
diff
changeset
 | 
1126  | 
from image_mset_eq_implies_permutes[OF _ this]  | 
| 65342 | 1127  | 
  obtain p where p: "p permutes {..<length ys}" and "\<forall>i\<in>{..<length ys}. xs ! i = ys ! (p i)"
 | 
1128  | 
by auto  | 
|
1129  | 
with length_eq have "permute_list p ys = xs"  | 
|
1130  | 
by (auto intro!: nth_equalityI simp: permute_list_nth)  | 
|
1131  | 
with p show thesis ..  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1132  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1133  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1134  | 
lemma permutes_natset_le:  | 
| 54681 | 1135  | 
fixes S :: "'a::wellorder set"  | 
| 65342 | 1136  | 
assumes "p permutes S"  | 
1137  | 
and "\<forall>i \<in> S. p i \<le> i"  | 
|
| 54681 | 1138  | 
shows "p = id"  | 
1139  | 
proof -  | 
|
| 65342 | 1140  | 
have "p n = n" for n  | 
1141  | 
using assms  | 
|
1142  | 
proof (induct n arbitrary: S rule: less_induct)  | 
|
1143  | 
case (less n)  | 
|
1144  | 
show ?case  | 
|
1145  | 
proof (cases "n \<in> S")  | 
|
1146  | 
case False  | 
|
1147  | 
with less(2) show ?thesis  | 
|
1148  | 
unfolding permutes_def by metis  | 
|
1149  | 
next  | 
|
1150  | 
case True  | 
|
1151  | 
with less(3) have "p n < n \<or> p n = n"  | 
|
1152  | 
by auto  | 
|
1153  | 
then show ?thesis  | 
|
1154  | 
proof  | 
|
1155  | 
assume "p n < n"  | 
|
1156  | 
with less have "p (p n) = p n"  | 
|
1157  | 
by metis  | 
|
1158  | 
with permutes_inj[OF less(2)] have "p n = n"  | 
|
1159  | 
unfolding inj_def by blast  | 
|
1160  | 
with \<open>p n < n\<close> have False  | 
|
1161  | 
by simp  | 
|
1162  | 
then show ?thesis ..  | 
|
1163  | 
qed  | 
|
| 54681 | 1164  | 
qed  | 
| 65342 | 1165  | 
qed  | 
1166  | 
then show ?thesis by (auto simp: fun_eq_iff)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1167  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1168  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1169  | 
lemma permutes_natset_ge:  | 
| 54681 | 1170  | 
fixes S :: "'a::wellorder set"  | 
1171  | 
assumes p: "p permutes S"  | 
|
1172  | 
and le: "\<forall>i \<in> S. p i \<ge> i"  | 
|
1173  | 
shows "p = id"  | 
|
1174  | 
proof -  | 
|
| 65342 | 1175  | 
have "i \<ge> inv p i" if "i \<in> S" for i  | 
1176  | 
proof -  | 
|
1177  | 
from that permutes_in_image[OF permutes_inv[OF p]] have "inv p i \<in> S"  | 
|
| 54681 | 1178  | 
by simp  | 
1179  | 
with le have "p (inv p i) \<ge> inv p i"  | 
|
1180  | 
by blast  | 
|
| 65342 | 1181  | 
with permutes_inverses[OF p] show ?thesis  | 
| 54681 | 1182  | 
by simp  | 
| 65342 | 1183  | 
qed  | 
1184  | 
then have "\<forall>i\<in>S. inv p i \<le> i"  | 
|
| 54681 | 1185  | 
by blast  | 
| 65342 | 1186  | 
from permutes_natset_le[OF permutes_inv[OF p] this] have "inv p = inv id"  | 
| 54681 | 1187  | 
by simp  | 
| 30488 | 1188  | 
then show ?thesis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1189  | 
apply (subst permutes_inv_inv[OF p, symmetric])  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1190  | 
apply (rule inv_unique_comp)  | 
| 65342 | 1191  | 
apply simp_all  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1192  | 
done  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1193  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1194  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1195  | 
lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}"
 | 
| 54681 | 1196  | 
apply (rule set_eqI)  | 
1197  | 
apply auto  | 
|
1198  | 
using permutes_inv_inv permutes_inv  | 
|
| 65342 | 1199  | 
apply auto  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1200  | 
apply (rule_tac x="inv x" in exI)  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1201  | 
apply auto  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1202  | 
done  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1203  | 
|
| 30488 | 1204  | 
lemma image_compose_permutations_left:  | 
| 65342 | 1205  | 
assumes "q permutes S"  | 
1206  | 
  shows "{q \<circ> p |p. p permutes S} = {p. p permutes S}"
 | 
|
| 54681 | 1207  | 
apply (rule set_eqI)  | 
1208  | 
apply auto  | 
|
| 65342 | 1209  | 
apply (rule permutes_compose)  | 
1210  | 
using assms  | 
|
1211  | 
apply auto  | 
|
| 54681 | 1212  | 
apply (rule_tac x = "inv q \<circ> x" in exI)  | 
1213  | 
apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o)  | 
|
1214  | 
done  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1215  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1216  | 
lemma image_compose_permutations_right:  | 
| 65342 | 1217  | 
assumes "q permutes S"  | 
| 54681 | 1218  | 
  shows "{p \<circ> q | p. p permutes S} = {p . p permutes S}"
 | 
1219  | 
apply (rule set_eqI)  | 
|
1220  | 
apply auto  | 
|
| 65342 | 1221  | 
apply (rule permutes_compose)  | 
1222  | 
using assms  | 
|
1223  | 
apply auto  | 
|
| 54681 | 1224  | 
apply (rule_tac x = "x \<circ> inv q" in exI)  | 
1225  | 
apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc)  | 
|
1226  | 
done  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1227  | 
|
| 54681 | 1228  | 
lemma permutes_in_seg: "p permutes {1 ..n} \<Longrightarrow> i \<in> {1..n} \<Longrightarrow> 1 \<le> p i \<and> p i \<le> n"
 | 
1229  | 
by (simp add: permutes_def) metis  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1230  | 
|
| 65342 | 1231  | 
lemma sum_permutations_inverse: "sum f {p. p permutes S} = sum (\<lambda>p. f(inv p)) {p. p permutes S}"
 | 
| 54681 | 1232  | 
(is "?lhs = ?rhs")  | 
1233  | 
proof -  | 
|
| 30036 | 1234  | 
  let ?S = "{p . p permutes S}"
 | 
| 65342 | 1235  | 
have *: "inj_on inv ?S"  | 
| 54681 | 1236  | 
proof (auto simp add: inj_on_def)  | 
1237  | 
fix q r  | 
|
1238  | 
assume q: "q permutes S"  | 
|
1239  | 
and r: "r permutes S"  | 
|
1240  | 
and qr: "inv q = inv r"  | 
|
1241  | 
then have "inv (inv q) = inv (inv r)"  | 
|
1242  | 
by simp  | 
|
1243  | 
with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r"  | 
|
1244  | 
by metis  | 
|
1245  | 
qed  | 
|
| 65342 | 1246  | 
have **: "inv ` ?S = ?S"  | 
| 54681 | 1247  | 
using image_inverse_permutations by blast  | 
| 65342 | 1248  | 
have ***: "?rhs = sum (f \<circ> inv) ?S"  | 
| 54681 | 1249  | 
by (simp add: o_def)  | 
| 65342 | 1250  | 
from sum.reindex[OF *, of f] show ?thesis  | 
1251  | 
by (simp only: ** ***)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1252  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1253  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1254  | 
lemma setum_permutations_compose_left:  | 
| 30036 | 1255  | 
assumes q: "q permutes S"  | 
| 64267 | 1256  | 
  shows "sum f {p. p permutes S} = sum (\<lambda>p. f(q \<circ> p)) {p. p permutes S}"
 | 
| 54681 | 1257  | 
(is "?lhs = ?rhs")  | 
1258  | 
proof -  | 
|
| 30036 | 1259  | 
  let ?S = "{p. p permutes S}"
 | 
| 67399 | 1260  | 
have *: "?rhs = sum (f \<circ> ((\<circ>) q)) ?S"  | 
| 54681 | 1261  | 
by (simp add: o_def)  | 
| 67399 | 1262  | 
have **: "inj_on ((\<circ>) q) ?S"  | 
| 54681 | 1263  | 
proof (auto simp add: inj_on_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1264  | 
fix p r  | 
| 54681 | 1265  | 
assume "p permutes S"  | 
1266  | 
and r: "r permutes S"  | 
|
1267  | 
and rp: "q \<circ> p = q \<circ> r"  | 
|
1268  | 
then have "inv q \<circ> q \<circ> p = inv q \<circ> q \<circ> r"  | 
|
1269  | 
by (simp add: comp_assoc)  | 
|
1270  | 
with permutes_inj[OF q, unfolded inj_iff] show "p = r"  | 
|
1271  | 
by simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1272  | 
qed  | 
| 67399 | 1273  | 
have "((\<circ>) q) ` ?S = ?S"  | 
| 54681 | 1274  | 
using image_compose_permutations_left[OF q] by auto  | 
| 65342 | 1275  | 
with * sum.reindex[OF **, of f] show ?thesis  | 
1276  | 
by (simp only:)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1277  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1278  | 
|
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1279  | 
lemma sum_permutations_compose_right:  | 
| 30036 | 1280  | 
assumes q: "q permutes S"  | 
| 64267 | 1281  | 
  shows "sum f {p. p permutes S} = sum (\<lambda>p. f(p \<circ> q)) {p. p permutes S}"
 | 
| 54681 | 1282  | 
(is "?lhs = ?rhs")  | 
1283  | 
proof -  | 
|
| 30036 | 1284  | 
  let ?S = "{p. p permutes S}"
 | 
| 65342 | 1285  | 
have *: "?rhs = sum (f \<circ> (\<lambda>p. p \<circ> q)) ?S"  | 
| 54681 | 1286  | 
by (simp add: o_def)  | 
| 65342 | 1287  | 
have **: "inj_on (\<lambda>p. p \<circ> q) ?S"  | 
| 54681 | 1288  | 
proof (auto simp add: inj_on_def)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1289  | 
fix p r  | 
| 54681 | 1290  | 
assume "p permutes S"  | 
1291  | 
and r: "r permutes S"  | 
|
1292  | 
and rp: "p \<circ> q = r \<circ> q"  | 
|
1293  | 
then have "p \<circ> (q \<circ> inv q) = r \<circ> (q \<circ> inv q)"  | 
|
1294  | 
by (simp add: o_assoc)  | 
|
1295  | 
with permutes_surj[OF q, unfolded surj_iff] show "p = r"  | 
|
1296  | 
by simp  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1297  | 
qed  | 
| 65342 | 1298  | 
from image_compose_permutations_right[OF q] have "(\<lambda>p. p \<circ> q) ` ?S = ?S"  | 
1299  | 
by auto  | 
|
1300  | 
with * sum.reindex[OF **, of f] show ?thesis  | 
|
1301  | 
by (simp only:)  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1302  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1303  | 
|
| 54681 | 1304  | 
|
| 60500 | 1305  | 
subsection \<open>Sum over a set of permutations (could generalize to iteration)\<close>  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1306  | 
|
| 64267 | 1307  | 
lemma sum_over_permutations_insert:  | 
| 54681 | 1308  | 
assumes fS: "finite S"  | 
1309  | 
and aS: "a \<notin> S"  | 
|
| 64267 | 1310  | 
  shows "sum f {p. p permutes (insert a S)} =
 | 
1311  | 
    sum (\<lambda>b. sum (\<lambda>q. f (Fun.swap a b id \<circ> q)) {p. p permutes S}) (insert a S)"
 | 
|
| 54681 | 1312  | 
proof -  | 
| 65342 | 1313  | 
have *: "\<And>f a b. (\<lambda>(b, p). f (Fun.swap a b id \<circ> p)) = f \<circ> (\<lambda>(b,p). Fun.swap a b id \<circ> p)"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
1314  | 
by (simp add: fun_eq_iff)  | 
| 65342 | 1315  | 
  have **: "\<And>P Q. {(a, b). a \<in> P \<and> b \<in> Q} = P \<times> Q"
 | 
| 54681 | 1316  | 
by blast  | 
| 30488 | 1317  | 
show ?thesis  | 
| 65342 | 1318  | 
unfolding * ** sum.cartesian_product permutes_insert  | 
| 64267 | 1319  | 
proof (rule sum.reindex)  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1320  | 
let ?f = "(\<lambda>(b, y). Fun.swap a b id \<circ> y)"  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1321  | 
    let ?P = "{p. p permutes S}"
 | 
| 54681 | 1322  | 
    {
 | 
1323  | 
fix b c p q  | 
|
1324  | 
assume b: "b \<in> insert a S"  | 
|
1325  | 
assume c: "c \<in> insert a S"  | 
|
1326  | 
assume p: "p permutes S"  | 
|
1327  | 
assume q: "q permutes S"  | 
|
1328  | 
assume eq: "Fun.swap a b id \<circ> p = Fun.swap a c id \<circ> q"  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1329  | 
from p q aS have pa: "p a = a" and qa: "q a = a"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32456 
diff
changeset
 | 
1330  | 
unfolding permutes_def by metis+  | 
| 54681 | 1331  | 
from eq have "(Fun.swap a b id \<circ> p) a = (Fun.swap a c id \<circ> q) a"  | 
1332  | 
by simp  | 
|
1333  | 
then have bc: "b = c"  | 
|
| 56545 | 1334  | 
by (simp add: permutes_def pa qa o_def fun_upd_def Fun.swap_def id_def  | 
| 62390 | 1335  | 
cong del: if_weak_cong split: if_split_asm)  | 
| 54681 | 1336  | 
from eq[unfolded bc] have "(\<lambda>p. Fun.swap a c id \<circ> p) (Fun.swap a c id \<circ> p) =  | 
1337  | 
(\<lambda>p. Fun.swap a c id \<circ> p) (Fun.swap a c id \<circ> q)" by simp  | 
|
1338  | 
then have "p = q"  | 
|
| 65342 | 1339  | 
unfolding o_assoc swap_id_idempotent by simp  | 
| 54681 | 1340  | 
with bc have "b = c \<and> p = q"  | 
1341  | 
by blast  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1342  | 
}  | 
| 30488 | 1343  | 
then show "inj_on ?f (insert a S \<times> ?P)"  | 
| 54681 | 1344  | 
unfolding inj_on_def by clarify metis  | 
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1345  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1346  | 
qed  | 
| 
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1347  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1348  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1349  | 
subsection \<open>Constructing permutations from association lists\<close>  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1350  | 
|
| 65342 | 1351  | 
definition list_permutes :: "('a \<times> 'a) list \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
1352  | 
where "list_permutes xs A \<longleftrightarrow>  | 
|
1353  | 
set (map fst xs) \<subseteq> A \<and>  | 
|
1354  | 
set (map snd xs) = set (map fst xs) \<and>  | 
|
1355  | 
distinct (map fst xs) \<and>  | 
|
1356  | 
distinct (map snd xs)"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1357  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1358  | 
lemma list_permutesI [simp]:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1359  | 
assumes "set (map fst xs) \<subseteq> A" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)"  | 
| 65342 | 1360  | 
shows "list_permutes xs A"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1361  | 
proof -  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1362  | 
from assms(2,3) have "distinct (map snd xs)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1363  | 
by (intro card_distinct) (simp_all add: distinct_card del: set_map)  | 
| 65342 | 1364  | 
with assms show ?thesis  | 
1365  | 
by (simp add: list_permutes_def)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1366  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1367  | 
|
| 65342 | 1368  | 
definition permutation_of_list :: "('a \<times> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
1369  | 
where "permutation_of_list xs x = (case map_of xs x of None \<Rightarrow> x | Some y \<Rightarrow> y)"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1370  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1371  | 
lemma permutation_of_list_Cons:  | 
| 65342 | 1372  | 
"permutation_of_list ((x, y) # xs) x' = (if x = x' then y else permutation_of_list xs x')"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1373  | 
by (simp add: permutation_of_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1374  | 
|
| 65342 | 1375  | 
fun inverse_permutation_of_list :: "('a \<times> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
1376  | 
where  | 
|
1377  | 
"inverse_permutation_of_list [] x = x"  | 
|
1378  | 
| "inverse_permutation_of_list ((y, x') # xs) x =  | 
|
1379  | 
(if x = x' then y else inverse_permutation_of_list xs x)"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1380  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1381  | 
declare inverse_permutation_of_list.simps [simp del]  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1382  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1383  | 
lemma inj_on_map_of:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1384  | 
assumes "distinct (map snd xs)"  | 
| 65342 | 1385  | 
shows "inj_on (map_of xs) (set (map fst xs))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1386  | 
proof (rule inj_onI)  | 
| 65342 | 1387  | 
fix x y  | 
1388  | 
assume xy: "x \<in> set (map fst xs)" "y \<in> set (map fst xs)"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1389  | 
assume eq: "map_of xs x = map_of xs y"  | 
| 65342 | 1390  | 
from xy obtain x' y' where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'"  | 
1391  | 
by (cases "map_of xs x"; cases "map_of xs y") (simp_all add: map_of_eq_None_iff)  | 
|
1392  | 
moreover from x'y' have *: "(x, x') \<in> set xs" "(y, y') \<in> set xs"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1393  | 
by (force dest: map_of_SomeD)+  | 
| 65342 | 1394  | 
moreover from * eq x'y' have "x' = y'"  | 
1395  | 
by simp  | 
|
1396  | 
ultimately show "x = y"  | 
|
1397  | 
using assms by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"])  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1398  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1399  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1400  | 
lemma inj_on_the: "None \<notin> A \<Longrightarrow> inj_on the A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1401  | 
by (auto simp: inj_on_def option.the_def split: option.splits)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1402  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1403  | 
lemma inj_on_map_of':  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1404  | 
assumes "distinct (map snd xs)"  | 
| 65342 | 1405  | 
shows "inj_on (the \<circ> map_of xs) (set (map fst xs))"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1406  | 
by (intro comp_inj_on inj_on_map_of assms inj_on_the)  | 
| 65342 | 1407  | 
(force simp: eq_commute[of None] map_of_eq_None_iff)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1408  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1409  | 
lemma image_map_of:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1410  | 
assumes "distinct (map fst xs)"  | 
| 65342 | 1411  | 
shows "map_of xs ` set (map fst xs) = Some ` set (map snd xs)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1412  | 
using assms by (auto simp: rev_image_eqI)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1413  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1414  | 
lemma the_Some_image [simp]: "the ` Some ` A = A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1415  | 
by (subst image_image) simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1416  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1417  | 
lemma image_map_of':  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1418  | 
assumes "distinct (map fst xs)"  | 
| 65342 | 1419  | 
shows "(the \<circ> map_of xs) ` set (map fst xs) = set (map snd xs)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1420  | 
by (simp only: image_comp [symmetric] image_map_of assms the_Some_image)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1421  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1422  | 
lemma permutation_of_list_permutes [simp]:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1423  | 
assumes "list_permutes xs A"  | 
| 65342 | 1424  | 
shows "permutation_of_list xs permutes A"  | 
1425  | 
(is "?f permutes _")  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1426  | 
proof (rule permutes_subset[OF bij_imp_permutes])  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1427  | 
from assms show "set (map fst xs) \<subseteq> A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1428  | 
by (simp add: list_permutes_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1429  | 
from assms have "inj_on (the \<circ> map_of xs) (set (map fst xs))" (is ?P)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1430  | 
by (intro inj_on_map_of') (simp_all add: list_permutes_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1431  | 
also have "?P \<longleftrightarrow> inj_on ?f (set (map fst xs))"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1432  | 
by (intro inj_on_cong)  | 
| 65342 | 1433  | 
(auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1434  | 
finally have "bij_betw ?f (set (map fst xs)) (?f ` set (map fst xs))"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1435  | 
by (rule inj_on_imp_bij_betw)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1436  | 
also from assms have "?f ` set (map fst xs) = (the \<circ> map_of xs) ` set (map fst xs)"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1437  | 
by (intro image_cong refl)  | 
| 65342 | 1438  | 
(auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64267 
diff
changeset
 | 
1439  | 
also from assms have "\<dots> = set (map fst xs)"  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1440  | 
by (subst image_map_of') (simp_all add: list_permutes_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1441  | 
finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" .  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1442  | 
qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1443  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1444  | 
lemma eval_permutation_of_list [simp]:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1445  | 
"permutation_of_list [] x = x"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1446  | 
"x = x' \<Longrightarrow> permutation_of_list ((x',y)#xs) x = y"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1447  | 
"x \<noteq> x' \<Longrightarrow> permutation_of_list ((x',y')#xs) x = permutation_of_list xs x"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1448  | 
by (simp_all add: permutation_of_list_def)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1449  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1450  | 
lemma eval_inverse_permutation_of_list [simp]:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1451  | 
"inverse_permutation_of_list [] x = x"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1452  | 
"x = x' \<Longrightarrow> inverse_permutation_of_list ((y,x')#xs) x = y"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1453  | 
"x \<noteq> x' \<Longrightarrow> inverse_permutation_of_list ((y',x')#xs) x = inverse_permutation_of_list xs x"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1454  | 
by (simp_all add: inverse_permutation_of_list.simps)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1455  | 
|
| 65342 | 1456  | 
lemma permutation_of_list_id: "x \<notin> set (map fst xs) \<Longrightarrow> permutation_of_list xs x = x"  | 
1457  | 
by (induct xs) (auto simp: permutation_of_list_Cons)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1458  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1459  | 
lemma permutation_of_list_unique':  | 
| 65342 | 1460  | 
"distinct (map fst xs) \<Longrightarrow> (x, y) \<in> set xs \<Longrightarrow> permutation_of_list xs x = y"  | 
1461  | 
by (induct xs) (force simp: permutation_of_list_Cons)+  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1462  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1463  | 
lemma permutation_of_list_unique:  | 
| 65342 | 1464  | 
"list_permutes xs A \<Longrightarrow> (x, y) \<in> set xs \<Longrightarrow> permutation_of_list xs x = y"  | 
1465  | 
by (intro permutation_of_list_unique') (simp_all add: list_permutes_def)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1466  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1467  | 
lemma inverse_permutation_of_list_id:  | 
| 65342 | 1468  | 
"x \<notin> set (map snd xs) \<Longrightarrow> inverse_permutation_of_list xs x = x"  | 
1469  | 
by (induct xs) auto  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1470  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1471  | 
lemma inverse_permutation_of_list_unique':  | 
| 65342 | 1472  | 
"distinct (map snd xs) \<Longrightarrow> (x, y) \<in> set xs \<Longrightarrow> inverse_permutation_of_list xs y = x"  | 
1473  | 
by (induct xs) (force simp: inverse_permutation_of_list.simps)+  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1474  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1475  | 
lemma inverse_permutation_of_list_unique:  | 
| 65342 | 1476  | 
"list_permutes xs A \<Longrightarrow> (x,y) \<in> set xs \<Longrightarrow> inverse_permutation_of_list xs y = x"  | 
1477  | 
by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def)  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1478  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1479  | 
lemma inverse_permutation_of_list_correct:  | 
| 65342 | 1480  | 
fixes A :: "'a set"  | 
1481  | 
assumes "list_permutes xs A"  | 
|
1482  | 
shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)"  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1483  | 
proof (rule ext, rule sym, subst permutes_inv_eq)  | 
| 65342 | 1484  | 
from assms show "permutation_of_list xs permutes A"  | 
1485  | 
by simp  | 
|
1486  | 
show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" for x  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1487  | 
proof (cases "x \<in> set (map snd xs)")  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1488  | 
case True  | 
| 65342 | 1489  | 
then obtain y where "(y, x) \<in> set xs" by auto  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1490  | 
with assms show ?thesis  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1491  | 
by (simp add: inverse_permutation_of_list_unique permutation_of_list_unique)  | 
| 65342 | 1492  | 
next  | 
1493  | 
case False  | 
|
1494  | 
with assms show ?thesis  | 
|
1495  | 
by (auto simp: list_permutes_def inverse_permutation_of_list_id permutation_of_list_id)  | 
|
1496  | 
qed  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1497  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
62390 
diff
changeset
 | 
1498  | 
|
| 
29840
 
cfab6a76aa13
Permutations, both general and specifically on finite sets.
 
chaieb 
parents:  
diff
changeset
 | 
1499  | 
end  |