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