--- a/src/HOL/Library/List_Permutation.thy Wed Feb 24 13:31:28 2021 +0000
+++ b/src/HOL/Library/List_Permutation.thy Wed Feb 24 13:31:33 2021 +0000
@@ -52,30 +52,6 @@
by induction simp_all
qed
-lemma list_permuted_induct [consumes 1, case_names Nil swap Cons trans]:
- \<open>P xs ys\<close>
- if \<open>mset xs = mset ys\<close>
- \<open>P [] []\<close>
- \<open>\<And>y x zs. P (y # x # zs) (x # y # zs)\<close>
- \<open>\<And>xs ys z. mset xs = mset ys \<Longrightarrow> P xs ys \<Longrightarrow> P (z # xs) (z # ys)\<close>
- \<open>\<And>xs ys zs. mset xs = mset ys \<Longrightarrow> mset ys = mset zs \<Longrightarrow> P xs ys \<Longrightarrow> P ys zs \<Longrightarrow> P xs zs\<close>
-proof -
- from \<open>mset xs = mset ys\<close> have \<open>xs <~~> ys\<close>
- 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
-
-
-subsection \<open>\<dots>that is equivalent to an already existing notion:\<close>
-
theorem mset_eq_perm: \<open>mset xs = mset ys \<longleftrightarrow> xs <~~> ys\<close>
by (simp add: perm_iff_eq_mset)
@@ -85,7 +61,7 @@
proposition perm_swap:
\<open>xs[i := xs ! j, j := xs ! i] <~~> xs\<close>
if \<open>i < length xs\<close> \<open>j < length xs\<close>
- using that by (cases \<open>i = j\<close>) (simp_all add: perm_iff_eq_mset mset_update)
+ using that by (simp add: perm_iff_eq_mset mset_swap)
proposition mset_le_perm_append: "mset xs \<subseteq># mset ys \<longleftrightarrow> (\<exists>zs. xs @ zs <~~> ys)"
by (auto simp add: perm_iff_eq_mset mset_subset_eq_exists_conv ex_mset dest: sym)
@@ -93,8 +69,8 @@
proposition perm_set_eq: "xs <~~> ys \<Longrightarrow> set xs = set ys"
by (rule mset_eq_setD) (simp add: perm_iff_eq_mset)
-proposition perm_distinct_iff: "xs <~~> ys \<Longrightarrow> distinct xs = distinct ys"
- by (auto simp add: perm_iff_eq_mset distinct_count_atmost_1 dest: mset_eq_setD)
+proposition perm_distinct_iff: "xs <~~> ys \<Longrightarrow> distinct xs \<longleftrightarrow> distinct ys"
+ by (rule mset_eq_imp_distinct_iff) (simp add: perm_iff_eq_mset)
theorem eq_set_perm_remdups: "set xs = set ys \<Longrightarrow> remdups xs <~~> remdups ys"
by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq)
@@ -105,72 +81,17 @@
theorem permutation_Ex_bij:
assumes "xs <~~> ys"
shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))"
- using assms
-proof induct
- case Nil
- then show ?case
- unfolding bij_betw_def by simp
-next
- case (swap y x l)
- show ?case
- proof (intro exI[of _ "Fun.swap 0 1 id"] conjI allI impI)
- show "bij_betw (Fun.swap 0 1 id) {..<length (y # x # l)} {..<length (x # y # l)}"
- by (auto simp: bij_betw_def)
- fix i
- assume "i < length (y # x # l)"
- show "(y # x # l) ! i = (x # y # l) ! (Fun.swap 0 1 id) i"
- by (cases i) (auto simp: Fun.swap_def gr0_conv_Suc)
- qed
-next
- case (Cons xs ys z)
- then obtain f where bij: "bij_betw f {..<length xs} {..<length ys}"
- and perm: "\<forall>i<length xs. xs ! i = ys ! (f i)"
- by blast
- let ?f = "\<lambda>i. case i of Suc n \<Rightarrow> Suc (f n) | 0 \<Rightarrow> 0"
- show ?case
- proof (intro exI[of _ ?f] allI conjI impI)
- have *: "{..<length (z#xs)} = {0} \<union> Suc ` {..<length xs}"
- "{..<length (z#ys)} = {0} \<union> Suc ` {..<length ys}"
- by (simp_all add: lessThan_Suc_eq_insert_0)
- show "bij_betw ?f {..<length (z#xs)} {..<length (z#ys)}"
- unfolding *
- proof (rule bij_betw_combine)
- show "bij_betw ?f (Suc ` {..<length xs}) (Suc ` {..<length ys})"
- using bij unfolding bij_betw_def
- by (auto intro!: inj_onI imageI dest: inj_onD simp: image_comp comp_def)
- qed (auto simp: bij_betw_def)
- fix i
- assume "i < length (z # xs)"
- then show "(z # xs) ! i = (z # ys) ! (?f i)"
- using perm by (cases i) auto
- qed
-next
- case (trans xs ys zs)
- then obtain f g
- where bij: "bij_betw f {..<length xs} {..<length ys}" "bij_betw g {..<length ys} {..<length zs}"
- and perm: "\<forall>i<length xs. xs ! i = ys ! (f i)" "\<forall>i<length ys. ys ! i = zs ! (g i)"
- by blast
- show ?case
- proof (intro exI[of _ "g \<circ> f"] conjI allI impI)
- show "bij_betw (g \<circ> f) {..<length xs} {..<length zs}"
- using bij by (rule bij_betw_trans)
- fix i
- assume "i < length xs"
- with bij have "f i < length ys"
- unfolding bij_betw_def by force
- with \<open>i < length xs\<close> show "xs ! i = zs ! (g \<circ> f) i"
- using trans(1,3) perm by auto
- qed
+proof -
+ from assms have \<open>mset ys = mset xs\<close>
+ by (simp add: perm_iff_eq_mset)
+ then obtain f where \<open>bij_betw f {..<length ys} {..<length xs}\<close>
+ \<open>xs = map (\<lambda>n. ys ! f n) [0..<length ys]\<close>
+ by (rule mset_eq_permutation)
+ then show ?thesis by auto
qed
proposition perm_finite: "finite {B. B <~~> A}"
-proof (rule finite_subset [where B="{xs. set xs \<subseteq> set A \<and> length xs \<le> length A}"])
- show "finite {xs. set xs \<subseteq> set A \<and> length xs \<le> length A}"
- using finite_lists_length_le by blast
-next
- show "{B. B <~~> A} \<subseteq> {xs. set xs \<subseteq> set A \<and> length xs \<le> length A}"
- by (auto simp add: perm_iff_eq_mset dest: mset_eq_setD mset_eq_length)
-qed
+ using mset_eq_finite by (auto simp add: perm_iff_eq_mset)
subsection \<open>Trivial conclusions:\<close>
--- a/src/HOL/Library/Multiset.thy Wed Feb 24 13:31:28 2021 +0000
+++ b/src/HOL/Library/Multiset.thy Wed Feb 24 13:31:33 2021 +0000
@@ -1863,12 +1863,11 @@
qed
lemma set_eq_iff_mset_eq_distinct:
- "distinct x \<Longrightarrow> distinct y \<Longrightarrow>
- (set x = set y) = (mset x = mset y)"
-by (auto simp: multiset_eq_iff distinct_count_atmost_1)
+ \<open>distinct x \<Longrightarrow> distinct y \<Longrightarrow> set x = set y \<longleftrightarrow> mset x = mset y\<close>
+ by (auto simp: multiset_eq_iff distinct_count_atmost_1)
lemma set_eq_iff_mset_remdups_eq:
- "(set x = set y) = (mset (remdups x) = mset (remdups y))"
+ \<open>set x = set y \<longleftrightarrow> mset (remdups x) = mset (remdups y)\<close>
apply (rule iffI)
apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1])
apply (drule distinct_remdups [THEN distinct_remdups
@@ -1876,6 +1875,10 @@
apply simp
done
+lemma mset_eq_imp_distinct_iff:
+ \<open>distinct xs \<longleftrightarrow> distinct ys\<close> if \<open>mset xs = mset ys\<close>
+ using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD)
+
lemma nth_mem_mset: "i < length ls \<Longrightarrow> (ls ! i) \<in># mset ls"
proof (induct ls arbitrary: i)
case Nil
@@ -2509,9 +2512,7 @@
qed
-subsection \<open>Alternative representations\<close>
-
-subsubsection \<open>Lists\<close>
+subsection \<open>Multiset as order-ignorant lists\<close>
context linorder
begin
@@ -2719,6 +2720,119 @@
mset (ls[j := ls ! i, i := ls ! j]) = mset ls"
by (cases "i = j") (simp_all add: mset_update nth_mem_mset)
+lemma mset_eq_permutation:
+ assumes \<open>mset xs = mset ys\<close>
+ obtains f where
+ \<open>bij_betw f {..<length xs} {..<length ys}\<close>
+ \<open>ys = map (\<lambda>n. xs ! f n) [0..<length xs]\<close>
+proof -
+ from assms have \<open>length ys = length xs\<close>
+ by (auto dest: mset_eq_length)
+ from assms have \<open>\<exists>f. f ` {..<length xs} = {..<length xs} \<and> ys = map (\<lambda>n. xs ! f n) [0..<length xs]\<close>
+ proof (induction xs arbitrary: ys rule: rev_induct)
+ case Nil
+ then show ?case by simp
+ next
+ case (snoc x xs)
+ from snoc.prems have \<open>x \<in> set ys\<close>
+ by (auto dest: union_single_eq_member)
+ then obtain zs ws where split: \<open>ys = zs @ x # ws\<close> and \<open>x \<notin> set zs\<close>
+ by (auto dest: split_list_first)
+ then have \<open>remove1 x ys = zs @ ws\<close>
+ by (simp add: remove1_append)
+ moreover from snoc.prems [symmetric] have \<open>mset xs = mset (remove1 x ys)\<close>
+ by simp
+ ultimately have \<open>mset xs = mset (zs @ ws)\<close>
+ by simp
+ then have \<open>\<exists>f. f ` {..<length xs} = {..<length xs} \<and> zs @ ws = map (\<lambda>n. xs ! f n) [0..<length xs]\<close>
+ by (rule snoc.IH)
+ then obtain f where
+ raw_surj: \<open>f ` {..<length xs} = {..<length xs}\<close>
+ and hyp: \<open>zs @ ws = map (\<lambda>n. xs ! f n) [0..<length xs]\<close> by blast
+ define l and k where \<open>l = length zs\<close> and \<open>k = length ws\<close>
+ then have [simp]: \<open>length zs = l\<close> \<open>length ws = k\<close>
+ by simp_all
+ from \<open>mset xs = mset (zs @ ws)\<close> have \<open>length xs = length (zs @ ws)\<close>
+ by (rule mset_eq_length)
+ then have [simp]: \<open>length xs = l + k\<close>
+ by simp
+ from raw_surj have f_surj: \<open>f ` {..<l + k} = {..<l + k}\<close>
+ by simp
+ have [simp]: \<open>[0..<l + k] = [0..<l] @ [l..<l + k]\<close>
+ by (rule nth_equalityI) (simp_all add: nth_append)
+ have [simp]: \<open>[l..<l + k] @ [l + k] = [l] @ [Suc l..<Suc (l + k)]\<close>
+ by (rule nth_equalityI)
+ (auto simp add: nth_append nth_Cons split: nat.split)
+ define g :: \<open>nat \<Rightarrow> nat\<close>
+ where \<open>g n = (if n < l then f n
+ else if n = l then l + k
+ else f (n - 1))\<close> for n
+ have \<open>{..<Suc (l + k)} = {..<l} \<union> {l} \<union> {Suc l..<Suc (l + k)}\<close>
+ by auto
+ then have \<open>g ` {..<Suc (l + k)} = g ` {..<l} \<union> {g l} \<union> g ` {Suc l..<Suc (l + k)}\<close>
+ by auto
+ also have \<open>g ` {Suc l..<Suc (l + k)} = f ` {l..<l + k}\<close>
+ apply (auto simp add: g_def Suc_le_lessD)
+ apply (auto simp add: image_def)
+ apply (metis Suc_le_mono atLeastLessThan_iff diff_Suc_Suc diff_zero lessI less_trans_Suc)
+ done
+ finally have \<open>g ` {..<Suc (l + k)} = f ` {..<l} \<union> {l + k} \<union> f ` {l..<l + k}\<close>
+ by (simp add: g_def)
+ also have \<open>\<dots> = {..<Suc (l + k)}\<close>
+ by simp (metis atLeastLessThan_add_Un f_surj image_Un le_add1 le_add_same_cancel1 lessThan_Suc lessThan_atLeast0)
+ finally have g_surj: \<open>g ` {..<Suc (l + k)} = {..<Suc (l + k)}\<close> .
+ from hyp have zs_f: \<open>zs = map (\<lambda>n. xs ! f n) [0..<l]\<close>
+ and ws_f: \<open>ws = map (\<lambda>n. xs ! f n) [l..<l + k]\<close>
+ by simp_all
+ have \<open>zs = map (\<lambda>n. (xs @ [x]) ! g n) [0..<l]\<close>
+ proof (rule sym, rule map_upt_eqI)
+ fix n
+ assume \<open>n < length zs\<close>
+ then have \<open>n < l\<close>
+ by simp
+ with f_surj have \<open>f n < l + k\<close>
+ by auto
+ with \<open>n < l\<close> show \<open>zs ! n = (xs @ [x]) ! g (0 + n)\<close>
+ by (simp add: zs_f g_def nth_append)
+ qed simp
+ moreover have \<open>x = (xs @ [x]) ! g l\<close>
+ by (simp add: g_def nth_append)
+ moreover have \<open>ws = map (\<lambda>n. (xs @ [x]) ! g n) [Suc l..<Suc (l + k)]\<close>
+ proof (rule sym, rule map_upt_eqI)
+ fix n
+ assume \<open>n < length ws\<close>
+ then have \<open>n < k\<close>
+ by simp
+ with f_surj have \<open>f (l + n) < l + k\<close>
+ by auto
+ with \<open>n < k\<close> show \<open>ws ! n = (xs @ [x]) ! g (Suc l + n)\<close>
+ by (simp add: ws_f g_def nth_append)
+ qed simp
+ ultimately have \<open>zs @ x # ws = map (\<lambda>n. (xs @ [x]) ! g n) [0..<length (xs @ [x])]\<close>
+ by simp
+ with g_surj show ?case
+ by (auto simp add: split)
+ qed
+ then obtain f where
+ surj: \<open>f ` {..<length xs} = {..<length xs}\<close>
+ and hyp: \<open>ys = map (\<lambda>n. xs ! f n) [0..<length xs]\<close> by blast
+ from surj have \<open>bij_betw f {..<length xs} {..<length ys}\<close>
+ by (simp add: bij_betw_def \<open>length ys = length xs\<close> eq_card_imp_inj_on)
+ then show thesis
+ using hyp ..
+qed
+
+proposition mset_eq_finite:
+ \<open>finite {ys. mset ys = mset xs}\<close>
+proof -
+ have \<open>{ys. mset ys = mset xs} \<subseteq> {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close>
+ by (auto simp add: dest: mset_eq_setD mset_eq_length)
+ moreover have \<open>finite {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close>
+ using finite_lists_length_le by blast
+ ultimately show ?thesis
+ by (rule finite_subset)
+qed
+
subsection \<open>The multiset order\<close>
@@ -3653,11 +3767,22 @@
lemma ex_mset: "\<exists>xs. mset xs = X"
by (induct X) (simp, metis mset.simps(2))
-inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool"
+inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool"
where
"pred_mset P {#}"
| "\<lbrakk>P a; pred_mset P M\<rbrakk> \<Longrightarrow> pred_mset P (add_mset a M)"
+lemma pred_mset_iff: \<comment> \<open>TODO: alias for \<^const>\<open>Multiset.Ball\<close>\<close>
+ \<open>pred_mset P M \<longleftrightarrow> Multiset.Ball M P\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+proof
+ assume ?P
+ then show ?Q by induction simp_all
+next
+ assume ?Q
+ then show ?P
+ by (induction M) (auto intro: pred_mset.intros)
+qed
+
bnf "'a multiset"
map: image_mset
sets: set_mset
@@ -3709,18 +3834,10 @@
show "z \<in> set_mset {#} \<Longrightarrow> False" for z
by auto
show "pred_mset P = (\<lambda>x. Ball (set_mset x) P)" for P
- proof (intro ext iffI)
- fix x
- assume "pred_mset P x"
- then show "Ball (set_mset x) P" by (induct pred: pred_mset; simp)
- next
- fix x
- assume "Ball (set_mset x) P"
- then show "pred_mset P x" by (induct x; auto intro: pred_mset.intros)
- qed
+ by (simp add: fun_eq_iff pred_mset_iff)
qed
-inductive rel_mset'
+inductive rel_mset' :: \<open>('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset \<Rightarrow> bool\<close>
where
Zero[intro]: "rel_mset' R {#} {#}"
| Plus[intro]: "\<lbrakk>R a b; rel_mset' R M N\<rbrakk> \<Longrightarrow> rel_mset' R (add_mset a M) (add_mset b N)"
--- a/src/HOL/List.thy Wed Feb 24 13:31:28 2021 +0000
+++ b/src/HOL/List.thy Wed Feb 24 13:31:33 2021 +0000
@@ -3389,7 +3389,24 @@
lemma nth_Cons_numeral [simp]:
"(x # xs) ! numeral v = xs ! (numeral v - 1)"
-by (simp add: nth_Cons')
+ by (simp add: nth_Cons')
+
+lemma map_upt_eqI:
+ \<open>map f [m..<n] = xs\<close> if \<open>length xs = n - m\<close>
+ \<open>\<And>i. i < length xs \<Longrightarrow> xs ! i = f (m + i)\<close>
+proof (rule nth_equalityI)
+ from \<open>length xs = n - m\<close> show \<open>length (map f [m..<n]) = length xs\<close>
+ by simp
+next
+ fix i
+ assume \<open>i < length (map f [m..<n])\<close>
+ then have \<open>i < n - m\<close>
+ by simp
+ with that have \<open>xs ! i = f (m + i)\<close>
+ by simp
+ with \<open>i < n - m\<close> show \<open>map f [m..<n] ! i = xs ! i\<close>
+ by simp
+qed
subsubsection \<open>\<open>upto\<close>: interval-list on \<^typ>\<open>int\<close>\<close>