# HG changeset patch # User paulson # Date 1582719708 0 # Node ID c213d067e60f5f7012e3a45c44d8b3aa9dff3145 # Parent c06604896c3db24d08027bfc694532ba0b02d86a Moved a number of general-purpose lemmas into HOL diff -r c06604896c3d -r c213d067e60f src/HOL/Analysis/Borel_Space.thy --- a/src/HOL/Analysis/Borel_Space.thy Tue Feb 25 17:37:22 2020 +0100 +++ b/src/HOL/Analysis/Borel_Space.thy Wed Feb 26 12:21:48 2020 +0000 @@ -28,78 +28,6 @@ by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) qed -definition\<^marker>\tag important\ "mono_on f A \ \r s. r \ A \ s \ A \ r \ s \ f r \ f s" - -lemma mono_onI: - "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on f A" - unfolding mono_on_def by simp - -lemma mono_onD: - "\mono_on f A; r \ A; s \ A; r \ s\ \ f r \ f s" - unfolding mono_on_def by simp - -lemma mono_imp_mono_on: "mono f \ mono_on f A" - unfolding mono_def mono_on_def by auto - -lemma mono_on_subset: "mono_on f A \ B \ A \ mono_on f B" - unfolding mono_on_def by auto - -definition\<^marker>\tag important\ "strict_mono_on f A \ \r s. r \ A \ s \ A \ r < s \ f r < f s" - -lemma strict_mono_onI: - "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on f A" - unfolding strict_mono_on_def by simp - -lemma strict_mono_onD: - "\strict_mono_on f A; r \ A; s \ A; r < s\ \ f r < f s" - unfolding strict_mono_on_def by simp - -lemma mono_on_greaterD: - assumes "mono_on g A" "x \ A" "y \ A" "g x > (g (y::_::linorder) :: _ :: linorder)" - shows "x > y" -proof (rule ccontr) - assume "\x > y" - hence "x \ y" by (simp add: not_less) - from assms(1-3) and this have "g x \ g y" by (rule mono_onD) - with assms(4) show False by simp -qed - -lemma strict_mono_inv: - fixes f :: "('a::linorder) \ ('b::linorder)" - assumes "strict_mono f" and "surj f" and inv: "\x. g (f x) = x" - shows "strict_mono g" -proof - fix x y :: 'b assume "x < y" - from \surj f\ obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast - with \x < y\ and \strict_mono f\ have "x' < y'" by (simp add: strict_mono_less) - with inv show "g x < g y" by simp -qed - -lemma strict_mono_on_imp_inj_on: - assumes "strict_mono_on (f :: (_ :: linorder) \ (_ :: preorder)) A" - shows "inj_on f A" -proof (rule inj_onI) - fix x y assume "x \ A" "y \ A" "f x = f y" - thus "x = y" - by (cases x y rule: linorder_cases) - (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) -qed - -lemma strict_mono_on_leD: - assumes "strict_mono_on (f :: (_ :: linorder) \ _ :: preorder) A" "x \ A" "y \ A" "x \ y" - shows "f x \ f y" -proof (insert le_less_linear[of y x], elim disjE) - assume "x < y" - with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all - thus ?thesis by (rule less_imp_le) -qed (insert assms, simp) - -lemma strict_mono_on_eqD: - fixes f :: "(_ :: linorder) \ (_ :: preorder)" - assumes "strict_mono_on f A" "f x = f y" "x \ A" "y \ A" - shows "y = x" - using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD) - proposition mono_on_imp_deriv_nonneg: assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)" assumes "x \ interior A" @@ -127,10 +55,6 @@ qed qed simp -lemma strict_mono_on_imp_mono_on: - "strict_mono_on (f :: (_ :: linorder) \ _ :: preorder) A \ mono_on f A" - by (rule mono_onI, rule strict_mono_on_leD) - proposition mono_on_ctble_discont: fixes f :: "real \ real" fixes A :: "real set" diff -r c06604896c3d -r c213d067e60f src/HOL/Fun.thy --- a/src/HOL/Fun.thy Tue Feb 25 17:37:22 2020 +0100 +++ b/src/HOL/Fun.thy Wed Feb 26 12:21:48 2020 +0000 @@ -913,6 +913,84 @@ then show False by blast qed +subsection \Monotonic functions over a set\ + +definition "mono_on f A \ \r s. r \ A \ s \ A \ r \ s \ f r \ f s" + +lemma mono_onI: + "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on f A" + unfolding mono_on_def by simp + +lemma mono_onD: + "\mono_on f A; r \ A; s \ A; r \ s\ \ f r \ f s" + unfolding mono_on_def by simp + +lemma mono_imp_mono_on: "mono f \ mono_on f A" + unfolding mono_def mono_on_def by auto + +lemma mono_on_subset: "mono_on f A \ B \ A \ mono_on f B" + unfolding mono_on_def by auto + +definition "strict_mono_on f A \ \r s. r \ A \ s \ A \ r < s \ f r < f s" + +lemma strict_mono_onI: + "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on f A" + unfolding strict_mono_on_def by simp + +lemma strict_mono_onD: + "\strict_mono_on f A; r \ A; s \ A; r < s\ \ f r < f s" + unfolding strict_mono_on_def by simp + +lemma mono_on_greaterD: + assumes "mono_on g A" "x \ A" "y \ A" "g x > (g (y::_::linorder) :: _ :: linorder)" + shows "x > y" +proof (rule ccontr) + assume "\x > y" + hence "x \ y" by (simp add: not_less) + from assms(1-3) and this have "g x \ g y" by (rule mono_onD) + with assms(4) show False by simp +qed + +lemma strict_mono_inv: + fixes f :: "('a::linorder) \ ('b::linorder)" + assumes "strict_mono f" and "surj f" and inv: "\x. g (f x) = x" + shows "strict_mono g" +proof + fix x y :: 'b assume "x < y" + from \surj f\ obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast + with \x < y\ and \strict_mono f\ have "x' < y'" by (simp add: strict_mono_less) + with inv show "g x < g y" by simp +qed + +lemma strict_mono_on_imp_inj_on: + assumes "strict_mono_on (f :: (_ :: linorder) \ (_ :: preorder)) A" + shows "inj_on f A" +proof (rule inj_onI) + fix x y assume "x \ A" "y \ A" "f x = f y" + thus "x = y" + by (cases x y rule: linorder_cases) + (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) +qed + +lemma strict_mono_on_leD: + assumes "strict_mono_on (f :: (_ :: linorder) \ _ :: preorder) A" "x \ A" "y \ A" "x \ y" + shows "f x \ f y" +proof (insert le_less_linear[of y x], elim disjE) + assume "x < y" + with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all + thus ?thesis by (rule less_imp_le) +qed (insert assms, simp) + +lemma strict_mono_on_eqD: + fixes f :: "(_ :: linorder) \ (_ :: preorder)" + assumes "strict_mono_on f A" "f x = f y" "x \ A" "y \ A" + shows "y = x" + using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD) + +lemma strict_mono_on_imp_mono_on: + "strict_mono_on (f :: (_ :: linorder) \ _ :: preorder) A \ mono_on f A" + by (rule mono_onI, rule strict_mono_on_leD) + subsection \Setup\ diff -r c06604896c3d -r c213d067e60f src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Tue Feb 25 17:37:22 2020 +0100 +++ b/src/HOL/Library/Ramsey.thy Wed Feb 26 12:21:48 2020 +0000 @@ -40,25 +40,25 @@ by (metis insert_iff singletonD) qed -lemma nsets_disjoint_2: +lemma nsets_disjoint_2: "X \ Y = {} \ [X \ Y]\<^bsup>2\<^esup> = [X]\<^bsup>2\<^esup> \ [Y]\<^bsup>2\<^esup> \ (\x\X. \y\Y. {{x,y}})" by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff) -lemma ordered_nsets_2_eq: - fixes A :: "'a::linorder set" +lemma ordered_nsets_2_eq: + fixes A :: "'a::linorder set" shows "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ x ?rhs" - unfolding numeral_nat + unfolding numeral_nat apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less) - by (metis antisym) + by (metis antisym) show "?rhs \ nsets A 2" unfolding numeral_nat by (auto simp: nsets_def card_Suc_eq) qed -lemma ordered_nsets_3_eq: - fixes A :: "'a::linorder set" +lemma ordered_nsets_3_eq: + fixes A :: "'a::linorder set" shows "nsets A 3 = {{x,y,z} | x y z. x \ A \ y \ A \ z \ A \ x y4\<^esup> = {U. \u x y z. U = {u,x,y,z} \ u \ A \ x \ A \ y \ A \ z \ A \ u < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - @@ -90,8 +90,8 @@ by auto qed -lemma ordered_nsets_5_eq: - fixes A :: "'a::linorder set" +lemma ordered_nsets_5_eq: + fixes A :: "'a::linorder set" shows "[A]\<^bsup>5\<^esup> = {U. \u v x y z. U = {u,v,x,y,z} \ u \ A \ v \ A \ x \ A \ y \ A \ z \ A \ u < v \ v < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - @@ -122,7 +122,7 @@ show "finite A" using infinite_arbitrarily_large that by auto then have "\ r \ card A" - using that by (simp add: set_eq_iff) (metis finite_subset get_smaller_card [of A r]) + using that by (simp add: set_eq_iff) (metis obtain_subset_with_card_n) then show "card A < r" using not_less by blast next @@ -189,7 +189,7 @@ have "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" using assms by (simp add: nsets_image_funcset) then show ?thesis - using f by fastforce + using f by fastforce qed subsubsection \Partition predicates\ @@ -198,7 +198,7 @@ where "partn \ \ \ \ \ \f \ nsets \ \ \ \. \H \ nsets \ \. \\\\. f ` (nsets H \) \ {\}" definition partn_lst :: "'a set \ nat list \ nat \ bool" - where "partn_lst \ \ \ \ \f \ nsets \ \ \ {..}. + where "partn_lst \ \ \ \ \f \ nsets \ \ \ {..}. \i < length \. \H \ nsets \ (\!i). f ` (nsets H \) \ {i}" lemma partn_lst_greater_resource: @@ -247,8 +247,8 @@ lemma partn_lst_eq_partn: "partn_lst {..Finite versions of Ramsey's theorem\ text \ @@ -266,7 +266,7 @@ lemma ramsey1: "\N::nat. partn_lst {..iH\nsets {.. {i}" - if "f \ nsets {.. {.. nsets {.. {.. \i. {q. q \ q0+q1 \ f {q} = i}" have "A 0 \ A 1 = {..q0 + q1}" @@ -280,7 +280,7 @@ then obtain i where "i < Suc (Suc 0)" "card (A i) \ [q0, q1] ! i" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) then obtain B where "B \ A i" "card B = [q0, q1] ! i" "finite B" - by (meson finite_subset get_smaller_card infinite_arbitrarily_large) + by (meson obtain_subset_with_card_n) then have "B \ nsets {.. f ` nsets B (Suc 0) \ {i}" by (auto simp: A_def nsets_def card_1_singleton_iff) then show ?thesis @@ -297,11 +297,11 @@ proof (induction r arbitrary: q1 q2) case 0 then show ?case - by (simp add: ramsey0) + by (simp add: ramsey0) next case (Suc r) note outer = this - show ?case + show ?case proof (cases "r = 0") case True then show ?thesis @@ -314,7 +314,7 @@ using Suc.prems proof (induct k \ "q1 + q2" arbitrary: q1 q2) case 0 - show ?case + show ?case proof show "partn_lst {..<1::nat} [q1, q2] (Suc r)" using nsets_empty_iff subset_insert 0 @@ -323,7 +323,7 @@ next case (Suc k) consider "q1 = 0 \ q2 = 0" | "q1 \ 0" "q2 \ 0" by auto - then show ?case + then show ?case proof cases case 1 then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)" @@ -335,12 +335,12 @@ with Suc have "k = (q1 - 1) + q2" "k = q1 + (q2 - 1)" by auto then obtain p1 p2::nat where p1: "partn_lst {..iH\nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \ {i}" - if f: "f \ nsets {..p} (Suc r) \ {.. nsets {..p} (Suc r) \ {.. \R. f (insert p R)" have "f (insert p i) \ {.. nsets {.. nsets {.. {.. {i}" - and U: "U \ nsets {.. nsets {.. {.. nsets {.. nsets {..p} n" if "X \ nsets {.. nsets {.. {.. {j}" - and V: "V \ nsets {.. nsets {.. {.. nsets {.. nsets {..p} q1" unfolding nsets_def using \q1 \ 0\ card_insert_if by fastforce - have invu_nsets: "inv_into {.. nsets V r" + have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" @@ -418,13 +418,13 @@ using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis - by (simp add: \f X = i\ \g (X - {p}) = i\) + by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q1" by (simp add: izero inq1) ultimately show ?thesis - by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) + by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) next case jone then have "u ` V \ nsets {..p} q2" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" - using hj + using hj by (force simp add: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD) ultimately show ?thesis using jone not_less_eq by fastforce qed - next + next case ione then have "U \ nsets {.. nsets {..p} n" if "X \ nsets {.. nsets {.. {.. {j}" - and V: "V \ nsets {.. nsets {.. {.. nsets {.. nsets {..p} q2" unfolding nsets_def using \q2 \ 0\ card_insert_if by fastforce - have invu_nsets: "inv_into {.. nsets V r" + have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" @@ -507,13 +507,13 @@ using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis - by (simp add: \f X = i\ \g (X - {p}) = i\) + by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q1" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" - using hj + using hj apply (clarsimp simp add: h_def image_subset_iff nsets_def) by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj) ultimately show ?thesis @@ -566,9 +566,9 @@ then obtain q1 q2 l where qs: "qs = q1#q2#l" by (metis Suc.hyps(2) length_Suc_conv) then obtain q::nat where q: "partn_lst {..qs = q1 # q2 # l\ by fastforce + using IH \qs = q1 # q2 # l\ by fastforce have keq: "Suc (length l) = k" using IH qs by auto show ?thesis @@ -582,21 +582,21 @@ unfolding g_def using f Suc IH by (auto simp: Pi_def not_less) then obtain i U where i: "i < k" and gi: "g ` nsets U r \ {i}" - and U: "U \ nsets {.. nsets {..iH\nsets {.. {i}" proof (cases "i = 0") case True then have "U \ nsets {.. {0, Suc 0}" using U gi unfolding g_def by (auto simp: image_subset_iff) - then obtain u where u: "bij_betw u {.. {..U \ nsets {.. mem_Collect_eq nsets_def) have u_nsets: "u ` X \ nsets {.. nsets {.. nsets {.. {.. {j}" - and V: "V \ nsets {.. nsets {.. Suc 0" and N: "partn_lst {..R\V. card R = m \ clique R E \ card R = n \ indep R E" + have "\R\V. card R = m \ clique R E \ card R = n \ indep R E" if "finite V" "N \ card V" for V :: "'a set" and E :: "'a set set" proof - from that @@ -673,7 +673,7 @@ have f: "f \ nsets {.. {.. {i}" - and U: "U \ nsets {.. nsets {.. (\x y. S = {x,y} \ x \ y)" -by (auto simp: card_Suc_eq numeral_eq_Suc) + by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" -by (auto simp: card_Suc_eq numeral_eq_Suc) + by (auto simp: card_Suc_eq numeral_eq_Suc) context ord begin @@ -1651,20 +1651,21 @@ apply (force intro: leI)+ done -lemma get_smaller_card: - assumes "finite A" "k \ card A" - obtains B where "B \ A" "card B = k" +lemma obtain_subset_with_card_n: + assumes "n \ card S" + obtains T where "T \ S" "card T = n" "finite T" proof - - obtain h where h: "bij_betw h {0..finite A\ ex_bij_betw_nat_finite by blast - show thesis - proof - show "h ` {0.. A" - by (metis \k \ card A\ bij_betw_def h image_mono ivl_subset zero_le) - have "inj_on h {0..k \ card A\ bij_betw_def h inj_on_subset ivl_subset zero_le) - then show "card (h ` {0.. *) -section \The Erdoes-Szekeres Theorem\ +section \The Erdös-Szekeres Theorem\ theory Erdoes_Szekeres imports Main begin -subsection \Addition to \<^theory>\HOL.Lattices_Big\ Theory\ - -lemma Max_gr: - assumes "finite A" - assumes "a \ A" "a > x" - shows "x < Max A" -using assms Max_ge less_le_trans by blast - -subsection \Additions to \<^theory>\HOL.Finite_Set\ Theory\ - -lemma obtain_subset_with_card_n: - assumes "n \ card S" - obtains T where "T \ S" "card T = n" -proof - - from assms obtain n' where "card S = n + n'" by (metis le_add_diff_inverse) - from this that show ?thesis - proof (induct n' arbitrary: S) - case 0 from this show ?case by auto - next - case Suc from this show ?case by (simp add: card_Suc_eq) (metis subset_insertI2) - qed -qed - lemma exists_set_with_max_card: assumes "finite S" "S \ {}" shows "\s \ S. card s = Max (card ` S)" -using assms -proof (induct S rule: finite.induct) - case (insertI S' s') - show ?case - proof (cases "S' \ {}") - case True - from this insertI.hyps(2) obtain s where s: "s \ S'" "card s = Max (card ` S')" by auto - from this(1) have that: "(if card s \ card s' then s else s') \ insert s' S'" by auto - have "card (if card s \ card s' then s else s') = Max (card ` insert s' S')" - using insertI(1) \S' \ {}\ s by auto - from this that show ?thesis by blast - qed (auto) -qed (auto) + by (metis (mono_tags, lifting) Max_in assms finite_imageI imageE image_is_empty) subsection \Definition of Monotonicity over a Carrier Set\ definition - "mono_on f R S = (\i\S. \j\S. i \ j \ R (f i) (f j))" + "gen_mono_on f R S = (\i\S. \j\S. i \ j \ R (f i) (f j))" -lemma mono_on_empty [simp]: "mono_on f R {}" -unfolding mono_on_def by auto +lemma gen_mono_on_empty [simp]: "gen_mono_on f R {}" + unfolding gen_mono_on_def by auto -lemma mono_on_singleton [simp]: "reflp R \ mono_on f R {x}" -unfolding mono_on_def reflp_def by auto +lemma gen_mono_on_singleton [simp]: "reflp R \ gen_mono_on f R {x}" + unfolding gen_mono_on_def reflp_def by auto -lemma mono_on_subset: "T \ S \ mono_on f R S \ mono_on f R T" -unfolding mono_on_def by (simp add: subset_iff) +lemma gen_mono_on_subset: "T \ S \ gen_mono_on f R S \ gen_mono_on f R T" + unfolding gen_mono_on_def by (simp add: subset_iff) -lemma not_mono_on_subset: "T \ S \ \ mono_on f R T \ \ mono_on f R S" -unfolding mono_on_def by blast +lemma not_gen_mono_on_subset: "T \ S \ \ gen_mono_on f R T \ \ gen_mono_on f R S" + unfolding gen_mono_on_def by blast lemma [simp]: "reflp ((\) :: 'a::order \ _ \ bool)" @@ -76,35 +41,35 @@ lemma Erdoes_Szekeres: fixes f :: "_ \ 'a::linorder" - shows "(\S. S \ {0..m * n} \ card S = m + 1 \ mono_on f (\) S) \ - (\S. S \ {0..m * n} \ card S = n + 1 \ mono_on f (\) S)" + shows "(\S. S \ {0..m * n} \ card S = m + 1 \ gen_mono_on f (\) S) \ + (\S. S \ {0..m * n} \ card S = n + 1 \ gen_mono_on f (\) S)" proof (rule ccontr) - let ?max_subseq = "\R k. Max (card ` {S. S \ {0..k} \ mono_on f R S \ k \ S})" + let ?max_subseq = "\R k. Max (card ` {S. S \ {0..k} \ gen_mono_on f R S \ k \ S})" define phi where "phi k = (?max_subseq (\) k, ?max_subseq (\) k)" for k - have one_member: "\R k. reflp R \ {k} \ {S. S \ {0..k} \ mono_on f R S \ k \ S}" by auto + have one_member: "\R k. reflp R \ {k} \ {S. S \ {0..k} \ gen_mono_on f R S \ k \ S}" by auto { fix R assume reflp: "reflp (R :: 'a::linorder \ _)" - from one_member[OF this] have non_empty: "\k. {S. S \ {0..k} \ mono_on f R S \ k \ S} \ {}" by force - from one_member[OF reflp] have "\k. card {k} \ card ` {S. S \ {0..k} \ mono_on f R S \ k \ S}" by blast + from one_member[OF this] have non_empty: "\k. {S. S \ {0..k} \ gen_mono_on f R S \ k \ S} \ {}" by force + from one_member[OF reflp] have "\k. card {k} \ card ` {S. S \ {0..k} \ gen_mono_on f R S \ k \ S}" by blast from this have lower_bound: "\k. k \ m * n \ ?max_subseq R k \ 1" by (auto intro!: Max_ge) fix b - assume not_mono_at: "\S. S \ {0..m * n} \ card S = b + 1 \ \ mono_on f R S" + assume not_mono_at: "\S. S \ {0..m * n} \ card S = b + 1 \ \ gen_mono_on f R S" { fix S assume "S \ {0..m * n}" "card S \ b + 1" moreover from \card S \ b + 1\ obtain T where "T \ S \ card T = Suc b" using obtain_subset_with_card_n by (metis Suc_eq_plus1) - ultimately have "\ mono_on f R S" using not_mono_at by (auto dest: not_mono_on_subset) + ultimately have "\ gen_mono_on f R S" using not_mono_at by (auto dest: not_gen_mono_on_subset) } - from this have "\S. S \ {0..m * n} \ mono_on f R S \ card S \ b" + from this have "\S. S \ {0..m * n} \ gen_mono_on f R S \ card S \ b" by (metis Suc_eq_plus1 Suc_leI not_le) - from this have "\k. k \ m * n \ \S. S \ {0..k} \ mono_on f R S \ card S \ b" + from this have "\k. k \ m * n \ \S. S \ {0..k} \ gen_mono_on f R S \ card S \ b" using order_trans by force from this non_empty have upper_bound: "\k. k \ m * n \ ?max_subseq R k \ b" by (auto intro: Max.boundedI) @@ -132,19 +97,19 @@ fix R assume R: "reflp (R :: 'a::linorder \ _)" "transp R" "R (f i) (f j)" from one_member[OF \reflp R\, of "i"] have - "\S \ {S. S \ {0..i} \ mono_on f R S \ i \ S}. card S = ?max_subseq R i" + "\S \ {S. S \ {0..i} \ gen_mono_on f R S \ i \ S}. card S = ?max_subseq R i" by (intro exists_set_with_max_card) auto - from this obtain S where S: "S \ {0..i} \ mono_on f R S \ i \ S" "card S = ?max_subseq R i" by auto + from this obtain S where S: "S \ {0..i} \ gen_mono_on f R S \ i \ S" "card S = ?max_subseq R i" by auto from S \i < j\ finite_subset have "j \ S" "finite S" "insert j S \ {0..j}" by auto - from S(1) R \i < j\ this have "mono_on f R (insert j S)" - unfolding mono_on_def reflp_def transp_def + from S(1) R \i < j\ this have "gen_mono_on f R (insert j S)" + unfolding gen_mono_on_def reflp_def transp_def by (metis atLeastAtMost_iff insert_iff le_antisym subsetCE) - from this have d: "insert j S \ {S. S \ {0..j} \ mono_on f R S \ j \ S}" + from this have d: "insert j S \ {S. S \ {0..j} \ gen_mono_on f R S \ j \ S}" using \insert j S \ {0..j}\ by blast from this \j \ S\ S(1) have "card (insert j S) \ - card ` {S. S \ {0..j} \ mono_on f R S \ j \ S} \ card S < card (insert j S)" + card ` {S. S \ {0..j} \ gen_mono_on f R S \ j \ S} \ card S < card (insert j S)" by (auto intro!: imageI) (auto simp add: \finite S\) - from this S(2) have "?max_subseq R i < ?max_subseq R j" by (auto intro: Max_gr) + from this S(2) have "?max_subseq R i < ?max_subseq R j" by (auto intro!: Max_gr_iff [THEN iffD2]) } note max_subseq_increase = this have "?max_subseq (\) i < ?max_subseq (\) j \ ?max_subseq (\) i < ?max_subseq (\) j" proof (cases "f j \ f i")