| author | wenzelm |
| Sun, 14 Mar 2021 22:55:52 +0100 | |
| changeset 73439 | cb127ce2c092 |
| parent 73329 | 2aef2de6b17c |
| permissions | -rw-r--r-- |
| 73297 | 1 |
(* Title: HOL/Library/List_Permutation.thy |
| 15005 | 2 |
Author: Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker |
| 11054 | 3 |
*) |
4 |
||
| 73297 | 5 |
section \<open>Permuted Lists\<close> |
| 11054 | 6 |
|
| 73297 | 7 |
theory List_Permutation |
|
73327
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
8 |
imports Permutations |
| 15131 | 9 |
begin |
| 11054 | 10 |
|
| 73329 | 11 |
text \<open> |
12 |
Note that multisets already provide the notion of permutated list and hence |
|
13 |
this theory mostly echoes material already logically present in theory |
|
14 |
\<^text>\<open>Permutations\<close>; it should be seldom needed. |
|
15 |
\<close> |
|
16 |
||
| 73305 | 17 |
subsection \<open>An inductive definition \ldots\<close> |
| 73300 | 18 |
|
| 70680 | 19 |
inductive perm :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infixr \<open><~~>\<close> 50) |
| 53238 | 20 |
where |
21 |
Nil [intro!]: "[] <~~> []" |
|
22 |
| swap [intro!]: "y # x # l <~~> x # y # l" |
|
23 |
| Cons [intro!]: "xs <~~> ys \<Longrightarrow> z # xs <~~> z # ys" |
|
24 |
| trans [intro]: "xs <~~> ys \<Longrightarrow> ys <~~> zs \<Longrightarrow> xs <~~> zs" |
|
| 11054 | 25 |
|
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
26 |
proposition perm_refl [iff]: "l <~~> l" |
| 17200 | 27 |
by (induct l) auto |
| 11054 | 28 |
|
| 73305 | 29 |
text \<open>\ldots that is equivalent to an already existing notion:\<close> |
| 11054 | 30 |
|
| 73300 | 31 |
lemma perm_iff_eq_mset: |
32 |
\<open>xs <~~> ys \<longleftrightarrow> mset xs = mset ys\<close> |
|
33 |
proof |
|
34 |
assume \<open>mset xs = mset ys\<close> |
|
35 |
then show \<open>xs <~~> ys\<close> |
|
36 |
proof (induction xs arbitrary: ys) |
|
37 |
case Nil |
|
38 |
then show ?case |
|
39 |
by simp |
|
40 |
next |
|
41 |
case (Cons x xs) |
|
42 |
from Cons.prems [symmetric] have \<open>mset xs = mset (remove1 x ys)\<close> |
|
43 |
by simp |
|
44 |
then have \<open>xs <~~> remove1 x ys\<close> |
|
45 |
by (rule Cons.IH) |
|
46 |
then have \<open>x # xs <~~> x # remove1 x ys\<close> |
|
47 |
by (rule perm.Cons) |
|
48 |
moreover from Cons.prems have \<open>x \<in> set ys\<close> |
|
49 |
by (auto dest: union_single_eq_member) |
|
50 |
then have \<open>x # remove1 x ys <~~> ys\<close> |
|
51 |
by (induction ys) auto |
|
52 |
ultimately show \<open>x # xs <~~> ys\<close> |
|
53 |
by (rule perm.trans) |
|
54 |
qed |
|
55 |
next |
|
56 |
assume \<open>xs <~~> ys\<close> |
|
57 |
then show \<open>mset xs = mset ys\<close> |
|
58 |
by induction simp_all |
|
59 |
qed |
|
| 11054 | 60 |
|
| 73300 | 61 |
theorem mset_eq_perm: \<open>mset xs = mset ys \<longleftrightarrow> xs <~~> ys\<close> |
62 |
by (simp add: perm_iff_eq_mset) |
|
| 11054 | 63 |
|
64 |
||
| 73300 | 65 |
subsection \<open>Nontrivial conclusions\<close> |
| 11054 | 66 |
|
| 73300 | 67 |
proposition perm_swap: |
68 |
\<open>xs[i := xs ! j, j := xs ! i] <~~> xs\<close> |
|
69 |
if \<open>i < length xs\<close> \<open>j < length xs\<close> |
|
| 73301 | 70 |
using that by (simp add: perm_iff_eq_mset mset_swap) |
| 15005 | 71 |
|
| 64587 | 72 |
proposition mset_le_perm_append: "mset xs \<subseteq># mset ys \<longleftrightarrow> (\<exists>zs. xs @ zs <~~> ys)" |
| 73300 | 73 |
by (auto simp add: perm_iff_eq_mset mset_subset_eq_exists_conv ex_mset dest: sym) |
| 15005 | 74 |
|
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
75 |
proposition perm_set_eq: "xs <~~> ys \<Longrightarrow> set xs = set ys" |
| 73300 | 76 |
by (rule mset_eq_setD) (simp add: perm_iff_eq_mset) |
| 25277 | 77 |
|
| 73301 | 78 |
proposition perm_distinct_iff: "xs <~~> ys \<Longrightarrow> distinct xs \<longleftrightarrow> distinct ys" |
79 |
by (rule mset_eq_imp_distinct_iff) (simp add: perm_iff_eq_mset) |
|
| 25277 | 80 |
|
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
81 |
theorem eq_set_perm_remdups: "set xs = set ys \<Longrightarrow> remdups xs <~~> remdups ys" |
| 73300 | 82 |
by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq) |
| 25287 | 83 |
|
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
84 |
proposition perm_remdups_iff_eq_set: "remdups x <~~> remdups y \<longleftrightarrow> set x = set y" |
| 73300 | 85 |
by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq) |
| 25287 | 86 |
|
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
87 |
theorem permutation_Ex_bij: |
| 39075 | 88 |
assumes "xs <~~> ys" |
89 |
shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))"
|
|
| 73301 | 90 |
proof - |
|
73327
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
91 |
from assms have \<open>mset xs = mset ys\<close> \<open>length xs = length ys\<close> |
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
92 |
by (auto simp add: perm_iff_eq_mset dest: mset_eq_length) |
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
93 |
from \<open>mset xs = mset ys\<close> obtain p where \<open>p permutes {..<length ys}\<close> \<open>permute_list p ys = xs\<close>
|
| 73301 | 94 |
by (rule mset_eq_permutation) |
|
73327
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
95 |
then have \<open>bij_betw p {..<length xs} {..<length ys}\<close>
|
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
96 |
by (simp add: \<open>length xs = length ys\<close> permutes_imp_bij) |
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
97 |
moreover have \<open>\<forall>i<length xs. xs ! i = ys ! (p i)\<close> |
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
98 |
using \<open>permute_list p ys = xs\<close> \<open>length xs = length ys\<close> \<open>p permutes {..<length ys}\<close> permute_list_nth
|
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
99 |
by auto |
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
100 |
ultimately show ?thesis |
|
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73305
diff
changeset
|
101 |
by blast |
| 39075 | 102 |
qed |
103 |
||
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
104 |
proposition perm_finite: "finite {B. B <~~> A}"
|
| 73301 | 105 |
using mset_eq_finite by (auto simp add: perm_iff_eq_mset) |
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
106 |
|
| 73300 | 107 |
|
108 |
subsection \<open>Trivial conclusions:\<close> |
|
109 |
||
110 |
proposition perm_empty_imp: "[] <~~> ys \<Longrightarrow> ys = []" |
|
111 |
by (simp add: perm_iff_eq_mset) |
|
112 |
||
113 |
||
114 |
text \<open>\medskip This more general theorem is easier to understand!\<close> |
|
115 |
||
116 |
proposition perm_length: "xs <~~> ys \<Longrightarrow> length xs = length ys" |
|
117 |
by (rule mset_eq_length) (simp add: perm_iff_eq_mset) |
|
118 |
||
119 |
proposition perm_sym: "xs <~~> ys \<Longrightarrow> ys <~~> xs" |
|
120 |
by (simp add: perm_iff_eq_mset) |
|
121 |
||
122 |
||
123 |
text \<open>We can insert the head anywhere in the list.\<close> |
|
124 |
||
125 |
proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" |
|
126 |
by (simp add: perm_iff_eq_mset) |
|
127 |
||
128 |
proposition perm_append_swap: "xs @ ys <~~> ys @ xs" |
|
129 |
by (simp add: perm_iff_eq_mset) |
|
130 |
||
131 |
proposition perm_append_single: "a # xs <~~> xs @ [a]" |
|
132 |
by (simp add: perm_iff_eq_mset) |
|
133 |
||
134 |
proposition perm_rev: "rev xs <~~> xs" |
|
135 |
by (simp add: perm_iff_eq_mset) |
|
136 |
||
137 |
proposition perm_append1: "xs <~~> ys \<Longrightarrow> l @ xs <~~> l @ ys" |
|
138 |
by (simp add: perm_iff_eq_mset) |
|
139 |
||
140 |
proposition perm_append2: "xs <~~> ys \<Longrightarrow> xs @ l <~~> ys @ l" |
|
141 |
by (simp add: perm_iff_eq_mset) |
|
142 |
||
143 |
proposition perm_empty [iff]: "[] <~~> xs \<longleftrightarrow> xs = []" |
|
144 |
by (simp add: perm_iff_eq_mset) |
|
145 |
||
146 |
proposition perm_empty2 [iff]: "xs <~~> [] \<longleftrightarrow> xs = []" |
|
147 |
by (simp add: perm_iff_eq_mset) |
|
148 |
||
149 |
proposition perm_sing_imp: "ys <~~> xs \<Longrightarrow> xs = [y] \<Longrightarrow> ys = [y]" |
|
150 |
by (simp add: perm_iff_eq_mset) |
|
151 |
||
152 |
proposition perm_sing_eq [iff]: "ys <~~> [y] \<longleftrightarrow> ys = [y]" |
|
153 |
by (simp add: perm_iff_eq_mset) |
|
154 |
||
155 |
proposition perm_sing_eq2 [iff]: "[y] <~~> ys \<longleftrightarrow> ys = [y]" |
|
156 |
by (simp add: perm_iff_eq_mset) |
|
157 |
||
158 |
proposition perm_remove: "x \<in> set ys \<Longrightarrow> ys <~~> x # remove1 x ys" |
|
159 |
by (simp add: perm_iff_eq_mset) |
|
160 |
||
161 |
||
162 |
text \<open>\medskip Congruence rule\<close> |
|
163 |
||
164 |
proposition perm_remove_perm: "xs <~~> ys \<Longrightarrow> remove1 z xs <~~> remove1 z ys" |
|
165 |
by (simp add: perm_iff_eq_mset) |
|
166 |
||
167 |
proposition remove_hd [simp]: "remove1 z (z # xs) = xs" |
|
168 |
by (simp add: perm_iff_eq_mset) |
|
169 |
||
170 |
proposition cons_perm_imp_perm: "z # xs <~~> z # ys \<Longrightarrow> xs <~~> ys" |
|
171 |
by (simp add: perm_iff_eq_mset) |
|
172 |
||
173 |
proposition cons_perm_eq [simp]: "z#xs <~~> z#ys \<longleftrightarrow> xs <~~> ys" |
|
174 |
by (simp add: perm_iff_eq_mset) |
|
175 |
||
176 |
proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \<Longrightarrow> xs <~~> ys" |
|
177 |
by (simp add: perm_iff_eq_mset) |
|
178 |
||
179 |
proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \<longleftrightarrow> xs <~~> ys" |
|
180 |
by (simp add: perm_iff_eq_mset) |
|
181 |
||
182 |
proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \<longleftrightarrow> xs <~~> ys" |
|
183 |
by (simp add: perm_iff_eq_mset) |
|
|
61699
a81dc5c4d6a9
New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents:
61585
diff
changeset
|
184 |
|
| 11054 | 185 |
end |