diff --git a/src/HOL/Library/List_Permutation.thy b/src/HOL/Library/List_Permutation.thy --- a/src/HOL/Library/List_Permutation.thy +++ b/src/HOL/Library/List_Permutation.thy @@ -1,230 +1,253 @@ (* Title: HOL/Library/List_Permutation.thy Author: Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker *) section \Permuted Lists\ theory List_Permutation imports Multiset begin +subsection \An inductive definition\\ + inductive perm :: "'a list \ 'a list \ bool" (infixr \<~~>\ 50) where Nil [intro!]: "[] <~~> []" | swap [intro!]: "y # x # l <~~> x # y # l" | Cons [intro!]: "xs <~~> ys \ z # xs <~~> z # ys" | trans [intro]: "xs <~~> ys \ ys <~~> zs \ xs <~~> zs" proposition perm_refl [iff]: "l <~~> l" by (induct l) auto - -subsection \Some examples of rule induction on permutations\ - -proposition perm_empty_imp: "[] <~~> ys \ ys = []" - by (induction "[] :: 'a list" ys pred: perm) simp_all - - -text \\medskip This more general theorem is easier to understand!\ - -proposition perm_length: "xs <~~> ys \ length xs = length ys" - by (induct pred: perm) simp_all - -proposition perm_sym: "xs <~~> ys \ ys <~~> xs" - by (induct pred: perm) auto - - -subsection \Ways of making new permutations\ - -text \We can insert the head anywhere in the list.\ - -proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" - by (induct xs) auto - -proposition perm_append_swap: "xs @ ys <~~> ys @ xs" - by (induct xs) (auto intro: perm_append_Cons) - -proposition perm_append_single: "a # xs <~~> xs @ [a]" - by (rule perm.trans [OF _ perm_append_swap]) simp +text \\that is equivalent to an already existing notion:\ -proposition perm_rev: "rev xs <~~> xs" - by (induct xs) (auto intro!: perm_append_single intro: perm_sym) - -proposition perm_append1: "xs <~~> ys \ l @ xs <~~> l @ ys" - by (induct l) auto - -proposition perm_append2: "xs <~~> ys \ xs @ l <~~> ys @ l" - by (blast intro!: perm_append_swap perm_append1) - - -subsection \Further results\ - -proposition perm_empty [iff]: "[] <~~> xs \ xs = []" - by (blast intro: perm_empty_imp) +lemma perm_iff_eq_mset: + \xs <~~> ys \ mset xs = mset ys\ +proof + assume \mset xs = mset ys\ + then show \xs <~~> ys\ + proof (induction xs arbitrary: ys) + case Nil + then show ?case + by simp + next + case (Cons x xs) + from Cons.prems [symmetric] have \mset xs = mset (remove1 x ys)\ + by simp + then have \xs <~~> remove1 x ys\ + by (rule Cons.IH) + then have \x # xs <~~> x # remove1 x ys\ + by (rule perm.Cons) + moreover from Cons.prems have \x \ set ys\ + by (auto dest: union_single_eq_member) + then have \x # remove1 x ys <~~> ys\ + by (induction ys) auto + ultimately show \x # xs <~~> ys\ + by (rule perm.trans) + qed +next + assume \xs <~~> ys\ + then show \mset xs = mset ys\ + by induction simp_all +qed -proposition perm_empty2 [iff]: "xs <~~> [] \ xs = []" - using perm_sym by auto - -proposition perm_sing_imp: "ys <~~> xs \ xs = [y] \ ys = [y]" - by (induct pred: perm) auto - -proposition perm_sing_eq [iff]: "ys <~~> [y] \ ys = [y]" - by (blast intro: perm_sing_imp) - -proposition perm_sing_eq2 [iff]: "[y] <~~> ys \ ys = [y]" - by (blast dest: perm_sym) - - -subsection \Removing elements\ - -proposition perm_remove: "x \ set ys \ ys <~~> x # remove1 x ys" - by (induct ys) auto +lemma list_permuted_induct [consumes 1, case_names Nil swap Cons trans]: + \P xs ys\ + if \mset xs = mset ys\ + \P [] []\ + \\y x zs. P (y # x # zs) (x # y # zs)\ + \\xs ys z. mset xs = mset ys \ P xs ys \ P (z # xs) (z # ys)\ + \\xs ys zs. mset xs = mset ys \ mset ys = mset zs \ P xs ys \ P ys zs \ P xs zs\ +proof - + from \mset xs = mset ys\ have \xs <~~> ys\ + by (simp add: perm_iff_eq_mset) + then show ?thesis + using that(2-3) apply (rule perm.induct) + apply (simp_all add: perm_iff_eq_mset) + apply (fact that(4)) + subgoal for xs ys zs + apply (rule that(5) [of xs ys zs]) + apply simp_all + done + done +qed -text \\medskip Congruence rule\ - -proposition perm_remove_perm: "xs <~~> ys \ remove1 z xs <~~> remove1 z ys" - by (induct pred: perm) auto - -proposition remove_hd [simp]: "remove1 z (z # xs) = xs" - by auto - -proposition cons_perm_imp_perm: "z # xs <~~> z # ys \ xs <~~> ys" - by (drule perm_remove_perm [where z = z]) auto - -proposition cons_perm_eq [iff]: "z#xs <~~> z#ys \ xs <~~> ys" - by (meson cons_perm_imp_perm perm.Cons) - -proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \ xs <~~> ys" - by (induct zs arbitrary: xs ys rule: rev_induct) auto +subsection \\that is equivalent to an already existing notion:\ -proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \ xs <~~> ys" - by (blast intro: append_perm_imp_perm perm_append1) - -proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \ xs <~~> ys" - by (meson perm.trans perm_append1_eq perm_append_swap) +theorem mset_eq_perm: \mset xs = mset ys \ xs <~~> ys\ + by (simp add: perm_iff_eq_mset) -theorem mset_eq_perm: "mset xs = mset ys \ xs <~~> ys" -proof - assume "mset xs = mset ys" - then show "xs <~~> ys" - proof (induction xs arbitrary: ys) - case (Cons x xs) - then have "x \ set ys" - using mset_eq_setD by fastforce - then show ?case - by (metis Cons.IH Cons.prems mset_remove1 perm.Cons perm.trans perm_remove perm_sym remove_hd) - qed auto -next - assume "xs <~~> ys" - then show "mset xs = mset ys" - by induction (simp_all add: union_ac) -qed + +subsection \Nontrivial conclusions\ + +proposition perm_swap: + \xs[i := xs ! j, j := xs ! i] <~~> xs\ + if \i < length xs\ \j < length xs\ + using that by (cases \i = j\) (simp_all add: perm_iff_eq_mset mset_update) proposition mset_le_perm_append: "mset xs \# mset ys \ (\zs. xs @ zs <~~> ys)" - apply (rule iffI) - apply (metis mset_append mset_eq_perm mset_subset_eq_exists_conv surjD surj_mset) - by (metis mset_append mset_eq_perm mset_subset_eq_exists_conv) + by (auto simp add: perm_iff_eq_mset mset_subset_eq_exists_conv ex_mset dest: sym) proposition perm_set_eq: "xs <~~> ys \ set xs = set ys" - by (metis mset_eq_perm mset_eq_setD) + by (rule mset_eq_setD) (simp add: perm_iff_eq_mset) proposition perm_distinct_iff: "xs <~~> ys \ distinct xs = distinct ys" - by (metis card_distinct distinct_card perm_length perm_set_eq) + by (auto simp add: perm_iff_eq_mset distinct_count_atmost_1 dest: mset_eq_setD) theorem eq_set_perm_remdups: "set xs = set ys \ remdups xs <~~> remdups ys" -proof (induction xs arbitrary: ys rule: length_induct) - case (1 xs) - show ?case - proof (cases "remdups xs") - case Nil - with "1.prems" show ?thesis - using "1.prems" by auto - next - case (Cons x us) - then have "x \ set (remdups ys)" - using "1.prems" set_remdups by fastforce - then show ?thesis - using "1.prems" mset_eq_perm set_eq_iff_mset_remdups_eq by blast - qed -qed + by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq) proposition perm_remdups_iff_eq_set: "remdups x <~~> remdups y \ set x = set y" - by (metis List.set_remdups perm_set_eq eq_set_perm_remdups) + by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq) theorem permutation_Ex_bij: assumes "xs <~~> ys" shows "\f. bij_betw f {.. (\ii Suc ` {.. Suc ` {..ii f"] conjI allI impI) show "bij_betw (g \ f) {..i < length xs\ show "xs ! i = zs ! (g \ f) i" - using trans(1,3)[THEN perm_length] perm by auto + using trans(1,3) perm by auto qed qed proposition perm_finite: "finite {B. B <~~> A}" -proof (rule finite_subset[where B="{xs. set xs \ set A \ length xs \ length A}"]) - show "finite {xs. set xs \ set A \ length xs \ length A}" - using finite_lists_length_le by blast +proof (rule finite_subset [where B="{xs. set xs \ set A \ length xs \ length A}"]) + show "finite {xs. set xs \ set A \ length xs \ length A}" + using finite_lists_length_le by blast next - show "{B. B <~~> A} \ {xs. set xs \ set A \ length xs \ length A}" - by (clarsimp simp add: perm_length perm_set_eq) + show "{B. B <~~> A} \ {xs. set xs \ set A \ length xs \ length A}" + by (auto simp add: perm_iff_eq_mset dest: mset_eq_setD mset_eq_length) qed -proposition perm_swap: - assumes "i < length xs" "j < length xs" - shows "xs[i := xs ! j, j := xs ! i] <~~> xs" - using assms by (simp add: mset_eq_perm[symmetric] mset_swap) + +subsection \Trivial conclusions:\ + +proposition perm_empty_imp: "[] <~~> ys \ ys = []" + by (simp add: perm_iff_eq_mset) + + +text \\medskip This more general theorem is easier to understand!\ + +proposition perm_length: "xs <~~> ys \ length xs = length ys" + by (rule mset_eq_length) (simp add: perm_iff_eq_mset) + +proposition perm_sym: "xs <~~> ys \ ys <~~> xs" + by (simp add: perm_iff_eq_mset) + + +text \We can insert the head anywhere in the list.\ + +proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" + by (simp add: perm_iff_eq_mset) + +proposition perm_append_swap: "xs @ ys <~~> ys @ xs" + by (simp add: perm_iff_eq_mset) + +proposition perm_append_single: "a # xs <~~> xs @ [a]" + by (simp add: perm_iff_eq_mset) + +proposition perm_rev: "rev xs <~~> xs" + by (simp add: perm_iff_eq_mset) + +proposition perm_append1: "xs <~~> ys \ l @ xs <~~> l @ ys" + by (simp add: perm_iff_eq_mset) + +proposition perm_append2: "xs <~~> ys \ xs @ l <~~> ys @ l" + by (simp add: perm_iff_eq_mset) + +proposition perm_empty [iff]: "[] <~~> xs \ xs = []" + by (simp add: perm_iff_eq_mset) + +proposition perm_empty2 [iff]: "xs <~~> [] \ xs = []" + by (simp add: perm_iff_eq_mset) + +proposition perm_sing_imp: "ys <~~> xs \ xs = [y] \ ys = [y]" + by (simp add: perm_iff_eq_mset) + +proposition perm_sing_eq [iff]: "ys <~~> [y] \ ys = [y]" + by (simp add: perm_iff_eq_mset) + +proposition perm_sing_eq2 [iff]: "[y] <~~> ys \ ys = [y]" + by (simp add: perm_iff_eq_mset) + +proposition perm_remove: "x \ set ys \ ys <~~> x # remove1 x ys" + by (simp add: perm_iff_eq_mset) + + +text \\medskip Congruence rule\ + +proposition perm_remove_perm: "xs <~~> ys \ remove1 z xs <~~> remove1 z ys" + by (simp add: perm_iff_eq_mset) + +proposition remove_hd [simp]: "remove1 z (z # xs) = xs" + by (simp add: perm_iff_eq_mset) + +proposition cons_perm_imp_perm: "z # xs <~~> z # ys \ xs <~~> ys" + by (simp add: perm_iff_eq_mset) + +proposition cons_perm_eq [simp]: "z#xs <~~> z#ys \ xs <~~> ys" + by (simp add: perm_iff_eq_mset) + +proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \ xs <~~> ys" + by (simp add: perm_iff_eq_mset) + +proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \ xs <~~> ys" + by (simp add: perm_iff_eq_mset) + +proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \ xs <~~> ys" + by (simp add: perm_iff_eq_mset) end