author | wenzelm |
Fri, 29 Nov 2024 17:40:15 +0100 | |
changeset 81507 | 08574da77b4a |
parent 73648 | 1bd3463e30b8 |
permissions | -rw-r--r-- |
63375 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
2 |
||
3 |
section \<open>Fragments on permuations\<close> |
|
4 |
||
5 |
theory Perm_Fragments |
|
73622 | 6 |
imports "HOL-Library.Dlist" "HOL-Combinatorics.Perm" |
63375 | 7 |
begin |
8 |
||
9 |
text \<open>On cycles\<close> |
|
10 |
||
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
11 |
context |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
12 |
includes permutation_syntax |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
13 |
begin |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
14 |
|
63882
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents:
63375
diff
changeset
|
15 |
lemma cycle_prod_list: |
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
16 |
"\<langle>a # as\<rangle> = prod_list (map (\<lambda>b. \<langle>a \<leftrightarrow> b\<rangle>) (rev as))" |
63375 | 17 |
by (induct as) simp_all |
18 |
||
19 |
lemma cycle_append [simp]: |
|
20 |
"\<langle>a # as @ bs\<rangle> = \<langle>a # bs\<rangle> * \<langle>a # as\<rangle>" |
|
21 |
proof (induct as rule: cycle.induct) |
|
22 |
case (3 b c as) |
|
23 |
then have "\<langle>a # (b # as) @ bs\<rangle> = \<langle>a # bs\<rangle> * \<langle>a # b # as\<rangle>" |
|
24 |
by simp |
|
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
25 |
then have "\<langle>a # as @ bs\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> = |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
26 |
\<langle>a # bs\<rangle> * \<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>" |
63375 | 27 |
by (simp add: ac_simps) |
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
28 |
then have "\<langle>a # as @ bs\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> = |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
29 |
\<langle>a # bs\<rangle> * \<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>" |
63375 | 30 |
by simp |
31 |
then have "\<langle>a # as @ bs\<rangle> = \<langle>a # bs\<rangle> * \<langle>a # as\<rangle>" |
|
32 |
by (simp add: ac_simps) |
|
33 |
then show "\<langle>a # (b # c # as) @ bs\<rangle> = |
|
34 |
\<langle>a # bs\<rangle> * \<langle>a # b # c # as\<rangle>" |
|
35 |
by (simp add: ac_simps) |
|
36 |
qed simp_all |
|
37 |
||
38 |
lemma affected_cycle: |
|
39 |
"affected \<langle>as\<rangle> \<subseteq> set as" |
|
40 |
proof (induct as rule: cycle.induct) |
|
41 |
case (3 a b as) |
|
42 |
from affected_times |
|
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
43 |
have "affected (\<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>) |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
44 |
\<subseteq> affected \<langle>a # as\<rangle> \<union> affected \<langle>a \<leftrightarrow> b\<rangle>" . |
63375 | 45 |
moreover from 3 |
46 |
have "affected (\<langle>a # as\<rangle>) \<subseteq> insert a (set as)" |
|
47 |
by simp |
|
48 |
moreover |
|
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
49 |
have "affected \<langle>a \<leftrightarrow> b\<rangle> \<subseteq> {a, b}" |
63375 | 50 |
by (cases "a = b") (simp_all add: affected_swap) |
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
51 |
ultimately have "affected (\<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>) |
63375 | 52 |
\<subseteq> insert a (insert b (set as))" |
53 |
by blast |
|
54 |
then show ?case by auto |
|
55 |
qed simp_all |
|
56 |
||
57 |
lemma orbit_cycle_non_elem: |
|
58 |
assumes "a \<notin> set as" |
|
59 |
shows "orbit \<langle>as\<rangle> a = {a}" |
|
60 |
unfolding not_in_affected_iff_orbit_eq_singleton [symmetric] |
|
61 |
using assms affected_cycle [of as] by blast |
|
62 |
||
63 |
lemma inverse_cycle: |
|
64 |
assumes "distinct as" |
|
65 |
shows "inverse \<langle>as\<rangle> = \<langle>rev as\<rangle>" |
|
66 |
using assms proof (induct as rule: cycle.induct) |
|
67 |
case (3 a b as) |
|
68 |
then have *: "inverse \<langle>a # as\<rangle> = \<langle>rev (a # as)\<rangle>" |
|
69 |
and distinct: "distinct (a # b # as)" |
|
70 |
by simp_all |
|
71 |
show " inverse \<langle>a # b # as\<rangle> = \<langle>rev (a # b # as)\<rangle>" |
|
72 |
proof (cases as rule: rev_cases) |
|
73 |
case Nil with * show ?thesis |
|
74 |
by (simp add: swap_sym) |
|
75 |
next |
|
76 |
case (snoc cs c) |
|
77 |
with distinct have "distinct (a # b # cs @ [c])" |
|
78 |
by simp |
|
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
79 |
then have **: "\<langle>a \<leftrightarrow> b\<rangle> * \<langle>c \<leftrightarrow> a\<rangle> = \<langle>c \<leftrightarrow> a\<rangle> * \<langle>c \<leftrightarrow> b\<rangle>" |
73648 | 80 |
by transfer (auto simp add: fun_eq_iff transpose_def) |
63375 | 81 |
with snoc * show ?thesis |
82 |
by (simp add: mult.assoc [symmetric]) |
|
83 |
qed |
|
84 |
qed simp_all |
|
85 |
||
86 |
lemma order_cycle_non_elem: |
|
87 |
assumes "a \<notin> set as" |
|
88 |
shows "order \<langle>as\<rangle> a = 1" |
|
89 |
proof - |
|
90 |
from assms have "orbit \<langle>as\<rangle> a = {a}" |
|
91 |
by (rule orbit_cycle_non_elem) |
|
92 |
then have "card (orbit \<langle>as\<rangle> a) = card {a}" |
|
93 |
by simp |
|
94 |
then show ?thesis |
|
95 |
by simp |
|
96 |
qed |
|
97 |
||
98 |
lemma orbit_cycle_elem: |
|
99 |
assumes "distinct as" and "a \<in> set as" |
|
100 |
shows "orbit \<langle>as\<rangle> a = set as" |
|
67226 | 101 |
oops \<comment> \<open>(we need rotation here\<close> |
63375 | 102 |
|
103 |
lemma order_cycle_elem: |
|
104 |
assumes "distinct as" and "a \<in> set as" |
|
105 |
shows "order \<langle>as\<rangle> a = length as" |
|
106 |
oops |
|
107 |
||
108 |
||
109 |
text \<open>Adding fixpoints\<close> |
|
110 |
||
111 |
definition fixate :: "'a \<Rightarrow> 'a perm \<Rightarrow> 'a perm" |
|
112 |
where |
|
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
113 |
"fixate a f = (if a \<in> affected f then f * \<langle>inverse f \<langle>$\<rangle> a \<leftrightarrow> a\<rangle> else f)" |
63375 | 114 |
|
115 |
lemma affected_fixate_trivial: |
|
116 |
assumes "a \<notin> affected f" |
|
117 |
shows "affected (fixate a f) = affected f" |
|
118 |
using assms by (simp add: fixate_def) |
|
119 |
||
120 |
lemma affected_fixate_binary_circle: |
|
121 |
assumes "order f a = 2" |
|
122 |
shows "affected (fixate a f) = affected f - {a, apply f a}" (is "?A = ?B") |
|
123 |
proof (rule set_eqI) |
|
124 |
interpret bijection "apply f" |
|
125 |
by standard simp |
|
126 |
fix b |
|
127 |
from assms order_greater_eq_two_iff [of f a] have "a \<in> affected f" |
|
128 |
by simp |
|
129 |
moreover have "apply (f ^ 2) a = a" |
|
130 |
by (simp add: assms [symmetric]) |
|
131 |
ultimately show "b \<in> ?A \<longleftrightarrow> b \<in> ?B" |
|
132 |
by (cases "b \<in> {a, apply (inverse f) a}") |
|
133 |
(auto simp add: in_affected power2_eq_square apply_inverse apply_times fixate_def) |
|
134 |
qed |
|
135 |
||
136 |
lemma affected_fixate_no_binary_circle: |
|
137 |
assumes "order f a > 2" |
|
138 |
shows "affected (fixate a f) = affected f - {a}" (is "?A = ?B") |
|
139 |
proof (rule set_eqI) |
|
140 |
interpret bijection "apply f" |
|
141 |
by standard simp |
|
142 |
fix b |
|
143 |
from assms order_greater_eq_two_iff [of f a] |
|
144 |
have "a \<in> affected f" |
|
145 |
by simp |
|
146 |
moreover from assms |
|
147 |
have "apply f (apply f a) \<noteq> a" |
|
148 |
using apply_power_eq_iff [of f 2 a 0] |
|
149 |
by (simp add: power2_eq_square apply_times) |
|
150 |
ultimately show "b \<in> ?A \<longleftrightarrow> b \<in> ?B" |
|
151 |
by (cases "b \<in> {a, apply (inverse f) a}") |
|
152 |
(auto simp add: in_affected apply_inverse apply_times fixate_def) |
|
153 |
qed |
|
154 |
||
155 |
lemma affected_fixate: |
|
156 |
"affected (fixate a f) \<subseteq> affected f - {a}" |
|
157 |
proof - |
|
158 |
have "a \<notin> affected f \<or> order f a = 2 \<or> order f a > 2" |
|
159 |
by (auto simp add: not_less dest: affected_order_greater_eq_two) |
|
160 |
then consider "a \<notin> affected f" | "order f a = 2" | "order f a > 2" |
|
161 |
by blast |
|
162 |
then show ?thesis apply cases |
|
163 |
using affected_fixate_trivial [of a f] |
|
164 |
affected_fixate_binary_circle [of f a] |
|
165 |
affected_fixate_no_binary_circle [of f a] |
|
166 |
by auto |
|
167 |
qed |
|
168 |
||
169 |
lemma orbit_fixate_self [simp]: |
|
170 |
"orbit (fixate a f) a = {a}" |
|
171 |
proof - |
|
172 |
have "apply (f * inverse f) a = a" |
|
173 |
by simp |
|
174 |
then have "apply f (apply (inverse f) a) = a" |
|
175 |
by (simp only: apply_times comp_apply) |
|
176 |
then show ?thesis |
|
177 |
by (simp add: fixate_def not_in_affected_iff_orbit_eq_singleton [symmetric] in_affected apply_times) |
|
178 |
qed |
|
179 |
||
180 |
lemma order_fixate_self [simp]: |
|
181 |
"order (fixate a f) a = 1" |
|
182 |
proof - |
|
183 |
have "card (orbit (fixate a f) a) = card {a}" |
|
184 |
by simp |
|
185 |
then show ?thesis |
|
186 |
by (simp only: card_orbit_eq) simp |
|
187 |
qed |
|
188 |
||
189 |
lemma |
|
190 |
assumes "b \<notin> orbit f a" |
|
191 |
shows "orbit (fixate b f) a = orbit f a" |
|
192 |
oops |
|
193 |
||
194 |
lemma |
|
195 |
assumes "b \<in> orbit f a" and "b \<noteq> a" |
|
196 |
shows "orbit (fixate b f) a = orbit f a - {b}" |
|
197 |
oops |
|
198 |
||
199 |
||
200 |
text \<open>Distilling cycles from permutations\<close> |
|
201 |
||
202 |
inductive_set orbits :: "'a perm \<Rightarrow> 'a set set" for f |
|
203 |
where |
|
204 |
in_orbitsI: "a \<in> affected f \<Longrightarrow> orbit f a \<in> orbits f" |
|
205 |
||
206 |
lemma orbits_unfold: |
|
207 |
"orbits f = orbit f ` affected f" |
|
208 |
by (auto intro: in_orbitsI elim: orbits.cases) |
|
209 |
||
210 |
lemma in_orbit_affected: |
|
211 |
assumes "b \<in> orbit f a" |
|
212 |
assumes "a \<in> affected f" |
|
213 |
shows "b \<in> affected f" |
|
214 |
proof (cases "a = b") |
|
215 |
case True with assms show ?thesis by simp |
|
216 |
next |
|
217 |
case False with assms have "{a, b} \<subseteq> orbit f a" |
|
218 |
by auto |
|
219 |
also from assms have "orbit f a \<subseteq> affected f" |
|
220 |
by (blast intro!: orbit_subset_eq_affected) |
|
221 |
finally show ?thesis by simp |
|
222 |
qed |
|
223 |
||
224 |
lemma Union_orbits [simp]: |
|
69745 | 225 |
"\<Union>(orbits f) = affected f" |
63375 | 226 |
by (auto simp add: orbits.simps intro: in_orbitsI in_orbit_affected) |
227 |
||
228 |
lemma finite_orbits [simp]: |
|
229 |
"finite (orbits f)" |
|
230 |
by (simp add: orbits_unfold) |
|
231 |
||
232 |
lemma card_in_orbits: |
|
233 |
assumes "A \<in> orbits f" |
|
234 |
shows "card A \<ge> 2" |
|
235 |
using assms by cases |
|
236 |
(auto dest: affected_order_greater_eq_two) |
|
237 |
||
238 |
lemma disjoint_orbits: |
|
239 |
assumes "A \<in> orbits f" and "B \<in> orbits f" and "A \<noteq> B" |
|
240 |
shows "A \<inter> B = {}" |
|
241 |
using \<open>A \<in> orbits f\<close> apply cases |
|
242 |
using \<open>B \<in> orbits f\<close> apply cases |
|
243 |
using \<open>A \<noteq> B\<close> apply (simp_all add: orbit_disjoint) |
|
244 |
done |
|
245 |
||
246 |
definition trace :: "'a \<Rightarrow> 'a perm \<Rightarrow> 'a list" |
|
247 |
where |
|
248 |
"trace a f = map (\<lambda>n. apply (f ^ n) a) [0..<order f a]" |
|
249 |
||
250 |
lemma set_trace_eq [simp]: |
|
251 |
"set (trace a f) = orbit f a" |
|
252 |
by (auto simp add: trace_def orbit_unfold_image) |
|
253 |
||
254 |
definition seeds :: "'a perm \<Rightarrow> 'a::linorder list" |
|
255 |
where |
|
256 |
"seeds f = sorted_list_of_set (Min ` orbits f)" |
|
257 |
||
258 |
definition cycles :: "'a perm \<Rightarrow> 'a::linorder list list" |
|
259 |
where |
|
260 |
"cycles f = map (\<lambda>a. trace a f) (seeds f)" |
|
261 |
||
72536
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
262 |
end |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
263 |
|
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
264 |
|
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
265 |
text \<open>Misc\<close> |
589645894305
bundle mixins for locale and class specifications
haftmann
parents:
69745
diff
changeset
|
266 |
|
63375 | 267 |
lemma (in comm_monoid_list_set) sorted_list_of_set: |
268 |
assumes "finite A" |
|
269 |
shows "list.F (map h (sorted_list_of_set A)) = set.F h A" |
|
270 |
proof - |
|
271 |
from distinct_sorted_list_of_set |
|
272 |
have "set.F h (set (sorted_list_of_set A)) = list.F (map h (sorted_list_of_set A))" |
|
273 |
by (rule distinct_set_conv_list) |
|
274 |
with \<open>finite A\<close> show ?thesis |
|
68084 | 275 |
by (simp) |
63375 | 276 |
qed |
277 |
||
278 |
primrec subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
279 |
where |
|
280 |
"subtract [] ys = ys" |
|
281 |
| "subtract (x # xs) ys = subtract xs (removeAll x ys)" |
|
282 |
||
283 |
lemma length_subtract_less_eq [simp]: |
|
284 |
"length (subtract xs ys) \<le> length ys" |
|
285 |
proof (induct xs arbitrary: ys) |
|
286 |
case Nil then show ?case by simp |
|
287 |
next |
|
288 |
case (Cons x xs) |
|
289 |
then have "length (subtract xs (removeAll x ys)) \<le> length (removeAll x ys)" . |
|
290 |
also have "length (removeAll x ys) \<le> length ys" |
|
291 |
by simp |
|
292 |
finally show ?case |
|
293 |
by simp |
|
294 |
qed |
|
295 |
||
296 |
end |