diff -r 8e5428ff35af -r bbe5d3ef2052 src/HOL/List.thy --- a/src/HOL/List.thy Sun Nov 22 18:26:54 2020 +0000 +++ b/src/HOL/List.thy Thu Nov 26 18:09:02 2020 +0000 @@ -4542,48 +4542,47 @@ "replicate (length (filter (\y. x = y) xs)) x = filter (\y. x = y) xs" by (induct xs) auto +text \This stronger version is thanks to Stepan Holub\ lemma comm_append_are_replicate: - "\ xs \ []; ys \ []; xs @ ys = ys @ xs \ + "xs @ ys = ys @ xs \ \m n zs. concat (replicate m zs) = xs \ concat (replicate n zs) = ys" -proof (induction "length (xs @ ys)" arbitrary: xs ys rule: less_induct) +proof (induction "length (xs @ ys) + length xs" arbitrary: xs ys rule: less_induct) case less - - define xs' ys' where "xs' = (if (length xs \ length ys) then xs else ys)" - and "ys' = (if (length xs \ length ys) then ys else xs)" - then have - prems': "length xs' \ length ys'" - "xs' @ ys' = ys' @ xs'" - and "xs' \ []" - and len: "length (xs @ ys) = length (xs' @ ys')" - using less by (auto intro: less.hyps) - - from prems' - obtain ws where "ys' = xs' @ ws" - by (auto simp: append_eq_append_conv2) - - have "\m n zs. concat (replicate m zs) = xs' \ concat (replicate n zs) = ys'" - proof (cases "ws = []") - case True - then have "concat (replicate 1 xs') = xs'" - and "concat (replicate 1 xs') = ys'" - using \ys' = xs' @ ws\ by auto - then show ?thesis by blast + consider "length ys < length xs" | "xs = []" | "length xs \ length ys \ xs \ []" + by linarith + then show ?case + proof (cases) + assume "length ys < length xs" + then show ?thesis + using less.hyps[OF _ less.prems[symmetric]] nat_add_left_cancel_less by auto + next + assume "xs = []" + then have "concat (replicate 0 ys) = xs \ concat (replicate 1 ys) = ys" + by simp + then show ?case + by blast next - case False - from \ys' = xs' @ ws\ and \xs' @ ys' = ys' @ xs'\ - have "xs' @ ws = ws @ xs'" by simp - then have "\m n zs. concat (replicate m zs) = xs' \ concat (replicate n zs) = ws" - using False and \xs' \ []\ and \ys' = xs' @ ws\ and len - by (intro less.hyps) auto - then obtain m n zs where *: "concat (replicate m zs) = xs'" - and "concat (replicate n zs) = ws" by blast - then have "concat (replicate (m + n) zs) = ys'" - using \ys' = xs' @ ws\ + assume "length xs \ length ys \ xs \ []" + then have "length xs \ length ys" and "xs \ []" + by blast+ + from \length xs \ length ys\ and \xs @ ys = ys @ xs\ + obtain ws where "ys = xs @ ws" + by (auto simp: append_eq_append_conv2) + from this and \xs \ []\ + have "length ws < length ys" + by simp + from \xs @ ys = ys @ xs\[unfolded \ys = xs @ ws\] + have "xs @ ws = ws @ xs" + by simp + from less.hyps[OF _ this] \length ws < length ys\ + obtain m n' zs where "concat (replicate m zs) = xs" and "concat (replicate n' zs) = ws" + by auto + then have "concat (replicate (m+n') zs) = ys" + using \ys = xs @ ws\ by (simp add: replicate_add) - with * show ?thesis by blast + then show ?case + using \concat (replicate m zs) = xs\ by blast qed - then show ?case - using xs'_def ys'_def by meson qed lemma comm_append_is_replicate: @@ -4594,7 +4593,7 @@ proof - obtain m n zs where "concat (replicate m zs) = xs" and "concat (replicate n zs) = ys" - using comm_append_are_replicate[of xs ys, OF assms] by blast + using comm_append_are_replicate[of xs ys] assms by blast then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys" using \xs \ []\ and \ys \ []\ by (auto simp: replicate_add)