# HG changeset patch # User haftmann # Date 1614112908 0 # Node ID beaff25452d28a2ba1fd375492f33b43c253241e # Parent 2ac92ba88d6bb1ded2db0a87f0b25bf2893b099e more specific name diff -r 2ac92ba88d6b -r beaff25452d2 NEWS --- a/NEWS Tue Feb 23 20:41:48 2021 +0000 +++ b/NEWS Tue Feb 23 20:41:48 2021 +0000 @@ -18,6 +18,9 @@ more small lemmas. Some theorems that were stated awkwardly before were corrected. Minor INCOMPATIBILITY. +* Theory "Permutation" in HOL-Library has been renamed to the more +specific "List_Permutation". + *** ML *** diff -r 2ac92ba88d6b -r beaff25452d2 src/HOL/Algebra/Divisibility.thy --- a/src/HOL/Algebra/Divisibility.thy Tue Feb 23 20:41:48 2021 +0000 +++ b/src/HOL/Algebra/Divisibility.thy Tue Feb 23 20:41:48 2021 +0000 @@ -6,7 +6,7 @@ section \Divisibility in monoids and rings\ theory Divisibility - imports "HOL-Library.Permutation" Coset Group + imports "HOL-Library.List_Permutation" Coset Group begin section \Factorial Monoids\ diff -r 2ac92ba88d6b -r beaff25452d2 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Tue Feb 23 20:41:48 2021 +0000 +++ b/src/HOL/Library/Library.thy Tue Feb 23 20:41:48 2021 +0000 @@ -50,6 +50,7 @@ Lattice_Constructions Linear_Temporal_Logic_on_Streams ListVector + List_Permutation Lub_Glb Mapping Monad_Syntax @@ -66,7 +67,6 @@ Pattern_Aliases Periodic_Fun Perm - Permutation Permutations Poly_Mapping Power_By_Squaring diff -r 2ac92ba88d6b -r beaff25452d2 src/HOL/Library/List_Permutation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/List_Permutation.thy Tue Feb 23 20:41:48 2021 +0000 @@ -0,0 +1,230 @@ +(* 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 + +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 + +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) + +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 + + +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 + +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" +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 + +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) + +proposition perm_set_eq: "xs <~~> ys \ set xs = set ys" + by (metis mset_eq_perm mset_eq_setD) + +proposition perm_distinct_iff: "xs <~~> ys \ distinct xs = distinct ys" + by (metis card_distinct distinct_card perm_length perm_set_eq) + +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 + +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) + +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 + 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 +next + show "{B. B <~~> A} \ {xs. set xs \ set A \ length xs \ length A}" + by (clarsimp simp add: perm_length perm_set_eq) +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) + +end diff -r 2ac92ba88d6b -r beaff25452d2 src/HOL/Library/Permutation.thy --- a/src/HOL/Library/Permutation.thy Tue Feb 23 20:41:48 2021 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,230 +0,0 @@ -(* Title: HOL/Library/Permutation.thy - Author: Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker -*) - -section \Permutations\ - -theory Permutation -imports Multiset -begin - -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 - -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) - -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 - - -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 - -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" -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 - -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) - -proposition perm_set_eq: "xs <~~> ys \ set xs = set ys" - by (metis mset_eq_perm mset_eq_setD) - -proposition perm_distinct_iff: "xs <~~> ys \ distinct xs = distinct ys" - by (metis card_distinct distinct_card perm_length perm_set_eq) - -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 - -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) - -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 - 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 -next - show "{B. B <~~> A} \ {xs. set xs \ set A \ length xs \ length A}" - by (clarsimp simp add: perm_length perm_set_eq) -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) - -end diff -r 2ac92ba88d6b -r beaff25452d2 src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Tue Feb 23 20:41:48 2021 +0000 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Tue Feb 23 20:41:48 2021 +0000 @@ -3,7 +3,7 @@ section \Formalization of a Countermeasure by Koepf \& Duermuth 2009\ theory Koepf_Duermuth_Countermeasure - imports "HOL-Probability.Information" "HOL-Library.Permutation" + imports "HOL-Probability.Information" "HOL-Library.List_Permutation" begin lemma SIGMA_image_vimage: