# HG changeset patch # User paulson # Date 1088092493 -7200 # Node ID 546c8e7e28d40c35101f44388e4d849c173f2574 # Parent 44ac09ba7213a9c3c8577b87115401e0a92ed587 Norbert Voelker diff -r 44ac09ba7213 -r 546c8e7e28d4 src/HOL/Library/Permutation.thy --- a/src/HOL/Library/Permutation.thy Thu Jun 24 17:52:55 2004 +0200 +++ b/src/HOL/Library/Permutation.thy Thu Jun 24 17:54:53 2004 +0200 @@ -1,15 +1,10 @@ (* Title: HOL/Library/Permutation.thy - ID: $Id$ - Author: Lawrence C Paulson and Thomas M Rasmussen - Copyright 1995 University of Cambridge - -TODO: it would be nice to prove (for "multiset", defined on -HOL/ex/Sorting.thy) xs <~~> ys = (\x. multiset xs x = multiset ys x) + Author: Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker *) header {* Permutations *} -theory Permutation = Main: +theory Permutation = Multiset: consts perm :: "('a list * 'a list) set" @@ -27,9 +22,7 @@ trans [intro]: "xs <~~> ys ==> ys <~~> zs ==> xs <~~> zs" lemma perm_refl [iff]: "l <~~> l" - apply (induct l) - apply auto - done +by (induct l, auto) subsection {* Some examples of rule induction on permutations *} @@ -41,9 +34,7 @@ done lemma xperm_empty_imp: "[] <~~> ys ==> ys = []" - apply (insert xperm_empty_imp_aux) - apply blast - done +by (insert xperm_empty_imp_aux, blast) text {* @@ -51,24 +42,16 @@ *} lemma perm_length: "xs <~~> ys ==> length xs = length ys" - apply (erule perm.induct) - apply simp_all - done +by (erule perm.induct, simp_all) lemma perm_empty_imp: "[] <~~> xs ==> xs = []" - apply (drule perm_length) - apply auto - done +by (drule perm_length, auto) lemma perm_sym: "xs <~~> ys ==> ys <~~> xs" - apply (erule perm.induct) - apply auto - done +by (erule perm.induct, auto) lemma perm_mem [rule_format]: "xs <~~> ys ==> x mem xs --> x mem ys" - apply (erule perm.induct) - apply auto - done +by (erule perm.induct, auto) subsection {* Ways of making new permutations *} @@ -78,44 +61,35 @@ *} lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" - apply (induct xs) - apply auto - done +by (induct xs, auto) lemma perm_append_swap: "xs @ ys <~~> ys @ xs" - apply (induct xs) - apply simp_all + apply (induct xs, simp_all) apply (blast intro: perm_append_Cons) done lemma perm_append_single: "a # xs <~~> xs @ [a]" apply (rule perm.trans) prefer 2 - apply (rule perm_append_swap) - apply simp + apply (rule perm_append_swap, simp) done lemma perm_rev: "rev xs <~~> xs" - apply (induct xs) - apply simp_all + apply (induct xs, simp_all) apply (blast intro!: perm_append_single intro: perm_sym) done lemma perm_append1: "xs <~~> ys ==> l @ xs <~~> l @ ys" - apply (induct l) - apply auto - done +by (induct l, auto) lemma perm_append2: "xs <~~> ys ==> xs @ l <~~> ys @ l" - apply (blast intro!: perm_append_swap perm_append1) - done +by (blast intro!: perm_append_swap perm_append1) subsection {* Further results *} lemma perm_empty [iff]: "([] <~~> xs) = (xs = [])" - apply (blast intro: perm_empty_imp) - done +by (blast intro: perm_empty_imp) lemma perm_empty2 [iff]: "(xs <~~> []) = (xs = [])" apply auto @@ -123,17 +97,13 @@ done lemma perm_sing_imp [rule_format]: "ys <~~> xs ==> xs = [y] --> ys = [y]" - apply (erule perm.induct) - apply auto - done +by (erule perm.induct, auto) lemma perm_sing_eq [iff]: "(ys <~~> [y]) = (ys = [y])" - apply (blast intro: perm_sing_imp) - done +by (blast intro: perm_sing_imp) lemma perm_sing_eq2 [iff]: "([y] <~~> ys) = (ys = [y])" - apply (blast dest: perm_sym) - done +by (blast dest: perm_sym) subsection {* Removing elements *} @@ -145,35 +115,26 @@ "remove x (y # ys) = (if x = y then ys else y # remove x ys)" lemma perm_remove: "x \ set ys ==> ys <~~> x # remove x ys" - apply (induct ys) - apply auto - done +by (induct ys, auto) lemma remove_commute: "remove x (remove y l) = remove y (remove x l)" - apply (induct l) - apply auto - done +by (induct l, auto) text {* \medskip Congruence rule *} lemma perm_remove_perm: "xs <~~> ys ==> remove z xs <~~> remove z ys" - apply (erule perm.induct) - apply auto - done +by (erule perm.induct, auto) lemma remove_hd [simp]: "remove z (z # xs) = xs" apply auto done lemma cons_perm_imp_perm: "z # xs <~~> z # ys ==> xs <~~> ys" - apply (drule_tac z = z in perm_remove_perm) - apply auto - done +by (drule_tac z = z in perm_remove_perm, auto) lemma cons_perm_eq [iff]: "(z#xs <~~> z#ys) = (xs <~~> ys)" - apply (blast intro: cons_perm_imp_perm) - done +by (blast intro: cons_perm_imp_perm) lemma append_perm_imp_perm: "!!xs ys. zs @ xs <~~> zs @ ys ==> xs <~~> ys" apply (induct zs rule: rev_induct) @@ -182,8 +143,7 @@ done lemma perm_append1_eq [iff]: "(zs @ xs <~~> zs @ ys) = (xs <~~> ys)" - apply (blast intro: append_perm_imp_perm perm_append1) - done +by (blast intro: append_perm_imp_perm perm_append1) lemma perm_append2_eq [iff]: "(xs @ zs <~~> ys @ zs) = (xs <~~> ys)" apply (safe intro!: perm_append2) @@ -193,4 +153,47 @@ apply (blast intro: perm_append_swap) done +(****************** Norbert Voelker 17 June 2004 **************) + +consts + multiset_of :: "'a list \ 'a multiset" +primrec + "multiset_of [] = {#}" + "multiset_of (a # x) = multiset_of x + {# a #}" + +lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" + by (induct_tac x, auto) + +lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" + by (induct_tac x, auto) + +lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" + by (induct_tac x, auto) + +lemma multiset_of_remove[simp]: + "multiset_of (remove a x) = multiset_of x - {#a#}" + by (induct_tac x, auto simp: multiset_eq_conv_count_eq) + +lemma multiset_of_eq_perm: "(multiset_of xs = multiset_of ys) = (xs <~~> ys) " + apply (rule iffI) + apply (erule_tac [2] perm.induct, simp_all add: union_ac) + apply (erule rev_mp, rule_tac x=ys in spec, induct_tac xs, auto) + apply (erule_tac x = "remove a x" in allE, drule sym, simp) + apply (subgoal_tac "a \ set x") + apply (drule_tac z=a in perm.Cons) + apply (erule perm.trans, rule perm_sym, erule perm_remove) + apply (drule_tac f=set_of in arg_cong, simp) + done + +lemma set_count_multiset_of: "set x = {a. 0 < count (multiset_of x) a}" + by (induct_tac x, auto) + +lemma distinct_count_multiset_of: + "distinct x \ count (multiset_of x) a = (if a \ set x then 1 else 0)" + by (erule rev_mp, induct_tac x, auto) + +lemma distinct_set_eq_iff_multiset_of_eq: + "\distinct x; distinct y\ \ (set x = set y) = (multiset_of x = multiset_of y)" + by (auto simp: multiset_eq_conv_count_eq distinct_count_multiset_of) + end