diff -r c52c5a5bf4e6 -r bfe92e4f6ea4 src/HOL/Library/Multiset.thy --- 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 \ distinct y \ - (set x = set y) = (mset x = mset y)" -by (auto simp: multiset_eq_iff distinct_count_atmost_1) + \distinct x \ distinct y \ set x = set y \ mset x = mset y\ + 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))" + \set x = set y \ mset (remdups x) = mset (remdups y)\ 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: + \distinct xs \ distinct ys\ if \mset xs = mset ys\ + using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD) + lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" proof (induct ls arbitrary: i) case Nil @@ -2509,9 +2512,7 @@ qed -subsection \Alternative representations\ - -subsubsection \Lists\ +subsection \Multiset as order-ignorant lists\ 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 \mset xs = mset ys\ + obtains f where + \bij_betw f {.. + \ys = map (\n. xs ! f n) [0.. +proof - + from assms have \length ys = length xs\ + by (auto dest: mset_eq_length) + from assms have \\f. f ` {.. ys = map (\n. xs ! f n) [0.. + proof (induction xs arbitrary: ys rule: rev_induct) + case Nil + then show ?case by simp + next + case (snoc x xs) + from snoc.prems have \x \ set ys\ + by (auto dest: union_single_eq_member) + then obtain zs ws where split: \ys = zs @ x # ws\ and \x \ set zs\ + by (auto dest: split_list_first) + then have \remove1 x ys = zs @ ws\ + by (simp add: remove1_append) + moreover from snoc.prems [symmetric] have \mset xs = mset (remove1 x ys)\ + by simp + ultimately have \mset xs = mset (zs @ ws)\ + by simp + then have \\f. f ` {.. zs @ ws = map (\n. xs ! f n) [0.. + by (rule snoc.IH) + then obtain f where + raw_surj: \f ` {.. + and hyp: \zs @ ws = map (\n. xs ! f n) [0.. by blast + define l and k where \l = length zs\ and \k = length ws\ + then have [simp]: \length zs = l\ \length ws = k\ + by simp_all + from \mset xs = mset (zs @ ws)\ have \length xs = length (zs @ ws)\ + by (rule mset_eq_length) + then have [simp]: \length xs = l + k\ + by simp + from raw_surj have f_surj: \f ` {.. + by simp + have [simp]: \[0.. + by (rule nth_equalityI) (simp_all add: nth_append) + have [simp]: \[l.. + by (rule nth_equalityI) + (auto simp add: nth_append nth_Cons split: nat.split) + define g :: \nat \ nat\ + where \g n = (if n < l then f n + else if n = l then l + k + else f (n - 1))\ for n + have \{.. {l} \ {Suc l.. + by auto + then have \g ` {.. {g l} \ g ` {Suc l.. + by auto + also have \g ` {Suc l.. + 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 \g ` {.. {l + k} \ f ` {l.. + by (simp add: g_def) + also have \\ = {.. + by simp (metis atLeastLessThan_add_Un f_surj image_Un le_add1 le_add_same_cancel1 lessThan_Suc lessThan_atLeast0) + finally have g_surj: \g ` {.. . + from hyp have zs_f: \zs = map (\n. xs ! f n) [0.. + and ws_f: \ws = map (\n. xs ! f n) [l.. + by simp_all + have \zs = map (\n. (xs @ [x]) ! g n) [0.. + proof (rule sym, rule map_upt_eqI) + fix n + assume \n < length zs\ + then have \n < l\ + by simp + with f_surj have \f n < l + k\ + by auto + with \n < l\ show \zs ! n = (xs @ [x]) ! g (0 + n)\ + by (simp add: zs_f g_def nth_append) + qed simp + moreover have \x = (xs @ [x]) ! g l\ + by (simp add: g_def nth_append) + moreover have \ws = map (\n. (xs @ [x]) ! g n) [Suc l.. + proof (rule sym, rule map_upt_eqI) + fix n + assume \n < length ws\ + then have \n < k\ + by simp + with f_surj have \f (l + n) < l + k\ + by auto + with \n < k\ show \ws ! n = (xs @ [x]) ! g (Suc l + n)\ + by (simp add: ws_f g_def nth_append) + qed simp + ultimately have \zs @ x # ws = map (\n. (xs @ [x]) ! g n) [0.. + by simp + with g_surj show ?case + by (auto simp add: split) + qed + then obtain f where + surj: \f ` {.. + and hyp: \ys = map (\n. xs ! f n) [0.. by blast + from surj have \bij_betw f {.. + by (simp add: bij_betw_def \length ys = length xs\ eq_card_imp_inj_on) + then show thesis + using hyp .. +qed + +proposition mset_eq_finite: + \finite {ys. mset ys = mset xs}\ +proof - + have \{ys. mset ys = mset xs} \ {ys. set ys \ set xs \ length ys \ length xs}\ + by (auto simp add: dest: mset_eq_setD mset_eq_length) + moreover have \finite {ys. set ys \ set xs \ length ys \ length xs}\ + using finite_lists_length_le by blast + ultimately show ?thesis + by (rule finite_subset) +qed + subsection \The multiset order\ @@ -3653,11 +3767,22 @@ lemma ex_mset: "\xs. mset xs = X" by (induct X) (simp, metis mset.simps(2)) -inductive pred_mset :: "('a \ bool) \ 'a multiset \ bool" +inductive pred_mset :: "('a \ bool) \ 'a multiset \ bool" where "pred_mset P {#}" | "\P a; pred_mset P M\ \ pred_mset P (add_mset a M)" +lemma pred_mset_iff: \ \TODO: alias for \<^const>\Multiset.Ball\\ + \pred_mset P M \ Multiset.Ball M P\ (is \?P \ ?Q\) +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 \ set_mset {#} \ False" for z by auto show "pred_mset P = (\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' :: \('a \ 'b \ bool) \ 'a multiset \ 'b multiset \ bool\ where Zero[intro]: "rel_mset' R {#} {#}" | Plus[intro]: "\R a b; rel_mset' R M N\ \ rel_mset' R (add_mset a M) (add_mset b N)"