# HG changeset patch # User Christian Sternagel # Date 1404374116 -7200 # Node ID 5a8b3e9d82a4b31c639b2d0593ee857f60d1287d # Parent 7e22776f2d3271ce5f7f9ddf6ba7489b6f93f798 weaker assumption for "list_emb_trans"; added lemma diff -r 7e22776f2d32 -r 5a8b3e9d82a4 src/HOL/Library/Sublist.thy --- a/src/HOL/Library/Sublist.thy Thu Jul 03 09:55:15 2014 +0200 +++ b/src/HOL/Library/Sublist.thy Thu Jul 03 09:55:16 2014 +0200 @@ -514,34 +514,31 @@ by (induct rule: list_emb.induct) auto lemma list_emb_trans: - assumes "\x y z. \x \ A; y \ A; z \ A; P x y; P y z\ \ P x z" - shows "\xs ys zs. \xs \ lists A; ys \ lists A; zs \ lists A; - list_emb P xs ys; list_emb P ys zs\ \ list_emb P xs zs" + assumes "\x y z. \x \ set xs; y \ set ys; z \ set zs; P x y; P y z\ \ P x z" + shows "\list_emb P xs ys; list_emb P ys zs\ \ list_emb P xs zs" proof - - fix xs ys zs assume "list_emb P xs ys" and "list_emb P ys zs" - and "xs \ lists A" and "ys \ lists A" and "zs \ lists A" - then show "list_emb P xs zs" + then show "list_emb P xs zs" using assms proof (induction arbitrary: zs) case list_emb_Nil show ?case by blast next case (list_emb_Cons xs ys y) from list_emb_ConsD [OF `list_emb P (y#ys) zs`] obtain us v vs - where zs: "zs = us @ v # vs" and "P y v" and "list_emb P ys vs" by blast + where zs: "zs = us @ v # vs" and "P\<^sup>=\<^sup>= y v" and "list_emb P ys vs" by blast then have "list_emb P ys (v#vs)" by blast then have "list_emb P ys zs" unfolding zs by (rule list_emb_append2) - from list_emb_Cons.IH [OF this] and list_emb_Cons.prems show ?case by simp + from list_emb_Cons.IH [OF this] and list_emb_Cons.prems show ?case by auto next case (list_emb_Cons2 x y xs ys) from list_emb_ConsD [OF `list_emb P (y#ys) zs`] obtain us v vs where zs: "zs = us @ v # vs" and "P y v" and "list_emb P ys vs" by blast - with list_emb_Cons2 have "list_emb P xs vs" by simp + with list_emb_Cons2 have "list_emb P xs vs" by auto moreover have "P x v" proof - - from zs and `zs \ lists A` have "v \ A" by auto - moreover have "x \ A" and "y \ A" using list_emb_Cons2 by simp_all + from zs have "v \ set zs" by auto + moreover have "x \ set (x#xs)" and "y \ set (y#ys)" by simp_all ultimately show ?thesis - using `P x y` and `P y v` and assms + using `P x y` and `P y v` and list_emb_Cons2 by blast qed ultimately have "list_emb P (x#xs) (v#vs)" by blast @@ -549,6 +546,11 @@ qed qed +lemma list_emb_set: + assumes "list_emb P xs ys" and "x \ set xs" + obtains y where "y \ set ys" and "P x y" + using assms by (induct) auto + subsection {* Sublists (special case of homeomorphic embedding) *} @@ -607,7 +609,7 @@ qed lemma sublisteq_trans: "sublisteq xs ys \ sublisteq ys zs \ sublisteq xs zs" - by (rule list_emb_trans [of UNIV "op ="]) auto + by (rule list_emb_trans [of _ _ _ "op ="]) auto lemma sublisteq_append_le_same_iff: "sublisteq (xs @ ys) ys \ xs = []" by (auto dest: list_emb_length)