# HG changeset patch # User paulson # Date 1575908016 0 # Node ID 09aee7f5b4475d35dd3f4839c1f525ccb4ece232 # Parent d67924987c34b19f2f6d525fad31488d2efed572 Ramsey with multiple colours and arbitrary exponents diff -r d67924987c34 -r 09aee7f5b447 NEWS --- a/NEWS Mon Dec 09 15:36:51 2019 +0000 +++ b/NEWS Mon Dec 09 16:13:36 2019 +0000 @@ -81,6 +81,9 @@ * Theory HOL-Library.Monad_Syntax: infix operation "bind" (>>=) associates to the left now as is customary. +* Theory HOL-Library.Ramsey: full finite Ramsey's theorem with +multiple colours and arbitrary exponents. + * Theory Complete_Lattices: renamed Inf_Sup -> Inf_eq_Sup and Sup_Inf -> Sup_eq_Inf diff -r d67924987c34 -r 09aee7f5b447 src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Mon Dec 09 15:36:51 2019 +0000 +++ b/src/HOL/Library/Ramsey.thy Mon Dec 09 16:13:36 2019 +0000 @@ -1,20 +1,526 @@ (* Title: HOL/Library/Ramsey.thy - Author: Tom Ridge. Converted to structured Isar by L C Paulson + Author: Tom Ridge. Full finite version by L C Paulson. *) section \Ramsey's Theorem\ theory Ramsey - imports Infinite_Set + imports Infinite_Set FuncSet begin -subsection \Finite Ramsey theorem(s)\ +subsection \Preliminary definitions\ + +subsubsection \The $n$-element subsets of a set $A$\ + +definition nsets :: "['a set, nat] \ 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999) + where "nsets A n \ {N. N \ A \ finite N \ card N = n}" + +lemma nsets_mono: "A \ B \ nsets A n \ nsets B n" + by (auto simp: nsets_def) + +lemma nsets_2_eq: "nsets A 2 = (\x\A. \y\A - {x}. {{x, y}})" + unfolding numeral_2_eq_2 + by (auto simp: nsets_def elim!: card_2_doubletonE) + +lemma binomial_eq_nsets: "n choose k = card (nsets {0.. finite A \ card A < r" + unfolding nsets_def +proof (intro iffI conjI) + assume that: "{N. N \ A \ finite N \ card N = r} = {}" + 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]) + then show "card A < r" + using not_less by blast +next + show "{N. N \ A \ finite N \ card N = r} = {}" + if "finite A \ card A < r" + using that card_mono leD by auto +qed + +lemma nsets_eq_empty: "n < r \ nsets {..x. {x}) ` A" + using card_eq_SucD by (force simp: nsets_def) + +subsubsection \Partition predicates\ + +definition partn :: "'a set \ nat \ nat \ 'b set \ bool" + where "partn \ \ \ \ \ \f \ nsets \ \ \ \. \H \ nsets \ \. \\\\. f ` (nsets H \) \ {\}" + +definition partn_lst :: "'a set \ nat list \ nat \ bool" + where "partn_lst \ \ \ \ \f \ nsets \ \ \ {..}. + \i < length \. \H \ nsets \ (\!i). f ` (nsets H \) \ {i}" + +lemma partn_lst_greater_resource: + fixes M::nat + assumes M: "partn_lst {.. \" and "M \ N" + shows "partn_lst {.. \" +proof (clarsimp simp: partn_lst_def) + fix f + assume "f \ nsets {.. \ {..}" + then have "f \ nsets {.. \ {..}" + by (meson Pi_anti_mono \M \ N\ lessThan_subset_iff nsets_mono subsetD) + then obtain i H where i: "i < length \" and H: "H \ nsets {.. ! i)" and subi: "f ` nsets H \ \ {i}" + using M partn_lst_def by blast + have "H \ nsets {.. ! i)" + using \M \ N\ H by (auto simp: nsets_def subset_iff) + then show "\i. \H\nsets {.. ! i). f ` nsets H \ \ {i}" + using i subi by blast +qed + +lemma partn_lst_fewer_colours: + assumes major: "partn_lst \ (n#\) \" and "n \ \" + shows "partn_lst \ \ \" +proof (clarsimp simp: partn_lst_def) + fix f :: "'a set \ nat" + assume f: "f \ [\]\<^bsup>\\<^esup> \ {..}" + then obtain i H where i: "i < Suc (length \)" + and H: "H \ [\]\<^bsup>((n # \) ! i)\<^esup>" + and hom: "\x\[H]\<^bsup>\\<^esup>. Suc (f x) = i" + using \n \ \\ major [unfolded partn_lst_def, rule_format, of "Suc o f"] + by (fastforce simp: image_subset_iff nsets_eq_empty_iff) + show "\i. \H\nsets \ (\ ! i). f ` [H]\<^bsup>\\<^esup> \ {i}" + proof (cases i) + case 0 + then have "[H]\<^bsup>\\<^esup> = {}" + using hom by blast + then show ?thesis + using 0 H \n \ \\ + by (simp add: nsets_eq_empty_iff) (simp add: nsets_def) + next + case (Suc i') + then show ?thesis + using i H hom by auto + qed +qed + +lemma partn_lst_eq_partn: "partn_lst {..Finite versions of Ramsey's theorem\ text \ To distinguish the finite and infinite ones, lower and upper case names are used. +\ - This is the most basic version in terms of cliques and independent +subsubsection \Trivial cases\ + +text \Vacuous, since we are dealing with 0-sets!\ +lemma ramsey0: "\N::nat. partn_lst {..Just the pigeon hole principle, since we are dealing with 1-sets\ +lemma ramsey1: "\N::nat. partn_lst {..iH\nsets {.. {i}" + if "f \ nsets {.. {.. \i. {q. q \ q0+q1 \ f {q} = i}" + have "A 0 \ A 1 = {..q0 + q1}" + using that by (auto simp: A_def PiE_iff nsets_one lessThan_Suc_atMost le_Suc_eq) + moreover have "A 0 \ A 1 = {}" + by (auto simp: A_def) + ultimately have "q0 + q1 \ card (A 0) + card (A 1)" + by (metis card_Un_le card_atMost eq_imp_le le_SucI le_trans) + then consider "card (A 0) \ q0" | "card (A 1) \ q1" + by linarith + 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) + then have "B \ nsets {.. f ` nsets B (Suc 0) \ {i}" + by (auto simp: A_def nsets_def card_1_singleton_iff) + then show ?thesis + using \i < Suc (Suc 0)\ by auto + qed + then show ?thesis + by (clarsimp simp: partn_lst_def) blast +qed + + +subsubsection \Ramsey's theorem with two colours and arbitrary exponents (hypergraph version)\ + +proposition ramsey2_full: "\N::nat. partn_lst {.. 0" + by simp + show ?thesis + using Suc.prems + proof (induct k \ "q1 + q2" arbitrary: q1 q2) + case 0 + show ?case + proof + show "partn_lst {..<1::nat} [q1, q2] (Suc r)" + using nsets_empty_iff subset_insert 0 + by (fastforce simp: partn_lst_def funcset_to_empty_iff nsets_eq_empty image_subset_iff) + qed + next + case (Suc k) + consider "q1 = 0 \ q2 = 0" | "q1 \ 0" "q2 \ 0" by auto + then show ?case + proof cases + case 1 + then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)" + unfolding partn_lst_def using \r > 0\ + by (fastforce simp add: nsets_empty_iff nsets_singleton_iff lessThan_Suc) + then show ?thesis by blast + next + case 2 + 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) \ {.. \R. f (insert p R)" + have "f (insert p i) \ {.. nsets {.. nsets {.. {.. {i}" + and U: "U \ nsets {.. {.. nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" + have "h \ nsets {.. {.. {j}" + and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q1" + unfolding nsets_def using \q1 \ 0\ card_insert_if by fastforce + 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" + using nsets_def that by auto + then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X + proof (cases "p \ X") + case True + then have Xp: "X - {p} \ nsets (u ` V) r" + using X by (auto simp: nsets_def) + moreover have "u ` V \ U" + using Vsub bij_betwE u by blast + ultimately have "X - {p} \ nsets U r" + by (meson in_mono nsets_mono) + then have "g (X - {p}) = i" + using gi by blast + have "f X = i" + 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\) + 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) + 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 + 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 + case ione + then have "U \ nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" + have "h \ nsets {.. {.. {j}" + and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q2" + unfolding nsets_def using \q2 \ 0\ card_insert_if by fastforce + 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" + using nsets_def that by auto + then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X + proof (cases "p \ X") + case True + then have Xp: "X - {p} \ nsets (u ` V) r" + using X by (auto simp: nsets_def) + moreover have "u ` V \ U" + using Vsub bij_betwE u by blast + ultimately have "X - {p} \ nsets U r" + by (meson in_mono nsets_mono) + then have "g (X - {p}) = i" + using gi by blast + have "f X = i" + 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\) + 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} q2" + by (simp add: ione inq1) + ultimately show ?thesis + by (metis ione image_subsetI insertI1 lessI nth_Cons_0 nth_Cons_Suc) + next + case jzero + then have "u ` V \ nsets {..p} q1" + using V u_nsets by auto + moreover have "f ` nsets (u ` V) (Suc r) \ {j}" + 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 + using jzero not_less_eq by fastforce + qed + qed + qed + then show "partn_lst {..Full Ramsey's theorem with multiple colours and arbitrary exponents\ + +theorem ramsey_full: "\N::nat. partn_lst {.. "length qs" arbitrary: qs) + case 0 + then show ?case + by (rule_tac x=" r" in exI) (simp add: partn_lst_def) +next + case (Suc k) + note IH = this + show ?case + proof (cases k) + case 0 + with Suc obtain q where "qs = [q]" + by (metis length_0_conv length_Suc_conv) + then show ?thesis + by (rule_tac x=q in exI) (auto simp: partn_lst_def funcset_to_empty_iff) + next + case (Suc k') + 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 + have keq: "Suc (length l) = k" + using IH qs by auto + show ?thesis + proof (intro exI conjI) + show "partn_lst {.. nsets {.. {.. \X. if f X < Suc (Suc 0) then 0 else f X - Suc 0" + have "g \ nsets {.. {.. {i}" + and U: "U \ 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 {.. \X. f (u ` X)" + have "f (u ` X) < Suc (Suc 0)" if "X \ nsets {.. nsets U r" + using u u_nsets that by (auto simp: nsets_def bij_betwE subset_eq) + then show ?thesis + using f01 by auto + qed + then have "h \ nsets {.. {.. {j}" + and V: "V \ nsets {.. (\x. (u ` x)) ` nsets V r" + apply (clarsimp simp add: nsets_def image_iff) + by (metis card_eq_0_iff card_image image_is_empty subset_image_inj) + then have "f ` nsets (u ` V) r \ h ` nsets V r" + by (auto simp: h_def) + then show "f ` nsets (u ` V) r \ {j}" + using hj by auto + show "(u ` V) \ nsets {.. {Suc i}" + using i gi False + apply (auto simp: g_def image_subset_iff) + by (metis Suc_lessD Suc_pred g_def gi image_subset_iff not_less_eq singleton_iff) + show "U \ nsets {..Simple graph version\ + +text \This is the most basic version in terms of cliques and independent sets, i.e. the version for graphs and 2 colours. \ @@ -24,90 +530,35 @@ lemma ramsey2: "\r\1. \(V::'a set) (E::'a set set). finite V \ card V \ r \ (\R \ V. card R = m \ clique R E \ card R = n \ indep R E)" - (is "\r\1. ?R m n r") -proof (induct k \ "m + n" arbitrary: m n) - case 0 - show ?case (is "\r. ?Q r") - proof - from 0 show "?Q 1" - by (clarsimp simp: indep_def) (metis card.empty emptyE empty_subsetI) +proof - + obtain N where "N \ Suc 0" and N: "partn_lst {..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 + obtain v where u: "inj_on v {.. V" + by (metis card_le_inj card_lessThan finite_lessThan) + define f where "f \ \e. if v ` e \ E then 0 else Suc 0" + have f: "f \ nsets {.. {.. {i}" + and U: "U \ nsets {.. V" + using U u by (auto simp: image_subset_iff nsets_def) + show "card (v ` U) = m \ clique (v ` U) E \ card (v ` U) = n \ indep (v ` U) E" + using i unfolding numeral_2_eq_2 + using gi U u + apply (simp add: image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq) + apply (auto simp: f_def nsets_def card_image inj_on_subset split: if_split_asm) + done + qed qed -next - case (Suc k) - consider "m = 0 \ n = 0" | "m \ 0" "n \ 0" by auto - then show ?case (is "\r. ?Q r") - proof cases - case 1 - then have "?Q 1" - by (simp add: clique_def) (meson card_empty empty_iff empty_subsetI indep_def) - then show ?thesis .. - next - case 2 - with Suc(2) have "k = (m - 1) + n" "k = m + (n - 1)" by auto - from this [THEN Suc(1)] - obtain r1 r2 where "r1 \ 1" "r2 \ 1" "?R (m - 1) n r1" "?R m (n - 1) r2" by auto - then have "r1 + r2 \ 1" by arith - moreover have "?R m n (r1 + r2)" (is "\V E. _ \ ?EX V E m n") - proof clarify - fix V :: "'a set" - fix E :: "'a set set" - assume "finite V" "r1 + r2 \ card V" - with \r1 \ 1\ have "V \ {}" by auto - then obtain v where "v \ V" by blast - let ?M = "{w \ V. w \ v \ {v, w} \ E}" - let ?N = "{w \ V. w \ v \ {v, w} \ E}" - from \v \ V\ have "V = insert v (?M \ ?N)" by auto - then have "card V = card (insert v (?M \ ?N))" by metis - also from \finite V\ have "\ = card ?M + card ?N + 1" - by (fastforce intro: card_Un_disjoint) - finally have "card V = card ?M + card ?N + 1" . - with \r1 + r2 \ card V\ have "r1 + r2 \ card ?M + card ?N + 1" by simp - then consider "r1 \ card ?M" | "r2 \ card ?N" by arith - then show "?EX V E m n" - proof cases - case 1 - from \finite V\ have "finite ?M" by auto - with \?R (m - 1) n r1\ and 1 have "?EX ?M E (m - 1) n" by blast - then obtain R where "R \ ?M" "v \ R" - and CI: "card R = m - 1 \ clique R E \ card R = n \ indep R E" (is "?C \ ?I") - by blast - from \R \ ?M\ have "R \ V" by auto - with \finite V\ have "finite R" by (metis finite_subset) - from CI show ?thesis - proof - assume "?I" - with \R \ V\ show ?thesis by blast - next - assume "?C" - with \R \ ?M\ have *: "clique (insert v R) E" - by (auto simp: clique_def insert_commute) - from \?C\ \finite R\ \v \ R\ \m \ 0\ have "card (insert v R) = m" by simp - with \R \ V\ \v \ V\ * show ?thesis by (metis insert_subset) - qed - next - case 2 - from \finite V\ have "finite ?N" by auto - with \?R m (n - 1) r2\ 2 have "?EX ?N E m (n - 1)" by blast - then obtain R where "R \ ?N" "v \ R" - and CI: "card R = m \ clique R E \ card R = n - 1 \ indep R E" (is "?C \ ?I") - by blast - from \R \ ?N\ have "R \ V" by auto - with \finite V\ have "finite R" by (metis finite_subset) - from CI show ?thesis - proof - assume "?C" - with \R \ V\ show ?thesis by blast - next - assume "?I" - with \R \ ?N\ have *: "indep (insert v R) E" - by (auto simp: indep_def insert_commute) - from \?I\ \finite R\ \v \ R\ \n \ 0\ have "card (insert v R) = n" by simp - with \R \ V\ \v \ V\ * show ?thesis by (metis insert_subset) - qed - qed - qed - ultimately show ?thesis by blast - qed + then show ?thesis + using \Suc 0 \ N\ by auto qed @@ -156,15 +607,15 @@ definition part :: "nat \ nat \ 'a set \ ('a set \ nat) \ bool" \ \the function \<^term>\f\ partitions the \<^term>\r\-subsets of the typically infinite set \<^term>\Y\ into \<^term>\s\ distinct categories.\ - where "part r s Y f \ (\X. X \ Y \ finite X \ card X = r \ f X < s)" + where "part r s Y f \ (\X \ nsets Y r. f X < s)" text \For induction, we decrease the value of \<^term>\r\ in partitions.\ lemma part_Suc_imp_part: "\infinite Y; part (Suc r) s Y f; y \ Y\ \ part r s (Y - {y}) (\u. f (insert y u))" - by (simp add: part_def subset_Diff_insert) + by (simp add: part_def nsets_def subset_Diff_insert) lemma part_subset: "part r s YY f \ Y \ YY \ part r s Y f" - unfolding part_def by blast + unfolding part_def nsets_def by blast subsection \Ramsey's Theorem: Infinitary Version\ @@ -299,7 +750,7 @@ \X. X \ Z \ finite X \ card X = r \ f X < s\ \ \Y t. Y \ Z \ infinite Y \ t < s \ (\X. X \ Y \ finite X \ card X = r \ f X = t)" - by (blast intro: Ramsey_induction [unfolded part_def]) + by (blast intro: Ramsey_induction [unfolded part_def nsets_def]) corollary Ramsey2: