# HG changeset patch # User paulson # Date 1707655934 0 # Node ID 9dee3b4fdb063b69ceb4378da94f951f76b50fa5 # Parent 9f22b71e209ed4875958b736d057bee4bee1fbbf new lemmas involving Ramsey numbers, infinite sets diff -r 9f22b71e209e -r 9dee3b4fdb06 src/HOL/Library/Equipollence.thy --- a/src/HOL/Library/Equipollence.thy Fri Feb 09 16:43:43 2024 +0000 +++ b/src/HOL/Library/Equipollence.thy Sun Feb 11 12:52:14 2024 +0000 @@ -747,6 +747,8 @@ "finite(I \\<^sub>E S) \ (\a. S \ {a}) \ I = {} \ finite I \ finite S" by (fastforce simp: finite_PiE_iff PiE_eq_empty_iff dest: subset_singletonD) +subsection \Misc other resultd\ + lemma lists_lepoll_mono: assumes "A \ B" shows "lists A \ lists B" proof - @@ -762,4 +764,37 @@ lemma lepoll_lists: "A \ lists A" unfolding lepoll_def inj_on_def by(rule_tac x="\x. [x]" in exI) auto +text \Dedekind's definition of infinite set\ + +lemma infinite_iff_psubset: "infinite A \ (\B. B \ A \ A\B)" +proof + assume "infinite A" + then obtain f :: "nat \ 'a" where "inj f" and f: "range f \ A" + by (meson infinite_countable_subset) + define C where "C \ A - range f" + have C: "A = range f \ C" "range f \ C = {}" + using f by (auto simp: C_def) + have *: "range (f \ Suc) \ range f" + using inj_eq [OF \inj f\] by (fastforce simp: set_eq_iff) + have "range f \ C \ range (f \ Suc) \ C" + proof (intro Un_eqpoll_cong) + show "range f \ range (f \ Suc)" + by (meson \inj f\ eqpoll_refl inj_Suc inj_compose inj_on_image_eqpoll_2) + show "disjnt (range f) C" + by (simp add: C disjnt_def) + then show "disjnt (range (f \ Suc)) C" + using "*" disjnt_subset1 by blast + qed auto + moreover have "range (f \ Suc) \ C \ A" + using "*" f C_def by blast + ultimately show "\B\A. A \ B" + by (metis C(1)) +next + assume "\B\A. A \ B" then show "infinite A" + by (metis card_subset_eq eqpoll_finite_iff eqpoll_iff_card psubsetE) +qed + +lemma infinite_iff_psubset_le: "infinite A \ (\B. B \ A \ A \ B)" + by (meson eqpoll_imp_lepoll infinite_iff_psubset lepoll_antisym psubsetE subset_imp_lepoll) + end diff -r 9f22b71e209e -r 9dee3b4fdb06 src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Fri Feb 09 16:43:43 2024 +0000 +++ b/src/HOL/Library/Ramsey.thy Sun Feb 11 12:52:14 2024 +0000 @@ -18,6 +18,9 @@ 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 finite_imp_finite_nsets: "finite A \ finite ([A]\<^bsup>k\<^esup>)" + by (simp add: nsets_def) + lemma nsets_mono: "A \ B \ nsets A n \ nsets B n" by (auto simp: nsets_def) @@ -27,6 +30,11 @@ lemma nsets_2_eq: "nsets A 2 = (\x\A. \y\A - {x}. {{x, y}})" by (auto simp: nsets_def card_2_iff) +lemma nsets2_E: + assumes "e \ [A]\<^bsup>2\<^esup>" + obtains x y where "e = {x,y}" "x \ A" "y \ A" "x\y" + using assms by (auto simp: nsets_def card_2_iff) + lemma nsets_doubleton_2_eq [simp]: "[{x, y}]\<^bsup>2\<^esup> = (if x=y then {} else {{x, y}})" by (auto simp: nsets_2_eq) @@ -254,6 +262,45 @@ apply (simp add: partn_lst_def partn_def numeral_2_eq_2) by (metis less_2_cases numeral_2_eq_2 lessThan_iff nth_Cons_0 nth_Cons_Suc) +lemma partn_lstE: + assumes "partn_lst \ \ \" "f \ nsets \ \ \ {.. = l" + obtains i H where "i < length \" "H \ nsets \ (\!i)" "f ` (nsets H \) \ {i}" + using partn_lst_def assms by blast + +lemma partn_lst_less: + assumes M: "partn_lst \ \ n" and eq: "length \' = length \" + and le: "\i. i < length \ \ \'!i \ \!i " + shows "partn_lst \ \' n" +proof (clarsimp simp: partn_lst_def) + fix f + assume "f \ [\]\<^bsup>n\<^esup> \ {..'}" + then obtain i H where i: "i < length \" + and "H \ \" and H: "card H = (\!i)" and "finite H" + and fi: "f ` nsets H n \ {i}" + using assms by (auto simp: partn_lst_def nsets_def) + then obtain bij where bij: "bij_betw bij H {0..<\!i}" + by (metis ex_bij_betw_finite_nat) + then have inj: "inj_on (inv_into H bij) {0..<\' ! i}" + by (metis bij_betw_def dual_order.refl i inj_on_inv_into ivl_subset le) + define H' where "H' = inv_into H bij ` {0..<\'!i}" + show "\i'. \H\[\]\<^bsup>(\' ! i)\<^esup>. f ` [H]\<^bsup>n\<^esup> \ {i}" + proof (intro exI bexI conjI) + show "i < length \'" + by (simp add: assms(2) i) + have "H' \ H" + using bij \i < length \\ bij_betw_imp_surj_on le + by (force simp: H'_def image_subset_iff intro: inv_into_into) + then have "finite H'" + by (simp add: \finite H\ finite_subset) + with \H' \ H\ have cardH': "card H' = (\'!i)" + unfolding H'_def by (simp add: inj card_image) + show "f ` [H']\<^bsup>n\<^esup> \ {i}" + by (meson \H' \ H\ dual_order.trans fi image_mono nsets_mono) + show "H' \ [\]\<^bsup>(\'! i)\<^esup>" + using \H \ \\ \H' \ H\ \finite H'\ cardH' nsets_def by fastforce + qed +qed + subsection \Finite versions of Ramsey's theorem\