# HG changeset patch # User nipkow # Date 1606415262 -3600 # Node ID e0ceaca7344a31dce6a3bdc8590c1b37f1f6bd71 # Parent 178de0e275a12252e2d5fee87b3cedd6b453ad7f# Parent 7b918b9f0122222b75bd80a20d92f4331430bd6e merged diff -r 178de0e275a1 -r e0ceaca7344a CONTRIBUTORS --- a/CONTRIBUTORS Thu Nov 26 18:32:06 2020 +0100 +++ b/CONTRIBUTORS Thu Nov 26 19:27:42 2020 +0100 @@ -5,6 +5,8 @@ Contributions to this Isabelle version -------------------------------------- +* November 2020: Stepan Holub + Removed preconditions from lemma comm_append_are_replicate * November 2020: Florian Haftmann Bundle mixins for locale and class expressions. diff -r 178de0e275a1 -r e0ceaca7344a src/HOL/Data_Structures/Interval_Tree.thy --- a/src/HOL/Data_Structures/Interval_Tree.thy Thu Nov 26 18:32:06 2020 +0100 +++ b/src/HOL/Data_Structures/Interval_Tree.thy Thu Nov 26 19:27:42 2020 +0100 @@ -265,10 +265,13 @@ fixes has_overlap :: "'t \ 'a::linorder ivl \ bool" assumes set_overlap: "invar s \ has_overlap s x = Interval_Tree.has_overlap (set s) x" +fun invar :: "('a::{linorder,order_bot}) ivl_tree \ bool" where +"invar t = (inv_max_hi t \ sorted(inorder t))" + interpretation S: Interval_Set where empty = Leaf and insert = insert and delete = delete and has_overlap = search and isin = isin and set = set_tree - and invar = "\t. inv_max_hi t \ sorted (inorder t)" + and invar = invar proof (standard, goal_cases) case 1 then show ?case by auto diff -r 178de0e275a1 -r e0ceaca7344a src/HOL/List.thy --- a/src/HOL/List.thy Thu Nov 26 18:32:06 2020 +0100 +++ b/src/HOL/List.thy Thu Nov 26 19:27:42 2020 +0100 @@ -4543,47 +4543,44 @@ by (induct xs) auto lemma comm_append_are_replicate: - "\ xs \ []; ys \ []; 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) + "xs @ ys = ys @ xs \ \m n zs. concat (replicate m zs) = xs \ concat (replicate n zs) = ys" +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 (1) "length ys < length xs" | (2) "xs = []" | (3) "length xs \ length ys \ xs \ []" + by linarith + then show ?case + proof (cases) + case 1 + then show ?thesis + using less.hyps[OF _ less.prems[symmetric]] nat_add_left_cancel_less by auto + next + case 2 + then have "concat (replicate 0 ys) = xs \ concat (replicate 1 ys) = ys" + by simp + then show ?thesis + 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\ + case 3 + 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 ?thesis + 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 +4591,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 assms(3)] by blast then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys" using \xs \ []\ and \ys \ []\ by (auto simp: replicate_add)