wenzelm@11054: (* Title: HOL/Library/Permutation.thy paulson@15005: Author: Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker wenzelm@11054: *) wenzelm@11054: wenzelm@14706: header {* Permutations *} wenzelm@11054: paulson@15005: theory Permutation = Multiset: wenzelm@11054: wenzelm@11054: consts wenzelm@11054: perm :: "('a list * 'a list) set" wenzelm@11054: wenzelm@11054: syntax wenzelm@11054: "_perm" :: "'a list => 'a list => bool" ("_ <~~> _" [50, 50] 50) wenzelm@11054: translations wenzelm@11054: "x <~~> y" == "(x, y) \ perm" wenzelm@11054: wenzelm@11054: inductive perm paulson@11153: intros paulson@11153: Nil [intro!]: "[] <~~> []" paulson@11153: swap [intro!]: "y # x # l <~~> x # y # l" paulson@11153: Cons [intro!]: "xs <~~> ys ==> z # xs <~~> z # ys" paulson@11153: trans [intro]: "xs <~~> ys ==> ys <~~> zs ==> xs <~~> zs" wenzelm@11054: wenzelm@11054: lemma perm_refl [iff]: "l <~~> l" paulson@15005: by (induct l, auto) wenzelm@11054: wenzelm@11054: wenzelm@11054: subsection {* Some examples of rule induction on permutations *} wenzelm@11054: wenzelm@11054: lemma xperm_empty_imp_aux: "xs <~~> ys ==> xs = [] --> ys = []" wenzelm@11054: -- {* the form of the premise lets the induction bind @{term xs} and @{term ys} *} wenzelm@11054: apply (erule perm.induct) wenzelm@11054: apply (simp_all (no_asm_simp)) wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma xperm_empty_imp: "[] <~~> ys ==> ys = []" paulson@15005: by (insert xperm_empty_imp_aux, blast) wenzelm@11054: wenzelm@11054: wenzelm@11054: text {* wenzelm@11054: \medskip This more general theorem is easier to understand! wenzelm@11054: *} wenzelm@11054: wenzelm@11054: lemma perm_length: "xs <~~> ys ==> length xs = length ys" paulson@15005: by (erule perm.induct, simp_all) wenzelm@11054: wenzelm@11054: lemma perm_empty_imp: "[] <~~> xs ==> xs = []" paulson@15005: by (drule perm_length, auto) wenzelm@11054: wenzelm@11054: lemma perm_sym: "xs <~~> ys ==> ys <~~> xs" paulson@15005: by (erule perm.induct, auto) wenzelm@11054: wenzelm@11054: lemma perm_mem [rule_format]: "xs <~~> ys ==> x mem xs --> x mem ys" paulson@15005: by (erule perm.induct, auto) wenzelm@11054: wenzelm@11054: wenzelm@11054: subsection {* Ways of making new permutations *} wenzelm@11054: wenzelm@11054: text {* wenzelm@11054: We can insert the head anywhere in the list. wenzelm@11054: *} wenzelm@11054: wenzelm@11054: lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" paulson@15005: by (induct xs, auto) wenzelm@11054: wenzelm@11054: lemma perm_append_swap: "xs @ ys <~~> ys @ xs" paulson@15005: apply (induct xs, simp_all) wenzelm@11054: apply (blast intro: perm_append_Cons) wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma perm_append_single: "a # xs <~~> xs @ [a]" wenzelm@11054: apply (rule perm.trans) wenzelm@11054: prefer 2 paulson@15005: apply (rule perm_append_swap, simp) wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma perm_rev: "rev xs <~~> xs" paulson@15005: apply (induct xs, simp_all) paulson@11153: apply (blast intro!: perm_append_single intro: perm_sym) wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma perm_append1: "xs <~~> ys ==> l @ xs <~~> l @ ys" paulson@15005: by (induct l, auto) wenzelm@11054: wenzelm@11054: lemma perm_append2: "xs <~~> ys ==> xs @ l <~~> ys @ l" paulson@15005: by (blast intro!: perm_append_swap perm_append1) wenzelm@11054: wenzelm@11054: wenzelm@11054: subsection {* Further results *} wenzelm@11054: wenzelm@11054: lemma perm_empty [iff]: "([] <~~> xs) = (xs = [])" paulson@15005: by (blast intro: perm_empty_imp) wenzelm@11054: wenzelm@11054: lemma perm_empty2 [iff]: "(xs <~~> []) = (xs = [])" wenzelm@11054: apply auto wenzelm@11054: apply (erule perm_sym [THEN perm_empty_imp]) wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma perm_sing_imp [rule_format]: "ys <~~> xs ==> xs = [y] --> ys = [y]" paulson@15005: by (erule perm.induct, auto) wenzelm@11054: wenzelm@11054: lemma perm_sing_eq [iff]: "(ys <~~> [y]) = (ys = [y])" paulson@15005: by (blast intro: perm_sing_imp) wenzelm@11054: wenzelm@11054: lemma perm_sing_eq2 [iff]: "([y] <~~> ys) = (ys = [y])" paulson@15005: by (blast dest: perm_sym) wenzelm@11054: wenzelm@11054: wenzelm@11054: subsection {* Removing elements *} wenzelm@11054: wenzelm@11054: consts wenzelm@11054: remove :: "'a => 'a list => 'a list" wenzelm@11054: primrec wenzelm@11054: "remove x [] = []" wenzelm@11054: "remove x (y # ys) = (if x = y then ys else y # remove x ys)" wenzelm@11054: wenzelm@11054: lemma perm_remove: "x \ set ys ==> ys <~~> x # remove x ys" paulson@15005: by (induct ys, auto) wenzelm@11054: wenzelm@11054: lemma remove_commute: "remove x (remove y l) = remove y (remove x l)" paulson@15005: by (induct l, auto) wenzelm@11054: wenzelm@11054: wenzelm@11054: text {* \medskip Congruence rule *} wenzelm@11054: wenzelm@11054: lemma perm_remove_perm: "xs <~~> ys ==> remove z xs <~~> remove z ys" paulson@15005: by (erule perm.induct, auto) wenzelm@11054: wenzelm@11054: lemma remove_hd [simp]: "remove z (z # xs) = xs" wenzelm@11054: apply auto wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma cons_perm_imp_perm: "z # xs <~~> z # ys ==> xs <~~> ys" paulson@15005: by (drule_tac z = z in perm_remove_perm, auto) wenzelm@11054: wenzelm@11054: lemma cons_perm_eq [iff]: "(z#xs <~~> z#ys) = (xs <~~> ys)" paulson@15005: by (blast intro: cons_perm_imp_perm) wenzelm@11054: wenzelm@11054: lemma append_perm_imp_perm: "!!xs ys. zs @ xs <~~> zs @ ys ==> xs <~~> ys" wenzelm@11054: apply (induct zs rule: rev_induct) wenzelm@11054: apply (simp_all (no_asm_use)) wenzelm@11054: apply blast wenzelm@11054: done wenzelm@11054: wenzelm@11054: lemma perm_append1_eq [iff]: "(zs @ xs <~~> zs @ ys) = (xs <~~> ys)" paulson@15005: by (blast intro: append_perm_imp_perm perm_append1) wenzelm@11054: wenzelm@11054: lemma perm_append2_eq [iff]: "(xs @ zs <~~> ys @ zs) = (xs <~~> ys)" wenzelm@11054: apply (safe intro!: perm_append2) wenzelm@11054: apply (rule append_perm_imp_perm) wenzelm@11054: apply (rule perm_append_swap [THEN perm.trans]) wenzelm@11054: -- {* the previous step helps this @{text blast} call succeed quickly *} wenzelm@11054: apply (blast intro: perm_append_swap) wenzelm@11054: done wenzelm@11054: paulson@15005: (****************** Norbert Voelker 17 June 2004 **************) paulson@15005: paulson@15005: consts paulson@15005: multiset_of :: "'a list \ 'a multiset" paulson@15005: primrec paulson@15005: "multiset_of [] = {#}" paulson@15005: "multiset_of (a # x) = multiset_of x + {# a #}" paulson@15005: paulson@15005: lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" paulson@15005: by (induct_tac x, auto) paulson@15005: paulson@15005: lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" paulson@15005: by (induct_tac x, auto) paulson@15005: paulson@15005: lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" paulson@15005: by (induct_tac x, auto) paulson@15005: paulson@15005: lemma multiset_of_remove[simp]: paulson@15005: "multiset_of (remove a x) = multiset_of x - {#a#}" paulson@15005: by (induct_tac x, auto simp: multiset_eq_conv_count_eq) paulson@15005: paulson@15005: lemma multiset_of_eq_perm: "(multiset_of xs = multiset_of ys) = (xs <~~> ys) " paulson@15005: apply (rule iffI) paulson@15005: apply (erule_tac [2] perm.induct, simp_all add: union_ac) paulson@15005: apply (erule rev_mp, rule_tac x=ys in spec, induct_tac xs, auto) paulson@15005: apply (erule_tac x = "remove a x" in allE, drule sym, simp) paulson@15005: apply (subgoal_tac "a \ set x") paulson@15005: apply (drule_tac z=a in perm.Cons) paulson@15005: apply (erule perm.trans, rule perm_sym, erule perm_remove) paulson@15005: apply (drule_tac f=set_of in arg_cong, simp) paulson@15005: done paulson@15005: paulson@15005: lemma set_count_multiset_of: "set x = {a. 0 < count (multiset_of x) a}" paulson@15005: by (induct_tac x, auto) paulson@15005: paulson@15005: lemma distinct_count_multiset_of: paulson@15005: "distinct x \ count (multiset_of x) a = (if a \ set x then 1 else 0)" paulson@15005: by (erule rev_mp, induct_tac x, auto) paulson@15005: paulson@15005: lemma distinct_set_eq_iff_multiset_of_eq: paulson@15005: "\distinct x; distinct y\ \ (set x = set y) = (multiset_of x = multiset_of y)" paulson@15005: by (auto simp: multiset_eq_conv_count_eq distinct_count_multiset_of) paulson@15005: wenzelm@11054: end