| 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 | 
 | 
|  |     11 | consts
 | 
|  |     12 |   perm :: "('a list * 'a list) set"
 | 
|  |     13 | 
 | 
|  |     14 | syntax
 | 
|  |     15 |   "_perm" :: "'a list => 'a list => bool"    ("_ <~~> _"  [50, 50] 50)
 | 
|  |     16 | translations
 | 
|  |     17 |   "x <~~> y" == "(x, y) \<in> perm"
 | 
|  |     18 | 
 | 
|  |     19 | inductive perm
 | 
| 11153 |     20 |   intros
 | 
|  |     21 |     Nil  [intro!]: "[] <~~> []"
 | 
|  |     22 |     swap [intro!]: "y # x # l <~~> x # y # l"
 | 
|  |     23 |     Cons [intro!]: "xs <~~> ys ==> z # xs <~~> z # ys"
 | 
|  |     24 |     trans [intro]: "xs <~~> ys ==> ys <~~> zs ==> xs <~~> zs"
 | 
| 11054 |     25 | 
 | 
|  |     26 | lemma perm_refl [iff]: "l <~~> l"
 | 
| 17200 |     27 |   by (induct l) auto
 | 
| 11054 |     28 | 
 | 
|  |     29 | 
 | 
|  |     30 | subsection {* Some examples of rule induction on permutations *}
 | 
|  |     31 | 
 | 
|  |     32 | lemma xperm_empty_imp_aux: "xs <~~> ys ==> xs = [] --> ys = []"
 | 
| 17200 |     33 |     -- {*the form of the premise lets the induction bind @{term xs}
 | 
| 15072 |     34 |          and @{term ys} *}
 | 
| 11054 |     35 |   apply (erule perm.induct)
 | 
|  |     36 |      apply (simp_all (no_asm_simp))
 | 
|  |     37 |   done
 | 
|  |     38 | 
 | 
|  |     39 | lemma xperm_empty_imp: "[] <~~> ys ==> ys = []"
 | 
| 17200 |     40 |   using xperm_empty_imp_aux by blast
 | 
| 11054 |     41 | 
 | 
|  |     42 | 
 | 
|  |     43 | text {*
 | 
|  |     44 |   \medskip This more general theorem is easier to understand!
 | 
|  |     45 |   *}
 | 
|  |     46 | 
 | 
|  |     47 | lemma perm_length: "xs <~~> ys ==> length xs = length ys"
 | 
| 17200 |     48 |   by (erule perm.induct) simp_all
 | 
| 11054 |     49 | 
 | 
|  |     50 | lemma perm_empty_imp: "[] <~~> xs ==> xs = []"
 | 
| 17200 |     51 |   by (drule perm_length) auto
 | 
| 11054 |     52 | 
 | 
|  |     53 | lemma perm_sym: "xs <~~> ys ==> ys <~~> xs"
 | 
| 17200 |     54 |   by (erule perm.induct) auto
 | 
| 11054 |     55 | 
 | 
|  |     56 | lemma perm_mem [rule_format]: "xs <~~> ys ==> x mem xs --> x mem ys"
 | 
| 17200 |     57 |   by (erule perm.induct) auto
 | 
| 11054 |     58 | 
 | 
|  |     59 | 
 | 
|  |     60 | subsection {* Ways of making new permutations *}
 | 
|  |     61 | 
 | 
|  |     62 | text {*
 | 
|  |     63 |   We can insert the head anywhere in the list.
 | 
|  |     64 | *}
 | 
|  |     65 | 
 | 
|  |     66 | lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys"
 | 
| 17200 |     67 |   by (induct xs) auto
 | 
| 11054 |     68 | 
 | 
|  |     69 | lemma perm_append_swap: "xs @ ys <~~> ys @ xs"
 | 
| 17200 |     70 |   apply (induct xs)
 | 
|  |     71 |     apply simp_all
 | 
| 11054 |     72 |   apply (blast intro: perm_append_Cons)
 | 
|  |     73 |   done
 | 
|  |     74 | 
 | 
|  |     75 | lemma perm_append_single: "a # xs <~~> xs @ [a]"
 | 
| 17200 |     76 |   by (rule perm.trans [OF _ perm_append_swap]) simp
 | 
| 11054 |     77 | 
 | 
|  |     78 | lemma perm_rev: "rev xs <~~> xs"
 | 
| 17200 |     79 |   apply (induct xs)
 | 
|  |     80 |    apply simp_all
 | 
| 11153 |     81 |   apply (blast intro!: perm_append_single intro: perm_sym)
 | 
| 11054 |     82 |   done
 | 
|  |     83 | 
 | 
|  |     84 | lemma perm_append1: "xs <~~> ys ==> l @ xs <~~> l @ ys"
 | 
| 17200 |     85 |   by (induct l) auto
 | 
| 11054 |     86 | 
 | 
|  |     87 | lemma perm_append2: "xs <~~> ys ==> xs @ l <~~> ys @ l"
 | 
| 17200 |     88 |   by (blast intro!: perm_append_swap perm_append1)
 | 
| 11054 |     89 | 
 | 
|  |     90 | 
 | 
|  |     91 | subsection {* Further results *}
 | 
|  |     92 | 
 | 
|  |     93 | lemma perm_empty [iff]: "([] <~~> xs) = (xs = [])"
 | 
| 17200 |     94 |   by (blast intro: perm_empty_imp)
 | 
| 11054 |     95 | 
 | 
|  |     96 | lemma perm_empty2 [iff]: "(xs <~~> []) = (xs = [])"
 | 
|  |     97 |   apply auto
 | 
|  |     98 |   apply (erule perm_sym [THEN perm_empty_imp])
 | 
|  |     99 |   done
 | 
|  |    100 | 
 | 
|  |    101 | lemma perm_sing_imp [rule_format]: "ys <~~> xs ==> xs = [y] --> ys = [y]"
 | 
| 17200 |    102 |   by (erule perm.induct) auto
 | 
| 11054 |    103 | 
 | 
|  |    104 | lemma perm_sing_eq [iff]: "(ys <~~> [y]) = (ys = [y])"
 | 
| 17200 |    105 |   by (blast intro: perm_sing_imp)
 | 
| 11054 |    106 | 
 | 
|  |    107 | lemma perm_sing_eq2 [iff]: "([y] <~~> ys) = (ys = [y])"
 | 
| 17200 |    108 |   by (blast dest: perm_sym)
 | 
| 11054 |    109 | 
 | 
|  |    110 | 
 | 
|  |    111 | subsection {* Removing elements *}
 | 
|  |    112 | 
 | 
|  |    113 | consts
 | 
|  |    114 |   remove :: "'a => 'a list => 'a list"
 | 
|  |    115 | primrec
 | 
|  |    116 |   "remove x [] = []"
 | 
|  |    117 |   "remove x (y # ys) = (if x = y then ys else y # remove x ys)"
 | 
|  |    118 | 
 | 
|  |    119 | lemma perm_remove: "x \<in> set ys ==> ys <~~> x # remove x ys"
 | 
| 17200 |    120 |   by (induct ys) auto
 | 
| 11054 |    121 | 
 | 
|  |    122 | lemma remove_commute: "remove x (remove y l) = remove y (remove x l)"
 | 
| 17200 |    123 |   by (induct l) auto
 | 
| 11054 |    124 | 
 | 
| 17200 |    125 | lemma multiset_of_remove[simp]:
 | 
|  |    126 |     "multiset_of (remove a x) = multiset_of x - {#a#}"
 | 
|  |    127 |   apply (induct x)
 | 
|  |    128 |    apply (auto simp: multiset_eq_conv_count_eq)
 | 
|  |    129 |   done
 | 
| 15072 |    130 | 
 | 
| 11054 |    131 | 
 | 
|  |    132 | text {* \medskip Congruence rule *}
 | 
|  |    133 | 
 | 
|  |    134 | lemma perm_remove_perm: "xs <~~> ys ==> remove z xs <~~> remove z ys"
 | 
| 17200 |    135 |   by (erule perm.induct) auto
 | 
| 11054 |    136 | 
 | 
|  |    137 | lemma remove_hd [simp]: "remove z (z # xs) = xs"
 | 
| 15072 |    138 |   by auto
 | 
| 11054 |    139 | 
 | 
|  |    140 | lemma cons_perm_imp_perm: "z # xs <~~> z # ys ==> xs <~~> ys"
 | 
| 17200 |    141 |   by (drule_tac z = z in perm_remove_perm) auto
 | 
| 11054 |    142 | 
 | 
|  |    143 | lemma cons_perm_eq [iff]: "(z#xs <~~> z#ys) = (xs <~~> ys)"
 | 
| 17200 |    144 |   by (blast intro: cons_perm_imp_perm)
 | 
| 11054 |    145 | 
 | 
|  |    146 | lemma append_perm_imp_perm: "!!xs ys. zs @ xs <~~> zs @ ys ==> xs <~~> ys"
 | 
|  |    147 |   apply (induct zs rule: rev_induct)
 | 
|  |    148 |    apply (simp_all (no_asm_use))
 | 
|  |    149 |   apply blast
 | 
|  |    150 |   done
 | 
|  |    151 | 
 | 
|  |    152 | lemma perm_append1_eq [iff]: "(zs @ xs <~~> zs @ ys) = (xs <~~> ys)"
 | 
| 17200 |    153 |   by (blast intro: append_perm_imp_perm perm_append1)
 | 
| 11054 |    154 | 
 | 
|  |    155 | lemma perm_append2_eq [iff]: "(xs @ zs <~~> ys @ zs) = (xs <~~> ys)"
 | 
|  |    156 |   apply (safe intro!: perm_append2)
 | 
|  |    157 |   apply (rule append_perm_imp_perm)
 | 
|  |    158 |   apply (rule perm_append_swap [THEN perm.trans])
 | 
|  |    159 |     -- {* the previous step helps this @{text blast} call succeed quickly *}
 | 
|  |    160 |   apply (blast intro: perm_append_swap)
 | 
|  |    161 |   done
 | 
|  |    162 | 
 | 
| 15072 |    163 | lemma multiset_of_eq_perm: "(multiset_of xs = multiset_of ys) = (xs <~~> ys) "
 | 
| 17200 |    164 |   apply (rule iffI)
 | 
|  |    165 |   apply (erule_tac [2] perm.induct, simp_all add: union_ac)
 | 
|  |    166 |   apply (erule rev_mp, rule_tac x=ys in spec)
 | 
|  |    167 |   apply (induct_tac xs, auto)
 | 
|  |    168 |   apply (erule_tac x = "remove a x" in allE, drule sym, simp)
 | 
|  |    169 |   apply (subgoal_tac "a \<in> set x")
 | 
|  |    170 |   apply (drule_tac z=a in perm.Cons)
 | 
|  |    171 |   apply (erule perm.trans, rule perm_sym, erule perm_remove)
 | 
| 15005 |    172 |   apply (drule_tac f=set_of in arg_cong, simp)
 | 
|  |    173 |   done
 | 
|  |    174 | 
 | 
| 17200 |    175 | lemma multiset_of_le_perm_append:
 | 
|  |    176 |     "(multiset_of xs \<le># multiset_of ys) = (\<exists>zs. xs @ zs <~~> ys)";
 | 
|  |    177 |   apply (auto simp: multiset_of_eq_perm[THEN sym] mset_le_exists_conv)
 | 
| 15072 |    178 |   apply (insert surj_multiset_of, drule surjD)
 | 
|  |    179 |   apply (blast intro: sym)+
 | 
|  |    180 |   done
 | 
| 15005 |    181 | 
 | 
| 11054 |    182 | end
 |