wenzelm@41959
|
1 |
(* Title: HOL/Library/Permutations.thy
|
wenzelm@41959
|
2 |
Author: Amine Chaieb, University of Cambridge
|
chaieb@29840
|
3 |
*)
|
chaieb@29840
|
4 |
|
chaieb@29840
|
5 |
header {* Permutations, both general and specifically on finite sets.*}
|
chaieb@29840
|
6 |
|
chaieb@29840
|
7 |
theory Permutations
|
huffman@36335
|
8 |
imports Parity Fact
|
chaieb@29840
|
9 |
begin
|
chaieb@29840
|
10 |
|
chaieb@29840
|
11 |
definition permutes (infixr "permutes" 41) where
|
chaieb@29840
|
12 |
"(p permutes S) \<longleftrightarrow> (\<forall>x. x \<notin> S \<longrightarrow> p x = x) \<and> (\<forall>y. \<exists>!x. p x = y)"
|
chaieb@29840
|
13 |
|
chaieb@29840
|
14 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
15 |
(* Transpositions. *)
|
chaieb@29840
|
16 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
17 |
|
huffman@30488
|
18 |
lemma swapid_sym: "Fun.swap a b id = Fun.swap b a id"
|
nipkow@39302
|
19 |
by (auto simp add: fun_eq_iff swap_def fun_upd_def)
|
chaieb@29840
|
20 |
lemma swap_id_refl: "Fun.swap a a id = id" by simp
|
chaieb@29840
|
21 |
lemma swap_id_sym: "Fun.swap a b id = Fun.swap b a id"
|
chaieb@29840
|
22 |
by (rule ext, simp add: swap_def)
|
chaieb@29840
|
23 |
lemma swap_id_idempotent[simp]: "Fun.swap a b id o Fun.swap a b id = id"
|
chaieb@29840
|
24 |
by (rule ext, auto simp add: swap_def)
|
chaieb@29840
|
25 |
|
chaieb@29840
|
26 |
lemma inv_unique_comp: assumes fg: "f o g = id" and gf: "g o f = id"
|
chaieb@29840
|
27 |
shows "inv f = g"
|
nipkow@39302
|
28 |
using fg gf inv_equality[of g f] by (auto simp add: fun_eq_iff)
|
chaieb@29840
|
29 |
|
chaieb@29840
|
30 |
lemma inverse_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id"
|
chaieb@29840
|
31 |
by (rule inv_unique_comp, simp_all)
|
chaieb@29840
|
32 |
|
chaieb@29840
|
33 |
lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)"
|
chaieb@29840
|
34 |
by (simp add: swap_def)
|
chaieb@29840
|
35 |
|
chaieb@29840
|
36 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
37 |
(* Basic consequences of the definition. *)
|
chaieb@29840
|
38 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
39 |
|
chaieb@29840
|
40 |
lemma permutes_in_image: "p permutes S \<Longrightarrow> p x \<in> S \<longleftrightarrow> x \<in> S"
|
chaieb@29840
|
41 |
unfolding permutes_def by metis
|
chaieb@29840
|
42 |
|
chaieb@29840
|
43 |
lemma permutes_image: assumes pS: "p permutes S" shows "p ` S = S"
|
chaieb@29840
|
44 |
using pS
|
huffman@30488
|
45 |
unfolding permutes_def
|
huffman@30488
|
46 |
apply -
|
nipkow@39302
|
47 |
apply (rule set_eqI)
|
chaieb@29840
|
48 |
apply (simp add: image_iff)
|
chaieb@29840
|
49 |
apply metis
|
chaieb@29840
|
50 |
done
|
chaieb@29840
|
51 |
|
huffman@30488
|
52 |
lemma permutes_inj: "p permutes S ==> inj p "
|
huffman@30488
|
53 |
unfolding permutes_def inj_on_def by blast
|
chaieb@29840
|
54 |
|
huffman@30488
|
55 |
lemma permutes_surj: "p permutes s ==> surj p"
|
huffman@30488
|
56 |
unfolding permutes_def surj_def by metis
|
chaieb@29840
|
57 |
|
chaieb@29840
|
58 |
lemma permutes_inv_o: assumes pS: "p permutes S"
|
chaieb@29840
|
59 |
shows " p o inv p = id"
|
chaieb@29840
|
60 |
and "inv p o p = id"
|
chaieb@29840
|
61 |
using permutes_inj[OF pS] permutes_surj[OF pS]
|
chaieb@29840
|
62 |
unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+
|
chaieb@29840
|
63 |
|
chaieb@29840
|
64 |
|
huffman@30488
|
65 |
lemma permutes_inverses:
|
chaieb@29840
|
66 |
fixes p :: "'a \<Rightarrow> 'a"
|
chaieb@29840
|
67 |
assumes pS: "p permutes S"
|
chaieb@29840
|
68 |
shows "p (inv p x) = x"
|
chaieb@29840
|
69 |
and "inv p (p x) = x"
|
nipkow@39302
|
70 |
using permutes_inv_o[OF pS, unfolded fun_eq_iff o_def] by auto
|
chaieb@29840
|
71 |
|
chaieb@29840
|
72 |
lemma permutes_subset: "p permutes S \<Longrightarrow> S \<subseteq> T ==> p permutes T"
|
chaieb@29840
|
73 |
unfolding permutes_def by blast
|
chaieb@29840
|
74 |
|
chaieb@29840
|
75 |
lemma permutes_empty[simp]: "p permutes {} \<longleftrightarrow> p = id"
|
nipkow@39302
|
76 |
unfolding fun_eq_iff permutes_def apply simp by metis
|
chaieb@29840
|
77 |
|
chaieb@29840
|
78 |
lemma permutes_sing[simp]: "p permutes {a} \<longleftrightarrow> p = id"
|
nipkow@39302
|
79 |
unfolding fun_eq_iff permutes_def apply simp by metis
|
huffman@30488
|
80 |
|
chaieb@29840
|
81 |
lemma permutes_univ: "p permutes UNIV \<longleftrightarrow> (\<forall>y. \<exists>!x. p x = y)"
|
chaieb@29840
|
82 |
unfolding permutes_def by simp
|
chaieb@29840
|
83 |
|
chaieb@29840
|
84 |
lemma permutes_inv_eq: "p permutes S ==> inv p y = x \<longleftrightarrow> p x = y"
|
nipkow@33057
|
85 |
unfolding permutes_def inv_def apply auto
|
chaieb@29840
|
86 |
apply (erule allE[where x=y])
|
chaieb@29840
|
87 |
apply (erule allE[where x=y])
|
chaieb@29840
|
88 |
apply (rule someI_ex) apply blast
|
chaieb@29840
|
89 |
apply (rule some1_equality)
|
chaieb@29840
|
90 |
apply blast
|
chaieb@29840
|
91 |
apply blast
|
chaieb@29840
|
92 |
done
|
chaieb@29840
|
93 |
|
chaieb@29840
|
94 |
lemma permutes_swap_id: "a \<in> S \<Longrightarrow> b \<in> S ==> Fun.swap a b id permutes S"
|
nipkow@32988
|
95 |
unfolding permutes_def swap_def fun_upd_def by auto metis
|
chaieb@29840
|
96 |
|
nipkow@32988
|
97 |
lemma permutes_superset:
|
nipkow@32988
|
98 |
"p permutes S \<Longrightarrow> (\<forall>x \<in> S - T. p x = x) \<Longrightarrow> p permutes T"
|
huffman@36361
|
99 |
by (simp add: Ball_def permutes_def) metis
|
chaieb@29840
|
100 |
|
chaieb@29840
|
101 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
102 |
(* Group properties. *)
|
chaieb@29840
|
103 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
104 |
|
huffman@30488
|
105 |
lemma permutes_id: "id permutes S" unfolding permutes_def by simp
|
chaieb@29840
|
106 |
|
chaieb@29840
|
107 |
lemma permutes_compose: "p permutes S \<Longrightarrow> q permutes S ==> q o p permutes S"
|
chaieb@29840
|
108 |
unfolding permutes_def o_def by metis
|
chaieb@29840
|
109 |
|
chaieb@29840
|
110 |
lemma permutes_inv: assumes pS: "p permutes S" shows "inv p permutes S"
|
huffman@30488
|
111 |
using pS unfolding permutes_def permutes_inv_eq[OF pS] by metis
|
chaieb@29840
|
112 |
|
chaieb@29840
|
113 |
lemma permutes_inv_inv: assumes pS: "p permutes S" shows "inv (inv p) = p"
|
nipkow@39302
|
114 |
unfolding fun_eq_iff permutes_inv_eq[OF pS] permutes_inv_eq[OF permutes_inv[OF pS]]
|
chaieb@29840
|
115 |
by blast
|
chaieb@29840
|
116 |
|
chaieb@29840
|
117 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
118 |
(* The number of permutations on a finite set. *)
|
chaieb@29840
|
119 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
120 |
|
huffman@30488
|
121 |
lemma permutes_insert_lemma:
|
chaieb@29840
|
122 |
assumes pS: "p permutes (insert a S)"
|
chaieb@29840
|
123 |
shows "Fun.swap a (p a) id o p permutes S"
|
chaieb@29840
|
124 |
apply (rule permutes_superset[where S = "insert a S"])
|
chaieb@29840
|
125 |
apply (rule permutes_compose[OF pS])
|
chaieb@29840
|
126 |
apply (rule permutes_swap_id, simp)
|
chaieb@29840
|
127 |
using permutes_in_image[OF pS, of a] apply simp
|
huffman@36361
|
128 |
apply (auto simp add: Ball_def swap_def)
|
chaieb@29840
|
129 |
done
|
chaieb@29840
|
130 |
|
chaieb@29840
|
131 |
lemma permutes_insert: "{p. p permutes (insert a S)} =
|
chaieb@29840
|
132 |
(\<lambda>(b,p). Fun.swap a b id o p) ` {(b,p). b \<in> insert a S \<and> p \<in> {p. p permutes S}}"
|
chaieb@29840
|
133 |
proof-
|
chaieb@29840
|
134 |
|
huffman@30488
|
135 |
{fix p
|
chaieb@29840
|
136 |
{assume pS: "p permutes insert a S"
|
chaieb@29840
|
137 |
let ?b = "p a"
|
chaieb@29840
|
138 |
let ?q = "Fun.swap a (p a) id o p"
|
nipkow@39302
|
139 |
have th0: "p = Fun.swap a ?b id o ?q" unfolding fun_eq_iff o_assoc by simp
|
huffman@30488
|
140 |
have th1: "?b \<in> insert a S " unfolding permutes_in_image[OF pS] by simp
|
chaieb@29840
|
141 |
from permutes_insert_lemma[OF pS] th0 th1
|
chaieb@29840
|
142 |
have "\<exists> b q. p = Fun.swap a b id o q \<and> b \<in> insert a S \<and> q permutes S" by blast}
|
chaieb@29840
|
143 |
moreover
|
chaieb@29840
|
144 |
{fix b q assume bq: "p = Fun.swap a b id o q" "b \<in> insert a S" "q permutes S"
|
huffman@30488
|
145 |
from permutes_subset[OF bq(3), of "insert a S"]
|
chaieb@29840
|
146 |
have qS: "q permutes insert a S" by auto
|
chaieb@29840
|
147 |
have aS: "a \<in> insert a S" by simp
|
chaieb@29840
|
148 |
from bq(1) permutes_compose[OF qS permutes_swap_id[OF aS bq(2)]]
|
chaieb@29840
|
149 |
have "p permutes insert a S" by simp }
|
chaieb@29840
|
150 |
ultimately have "p permutes insert a S \<longleftrightarrow> (\<exists> b q. p = Fun.swap a b id o q \<and> b \<in> insert a S \<and> q permutes S)" by blast}
|
chaieb@29840
|
151 |
thus ?thesis by auto
|
chaieb@29840
|
152 |
qed
|
chaieb@29840
|
153 |
|
hoelzl@33715
|
154 |
lemma card_permutations: assumes Sn: "card S = n" and fS: "finite S"
|
hoelzl@33715
|
155 |
shows "card {p. p permutes S} = fact n"
|
hoelzl@33715
|
156 |
using fS Sn proof (induct arbitrary: n)
|
huffman@36361
|
157 |
case empty thus ?case by simp
|
hoelzl@33715
|
158 |
next
|
hoelzl@33715
|
159 |
case (insert x F)
|
hoelzl@33715
|
160 |
{ fix n assume H0: "card (insert x F) = n"
|
hoelzl@33715
|
161 |
let ?xF = "{p. p permutes insert x F}"
|
hoelzl@33715
|
162 |
let ?pF = "{p. p permutes F}"
|
hoelzl@33715
|
163 |
let ?pF' = "{(b, p). b \<in> insert x F \<and> p \<in> ?pF}"
|
hoelzl@33715
|
164 |
let ?g = "(\<lambda>(b, p). Fun.swap x b id \<circ> p)"
|
hoelzl@33715
|
165 |
from permutes_insert[of x F]
|
hoelzl@33715
|
166 |
have xfgpF': "?xF = ?g ` ?pF'" .
|
hoelzl@33715
|
167 |
have Fs: "card F = n - 1" using `x \<notin> F` H0 `finite F` by auto
|
hoelzl@33715
|
168 |
from insert.hyps Fs have pFs: "card ?pF = fact (n - 1)" using `finite F` by auto
|
wenzelm@53374
|
169 |
then have "finite ?pF" using fact_gt_zero_nat by (auto intro: card_ge_0_finite)
|
wenzelm@53374
|
170 |
then have pF'f: "finite ?pF'" using H0 `finite F`
|
hoelzl@33715
|
171 |
apply (simp only: Collect_split Collect_mem_eq)
|
hoelzl@33715
|
172 |
apply (rule finite_cartesian_product)
|
hoelzl@33715
|
173 |
apply simp_all
|
hoelzl@33715
|
174 |
done
|
chaieb@29840
|
175 |
|
hoelzl@33715
|
176 |
have ginj: "inj_on ?g ?pF'"
|
hoelzl@33715
|
177 |
proof-
|
hoelzl@33715
|
178 |
{
|
hoelzl@33715
|
179 |
fix b p c q assume bp: "(b,p) \<in> ?pF'" and cq: "(c,q) \<in> ?pF'"
|
hoelzl@33715
|
180 |
and eq: "?g (b,p) = ?g (c,q)"
|
hoelzl@33715
|
181 |
from bp cq have ths: "b \<in> insert x F" "c \<in> insert x F" "x \<in> insert x F" "p permutes F" "q permutes F" by auto
|
hoelzl@33715
|
182 |
from ths(4) `x \<notin> F` eq have "b = ?g (b,p) x" unfolding permutes_def
|
nipkow@39302
|
183 |
by (auto simp add: swap_def fun_upd_def fun_eq_iff)
|
hoelzl@33715
|
184 |
also have "\<dots> = ?g (c,q) x" using ths(5) `x \<notin> F` eq
|
nipkow@39302
|
185 |
by (auto simp add: swap_def fun_upd_def fun_eq_iff)
|
hoelzl@33715
|
186 |
also have "\<dots> = c"using ths(5) `x \<notin> F` unfolding permutes_def
|
nipkow@39302
|
187 |
by (auto simp add: swap_def fun_upd_def fun_eq_iff)
|
hoelzl@33715
|
188 |
finally have bc: "b = c" .
|
hoelzl@33715
|
189 |
hence "Fun.swap x b id = Fun.swap x c id" by simp
|
hoelzl@33715
|
190 |
with eq have "Fun.swap x b id o p = Fun.swap x b id o q" by simp
|
hoelzl@33715
|
191 |
hence "Fun.swap x b id o (Fun.swap x b id o p) = Fun.swap x b id o (Fun.swap x b id o q)" by simp
|
hoelzl@33715
|
192 |
hence "p = q" by (simp add: o_assoc)
|
hoelzl@33715
|
193 |
with bc have "(b,p) = (c,q)" by simp
|
hoelzl@33715
|
194 |
}
|
hoelzl@33715
|
195 |
thus ?thesis unfolding inj_on_def by blast
|
hoelzl@33715
|
196 |
qed
|
hoelzl@33715
|
197 |
from `x \<notin> F` H0 have n0: "n \<noteq> 0 " using `finite F` by auto
|
wenzelm@53374
|
198 |
hence "\<exists>m. n = Suc m" by presburger
|
hoelzl@33715
|
199 |
then obtain m where n[simp]: "n = Suc m" by blast
|
hoelzl@33715
|
200 |
from pFs H0 have xFc: "card ?xF = fact n"
|
hoelzl@33715
|
201 |
unfolding xfgpF' card_image[OF ginj] using `finite F` `finite ?pF`
|
hoelzl@33715
|
202 |
apply (simp only: Collect_split Collect_mem_eq card_cartesian_product)
|
hoelzl@33715
|
203 |
by simp
|
hoelzl@33715
|
204 |
from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" unfolding xfgpF' by simp
|
hoelzl@33715
|
205 |
have "card ?xF = fact n"
|
hoelzl@33715
|
206 |
using xFf xFc unfolding xFf by blast
|
hoelzl@33715
|
207 |
}
|
hoelzl@33715
|
208 |
thus ?case using insert by simp
|
chaieb@29840
|
209 |
qed
|
chaieb@29840
|
210 |
|
hoelzl@33715
|
211 |
lemma finite_permutations: assumes fS: "finite S" shows "finite {p. p permutes S}"
|
hoelzl@33715
|
212 |
using card_permutations[OF refl fS] fact_gt_zero_nat
|
hoelzl@33715
|
213 |
by (auto intro: card_ge_0_finite)
|
chaieb@29840
|
214 |
|
chaieb@29840
|
215 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
216 |
(* Permutations of index set for iterated operations. *)
|
chaieb@29840
|
217 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
218 |
|
haftmann@51489
|
219 |
lemma (in comm_monoid_set) permute:
|
haftmann@51489
|
220 |
assumes "p permutes S"
|
haftmann@51489
|
221 |
shows "F g S = F (g o p) S"
|
haftmann@51489
|
222 |
proof -
|
haftmann@51489
|
223 |
from `p permutes S` have "inj p" by (rule permutes_inj)
|
haftmann@51489
|
224 |
then have "inj_on p S" by (auto intro: subset_inj_on)
|
haftmann@51489
|
225 |
then have "F g (p ` S) = F (g o p) S" by (rule reindex)
|
haftmann@51489
|
226 |
moreover from `p permutes S` have "p ` S = S" by (rule permutes_image)
|
haftmann@51489
|
227 |
ultimately show ?thesis by simp
|
chaieb@29840
|
228 |
qed
|
chaieb@29840
|
229 |
|
haftmann@51489
|
230 |
lemma setsum_permute:
|
haftmann@51489
|
231 |
assumes "p permutes S"
|
chaieb@29840
|
232 |
shows "setsum f S = setsum (f o p) S"
|
haftmann@51489
|
233 |
using assms by (fact setsum.permute)
|
chaieb@29840
|
234 |
|
haftmann@51489
|
235 |
lemma setsum_permute_natseg:
|
haftmann@51489
|
236 |
assumes pS: "p permutes {m .. n}"
|
chaieb@29840
|
237 |
shows "setsum f {m .. n} = setsum (f o p) {m .. n}"
|
haftmann@51489
|
238 |
using setsum_permute [OF pS, of f ] pS by blast
|
chaieb@29840
|
239 |
|
haftmann@51489
|
240 |
lemma setprod_permute:
|
haftmann@51489
|
241 |
assumes "p permutes S"
|
chaieb@29840
|
242 |
shows "setprod f S = setprod (f o p) S"
|
haftmann@51489
|
243 |
using assms by (fact setprod.permute)
|
chaieb@29840
|
244 |
|
haftmann@51489
|
245 |
lemma setprod_permute_natseg:
|
haftmann@51489
|
246 |
assumes pS: "p permutes {m .. n}"
|
chaieb@29840
|
247 |
shows "setprod f {m .. n} = setprod (f o p) {m .. n}"
|
haftmann@51489
|
248 |
using setprod_permute [OF pS, of f ] pS by blast
|
chaieb@29840
|
249 |
|
chaieb@29840
|
250 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
251 |
(* Various combinations of transpositions with 2, 1 and 0 common elements. *)
|
chaieb@29840
|
252 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
253 |
|
nipkow@39302
|
254 |
lemma swap_id_common:" a \<noteq> c \<Longrightarrow> b \<noteq> c \<Longrightarrow> Fun.swap a b id o Fun.swap a c id = Fun.swap b c id o Fun.swap a b id" by (simp add: fun_eq_iff swap_def)
|
chaieb@29840
|
255 |
|
nipkow@39302
|
256 |
lemma swap_id_common': "~(a = b) \<Longrightarrow> ~(a = c) \<Longrightarrow> Fun.swap a c id o Fun.swap b c id = Fun.swap b c id o Fun.swap a b id" by (simp add: fun_eq_iff swap_def)
|
chaieb@29840
|
257 |
|
chaieb@29840
|
258 |
lemma swap_id_independent: "~(a = c) \<Longrightarrow> ~(a = d) \<Longrightarrow> ~(b = c) \<Longrightarrow> ~(b = d) ==> Fun.swap a b id o Fun.swap c d id = Fun.swap c d id o Fun.swap a b id"
|
nipkow@39302
|
259 |
by (simp add: swap_def fun_eq_iff)
|
chaieb@29840
|
260 |
|
chaieb@29840
|
261 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
262 |
(* Permutations as transposition sequences. *)
|
chaieb@29840
|
263 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
264 |
|
chaieb@29840
|
265 |
|
chaieb@29840
|
266 |
inductive swapidseq :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
|
chaieb@29840
|
267 |
id[simp]: "swapidseq 0 id"
|
chaieb@29840
|
268 |
| comp_Suc: "swapidseq n p \<Longrightarrow> a \<noteq> b \<Longrightarrow> swapidseq (Suc n) (Fun.swap a b id o p)"
|
chaieb@29840
|
269 |
|
chaieb@29840
|
270 |
declare id[unfolded id_def, simp]
|
chaieb@29840
|
271 |
definition "permutation p \<longleftrightarrow> (\<exists>n. swapidseq n p)"
|
chaieb@29840
|
272 |
|
chaieb@29840
|
273 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
274 |
(* Some closure properties of the set of permutations, with lengths. *)
|
chaieb@29840
|
275 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
276 |
|
chaieb@29840
|
277 |
lemma permutation_id[simp]: "permutation id"unfolding permutation_def
|
chaieb@29840
|
278 |
by (rule exI[where x=0], simp)
|
chaieb@29840
|
279 |
declare permutation_id[unfolded id_def, simp]
|
chaieb@29840
|
280 |
|
chaieb@29840
|
281 |
lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)"
|
chaieb@29840
|
282 |
apply clarsimp
|
chaieb@29840
|
283 |
using comp_Suc[of 0 id a b] by simp
|
chaieb@29840
|
284 |
|
chaieb@29840
|
285 |
lemma permutation_swap_id: "permutation (Fun.swap a b id)"
|
chaieb@29840
|
286 |
apply (cases "a=b", simp_all)
|
huffman@30488
|
287 |
unfolding permutation_def using swapidseq_swap[of a b] by blast
|
chaieb@29840
|
288 |
|
chaieb@29840
|
289 |
lemma swapidseq_comp_add: "swapidseq n p \<Longrightarrow> swapidseq m q ==> swapidseq (n + m) (p o q)"
|
chaieb@29840
|
290 |
proof (induct n p arbitrary: m q rule: swapidseq.induct)
|
chaieb@29840
|
291 |
case (id m q) thus ?case by simp
|
chaieb@29840
|
292 |
next
|
huffman@30488
|
293 |
case (comp_Suc n p a b m q)
|
chaieb@29840
|
294 |
have th: "Suc n + m = Suc (n + m)" by arith
|
haftmann@49739
|
295 |
show ?case unfolding th comp_assoc
|
huffman@30488
|
296 |
apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3) by blast+
|
chaieb@29840
|
297 |
qed
|
chaieb@29840
|
298 |
|
chaieb@29840
|
299 |
lemma permutation_compose: "permutation p \<Longrightarrow> permutation q ==> permutation(p o q)"
|
chaieb@29840
|
300 |
unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis
|
chaieb@29840
|
301 |
|
chaieb@29840
|
302 |
lemma swapidseq_endswap: "swapidseq n p \<Longrightarrow> a \<noteq> b ==> swapidseq (Suc n) (p o Fun.swap a b id)"
|
chaieb@29840
|
303 |
apply (induct n p rule: swapidseq.induct)
|
chaieb@29840
|
304 |
using swapidseq_swap[of a b]
|
haftmann@49739
|
305 |
by (auto simp add: comp_assoc intro: swapidseq.comp_Suc)
|
chaieb@29840
|
306 |
|
chaieb@29840
|
307 |
lemma swapidseq_inverse_exists: "swapidseq n p ==> \<exists>q. swapidseq n q \<and> p o q = id \<and> q o p = id"
|
chaieb@29840
|
308 |
proof(induct n p rule: swapidseq.induct)
|
chaieb@29840
|
309 |
case id thus ?case by (rule exI[where x=id], simp)
|
huffman@30488
|
310 |
next
|
chaieb@29840
|
311 |
case (comp_Suc n p a b)
|
chaieb@29840
|
312 |
from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
|
chaieb@29840
|
313 |
let ?q = "q o Fun.swap a b id"
|
chaieb@29840
|
314 |
note H = comp_Suc.hyps
|
chaieb@29840
|
315 |
from swapidseq_swap[of a b] H(3) have th0: "swapidseq 1 (Fun.swap a b id)" by simp
|
huffman@30488
|
316 |
from swapidseq_comp_add[OF q(1) th0] have th1:"swapidseq (Suc n) ?q" by simp
|
chaieb@29840
|
317 |
have "Fun.swap a b id o p o ?q = Fun.swap a b id o (p o q) o Fun.swap a b id" by (simp add: o_assoc)
|
chaieb@29840
|
318 |
also have "\<dots> = id" by (simp add: q(2))
|
chaieb@29840
|
319 |
finally have th2: "Fun.swap a b id o p o ?q = id" .
|
huffman@30488
|
320 |
have "?q \<circ> (Fun.swap a b id \<circ> p) = q \<circ> (Fun.swap a b id o Fun.swap a b id) \<circ> p" by (simp only: o_assoc)
|
chaieb@29840
|
321 |
hence "?q \<circ> (Fun.swap a b id \<circ> p) = id" by (simp add: q(3))
|
chaieb@29840
|
322 |
with th1 th2 show ?case by blast
|
chaieb@29840
|
323 |
qed
|
chaieb@29840
|
324 |
|
chaieb@29840
|
325 |
|
chaieb@29840
|
326 |
lemma swapidseq_inverse: assumes H: "swapidseq n p" shows "swapidseq n (inv p)"
|
chaieb@29840
|
327 |
using swapidseq_inverse_exists[OF H] inv_unique_comp[of p] by auto
|
chaieb@29840
|
328 |
|
chaieb@29840
|
329 |
lemma permutation_inverse: "permutation p ==> permutation (inv p)"
|
chaieb@29840
|
330 |
using permutation_def swapidseq_inverse by blast
|
chaieb@29840
|
331 |
|
chaieb@29840
|
332 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
333 |
(* The identity map only has even transposition sequences. *)
|
chaieb@29840
|
334 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
335 |
|
chaieb@29840
|
336 |
lemma symmetry_lemma:"(\<And>a b c d. P a b c d ==> P a b d c) \<Longrightarrow>
|
chaieb@29840
|
337 |
(\<And>a b c d. a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> (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) ==> P a b c d)
|
chaieb@29840
|
338 |
==> (\<And>a b c d. a \<noteq> b --> c \<noteq> d \<longrightarrow> P a b c d)" by metis
|
chaieb@29840
|
339 |
|
huffman@30488
|
340 |
lemma swap_general: "a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> Fun.swap a b id o Fun.swap c d id = id \<or>
|
huffman@30488
|
341 |
(\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id)"
|
chaieb@29840
|
342 |
proof-
|
chaieb@29840
|
343 |
assume H: "a\<noteq>b" "c\<noteq>d"
|
huffman@30488
|
344 |
have "a \<noteq> b \<longrightarrow> c \<noteq> d \<longrightarrow>
|
huffman@30488
|
345 |
( Fun.swap a b id o Fun.swap c d id = id \<or>
|
huffman@30488
|
346 |
(\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id))"
|
chaieb@29840
|
347 |
apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d])
|
huffman@30488
|
348 |
apply (simp_all only: swapid_sym)
|
chaieb@29840
|
349 |
apply (case_tac "a = c \<and> b = d", clarsimp simp only: swapid_sym swap_id_idempotent)
|
chaieb@29840
|
350 |
apply (case_tac "a = c \<and> b \<noteq> d")
|
chaieb@29840
|
351 |
apply (rule disjI2)
|
chaieb@29840
|
352 |
apply (rule_tac x="b" in exI)
|
chaieb@29840
|
353 |
apply (rule_tac x="d" in exI)
|
chaieb@29840
|
354 |
apply (rule_tac x="b" in exI)
|
nipkow@39302
|
355 |
apply (clarsimp simp add: fun_eq_iff swap_def)
|
chaieb@29840
|
356 |
apply (case_tac "a \<noteq> c \<and> b = d")
|
chaieb@29840
|
357 |
apply (rule disjI2)
|
chaieb@29840
|
358 |
apply (rule_tac x="c" in exI)
|
chaieb@29840
|
359 |
apply (rule_tac x="d" in exI)
|
chaieb@29840
|
360 |
apply (rule_tac x="c" in exI)
|
nipkow@39302
|
361 |
apply (clarsimp simp add: fun_eq_iff swap_def)
|
chaieb@29840
|
362 |
apply (rule disjI2)
|
chaieb@29840
|
363 |
apply (rule_tac x="c" in exI)
|
chaieb@29840
|
364 |
apply (rule_tac x="d" in exI)
|
chaieb@29840
|
365 |
apply (rule_tac x="b" in exI)
|
nipkow@39302
|
366 |
apply (clarsimp simp add: fun_eq_iff swap_def)
|
chaieb@29840
|
367 |
done
|
huffman@30488
|
368 |
with H show ?thesis by metis
|
chaieb@29840
|
369 |
qed
|
chaieb@29840
|
370 |
|
chaieb@29840
|
371 |
lemma swapidseq_id_iff[simp]: "swapidseq 0 p \<longleftrightarrow> p = id"
|
chaieb@29840
|
372 |
using swapidseq.cases[of 0 p "p = id"]
|
chaieb@29840
|
373 |
by auto
|
chaieb@29840
|
374 |
|
chaieb@29840
|
375 |
lemma swapidseq_cases: "swapidseq n p \<longleftrightarrow> (n=0 \<and> p = id \<or> (\<exists>a b q m. n = Suc m \<and> p = Fun.swap a b id o q \<and> swapidseq m q \<and> a\<noteq> b))"
|
chaieb@29840
|
376 |
apply (rule iffI)
|
chaieb@29840
|
377 |
apply (erule swapidseq.cases[of n p])
|
chaieb@29840
|
378 |
apply simp
|
chaieb@29840
|
379 |
apply (rule disjI2)
|
chaieb@29840
|
380 |
apply (rule_tac x= "a" in exI)
|
chaieb@29840
|
381 |
apply (rule_tac x= "b" in exI)
|
chaieb@29840
|
382 |
apply (rule_tac x= "pa" in exI)
|
chaieb@29840
|
383 |
apply (rule_tac x= "na" in exI)
|
chaieb@29840
|
384 |
apply simp
|
chaieb@29840
|
385 |
apply auto
|
chaieb@29840
|
386 |
apply (rule comp_Suc, simp_all)
|
chaieb@29840
|
387 |
done
|
chaieb@29840
|
388 |
lemma fixing_swapidseq_decrease:
|
chaieb@29840
|
389 |
assumes spn: "swapidseq n p" and ab: "a\<noteq>b" and pa: "(Fun.swap a b id o p) a = a"
|
chaieb@29840
|
390 |
shows "n \<noteq> 0 \<and> swapidseq (n - 1) (Fun.swap a b id o p)"
|
chaieb@29840
|
391 |
using spn ab pa
|
chaieb@29840
|
392 |
proof(induct n arbitrary: p a b)
|
chaieb@29840
|
393 |
case 0 thus ?case by (auto simp add: swap_def fun_upd_def)
|
chaieb@29840
|
394 |
next
|
chaieb@29840
|
395 |
case (Suc n p a b)
|
chaieb@29840
|
396 |
from Suc.prems(1) swapidseq_cases[of "Suc n" p] obtain
|
chaieb@29840
|
397 |
c d q m where cdqm: "Suc n = Suc m" "p = Fun.swap c d id o q" "swapidseq m q" "c \<noteq> d" "n = m"
|
chaieb@29840
|
398 |
by auto
|
chaieb@29840
|
399 |
{assume H: "Fun.swap a b id o Fun.swap c d id = id"
|
huffman@30488
|
400 |
|
huffman@30488
|
401 |
have ?case apply (simp only: cdqm o_assoc H)
|
chaieb@29840
|
402 |
by (simp add: cdqm)}
|
chaieb@29840
|
403 |
moreover
|
chaieb@29840
|
404 |
{ fix x y z
|
huffman@30488
|
405 |
assume H: "x\<noteq>a" "y\<noteq>a" "z \<noteq>a" "x \<noteq>y"
|
chaieb@29840
|
406 |
"Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id"
|
chaieb@29840
|
407 |
from H have az: "a \<noteq> z" by simp
|
chaieb@29840
|
408 |
|
chaieb@29840
|
409 |
{fix h have "(Fun.swap x y id o h) a = a \<longleftrightarrow> h a = a"
|
chaieb@29840
|
410 |
using H by (simp add: swap_def)}
|
chaieb@29840
|
411 |
note th3 = this
|
chaieb@29840
|
412 |
from cdqm(2) have "Fun.swap a b id o p = Fun.swap a b id o (Fun.swap c d id o q)" by simp
|
chaieb@29840
|
413 |
hence "Fun.swap a b id o p = Fun.swap x y id o (Fun.swap a z id o q)" by (simp add: o_assoc H)
|
chaieb@29840
|
414 |
hence "(Fun.swap a b id o p) a = (Fun.swap x y id o (Fun.swap a z id o q)) a" by simp
|
chaieb@29840
|
415 |
hence "(Fun.swap x y id o (Fun.swap a z id o q)) a = a" unfolding Suc by metis
|
chaieb@29840
|
416 |
hence th1: "(Fun.swap a z id o q) a = a" unfolding th3 .
|
chaieb@29840
|
417 |
from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az th1]
|
chaieb@29840
|
418 |
have th2: "swapidseq (n - 1) (Fun.swap a z id o q)" "n \<noteq> 0" by blast+
|
huffman@30488
|
419 |
have th: "Suc n - 1 = Suc (n - 1)" using th2(2) by auto
|
chaieb@29840
|
420 |
have ?case unfolding cdqm(2) H o_assoc th
|
haftmann@49739
|
421 |
apply (simp only: Suc_not_Zero simp_thms comp_assoc)
|
chaieb@29840
|
422 |
apply (rule comp_Suc)
|
chaieb@29840
|
423 |
using th2 H apply blast+
|
chaieb@29840
|
424 |
done}
|
huffman@30488
|
425 |
ultimately show ?case using swap_general[OF Suc.prems(2) cdqm(4)] by metis
|
chaieb@29840
|
426 |
qed
|
chaieb@29840
|
427 |
|
huffman@30488
|
428 |
lemma swapidseq_identity_even:
|
chaieb@29840
|
429 |
assumes "swapidseq n (id :: 'a \<Rightarrow> 'a)" shows "even n"
|
chaieb@29840
|
430 |
using `swapidseq n id`
|
chaieb@29840
|
431 |
proof(induct n rule: nat_less_induct)
|
chaieb@29840
|
432 |
fix n
|
chaieb@29840
|
433 |
assume H: "\<forall>m<n. swapidseq m (id::'a \<Rightarrow> 'a) \<longrightarrow> even m" "swapidseq n (id :: 'a \<Rightarrow> 'a)"
|
huffman@30488
|
434 |
{assume "n = 0" hence "even n" by arith}
|
huffman@30488
|
435 |
moreover
|
chaieb@29840
|
436 |
{fix a b :: 'a and q m
|
chaieb@29840
|
437 |
assume h: "n = Suc m" "(id :: 'a \<Rightarrow> 'a) = Fun.swap a b id \<circ> q" "swapidseq m q" "a \<noteq> b"
|
chaieb@29840
|
438 |
from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]]
|
chaieb@29840
|
439 |
have m: "m \<noteq> 0" "swapidseq (m - 1) (id :: 'a \<Rightarrow> 'a)" by auto
|
chaieb@29840
|
440 |
from h m have mn: "m - 1 < n" by arith
|
chaieb@29840
|
441 |
from H(1)[rule_format, OF mn m(2)] h(1) m(1) have "even n" apply arith done}
|
chaieb@29840
|
442 |
ultimately show "even n" using H(2)[unfolded swapidseq_cases[of n id]] by auto
|
chaieb@29840
|
443 |
qed
|
chaieb@29840
|
444 |
|
chaieb@29840
|
445 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
446 |
(* Therefore we have a welldefined notion of parity. *)
|
chaieb@29840
|
447 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
448 |
|
chaieb@29840
|
449 |
definition "evenperm p = even (SOME n. swapidseq n p)"
|
chaieb@29840
|
450 |
|
huffman@30488
|
451 |
lemma swapidseq_even_even: assumes
|
chaieb@29840
|
452 |
m: "swapidseq m p" and n: "swapidseq n p"
|
chaieb@29840
|
453 |
shows "even m \<longleftrightarrow> even n"
|
chaieb@29840
|
454 |
proof-
|
chaieb@29840
|
455 |
from swapidseq_inverse_exists[OF n]
|
chaieb@29840
|
456 |
obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
|
huffman@30488
|
457 |
|
chaieb@29840
|
458 |
from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]]
|
chaieb@29840
|
459 |
show ?thesis by arith
|
chaieb@29840
|
460 |
qed
|
chaieb@29840
|
461 |
|
chaieb@29840
|
462 |
lemma evenperm_unique: assumes p: "swapidseq n p" and n:"even n = b"
|
chaieb@29840
|
463 |
shows "evenperm p = b"
|
chaieb@29840
|
464 |
unfolding n[symmetric] evenperm_def
|
chaieb@29840
|
465 |
apply (rule swapidseq_even_even[where p = p])
|
chaieb@29840
|
466 |
apply (rule someI[where x = n])
|
chaieb@29840
|
467 |
using p by blast+
|
chaieb@29840
|
468 |
|
chaieb@29840
|
469 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
470 |
(* And it has the expected composition properties. *)
|
chaieb@29840
|
471 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
472 |
|
chaieb@29840
|
473 |
lemma evenperm_id[simp]: "evenperm id = True"
|
chaieb@29840
|
474 |
apply (rule evenperm_unique[where n = 0]) by simp_all
|
chaieb@29840
|
475 |
|
chaieb@29840
|
476 |
lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)"
|
chaieb@29840
|
477 |
apply (rule evenperm_unique[where n="if a = b then 0 else 1"])
|
chaieb@29840
|
478 |
by (simp_all add: swapidseq_swap)
|
chaieb@29840
|
479 |
|
huffman@30488
|
480 |
lemma evenperm_comp:
|
chaieb@29840
|
481 |
assumes p: "permutation p" and q:"permutation q"
|
chaieb@29840
|
482 |
shows "evenperm (p o q) = (evenperm p = evenperm q)"
|
chaieb@29840
|
483 |
proof-
|
huffman@30488
|
484 |
from p q obtain
|
huffman@30488
|
485 |
n m where n: "swapidseq n p" and m: "swapidseq m q"
|
chaieb@29840
|
486 |
unfolding permutation_def by blast
|
chaieb@29840
|
487 |
note nm = swapidseq_comp_add[OF n m]
|
chaieb@29840
|
488 |
have th: "even (n + m) = (even n \<longleftrightarrow> even m)" by arith
|
chaieb@29840
|
489 |
from evenperm_unique[OF n refl] evenperm_unique[OF m refl]
|
chaieb@29840
|
490 |
evenperm_unique[OF nm th]
|
chaieb@29840
|
491 |
show ?thesis by blast
|
chaieb@29840
|
492 |
qed
|
chaieb@29840
|
493 |
|
chaieb@29840
|
494 |
lemma evenperm_inv: assumes p: "permutation p"
|
chaieb@29840
|
495 |
shows "evenperm (inv p) = evenperm p"
|
chaieb@29840
|
496 |
proof-
|
chaieb@29840
|
497 |
from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
|
chaieb@29840
|
498 |
from evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]
|
chaieb@29840
|
499 |
show ?thesis .
|
chaieb@29840
|
500 |
qed
|
chaieb@29840
|
501 |
|
chaieb@29840
|
502 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
503 |
(* A more abstract characterization of permutations. *)
|
chaieb@29840
|
504 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
505 |
|
chaieb@29840
|
506 |
|
chaieb@29840
|
507 |
lemma bij_iff: "bij f \<longleftrightarrow> (\<forall>x. \<exists>!y. f y = x)"
|
chaieb@29840
|
508 |
unfolding bij_def inj_on_def surj_def
|
chaieb@29840
|
509 |
apply auto
|
chaieb@29840
|
510 |
apply metis
|
chaieb@29840
|
511 |
apply metis
|
chaieb@29840
|
512 |
done
|
chaieb@29840
|
513 |
|
huffman@30488
|
514 |
lemma permutation_bijective:
|
huffman@30488
|
515 |
assumes p: "permutation p"
|
chaieb@29840
|
516 |
shows "bij p"
|
chaieb@29840
|
517 |
proof-
|
chaieb@29840
|
518 |
from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
|
huffman@30488
|
519 |
from swapidseq_inverse_exists[OF n] obtain q where
|
chaieb@29840
|
520 |
q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
|
nipkow@39302
|
521 |
thus ?thesis unfolding bij_iff apply (auto simp add: fun_eq_iff) apply metis done
|
huffman@30488
|
522 |
qed
|
chaieb@29840
|
523 |
|
chaieb@29840
|
524 |
lemma permutation_finite_support: assumes p: "permutation p"
|
chaieb@29840
|
525 |
shows "finite {x. p x \<noteq> x}"
|
chaieb@29840
|
526 |
proof-
|
chaieb@29840
|
527 |
from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
|
chaieb@29840
|
528 |
from n show ?thesis
|
chaieb@29840
|
529 |
proof(induct n p rule: swapidseq.induct)
|
chaieb@29840
|
530 |
case id thus ?case by simp
|
chaieb@29840
|
531 |
next
|
chaieb@29840
|
532 |
case (comp_Suc n p a b)
|
chaieb@29840
|
533 |
let ?S = "insert a (insert b {x. p x \<noteq> x})"
|
chaieb@29840
|
534 |
from comp_Suc.hyps(2) have fS: "finite ?S" by simp
|
chaieb@29840
|
535 |
from `a \<noteq> b` have th: "{x. (Fun.swap a b id o p) x \<noteq> x} \<subseteq> ?S"
|
chaieb@29840
|
536 |
by (auto simp add: swap_def)
|
chaieb@29840
|
537 |
from finite_subset[OF th fS] show ?case .
|
chaieb@29840
|
538 |
qed
|
chaieb@29840
|
539 |
qed
|
chaieb@29840
|
540 |
|
chaieb@29840
|
541 |
lemma bij_inv_eq_iff: "bij p ==> x = inv p y \<longleftrightarrow> p x = y"
|
chaieb@29840
|
542 |
using surj_f_inv_f[of p] inv_f_f[of f] by (auto simp add: bij_def)
|
chaieb@29840
|
543 |
|
huffman@30488
|
544 |
lemma bij_swap_comp:
|
chaieb@29840
|
545 |
assumes bp: "bij p" shows "Fun.swap a b id o p = Fun.swap (inv p a) (inv p b) p"
|
chaieb@29840
|
546 |
using surj_f_inv_f[OF bij_is_surj[OF bp]]
|
nipkow@39302
|
547 |
by (simp add: fun_eq_iff swap_def bij_inv_eq_iff[OF bp])
|
chaieb@29840
|
548 |
|
chaieb@29840
|
549 |
lemma bij_swap_ompose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id o p)"
|
chaieb@29840
|
550 |
proof-
|
chaieb@29840
|
551 |
assume H: "bij p"
|
huffman@30488
|
552 |
show ?thesis
|
chaieb@29840
|
553 |
unfolding bij_swap_comp[OF H] bij_swap_iff
|
chaieb@29840
|
554 |
using H .
|
chaieb@29840
|
555 |
qed
|
chaieb@29840
|
556 |
|
huffman@30488
|
557 |
lemma permutation_lemma:
|
chaieb@29840
|
558 |
assumes fS: "finite S" and p: "bij p" and pS: "\<forall>x. x\<notin> S \<longrightarrow> p x = x"
|
chaieb@29840
|
559 |
shows "permutation p"
|
chaieb@29840
|
560 |
using fS p pS
|
chaieb@29840
|
561 |
proof(induct S arbitrary: p rule: finite_induct)
|
chaieb@29840
|
562 |
case (empty p) thus ?case by simp
|
chaieb@29840
|
563 |
next
|
chaieb@29840
|
564 |
case (insert a F p)
|
chaieb@29840
|
565 |
let ?r = "Fun.swap a (p a) id o p"
|
chaieb@29840
|
566 |
let ?q = "Fun.swap a (p a) id o ?r "
|
chaieb@29840
|
567 |
have raa: "?r a = a" by (simp add: swap_def)
|
chaieb@29840
|
568 |
from bij_swap_ompose_bij[OF insert(4)]
|
huffman@30488
|
569 |
have br: "bij ?r" .
|
huffman@30488
|
570 |
|
huffman@30488
|
571 |
from insert raa have th: "\<forall>x. x \<notin> F \<longrightarrow> ?r x = x"
|
chaieb@29840
|
572 |
apply (clarsimp simp add: swap_def)
|
chaieb@29840
|
573 |
apply (erule_tac x="x" in allE)
|
chaieb@29840
|
574 |
apply auto
|
chaieb@29840
|
575 |
unfolding bij_iff apply metis
|
chaieb@29840
|
576 |
done
|
chaieb@29840
|
577 |
from insert(3)[OF br th]
|
chaieb@29840
|
578 |
have rp: "permutation ?r" .
|
chaieb@29840
|
579 |
have "permutation ?q" by (simp add: permutation_compose permutation_swap_id rp)
|
chaieb@29840
|
580 |
thus ?case by (simp add: o_assoc)
|
chaieb@29840
|
581 |
qed
|
chaieb@29840
|
582 |
|
huffman@30488
|
583 |
lemma permutation: "permutation p \<longleftrightarrow> bij p \<and> finite {x. p x \<noteq> x}"
|
chaieb@29840
|
584 |
(is "?lhs \<longleftrightarrow> ?b \<and> ?f")
|
chaieb@29840
|
585 |
proof
|
chaieb@29840
|
586 |
assume p: ?lhs
|
chaieb@29840
|
587 |
from p permutation_bijective permutation_finite_support show "?b \<and> ?f" by auto
|
chaieb@29840
|
588 |
next
|
chaieb@29840
|
589 |
assume bf: "?b \<and> ?f"
|
chaieb@29840
|
590 |
hence bf: "?f" "?b" by blast+
|
chaieb@29840
|
591 |
from permutation_lemma[OF bf] show ?lhs by blast
|
chaieb@29840
|
592 |
qed
|
chaieb@29840
|
593 |
|
chaieb@29840
|
594 |
lemma permutation_inverse_works: assumes p: "permutation p"
|
chaieb@29840
|
595 |
shows "inv p o p = id" "p o inv p = id"
|
huffman@44227
|
596 |
using permutation_bijective [OF p]
|
huffman@44227
|
597 |
unfolding bij_def inj_iff surj_iff by auto
|
chaieb@29840
|
598 |
|
chaieb@29840
|
599 |
lemma permutation_inverse_compose:
|
chaieb@29840
|
600 |
assumes p: "permutation p" and q: "permutation q"
|
chaieb@29840
|
601 |
shows "inv (p o q) = inv q o inv p"
|
chaieb@29840
|
602 |
proof-
|
chaieb@29840
|
603 |
note ps = permutation_inverse_works[OF p]
|
chaieb@29840
|
604 |
note qs = permutation_inverse_works[OF q]
|
chaieb@29840
|
605 |
have "p o q o (inv q o inv p) = p o (q o inv q) o inv p" by (simp add: o_assoc)
|
chaieb@29840
|
606 |
also have "\<dots> = id" by (simp add: ps qs)
|
chaieb@29840
|
607 |
finally have th0: "p o q o (inv q o inv p) = id" .
|
chaieb@29840
|
608 |
have "inv q o inv p o (p o q) = inv q o (inv p o p) o q" by (simp add: o_assoc)
|
chaieb@29840
|
609 |
also have "\<dots> = id" by (simp add: ps qs)
|
huffman@30488
|
610 |
finally have th1: "inv q o inv p o (p o q) = id" .
|
chaieb@29840
|
611 |
from inv_unique_comp[OF th0 th1] show ?thesis .
|
chaieb@29840
|
612 |
qed
|
chaieb@29840
|
613 |
|
chaieb@29840
|
614 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
615 |
(* Relation to "permutes". *)
|
chaieb@29840
|
616 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
617 |
|
chaieb@29840
|
618 |
lemma permutation_permutes: "permutation p \<longleftrightarrow> (\<exists>S. finite S \<and> p permutes S)"
|
chaieb@29840
|
619 |
unfolding permutation permutes_def bij_iff[symmetric]
|
chaieb@29840
|
620 |
apply (rule iffI, clarify)
|
chaieb@29840
|
621 |
apply (rule exI[where x="{x. p x \<noteq> x}"])
|
chaieb@29840
|
622 |
apply simp
|
chaieb@29840
|
623 |
apply clarsimp
|
chaieb@29840
|
624 |
apply (rule_tac B="S" in finite_subset)
|
chaieb@29840
|
625 |
apply auto
|
chaieb@29840
|
626 |
done
|
chaieb@29840
|
627 |
|
chaieb@29840
|
628 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
629 |
(* Hence a sort of induction principle composing by swaps. *)
|
chaieb@29840
|
630 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
631 |
|
chaieb@29840
|
632 |
lemma permutes_induct: "finite S \<Longrightarrow> P id \<Longrightarrow> (\<And> a b p. a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> P p \<Longrightarrow> P p \<Longrightarrow> permutation p ==> P (Fun.swap a b id o p))
|
chaieb@29840
|
633 |
==> (\<And>p. p permutes S ==> P p)"
|
chaieb@29840
|
634 |
proof(induct S rule: finite_induct)
|
chaieb@29840
|
635 |
case empty thus ?case by auto
|
huffman@30488
|
636 |
next
|
chaieb@29840
|
637 |
case (insert x F p)
|
chaieb@29840
|
638 |
let ?r = "Fun.swap x (p x) id o p"
|
chaieb@29840
|
639 |
let ?q = "Fun.swap x (p x) id o ?r"
|
chaieb@29840
|
640 |
have qp: "?q = p" by (simp add: o_assoc)
|
chaieb@29840
|
641 |
from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r" by blast
|
huffman@30488
|
642 |
from permutes_in_image[OF insert.prems(3), of x]
|
chaieb@29840
|
643 |
have pxF: "p x \<in> insert x F" by simp
|
chaieb@29840
|
644 |
have xF: "x \<in> insert x F" by simp
|
chaieb@29840
|
645 |
have rp: "permutation ?r"
|
huffman@30488
|
646 |
unfolding permutation_permutes using insert.hyps(1)
|
chaieb@29840
|
647 |
permutes_insert_lemma[OF insert.prems(3)] by blast
|
huffman@30488
|
648 |
from insert.prems(2)[OF xF pxF Pr Pr rp]
|
huffman@30488
|
649 |
show ?case unfolding qp .
|
chaieb@29840
|
650 |
qed
|
chaieb@29840
|
651 |
|
chaieb@29840
|
652 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
653 |
(* Sign of a permutation as a real number. *)
|
chaieb@29840
|
654 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
655 |
|
chaieb@29840
|
656 |
definition "sign p = (if evenperm p then (1::int) else -1)"
|
chaieb@29840
|
657 |
|
huffman@30488
|
658 |
lemma sign_nz: "sign p \<noteq> 0" by (simp add: sign_def)
|
chaieb@29840
|
659 |
lemma sign_id: "sign id = 1" by (simp add: sign_def)
|
chaieb@29840
|
660 |
lemma sign_inverse: "permutation p ==> sign (inv p) = sign p"
|
chaieb@29840
|
661 |
by (simp add: sign_def evenperm_inv)
|
chaieb@29840
|
662 |
lemma sign_compose: "permutation p \<Longrightarrow> permutation q ==> sign (p o q) = sign(p) * sign(q)" by (simp add: sign_def evenperm_comp)
|
chaieb@29840
|
663 |
lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)"
|
chaieb@29840
|
664 |
by (simp add: sign_def evenperm_swap)
|
chaieb@29840
|
665 |
lemma sign_idempotent: "sign p * sign p = 1" by (simp add: sign_def)
|
chaieb@29840
|
666 |
|
chaieb@29840
|
667 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
668 |
(* More lemmas about permutations. *)
|
chaieb@29840
|
669 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
670 |
|
chaieb@29840
|
671 |
lemma permutes_natset_le:
|
huffman@30037
|
672 |
assumes p: "p permutes (S::'a::wellorder set)" and le: "\<forall>i \<in> S. p i <= i" shows "p = id"
|
chaieb@29840
|
673 |
proof-
|
chaieb@29840
|
674 |
{fix n
|
huffman@30488
|
675 |
have "p n = n"
|
chaieb@29840
|
676 |
using p le
|
huffman@30037
|
677 |
proof(induct n arbitrary: S rule: less_induct)
|
huffman@30488
|
678 |
fix n S assume H: "\<And>m S. \<lbrakk>m < n; p permutes S; \<forall>i\<in>S. p i \<le> i\<rbrakk> \<Longrightarrow> p m = m"
|
wenzelm@32960
|
679 |
"p permutes S" "\<forall>i \<in>S. p i \<le> i"
|
chaieb@29840
|
680 |
{assume "n \<notin> S"
|
wenzelm@32960
|
681 |
with H(2) have "p n = n" unfolding permutes_def by metis}
|
chaieb@29840
|
682 |
moreover
|
chaieb@29840
|
683 |
{assume ns: "n \<in> S"
|
wenzelm@32960
|
684 |
from H(3) ns have "p n < n \<or> p n = n" by auto
|
wenzelm@32960
|
685 |
moreover{assume h: "p n < n"
|
wenzelm@32960
|
686 |
from H h have "p (p n) = p n" by metis
|
wenzelm@32960
|
687 |
with permutes_inj[OF H(2)] have "p n = n" unfolding inj_on_def by blast
|
wenzelm@32960
|
688 |
with h have False by simp}
|
wenzelm@32960
|
689 |
ultimately have "p n = n" by blast }
|
chaieb@29840
|
690 |
ultimately show "p n = n" by blast
|
chaieb@29840
|
691 |
qed}
|
nipkow@39302
|
692 |
thus ?thesis by (auto simp add: fun_eq_iff)
|
chaieb@29840
|
693 |
qed
|
chaieb@29840
|
694 |
|
chaieb@29840
|
695 |
lemma permutes_natset_ge:
|
huffman@30037
|
696 |
assumes p: "p permutes (S::'a::wellorder set)" and le: "\<forall>i \<in> S. p i \<ge> i" shows "p = id"
|
chaieb@29840
|
697 |
proof-
|
chaieb@29840
|
698 |
{fix i assume i: "i \<in> S"
|
chaieb@29840
|
699 |
from i permutes_in_image[OF permutes_inv[OF p]] have "inv p i \<in> S" by simp
|
chaieb@29840
|
700 |
with le have "p (inv p i) \<ge> inv p i" by blast
|
chaieb@29840
|
701 |
with permutes_inverses[OF p] have "i \<ge> inv p i" by simp}
|
chaieb@29840
|
702 |
then have th: "\<forall>i\<in>S. inv p i \<le> i" by blast
|
huffman@30488
|
703 |
from permutes_natset_le[OF permutes_inv[OF p] th]
|
chaieb@29840
|
704 |
have "inv p = inv id" by simp
|
huffman@30488
|
705 |
then show ?thesis
|
chaieb@29840
|
706 |
apply (subst permutes_inv_inv[OF p, symmetric])
|
chaieb@29840
|
707 |
apply (rule inv_unique_comp)
|
chaieb@29840
|
708 |
apply simp_all
|
chaieb@29840
|
709 |
done
|
chaieb@29840
|
710 |
qed
|
chaieb@29840
|
711 |
|
chaieb@29840
|
712 |
lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}"
|
nipkow@39302
|
713 |
apply (rule set_eqI)
|
chaieb@29840
|
714 |
apply auto
|
chaieb@29840
|
715 |
using permutes_inv_inv permutes_inv apply auto
|
chaieb@29840
|
716 |
apply (rule_tac x="inv x" in exI)
|
chaieb@29840
|
717 |
apply auto
|
chaieb@29840
|
718 |
done
|
chaieb@29840
|
719 |
|
huffman@30488
|
720 |
lemma image_compose_permutations_left:
|
chaieb@29840
|
721 |
assumes q: "q permutes S" shows "{q o p | p. p permutes S} = {p . p permutes S}"
|
nipkow@39302
|
722 |
apply (rule set_eqI)
|
chaieb@29840
|
723 |
apply auto
|
chaieb@29840
|
724 |
apply (rule permutes_compose)
|
chaieb@29840
|
725 |
using q apply auto
|
chaieb@29840
|
726 |
apply (rule_tac x = "inv q o x" in exI)
|
chaieb@29840
|
727 |
by (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o)
|
chaieb@29840
|
728 |
|
chaieb@29840
|
729 |
lemma image_compose_permutations_right:
|
chaieb@29840
|
730 |
assumes q: "q permutes S"
|
chaieb@29840
|
731 |
shows "{p o q | p. p permutes S} = {p . p permutes S}"
|
nipkow@39302
|
732 |
apply (rule set_eqI)
|
chaieb@29840
|
733 |
apply auto
|
chaieb@29840
|
734 |
apply (rule permutes_compose)
|
chaieb@29840
|
735 |
using q apply auto
|
chaieb@29840
|
736 |
apply (rule_tac x = "x o inv q" in exI)
|
haftmann@49739
|
737 |
by (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc)
|
chaieb@29840
|
738 |
|
chaieb@29840
|
739 |
lemma permutes_in_seg: "p permutes {1 ..n} \<Longrightarrow> i \<in> {1..n} ==> 1 <= p i \<and> p i <= n"
|
chaieb@29840
|
740 |
|
chaieb@29840
|
741 |
apply (simp add: permutes_def)
|
chaieb@29840
|
742 |
apply metis
|
chaieb@29840
|
743 |
done
|
chaieb@29840
|
744 |
|
huffman@30036
|
745 |
lemma setsum_permutations_inverse: "setsum f {p. p permutes S} = setsum (\<lambda>p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs")
|
chaieb@29840
|
746 |
proof-
|
huffman@30036
|
747 |
let ?S = "{p . p permutes S}"
|
huffman@30488
|
748 |
have th0: "inj_on inv ?S"
|
chaieb@29840
|
749 |
proof(auto simp add: inj_on_def)
|
chaieb@29840
|
750 |
fix q r
|
huffman@30036
|
751 |
assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r"
|
chaieb@29840
|
752 |
hence "inv (inv q) = inv (inv r)" by simp
|
chaieb@29840
|
753 |
with permutes_inv_inv[OF q] permutes_inv_inv[OF r]
|
chaieb@29840
|
754 |
show "q = r" by metis
|
chaieb@29840
|
755 |
qed
|
chaieb@29840
|
756 |
have th1: "inv ` ?S = ?S" using image_inverse_permutations by blast
|
chaieb@29840
|
757 |
have th2: "?rhs = setsum (f o inv) ?S" by (simp add: o_def)
|
chaieb@29840
|
758 |
from setsum_reindex[OF th0, of f] show ?thesis unfolding th1 th2 .
|
chaieb@29840
|
759 |
qed
|
chaieb@29840
|
760 |
|
chaieb@29840
|
761 |
lemma setum_permutations_compose_left:
|
huffman@30036
|
762 |
assumes q: "q permutes S"
|
huffman@30036
|
763 |
shows "setsum f {p. p permutes S} =
|
huffman@30036
|
764 |
setsum (\<lambda>p. f(q o p)) {p. p permutes S}" (is "?lhs = ?rhs")
|
chaieb@29840
|
765 |
proof-
|
huffman@30036
|
766 |
let ?S = "{p. p permutes S}"
|
chaieb@29840
|
767 |
have th0: "?rhs = setsum (f o (op o q)) ?S" by (simp add: o_def)
|
chaieb@29840
|
768 |
have th1: "inj_on (op o q) ?S"
|
chaieb@29840
|
769 |
apply (auto simp add: inj_on_def)
|
chaieb@29840
|
770 |
proof-
|
chaieb@29840
|
771 |
fix p r
|
huffman@30036
|
772 |
assume "p permutes S" and r:"r permutes S" and rp: "q \<circ> p = q \<circ> r"
|
haftmann@49739
|
773 |
hence "inv q o q o p = inv q o q o r" by (simp add: comp_assoc)
|
chaieb@29840
|
774 |
with permutes_inj[OF q, unfolded inj_iff]
|
chaieb@29840
|
775 |
|
chaieb@29840
|
776 |
show "p = r" by simp
|
chaieb@29840
|
777 |
qed
|
chaieb@29840
|
778 |
have th3: "(op o q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto
|
chaieb@29840
|
779 |
from setsum_reindex[OF th1, of f]
|
chaieb@29840
|
780 |
show ?thesis unfolding th0 th1 th3 .
|
chaieb@29840
|
781 |
qed
|
chaieb@29840
|
782 |
|
chaieb@29840
|
783 |
lemma sum_permutations_compose_right:
|
huffman@30036
|
784 |
assumes q: "q permutes S"
|
huffman@30036
|
785 |
shows "setsum f {p. p permutes S} =
|
huffman@30036
|
786 |
setsum (\<lambda>p. f(p o q)) {p. p permutes S}" (is "?lhs = ?rhs")
|
chaieb@29840
|
787 |
proof-
|
huffman@30036
|
788 |
let ?S = "{p. p permutes S}"
|
chaieb@29840
|
789 |
have th0: "?rhs = setsum (f o (\<lambda>p. p o q)) ?S" by (simp add: o_def)
|
chaieb@29840
|
790 |
have th1: "inj_on (\<lambda>p. p o q) ?S"
|
chaieb@29840
|
791 |
apply (auto simp add: inj_on_def)
|
chaieb@29840
|
792 |
proof-
|
chaieb@29840
|
793 |
fix p r
|
huffman@30036
|
794 |
assume "p permutes S" and r:"r permutes S" and rp: "p o q = r o q"
|
chaieb@29840
|
795 |
hence "p o (q o inv q) = r o (q o inv q)" by (simp add: o_assoc)
|
chaieb@29840
|
796 |
with permutes_surj[OF q, unfolded surj_iff]
|
chaieb@29840
|
797 |
|
chaieb@29840
|
798 |
show "p = r" by simp
|
chaieb@29840
|
799 |
qed
|
chaieb@29840
|
800 |
have th3: "(\<lambda>p. p o q) ` ?S = ?S" using image_compose_permutations_right[OF q] by auto
|
chaieb@29840
|
801 |
from setsum_reindex[OF th1, of f]
|
chaieb@29840
|
802 |
show ?thesis unfolding th0 th1 th3 .
|
chaieb@29840
|
803 |
qed
|
chaieb@29840
|
804 |
|
chaieb@29840
|
805 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
806 |
(* Sum over a set of permutations (could generalize to iteration). *)
|
chaieb@29840
|
807 |
(* ------------------------------------------------------------------------- *)
|
chaieb@29840
|
808 |
|
chaieb@29840
|
809 |
lemma setsum_over_permutations_insert:
|
chaieb@29840
|
810 |
assumes fS: "finite S" and aS: "a \<notin> S"
|
chaieb@29840
|
811 |
shows "setsum f {p. p permutes (insert a S)} = setsum (\<lambda>b. setsum (\<lambda>q. f (Fun.swap a b id o q)) {p. p permutes S}) (insert a S)"
|
chaieb@29840
|
812 |
proof-
|
chaieb@29840
|
813 |
have th0: "\<And>f a b. (\<lambda>(b,p). f (Fun.swap a b id o p)) = f o (\<lambda>(b,p). Fun.swap a b id o p)"
|
nipkow@39302
|
814 |
by (simp add: fun_eq_iff)
|
chaieb@29840
|
815 |
have th1: "\<And>P Q. P \<times> Q = {(a,b). a \<in> P \<and> b \<in> Q}" by blast
|
chaieb@29840
|
816 |
have th2: "\<And>P Q. P \<Longrightarrow> (P \<Longrightarrow> Q) \<Longrightarrow> P \<and> Q" by blast
|
huffman@30488
|
817 |
show ?thesis
|
huffman@30488
|
818 |
unfolding permutes_insert
|
chaieb@29840
|
819 |
unfolding setsum_cartesian_product
|
chaieb@29840
|
820 |
unfolding th1[symmetric]
|
chaieb@29840
|
821 |
unfolding th0
|
chaieb@29840
|
822 |
proof(rule setsum_reindex)
|
chaieb@29840
|
823 |
let ?f = "(\<lambda>(b, y). Fun.swap a b id \<circ> y)"
|
chaieb@29840
|
824 |
let ?P = "{p. p permutes S}"
|
huffman@30488
|
825 |
{fix b c p q assume b: "b \<in> insert a S" and c: "c \<in> insert a S"
|
huffman@30488
|
826 |
and p: "p permutes S" and q: "q permutes S"
|
chaieb@29840
|
827 |
and eq: "Fun.swap a b id o p = Fun.swap a c id o q"
|
chaieb@29840
|
828 |
from p q aS have pa: "p a = a" and qa: "q a = a"
|
wenzelm@32960
|
829 |
unfolding permutes_def by metis+
|
chaieb@29840
|
830 |
from eq have "(Fun.swap a b id o p) a = (Fun.swap a c id o q) a" by simp
|
chaieb@29840
|
831 |
hence bc: "b = c"
|
wenzelm@32960
|
832 |
by (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong split: split_if_asm)
|
chaieb@29840
|
833 |
from eq[unfolded bc] have "(\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o p) = (\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o q)" by simp
|
chaieb@29840
|
834 |
hence "p = q" unfolding o_assoc swap_id_idempotent
|
wenzelm@32960
|
835 |
by (simp add: o_def)
|
chaieb@29840
|
836 |
with bc have "b = c \<and> p = q" by blast
|
chaieb@29840
|
837 |
}
|
huffman@30488
|
838 |
then show "inj_on ?f (insert a S \<times> ?P)"
|
chaieb@29840
|
839 |
unfolding inj_on_def
|
chaieb@29840
|
840 |
apply clarify by metis
|
chaieb@29840
|
841 |
qed
|
chaieb@29840
|
842 |
qed
|
chaieb@29840
|
843 |
|
chaieb@29840
|
844 |
end
|
haftmann@51489
|
845 |
|