| 11054 |      1 | (*  Title:      HOL/Library/Permutation.thy
 | 
| 15005 |      2 |     Author:     Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker
 | 
| 11054 |      3 | *)
 | 
|  |      4 | 
 | 
| 14706 |      5 | header {* Permutations *}
 | 
| 11054 |      6 | 
 | 
| 15131 |      7 | theory Permutation
 | 
| 15140 |      8 | imports Multiset
 | 
| 15131 |      9 | begin
 | 
| 11054 |     10 | 
 | 
| 23755 |     11 | inductive
 | 
|  |     12 |   perm :: "'a list => 'a list => bool"  ("_ <~~> _"  [50, 50] 50)
 | 
|  |     13 |   where
 | 
| 11153 |     14 |     Nil  [intro!]: "[] <~~> []"
 | 
| 23755 |     15 |   | swap [intro!]: "y # x # l <~~> x # y # l"
 | 
|  |     16 |   | Cons [intro!]: "xs <~~> ys ==> z # xs <~~> z # ys"
 | 
|  |     17 |   | trans [intro]: "xs <~~> ys ==> ys <~~> zs ==> xs <~~> zs"
 | 
| 11054 |     18 | 
 | 
|  |     19 | lemma perm_refl [iff]: "l <~~> l"
 | 
| 17200 |     20 |   by (induct l) auto
 | 
| 11054 |     21 | 
 | 
|  |     22 | 
 | 
|  |     23 | subsection {* Some examples of rule induction on permutations *}
 | 
|  |     24 | 
 | 
|  |     25 | lemma xperm_empty_imp: "[] <~~> ys ==> ys = []"
 | 
| 25379 |     26 |   by (induct xs == "[]::'a list" ys pred: perm) simp_all
 | 
| 11054 |     27 | 
 | 
|  |     28 | 
 | 
|  |     29 | text {*
 | 
|  |     30 |   \medskip This more general theorem is easier to understand!
 | 
|  |     31 |   *}
 | 
|  |     32 | 
 | 
|  |     33 | lemma perm_length: "xs <~~> ys ==> length xs = length ys"
 | 
| 25379 |     34 |   by (induct pred: perm) simp_all
 | 
| 11054 |     35 | 
 | 
|  |     36 | lemma perm_empty_imp: "[] <~~> xs ==> xs = []"
 | 
| 17200 |     37 |   by (drule perm_length) auto
 | 
| 11054 |     38 | 
 | 
|  |     39 | lemma perm_sym: "xs <~~> ys ==> ys <~~> xs"
 | 
| 25379 |     40 |   by (induct pred: perm) auto
 | 
| 11054 |     41 | 
 | 
|  |     42 | 
 | 
|  |     43 | subsection {* Ways of making new permutations *}
 | 
|  |     44 | 
 | 
|  |     45 | text {*
 | 
|  |     46 |   We can insert the head anywhere in the list.
 | 
|  |     47 | *}
 | 
|  |     48 | 
 | 
|  |     49 | lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys"
 | 
| 17200 |     50 |   by (induct xs) auto
 | 
| 11054 |     51 | 
 | 
|  |     52 | lemma perm_append_swap: "xs @ ys <~~> ys @ xs"
 | 
| 17200 |     53 |   apply (induct xs)
 | 
|  |     54 |     apply simp_all
 | 
| 11054 |     55 |   apply (blast intro: perm_append_Cons)
 | 
|  |     56 |   done
 | 
|  |     57 | 
 | 
|  |     58 | lemma perm_append_single: "a # xs <~~> xs @ [a]"
 | 
| 17200 |     59 |   by (rule perm.trans [OF _ perm_append_swap]) simp
 | 
| 11054 |     60 | 
 | 
|  |     61 | lemma perm_rev: "rev xs <~~> xs"
 | 
| 17200 |     62 |   apply (induct xs)
 | 
|  |     63 |    apply simp_all
 | 
| 11153 |     64 |   apply (blast intro!: perm_append_single intro: perm_sym)
 | 
| 11054 |     65 |   done
 | 
|  |     66 | 
 | 
|  |     67 | lemma perm_append1: "xs <~~> ys ==> l @ xs <~~> l @ ys"
 | 
| 17200 |     68 |   by (induct l) auto
 | 
| 11054 |     69 | 
 | 
|  |     70 | lemma perm_append2: "xs <~~> ys ==> xs @ l <~~> ys @ l"
 | 
| 17200 |     71 |   by (blast intro!: perm_append_swap perm_append1)
 | 
| 11054 |     72 | 
 | 
|  |     73 | 
 | 
|  |     74 | subsection {* Further results *}
 | 
|  |     75 | 
 | 
|  |     76 | lemma perm_empty [iff]: "([] <~~> xs) = (xs = [])"
 | 
| 17200 |     77 |   by (blast intro: perm_empty_imp)
 | 
| 11054 |     78 | 
 | 
|  |     79 | lemma perm_empty2 [iff]: "(xs <~~> []) = (xs = [])"
 | 
|  |     80 |   apply auto
 | 
|  |     81 |   apply (erule perm_sym [THEN perm_empty_imp])
 | 
|  |     82 |   done
 | 
|  |     83 | 
 | 
| 25379 |     84 | lemma perm_sing_imp: "ys <~~> xs ==> xs = [y] ==> ys = [y]"
 | 
|  |     85 |   by (induct pred: perm) auto
 | 
| 11054 |     86 | 
 | 
|  |     87 | lemma perm_sing_eq [iff]: "(ys <~~> [y]) = (ys = [y])"
 | 
| 17200 |     88 |   by (blast intro: perm_sing_imp)
 | 
| 11054 |     89 | 
 | 
|  |     90 | lemma perm_sing_eq2 [iff]: "([y] <~~> ys) = (ys = [y])"
 | 
| 17200 |     91 |   by (blast dest: perm_sym)
 | 
| 11054 |     92 | 
 | 
|  |     93 | 
 | 
|  |     94 | subsection {* Removing elements *}
 | 
|  |     95 | 
 | 
|  |     96 | consts
 | 
|  |     97 |   remove :: "'a => 'a list => 'a list"
 | 
|  |     98 | primrec
 | 
|  |     99 |   "remove x [] = []"
 | 
|  |    100 |   "remove x (y # ys) = (if x = y then ys else y # remove x ys)"
 | 
|  |    101 | 
 | 
|  |    102 | lemma perm_remove: "x \<in> set ys ==> ys <~~> x # remove x ys"
 | 
| 17200 |    103 |   by (induct ys) auto
 | 
| 11054 |    104 | 
 | 
|  |    105 | lemma remove_commute: "remove x (remove y l) = remove y (remove x l)"
 | 
| 17200 |    106 |   by (induct l) auto
 | 
| 11054 |    107 | 
 | 
| 25379 |    108 | lemma multiset_of_remove [simp]:
 | 
| 17200 |    109 |     "multiset_of (remove a x) = multiset_of x - {#a#}"
 | 
|  |    110 |   apply (induct x)
 | 
|  |    111 |    apply (auto simp: multiset_eq_conv_count_eq)
 | 
|  |    112 |   done
 | 
| 15072 |    113 | 
 | 
| 11054 |    114 | 
 | 
|  |    115 | text {* \medskip Congruence rule *}
 | 
|  |    116 | 
 | 
|  |    117 | lemma perm_remove_perm: "xs <~~> ys ==> remove z xs <~~> remove z ys"
 | 
| 25379 |    118 |   by (induct pred: perm) auto
 | 
| 11054 |    119 | 
 | 
|  |    120 | lemma remove_hd [simp]: "remove z (z # xs) = xs"
 | 
| 15072 |    121 |   by auto
 | 
| 11054 |    122 | 
 | 
|  |    123 | lemma cons_perm_imp_perm: "z # xs <~~> z # ys ==> xs <~~> ys"
 | 
| 17200 |    124 |   by (drule_tac z = z in perm_remove_perm) auto
 | 
| 11054 |    125 | 
 | 
|  |    126 | lemma cons_perm_eq [iff]: "(z#xs <~~> z#ys) = (xs <~~> ys)"
 | 
| 17200 |    127 |   by (blast intro: cons_perm_imp_perm)
 | 
| 11054 |    128 | 
 | 
| 25379 |    129 | lemma append_perm_imp_perm: "zs @ xs <~~> zs @ ys ==> xs <~~> ys"
 | 
|  |    130 |   apply (induct zs arbitrary: xs ys rule: rev_induct)
 | 
| 11054 |    131 |    apply (simp_all (no_asm_use))
 | 
|  |    132 |   apply blast
 | 
|  |    133 |   done
 | 
|  |    134 | 
 | 
|  |    135 | lemma perm_append1_eq [iff]: "(zs @ xs <~~> zs @ ys) = (xs <~~> ys)"
 | 
| 17200 |    136 |   by (blast intro: append_perm_imp_perm perm_append1)
 | 
| 11054 |    137 | 
 | 
|  |    138 | lemma perm_append2_eq [iff]: "(xs @ zs <~~> ys @ zs) = (xs <~~> ys)"
 | 
|  |    139 |   apply (safe intro!: perm_append2)
 | 
|  |    140 |   apply (rule append_perm_imp_perm)
 | 
|  |    141 |   apply (rule perm_append_swap [THEN perm.trans])
 | 
|  |    142 |     -- {* the previous step helps this @{text blast} call succeed quickly *}
 | 
|  |    143 |   apply (blast intro: perm_append_swap)
 | 
|  |    144 |   done
 | 
|  |    145 | 
 | 
| 15072 |    146 | lemma multiset_of_eq_perm: "(multiset_of xs = multiset_of ys) = (xs <~~> ys) "
 | 
| 17200 |    147 |   apply (rule iffI)
 | 
|  |    148 |   apply (erule_tac [2] perm.induct, simp_all add: union_ac)
 | 
|  |    149 |   apply (erule rev_mp, rule_tac x=ys in spec)
 | 
|  |    150 |   apply (induct_tac xs, auto)
 | 
|  |    151 |   apply (erule_tac x = "remove a x" in allE, drule sym, simp)
 | 
|  |    152 |   apply (subgoal_tac "a \<in> set x")
 | 
|  |    153 |   apply (drule_tac z=a in perm.Cons)
 | 
|  |    154 |   apply (erule perm.trans, rule perm_sym, erule perm_remove)
 | 
| 15005 |    155 |   apply (drule_tac f=set_of in arg_cong, simp)
 | 
|  |    156 |   done
 | 
|  |    157 | 
 | 
| 17200 |    158 | lemma multiset_of_le_perm_append:
 | 
|  |    159 |     "(multiset_of xs \<le># multiset_of ys) = (\<exists>zs. xs @ zs <~~> ys)";
 | 
|  |    160 |   apply (auto simp: multiset_of_eq_perm[THEN sym] mset_le_exists_conv)
 | 
| 15072 |    161 |   apply (insert surj_multiset_of, drule surjD)
 | 
|  |    162 |   apply (blast intro: sym)+
 | 
|  |    163 |   done
 | 
| 15005 |    164 | 
 | 
| 25277 |    165 | lemma perm_set_eq: "xs <~~> ys ==> set xs = set ys"
 | 
| 25379 |    166 |   by (metis multiset_of_eq_perm multiset_of_eq_setD)
 | 
| 25277 |    167 | 
 | 
|  |    168 | lemma perm_distinct_iff: "xs <~~> ys ==> distinct xs = distinct ys"
 | 
| 25379 |    169 |   apply (induct pred: perm)
 | 
|  |    170 |      apply simp_all
 | 
|  |    171 |    apply fastsimp
 | 
|  |    172 |   apply (metis perm_set_eq)
 | 
|  |    173 |   done
 | 
| 25277 |    174 | 
 | 
| 25287 |    175 | lemma eq_set_perm_remdups: "set xs = set ys ==> remdups xs <~~> remdups ys"
 | 
| 25379 |    176 |   apply (induct xs arbitrary: ys rule: length_induct)
 | 
|  |    177 |   apply (case_tac "remdups xs", simp, simp)
 | 
|  |    178 |   apply (subgoal_tac "a : set (remdups ys)")
 | 
|  |    179 |    prefer 2 apply (metis set.simps(2) insert_iff set_remdups)
 | 
|  |    180 |   apply (drule split_list) apply(elim exE conjE)
 | 
|  |    181 |   apply (drule_tac x=list in spec) apply(erule impE) prefer 2
 | 
|  |    182 |    apply (drule_tac x="ysa@zs" in spec) apply(erule impE) prefer 2
 | 
|  |    183 |     apply simp
 | 
|  |    184 |     apply (subgoal_tac "a#list <~~> a#ysa@zs")
 | 
|  |    185 |      apply (metis Cons_eq_appendI perm_append_Cons trans)
 | 
|  |    186 |     apply (metis Cons Cons_eq_appendI distinct.simps(2)
 | 
|  |    187 |       distinct_remdups distinct_remdups_id perm_append_swap perm_distinct_iff)
 | 
|  |    188 |    apply (subgoal_tac "set (a#list) = set (ysa@a#zs) & distinct (a#list) & distinct (ysa@a#zs)")
 | 
|  |    189 |     apply (fastsimp simp add: insert_ident)
 | 
|  |    190 |    apply (metis distinct_remdups set_remdups)
 | 
|  |    191 |   apply (metis Nat.le_less_trans Suc_length_conv le_def length_remdups_leq less_Suc_eq)
 | 
|  |    192 |   done
 | 
| 25287 |    193 | 
 | 
|  |    194 | lemma perm_remdups_iff_eq_set: "remdups x <~~> remdups y = (set x = set y)"
 | 
| 25379 |    195 |   by (metis List.set_remdups perm_set_eq eq_set_perm_remdups)
 | 
| 25287 |    196 | 
 | 
| 11054 |    197 | end
 |