| author | nipkow | 
| Tue, 11 Oct 2022 10:45:42 +0200 | |
| changeset 76259 | d1c26efb7a47 | 
| parent 73709 | 58bd53caf800 | 
| child 76987 | 4c275405faae | 
| permissions | -rw-r--r-- | 
| 19944 | 1 | (* Title: HOL/Library/Ramsey.thy | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 2 | Author: Tom Ridge. Full finite version by L C Paulson. | 
| 19944 | 3 | *) | 
| 4 | ||
| 65075 | 5 | section \<open>Ramsey's Theorem\<close> | 
| 19944 | 6 | |
| 25594 | 7 | theory Ramsey | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 8 | imports Infinite_Set FuncSet | 
| 25594 | 9 | begin | 
| 19944 | 10 | |
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 11 | subsection \<open>Preliminary definitions\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 12 | |
| 73709 | 13 | abbreviation strict_sorted :: "'a::linorder list \<Rightarrow> bool" where | 
| 14 | "strict_sorted \<equiv> sorted_wrt (<)" | |
| 15 | ||
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 16 | subsubsection \<open>The $n$-element subsets of a set $A$\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 17 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 18 | definition nsets :: "['a set, nat] \<Rightarrow> 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999)
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 19 |   where "nsets A n \<equiv> {N. N \<subseteq> A \<and> finite N \<and> card N = n}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 20 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 21 | lemma nsets_mono: "A \<subseteq> B \<Longrightarrow> nsets A n \<subseteq> nsets B n" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 22 | by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 23 | |
| 72231 | 24 | lemma nsets_Pi_contra: "A' \<subseteq> A \<Longrightarrow> Pi ([A]\<^bsup>n\<^esup>) B \<subseteq> Pi ([A']\<^bsup>n\<^esup>) B" | 
| 25 | by (auto simp: nsets_def) | |
| 26 | ||
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 27 | lemma nsets_2_eq: "nsets A 2 = (\<Union>x\<in>A. \<Union>y\<in>A - {x}. {{x, y}})"
 | 
| 72231 | 28 | by (auto simp: nsets_def card_2_iff) | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 29 | |
| 71452 | 30 | lemma nsets_doubleton_2_eq [simp]: "[{x, y}]\<^bsup>2\<^esup> = (if x=y then {} else {{x, y}})"
 | 
| 31 | by (auto simp: nsets_2_eq) | |
| 32 | ||
| 33 | lemma doubleton_in_nsets_2 [simp]: "{x,y} \<in> [A]\<^bsup>2\<^esup> \<longleftrightarrow> x \<in> A \<and> y \<in> A \<and> x \<noteq> y"
 | |
| 34 | by (auto simp: nsets_2_eq Set.doubleton_eq_iff) | |
| 35 | ||
| 71464 | 36 | lemma nsets_3_eq: "nsets A 3 = (\<Union>x\<in>A. \<Union>y\<in>A - {x}. \<Union>z\<in>A - {x,y}. {{x,y,z}})"
 | 
| 37 | by (simp add: eval_nat_numeral nsets_def card_Suc_eq) blast | |
| 38 | ||
| 39 | lemma nsets_4_eq: "[A]\<^bsup>4\<^esup> = (\<Union>u\<in>A. \<Union>x\<in>A - {u}. \<Union>y\<in>A - {u,x}. \<Union>z\<in>A - {u,x,y}. {{u,x,y,z}})"
 | |
| 40 | (is "_ = ?rhs") | |
| 41 | proof | |
| 42 | show "[A]\<^bsup>4\<^esup> \<subseteq> ?rhs" | |
| 43 | by (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) blast | |
| 44 | show "?rhs \<subseteq> [A]\<^bsup>4\<^esup>" | |
| 45 | apply (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) | |
| 46 | by (metis insert_iff singletonD) | |
| 47 | qed | |
| 48 | ||
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 49 | lemma nsets_disjoint_2: | 
| 71464 | 50 |   "X \<inter> Y = {} \<Longrightarrow> [X \<union> Y]\<^bsup>2\<^esup> = [X]\<^bsup>2\<^esup> \<union> [Y]\<^bsup>2\<^esup> \<union> (\<Union>x\<in>X. \<Union>y\<in>Y. {{x,y}})"
 | 
| 51 | by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff) | |
| 52 | ||
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 53 | lemma ordered_nsets_2_eq: | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 54 | fixes A :: "'a::linorder set" | 
| 71464 | 55 |   shows "nsets A 2 = {{x,y} | x y. x \<in> A \<and> y \<in> A \<and> x<y}"
 | 
| 56 | (is "_ = ?rhs") | |
| 57 | proof | |
| 58 | show "nsets A 2 \<subseteq> ?rhs" | |
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 59 | unfolding numeral_nat | 
| 71464 | 60 | apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 61 | by (metis antisym) | 
| 71464 | 62 | show "?rhs \<subseteq> nsets A 2" | 
| 63 | unfolding numeral_nat by (auto simp: nsets_def card_Suc_eq) | |
| 64 | qed | |
| 65 | ||
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 66 | lemma ordered_nsets_3_eq: | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 67 | fixes A :: "'a::linorder set" | 
| 71464 | 68 |   shows "nsets A 3 = {{x,y,z} | x y z. x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> x<y \<and> y<z}"
 | 
| 69 | (is "_ = ?rhs") | |
| 70 | proof | |
| 71 | show "nsets A 3 \<subseteq> ?rhs" | |
| 72 | apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral) | |
| 73 | by (metis insert_commute linorder_cases) | |
| 74 | show "?rhs \<subseteq> nsets A 3" | |
| 75 | apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral) | |
| 76 | by (metis empty_iff insert_iff not_less_iff_gr_or_eq) | |
| 77 | qed | |
| 78 | ||
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 79 | lemma ordered_nsets_4_eq: | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 80 | fixes A :: "'a::linorder set" | 
| 71464 | 81 |   shows "[A]\<^bsup>4\<^esup> = {U. \<exists>u x y z. U = {u,x,y,z} \<and> u \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < x \<and> x < y \<and> y < z}"
 | 
| 82 | (is "_ = Collect ?RHS") | |
| 83 | proof - | |
| 84 |   { fix U
 | |
| 85 | assume "U \<in> [A]\<^bsup>4\<^esup>" | |
| 86 | then obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \<subseteq> A" | |
| 87 | by (simp add: nsets_def) (metis finite_set_strict_sorted) | |
| 88 | then have "?RHS U" | |
| 89 | unfolding numeral_nat length_Suc_conv by auto blast } | |
| 90 | moreover | |
| 91 | have "Collect ?RHS \<subseteq> [A]\<^bsup>4\<^esup>" | |
| 92 | apply (clarsimp simp add: nsets_def eval_nat_numeral) | |
| 93 | apply (subst card_insert_disjoint, auto)+ | |
| 94 | done | |
| 95 | ultimately show ?thesis | |
| 96 | by auto | |
| 97 | qed | |
| 98 | ||
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 99 | lemma ordered_nsets_5_eq: | 
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 100 | fixes A :: "'a::linorder set" | 
| 71464 | 101 |   shows "[A]\<^bsup>5\<^esup> = {U. \<exists>u v x y z. U = {u,v,x,y,z} \<and> u \<in> A \<and> v \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < v \<and> v < x \<and> x < y \<and> y < z}"
 | 
| 102 | (is "_ = Collect ?RHS") | |
| 103 | proof - | |
| 104 |   { fix U
 | |
| 105 | assume "U \<in> [A]\<^bsup>5\<^esup>" | |
| 106 | then obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \<subseteq> A" | |
| 107 | apply (simp add: nsets_def) | |
| 108 | by (metis finite_set_strict_sorted) | |
| 109 | then have "?RHS U" | |
| 110 | unfolding numeral_nat length_Suc_conv by auto blast } | |
| 111 | moreover | |
| 112 | have "Collect ?RHS \<subseteq> [A]\<^bsup>5\<^esup>" | |
| 113 | apply (clarsimp simp add: nsets_def eval_nat_numeral) | |
| 114 | apply (subst card_insert_disjoint, auto)+ | |
| 115 | done | |
| 116 | ultimately show ?thesis | |
| 117 | by auto | |
| 118 | qed | |
| 119 | ||
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 120 | lemma binomial_eq_nsets: "n choose k = card (nsets {0..<n} k)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 121 | apply (simp add: binomial_def nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 122 | by (meson subset_eq_atLeast0_lessThan_finite) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 123 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 124 | lemma nsets_eq_empty_iff: "nsets A r = {} \<longleftrightarrow> finite A \<and> card A < r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 125 | unfolding nsets_def | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 126 | proof (intro iffI conjI) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 127 |   assume that: "{N. N \<subseteq> A \<and> finite N \<and> card N = r} = {}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 128 | show "finite A" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 129 | using infinite_arbitrarily_large that by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 130 | then have "\<not> r \<le> card A" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 131 | using that by (simp add: set_eq_iff) (metis obtain_subset_with_card_n) | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 132 | then show "card A < r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 133 | using not_less by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 134 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 135 |   show "{N. N \<subseteq> A \<and> finite N \<and> card N = r} = {}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 136 | if "finite A \<and> card A < r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 137 | using that card_mono leD by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 138 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 139 | |
| 71452 | 140 | lemma nsets_eq_empty: "\<lbrakk>finite A; card A < r\<rbrakk> \<Longrightarrow> nsets A r = {}"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 141 | by (simp add: nsets_eq_empty_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 142 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 143 | lemma nsets_empty_iff: "nsets {} r = (if r=0 then {{}} else {})"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 144 | by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 145 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 146 | lemma nsets_singleton_iff: "nsets {a} r = (if r=0 then {{}} else if r=1 then {{a}} else {})"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 147 | by (auto simp: nsets_def card_gt_0_iff subset_singleton_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 148 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 149 | lemma nsets_self [simp]: "nsets {..<m} m = {{..<m}}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 150 | unfolding nsets_def | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 151 | apply auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 152 | by (metis add.left_neutral lessThan_atLeast0 lessThan_iff subset_card_intvl_is_intvl) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 153 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 154 | lemma nsets_zero [simp]: "nsets A 0 = {{}}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 155 | by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 156 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 157 | lemma nsets_one: "nsets A (Suc 0) = (\<lambda>x. {x}) ` A"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 158 | using card_eq_SucD by (force simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 159 | |
| 71405 | 160 | lemma inj_on_nsets: | 
| 161 | assumes "inj_on f A" | |
| 162 | shows "inj_on (\<lambda>X. f ` X) ([A]\<^bsup>n\<^esup>)" | |
| 163 | using assms unfolding nsets_def | |
| 164 | by (metis (no_types, lifting) inj_on_inverseI inv_into_image_cancel mem_Collect_eq) | |
| 165 | ||
| 166 | lemma bij_betw_nsets: | |
| 167 | assumes "bij_betw f A B" | |
| 168 | shows "bij_betw (\<lambda>X. f ` X) ([A]\<^bsup>n\<^esup>) ([B]\<^bsup>n\<^esup>)" | |
| 169 | proof - | |
| 170 | have "(`) f ` [A]\<^bsup>n\<^esup> = [f ` A]\<^bsup>n\<^esup>" | |
| 171 | using assms | |
| 172 | apply (auto simp: nsets_def bij_betw_def image_iff card_image inj_on_subset) | |
| 173 | by (metis card_image inj_on_finite order_refl subset_image_inj) | |
| 174 | with assms show ?thesis | |
| 175 | by (auto simp: bij_betw_def inj_on_nsets) | |
| 176 | qed | |
| 177 | ||
| 71452 | 178 | lemma nset_image_obtains: | 
| 179 | assumes "X \<in> [f`A]\<^bsup>k\<^esup>" "inj_on f A" | |
| 180 | obtains Y where "Y \<in> [A]\<^bsup>k\<^esup>" "X = f ` Y" | |
| 181 | using assms | |
| 182 | apply (clarsimp simp add: nsets_def subset_image_iff) | |
| 183 | by (metis card_image finite_imageD inj_on_subset) | |
| 184 | ||
| 185 | lemma nsets_image_funcset: | |
| 186 | assumes "g \<in> S \<rightarrow> T" and "inj_on g S" | |
| 187 | shows "(\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> [T]\<^bsup>k\<^esup>" | |
| 188 | using assms | |
| 189 | by (fastforce simp add: nsets_def card_image inj_on_subset subset_iff simp flip: image_subset_iff_funcset) | |
| 190 | ||
| 191 | lemma nsets_compose_image_funcset: | |
| 192 | assumes f: "f \<in> [T]\<^bsup>k\<^esup> \<rightarrow> D" and "g \<in> S \<rightarrow> T" and "inj_on g S" | |
| 193 | shows "f \<circ> (\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> D" | |
| 194 | proof - | |
| 195 | have "(\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> [T]\<^bsup>k\<^esup>" | |
| 196 | using assms by (simp add: nsets_image_funcset) | |
| 197 | then show ?thesis | |
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 198 | using f by fastforce | 
| 71452 | 199 | qed | 
| 200 | ||
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 201 | subsubsection \<open>Partition predicates\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 202 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 203 | definition partn :: "'a set \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'b set \<Rightarrow> bool" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 204 |   where "partn \<beta> \<alpha> \<gamma> \<delta> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  \<delta>. \<exists>H \<in> nsets \<beta> \<alpha>. \<exists>\<xi>\<in>\<delta>. f ` (nsets H \<gamma>) \<subseteq> {\<xi>}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 205 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 206 | definition partn_lst :: "'a set \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 207 |   where "partn_lst \<beta> \<alpha> \<gamma> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  {..<length \<alpha>}.
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 208 |               \<exists>i < length \<alpha>. \<exists>H \<in> nsets \<beta> (\<alpha>!i). f ` (nsets H \<gamma>) \<subseteq> {i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 209 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 210 | lemma partn_lst_greater_resource: | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 211 | fixes M::nat | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 212 |   assumes M: "partn_lst {..<M} \<alpha> \<gamma>" and "M \<le> N"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 213 |   shows "partn_lst {..<N} \<alpha> \<gamma>"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 214 | proof (clarsimp simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 215 | fix f | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 216 |   assume "f \<in> nsets {..<N} \<gamma> \<rightarrow> {..<length \<alpha>}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 217 |   then have "f \<in> nsets {..<M} \<gamma> \<rightarrow> {..<length \<alpha>}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 218 | by (meson Pi_anti_mono \<open>M \<le> N\<close> lessThan_subset_iff nsets_mono subsetD) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 219 |   then obtain i H where i: "i < length \<alpha>" and H: "H \<in> nsets {..<M} (\<alpha> ! i)" and subi: "f ` nsets H \<gamma> \<subseteq> {i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 220 | using M partn_lst_def by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 221 |   have "H \<in> nsets {..<N} (\<alpha> ! i)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 222 | using \<open>M \<le> N\<close> H by (auto simp: nsets_def subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 223 |   then show "\<exists>i<length \<alpha>. \<exists>H\<in>nsets {..<N} (\<alpha> ! i). f ` nsets H \<gamma> \<subseteq> {i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 224 | using i subi by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 225 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 226 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 227 | lemma partn_lst_fewer_colours: | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 228 | assumes major: "partn_lst \<beta> (n#\<alpha>) \<gamma>" and "n \<ge> \<gamma>" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 229 | shows "partn_lst \<beta> \<alpha> \<gamma>" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 230 | proof (clarsimp simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 231 | fix f :: "'a set \<Rightarrow> nat" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 232 |   assume f: "f \<in> [\<beta>]\<^bsup>\<gamma>\<^esup> \<rightarrow> {..<length \<alpha>}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 233 | then obtain i H where i: "i < Suc (length \<alpha>)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 234 | and H: "H \<in> [\<beta>]\<^bsup>((n # \<alpha>) ! i)\<^esup>" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 235 | and hom: "\<forall>x\<in>[H]\<^bsup>\<gamma>\<^esup>. Suc (f x) = i" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 236 | using \<open>n \<ge> \<gamma>\<close> major [unfolded partn_lst_def, rule_format, of "Suc o f"] | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 237 | by (fastforce simp: image_subset_iff nsets_eq_empty_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 238 |   show "\<exists>i<length \<alpha>. \<exists>H\<in>nsets \<beta> (\<alpha> ! i). f ` [H]\<^bsup>\<gamma>\<^esup> \<subseteq> {i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 239 | proof (cases i) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 240 | case 0 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 241 |     then have "[H]\<^bsup>\<gamma>\<^esup> = {}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 242 | using hom by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 243 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 244 | using 0 H \<open>n \<ge> \<gamma>\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 245 | by (simp add: nsets_eq_empty_iff) (simp add: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 246 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 247 | case (Suc i') | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 248 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 249 | using i H hom by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 250 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 251 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 252 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 253 | lemma partn_lst_eq_partn: "partn_lst {..<n} [m,m] 2 = partn {..<n} m 2 {..<2::nat}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 254 | apply (simp add: partn_lst_def partn_def numeral_2_eq_2) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 255 | by (metis less_2_cases numeral_2_eq_2 lessThan_iff nth_Cons_0 nth_Cons_Suc) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 256 | |
| 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 257 | |
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 258 | subsection \<open>Finite versions of Ramsey's theorem\<close> | 
| 40695 | 259 | |
| 65075 | 260 | text \<open> | 
| 261 | To distinguish the finite and infinite ones, lower and upper case | |
| 262 | names are used. | |
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 263 | \<close> | 
| 40695 | 264 | |
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 265 | subsubsection \<open>Trivial cases\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 266 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 267 | text \<open>Vacuous, since we are dealing with 0-sets!\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 268 | lemma ramsey0: "\<exists>N::nat. partn_lst {..<N} [q1,q2] 0"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 269 | by (force simp: partn_lst_def ex_in_conv less_Suc_eq nsets_eq_empty_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 270 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 271 | text \<open>Just the pigeon hole principle, since we are dealing with 1-sets\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 272 | lemma ramsey1: "\<exists>N::nat. partn_lst {..<N} [q0,q1] 1"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 273 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 274 |   have "\<exists>i<Suc (Suc 0). \<exists>H\<in>nsets {..<Suc (q0 + q1)} ([q0, q1] ! i). f ` nsets H (Suc 0) \<subseteq> {i}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 275 |     if "f \<in> nsets {..<Suc (q0 + q1)} (Suc 0) \<rightarrow> {..<Suc (Suc 0)}" for f
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 276 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 277 |     define A where "A \<equiv> \<lambda>i. {q. q \<le> q0+q1 \<and> f {q} = i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 278 |     have "A 0 \<union> A 1 = {..q0 + q1}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 279 | using that by (auto simp: A_def PiE_iff nsets_one lessThan_Suc_atMost le_Suc_eq) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 280 |     moreover have "A 0 \<inter> A 1 = {}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 281 | by (auto simp: A_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 282 | ultimately have "q0 + q1 \<le> card (A 0) + card (A 1)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 283 | by (metis card_Un_le card_atMost eq_imp_le le_SucI le_trans) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 284 | then consider "card (A 0) \<ge> q0" | "card (A 1) \<ge> q1" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 285 | by linarith | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 286 | then obtain i where "i < Suc (Suc 0)" "card (A i) \<ge> [q0, q1] ! i" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 287 | by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 288 | then obtain B where "B \<subseteq> A i" "card B = [q0, q1] ! i" "finite B" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 289 | by (meson obtain_subset_with_card_n) | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 290 |     then have "B \<in> nsets {..<Suc (q0 + q1)} ([q0, q1] ! i) \<and> f ` nsets B (Suc 0) \<subseteq> {i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 291 | by (auto simp: A_def nsets_def card_1_singleton_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 292 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 293 | using \<open>i < Suc (Suc 0)\<close> by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 294 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 295 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 296 | by (clarsimp simp: partn_lst_def) blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 297 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 298 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 299 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 300 | subsubsection \<open>Ramsey's theorem with two colours and arbitrary exponents (hypergraph version)\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 301 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 302 | proposition ramsey2_full: "\<exists>N::nat. partn_lst {..<N} [q1,q2] r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 303 | proof (induction r arbitrary: q1 q2) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 304 | case 0 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 305 | then show ?case | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 306 | by (simp add: ramsey0) | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 307 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 308 | case (Suc r) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 309 | note outer = this | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 310 | show ?case | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 311 | proof (cases "r = 0") | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 312 | case True | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 313 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 314 | using ramsey1 by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 315 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 316 | case False | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 317 | then have "r > 0" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 318 | by simp | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 319 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 320 | using Suc.prems | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 321 | proof (induct k \<equiv> "q1 + q2" arbitrary: q1 q2) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 322 | case 0 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 323 | show ?case | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 324 | proof | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 325 |         show "partn_lst {..<1::nat} [q1, q2] (Suc r)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 326 | using nsets_empty_iff subset_insert 0 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 327 | by (fastforce simp: partn_lst_def funcset_to_empty_iff nsets_eq_empty image_subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 328 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 329 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 330 | case (Suc k) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 331 | consider "q1 = 0 \<or> q2 = 0" | "q1 \<noteq> 0" "q2 \<noteq> 0" by auto | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 332 | then show ?case | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 333 | proof cases | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 334 | case 1 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 335 |         then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 336 | unfolding partn_lst_def using \<open>r > 0\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 337 | by (fastforce simp add: nsets_empty_iff nsets_singleton_iff lessThan_Suc) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 338 | then show ?thesis by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 339 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 340 | case 2 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 341 | with Suc have "k = (q1 - 1) + q2" "k = q1 + (q2 - 1)" by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 342 |         then obtain p1 p2::nat where  p1: "partn_lst {..<p1} [q1-1,q2] (Suc r)" and p2: "partn_lst {..<p2} [q1,q2-1] (Suc r)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 343 | using Suc.hyps by blast | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 344 |         then obtain p::nat where p: "partn_lst {..<p} [p1,p2] r"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 345 | using outer Suc.prems by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 346 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 347 | proof (intro exI conjI) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 348 |           have "\<exists>i<Suc (Suc 0). \<exists>H\<in>nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \<subseteq> {i}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 349 |             if f: "f \<in> nsets {..p} (Suc r) \<rightarrow> {..<Suc (Suc 0)}" for f
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 350 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 351 | define g where "g \<equiv> \<lambda>R. f (insert p R)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 352 |             have "f (insert p i) \<in> {..<Suc (Suc 0)}" if "i \<in> nsets {..<p} r" for i
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 353 | using that card_insert_if by (fastforce simp: nsets_def intro!: Pi_mem [OF f]) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 354 |             then have g: "g \<in> nsets {..<p} r \<rightarrow> {..<Suc (Suc 0)}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 355 | by (force simp: g_def PiE_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 356 |             then obtain i U where i: "i < Suc (Suc 0)" and gi: "g ` nsets U r \<subseteq> {i}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 357 |               and U: "U \<in> nsets {..<p} ([p1, p2] ! i)"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 358 | using p by (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 359 |             then have Usub: "U \<subseteq> {..<p}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 360 | by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 361 | consider (izero) "i = 0" | (ione) "i = Suc 0" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 362 | using i by linarith | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 363 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 364 | proof cases | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 365 | case izero | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 366 |               then have "U \<in> nsets {..<p} p1"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 367 | using U by simp | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 368 |               then obtain u where u: "bij_betw u {..<p1} U"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 369 | using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp add: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 370 |               have u_nsets: "u ` X \<in> nsets {..p} n" if "X \<in> nsets {..<p1} n" for X n
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 371 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 372 | have "inj_on u X" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 373 | using u that bij_betw_imp_inj_on inj_on_subset by (force simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 374 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 375 | using Usub u that bij_betwE | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 376 | by (fastforce simp add: nsets_def card_image) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 377 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 378 | define h where "h \<equiv> \<lambda>R. f (u ` R)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 379 |               have "h \<in> nsets {..<p1} (Suc r) \<rightarrow> {..<Suc (Suc 0)}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 380 | unfolding h_def using f u_nsets by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 381 |               then obtain j V where j: "j <Suc (Suc 0)" and hj: "h ` nsets V (Suc r) \<subseteq> {j}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 382 |                 and V: "V \<in> nsets {..<p1} ([q1 - Suc 0, q2] ! j)"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 383 | using p1 by (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 384 |               then have Vsub: "V \<subseteq> {..<p1}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 385 | by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 386 |               have invinv_eq: "u ` inv_into {..<p1} u ` X = X" if "X \<subseteq> u ` {..<p1}" for X
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 387 | by (simp add: image_inv_into_cancel that) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 388 | let ?W = "insert p (u ` V)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 389 | consider (jzero) "j = 0" | (jone) "j = Suc 0" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 390 | using j by linarith | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 391 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 392 | proof cases | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 393 | case jzero | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 394 |                 then have "V \<in> nsets {..<p1} (q1 - Suc 0)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 395 | using V by simp | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 396 |                 then have "u ` V \<in> nsets {..<p} (q1 - Suc 0)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 397 | using u_nsets [of _ "q1 - Suc 0"] nsets_mono [OF Vsub] Usub u | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 398 | unfolding bij_betw_def nsets_def | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 399 | by (fastforce elim!: subsetD) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 400 |                 then have inq1: "?W \<in> nsets {..p} q1"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 401 | unfolding nsets_def using \<open>q1 \<noteq> 0\<close> card_insert_if by fastforce | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 402 |                 have invu_nsets: "inv_into {..<p1} u ` X \<in> nsets V r"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 403 | if "X \<in> nsets (u ` V) r" for X r | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 404 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 405 | have "X \<subseteq> u ` V \<and> finite X \<and> card X = r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 406 | using nsets_def that by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 407 |                   then have [simp]: "card (inv_into {..<p1} u ` X) = card X"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 408 | by (meson Vsub bij_betw_def bij_betw_inv_into card_image image_mono inj_on_subset u) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 409 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 410 | using that u Vsub by (fastforce simp: nsets_def bij_betw_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 411 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 412 | have "f X = i" if X: "X \<in> nsets ?W (Suc r)" for X | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 413 | proof (cases "p \<in> X") | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 414 | case True | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 415 |                   then have Xp: "X - {p} \<in> nsets (u ` V) r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 416 | using X by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 417 | moreover have "u ` V \<subseteq> U" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 418 | using Vsub bij_betwE u by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 419 |                   ultimately have "X - {p} \<in> nsets U r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 420 | by (meson in_mono nsets_mono) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 421 |                   then have "g (X - {p}) = i"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 422 | using gi by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 423 | have "f X = i" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 424 |                     using gi True \<open>X - {p} \<in> nsets U r\<close> insert_Diff
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 425 | by (fastforce simp add: g_def image_subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 426 | then show ?thesis | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 427 |                     by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>)
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 428 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 429 | case False | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 430 | then have Xim: "X \<in> nsets (u ` V) (Suc r)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 431 | using X by (auto simp: nsets_def subset_insert) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 432 |                   then have "u ` inv_into {..<p1} u ` X = X"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 433 | using Vsub bij_betw_imp_inj_on u | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 434 | by (fastforce simp: nsets_def image_mono invinv_eq subset_trans) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 435 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 436 | using izero jzero hj Xim invu_nsets unfolding h_def | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 437 | by (fastforce simp add: image_subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 438 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 439 |                 moreover have "insert p (u ` V) \<in> nsets {..p} q1"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 440 | by (simp add: izero inq1) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 441 | ultimately show ?thesis | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 442 | by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 443 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 444 | case jone | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 445 |                 then have "u ` V \<in> nsets {..p} q2"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 446 | using V u_nsets by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 447 |                 moreover have "f ` nsets (u ` V) (Suc r) \<subseteq> {j}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 448 | using hj | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 449 | by (force simp add: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 450 | ultimately show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 451 | using jone not_less_eq by fastforce | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 452 | qed | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 453 | next | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 454 | case ione | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 455 |               then have "U \<in> nsets {..<p} p2"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 456 | using U by simp | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 457 |               then obtain u where u: "bij_betw u {..<p2} U"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 458 | using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp add: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 459 |               have u_nsets: "u ` X \<in> nsets {..p} n" if "X \<in> nsets {..<p2} n" for X n
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 460 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 461 | have "inj_on u X" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 462 | using u that bij_betw_imp_inj_on inj_on_subset by (force simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 463 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 464 | using Usub u that bij_betwE | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 465 | by (fastforce simp add: nsets_def card_image) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 466 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 467 | define h where "h \<equiv> \<lambda>R. f (u ` R)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 468 |               have "h \<in> nsets {..<p2} (Suc r) \<rightarrow> {..<Suc (Suc 0)}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 469 | unfolding h_def using f u_nsets by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 470 |               then obtain j V where j: "j <Suc (Suc 0)" and hj: "h ` nsets V (Suc r) \<subseteq> {j}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 471 |                 and V: "V \<in> nsets {..<p2} ([q1, q2 - Suc 0] ! j)"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 472 | using p2 by (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 473 |               then have Vsub: "V \<subseteq> {..<p2}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 474 | by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 475 |               have invinv_eq: "u ` inv_into {..<p2} u ` X = X" if "X \<subseteq> u ` {..<p2}" for X
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 476 | by (simp add: image_inv_into_cancel that) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 477 | let ?W = "insert p (u ` V)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 478 | consider (jzero) "j = 0" | (jone) "j = Suc 0" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 479 | using j by linarith | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 480 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 481 | proof cases | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 482 | case jone | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 483 |                 then have "V \<in> nsets {..<p2} (q2 - Suc 0)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 484 | using V by simp | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 485 |                 then have "u ` V \<in> nsets {..<p} (q2 - Suc 0)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 486 | using u_nsets [of _ "q2 - Suc 0"] nsets_mono [OF Vsub] Usub u | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 487 | unfolding bij_betw_def nsets_def | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 488 | by (fastforce elim!: subsetD) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 489 |                 then have inq1: "?W \<in> nsets {..p} q2"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 490 | unfolding nsets_def using \<open>q2 \<noteq> 0\<close> card_insert_if by fastforce | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 491 |                 have invu_nsets: "inv_into {..<p2} u ` X \<in> nsets V r"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 492 | if "X \<in> nsets (u ` V) r" for X r | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 493 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 494 | have "X \<subseteq> u ` V \<and> finite X \<and> card X = r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 495 | using nsets_def that by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 496 |                   then have [simp]: "card (inv_into {..<p2} u ` X) = card X"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 497 | by (meson Vsub bij_betw_def bij_betw_inv_into card_image image_mono inj_on_subset u) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 498 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 499 | using that u Vsub by (fastforce simp: nsets_def bij_betw_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 500 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 501 | have "f X = i" if X: "X \<in> nsets ?W (Suc r)" for X | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 502 | proof (cases "p \<in> X") | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 503 | case True | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 504 |                   then have Xp: "X - {p} \<in> nsets (u ` V) r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 505 | using X by (auto simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 506 | moreover have "u ` V \<subseteq> U" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 507 | using Vsub bij_betwE u by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 508 |                   ultimately have "X - {p} \<in> nsets U r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 509 | by (meson in_mono nsets_mono) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 510 |                   then have "g (X - {p}) = i"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 511 | using gi by blast | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 512 | have "f X = i" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 513 |                     using gi True \<open>X - {p} \<in> nsets U r\<close> insert_Diff
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 514 | by (fastforce simp add: g_def image_subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 515 | then show ?thesis | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 516 |                     by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>)
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 517 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 518 | case False | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 519 | then have Xim: "X \<in> nsets (u ` V) (Suc r)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 520 | using X by (auto simp: nsets_def subset_insert) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 521 |                   then have "u ` inv_into {..<p2} u ` X = X"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 522 | using Vsub bij_betw_imp_inj_on u | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 523 | by (fastforce simp: nsets_def image_mono invinv_eq subset_trans) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 524 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 525 | using ione jone hj Xim invu_nsets unfolding h_def | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 526 | by (fastforce simp add: image_subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 527 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 528 |                 moreover have "insert p (u ` V) \<in> nsets {..p} q2"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 529 | by (simp add: ione inq1) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 530 | ultimately show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 531 | by (metis ione image_subsetI insertI1 lessI nth_Cons_0 nth_Cons_Suc) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 532 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 533 | case jzero | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 534 |                 then have "u ` V \<in> nsets {..p} q1"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 535 | using V u_nsets by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 536 |                 moreover have "f ` nsets (u ` V) (Suc r) \<subseteq> {j}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 537 | using hj | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 538 | apply (clarsimp simp add: h_def image_subset_iff nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 539 | by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 540 | ultimately show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 541 | using jzero not_less_eq by fastforce | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 542 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 543 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 544 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 545 |           then show "partn_lst {..<Suc p} [q1,q2] (Suc r)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 546 | using lessThan_Suc lessThan_Suc_atMost by (auto simp: partn_lst_def insert_commute) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 547 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 548 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 549 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 550 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 551 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 552 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 553 | subsubsection \<open>Full Ramsey's theorem with multiple colours and arbitrary exponents\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 554 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 555 | theorem ramsey_full: "\<exists>N::nat. partn_lst {..<N} qs r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 556 | proof (induction k \<equiv> "length qs" arbitrary: qs) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 557 | case 0 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 558 | then show ?case | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 559 | by (rule_tac x=" r" in exI) (simp add: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 560 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 561 | case (Suc k) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 562 | note IH = this | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 563 | show ?case | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 564 | proof (cases k) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 565 | case 0 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 566 | with Suc obtain q where "qs = [q]" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 567 | by (metis length_0_conv length_Suc_conv) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 568 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 569 | by (rule_tac x=q in exI) (auto simp: partn_lst_def funcset_to_empty_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 570 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 571 | case (Suc k') | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 572 | then obtain q1 q2 l where qs: "qs = q1#q2#l" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 573 | by (metis Suc.hyps(2) length_Suc_conv) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 574 |     then obtain q::nat where q: "partn_lst {..<q} [q1,q2] r"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 575 | using ramsey2_full by blast | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 576 |     then obtain p::nat where p: "partn_lst {..<p} (q#l) r"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 577 | using IH \<open>qs = q1 # q2 # l\<close> by fastforce | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 578 | have keq: "Suc (length l) = k" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 579 | using IH qs by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 580 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 581 | proof (intro exI conjI) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 582 |       show "partn_lst {..<p} qs r"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 583 | proof (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 584 | fix f | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 585 |         assume f: "f \<in> nsets {..<p} r \<rightarrow> {..<length qs}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 586 | define g where "g \<equiv> \<lambda>X. if f X < Suc (Suc 0) then 0 else f X - Suc 0" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 587 |         have "g \<in> nsets {..<p} r \<rightarrow> {..<k}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 588 | unfolding g_def using f Suc IH | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 589 | by (auto simp: Pi_def not_less) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 590 |         then obtain i U where i: "i < k" and gi: "g ` nsets U r \<subseteq> {i}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 591 |                 and U: "U \<in> nsets {..<p} ((q#l) ! i)"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 592 | using p keq by (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 593 |         show "\<exists>i<length qs. \<exists>H\<in>nsets {..<p} (qs ! i). f ` nsets H r \<subseteq> {i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 594 | proof (cases "i = 0") | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 595 | case True | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 596 |           then have "U \<in> nsets {..<p} q" and f01: "f ` nsets U r \<subseteq> {0, Suc 0}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 597 | using U gi unfolding g_def by (auto simp: image_subset_iff) | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 598 |           then obtain u where u: "bij_betw u {..<q} U"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 599 | using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp add: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 600 |           then have Usub: "U \<subseteq> {..<p}"
 | 
| 73655 
26a1d66b9077
tuned proofs --- avoid z3, which is absent on arm64-linux;
 wenzelm parents: 
72231diff
changeset | 601 | by (smt (verit) U mem_Collect_eq nsets_def) | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 602 |           have u_nsets: "u ` X \<in> nsets {..<p} n" if "X \<in> nsets {..<q} n" for X n
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 603 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 604 | have "inj_on u X" | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 605 | using u that bij_betw_imp_inj_on inj_on_subset | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 606 | by (force simp: nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 607 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 608 | using Usub u that bij_betwE | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 609 | by (fastforce simp add: nsets_def card_image) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 610 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 611 | define h where "h \<equiv> \<lambda>X. f (u ` X)" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 612 |           have "f (u ` X) < Suc (Suc 0)" if "X \<in> nsets {..<q} r" for X
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 613 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 614 | have "u ` X \<in> nsets U r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 615 | using u u_nsets that by (auto simp: nsets_def bij_betwE subset_eq) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 616 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 617 | using f01 by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 618 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 619 |           then have "h \<in> nsets {..<q} r \<rightarrow> {..<Suc (Suc 0)}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 620 | unfolding h_def by blast | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 621 |           then obtain j V where j: "j < Suc (Suc 0)" and hj: "h ` nsets V r \<subseteq> {j}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 622 |             and V: "V \<in> nsets {..<q} ([q1,q2] ! j)"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 623 | using q by (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 624 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 625 | proof (intro exI conjI bexI) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 626 | show "j < length qs" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 627 | using Suc Suc.hyps(2) j by linarith | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 628 | have "nsets (u ` V) r \<subseteq> (\<lambda>x. (u ` x)) ` nsets V r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 629 | apply (clarsimp simp add: nsets_def image_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 630 | by (metis card_eq_0_iff card_image image_is_empty subset_image_inj) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 631 | then have "f ` nsets (u ` V) r \<subseteq> h ` nsets V r" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 632 | by (auto simp: h_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 633 |             then show "f ` nsets (u ` V) r \<subseteq> {j}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 634 | using hj by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 635 |             show "(u ` V) \<in> nsets {..<p} (qs ! j)"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 636 | using V j less_2_cases numeral_2_eq_2 qs u_nsets by fastforce | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 637 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 638 | next | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 639 | case False | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 640 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 641 | proof (intro exI conjI bexI) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 642 | show "Suc i < length qs" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 643 | using Suc.hyps(2) i by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 644 |             show "f ` nsets U r \<subseteq> {Suc i}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 645 | using i gi False | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 646 | apply (auto simp: g_def image_subset_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 647 | by (metis Suc_lessD Suc_pred g_def gi image_subset_iff not_less_eq singleton_iff) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 648 |             show "U \<in> nsets {..<p} (qs ! (Suc i))"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 649 | using False U qs by auto | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 650 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 651 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 652 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 653 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 654 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 655 | qed | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 656 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 657 | subsubsection \<open>Simple graph version\<close> | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 658 | |
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 659 | text \<open>This is the most basic version in terms of cliques and independent | 
| 65075 | 660 | sets, i.e. the version for graphs and 2 colours. | 
| 661 | \<close> | |
| 40695 | 662 | |
| 65075 | 663 | definition "clique V E \<longleftrightarrow> (\<forall>v\<in>V. \<forall>w\<in>V. v \<noteq> w \<longrightarrow> {v, w} \<in> E)"
 | 
| 664 | definition "indep V E \<longleftrightarrow> (\<forall>v\<in>V. \<forall>w\<in>V. v \<noteq> w \<longrightarrow> {v, w} \<notin> E)"
 | |
| 40695 | 665 | |
| 666 | lemma ramsey2: | |
| 65075 | 667 | "\<exists>r\<ge>1. \<forall>(V::'a set) (E::'a set set). finite V \<and> card V \<ge> r \<longrightarrow> | 
| 668 | (\<exists>R \<subseteq> V. card R = m \<and> clique R E \<or> card R = n \<and> indep R E)" | |
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 669 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 670 |   obtain N where "N \<ge> Suc 0" and N: "partn_lst {..<N} [m,n] 2"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 671 | using ramsey2_full nat_le_linear partn_lst_greater_resource by blast | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 672 | have "\<exists>R\<subseteq>V. card R = m \<and> clique R E \<or> card R = n \<and> indep R E" | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 673 | if "finite V" "N \<le> card V" for V :: "'a set" and E :: "'a set set" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 674 | proof - | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 675 | from that | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 676 |     obtain v where u: "inj_on v {..<N}" "v ` {..<N} \<subseteq> V"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 677 | by (metis card_le_inj card_lessThan finite_lessThan) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 678 | define f where "f \<equiv> \<lambda>e. if v ` e \<in> E then 0 else Suc 0" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 679 |     have f: "f \<in> nsets {..<N} 2 \<rightarrow> {..<Suc (Suc 0)}"
 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 680 | by (simp add: f_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 681 |     then obtain i U where i: "i < 2" and gi: "f ` nsets U 2 \<subseteq> {i}"
 | 
| 71472 
c213d067e60f
Moved a number of general-purpose lemmas into HOL
 paulson <lp15@cam.ac.uk> parents: 
71464diff
changeset | 682 |       and U: "U \<in> nsets {..<N} ([m,n] ! i)"
 | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 683 | using N numeral_2_eq_2 by (auto simp: partn_lst_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 684 | show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 685 | proof (intro exI conjI) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 686 | show "v ` U \<subseteq> V" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 687 | using U u by (auto simp: image_subset_iff nsets_def) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 688 | show "card (v ` U) = m \<and> clique (v ` U) E \<or> card (v ` U) = n \<and> indep (v ` U) E" | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 689 | using i unfolding numeral_2_eq_2 | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 690 | using gi U u | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 691 | apply (simp add: image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 692 | apply (auto simp: f_def nsets_def card_image inj_on_subset split: if_split_asm) | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 693 | done | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 694 | qed | 
| 40695 | 695 | qed | 
| 71259 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 696 | then show ?thesis | 
| 
09aee7f5b447
Ramsey with multiple colours and arbitrary exponents
 paulson <lp15@cam.ac.uk> parents: 
71083diff
changeset | 697 | using \<open>Suc 0 \<le> N\<close> by auto | 
| 40695 | 698 | qed | 
| 699 | ||
| 700 | ||
| 60500 | 701 | subsection \<open>Preliminaries\<close> | 
| 19944 | 702 | |
| 60500 | 703 | subsubsection \<open>``Axiom'' of Dependent Choice\<close> | 
| 19944 | 704 | |
| 65075 | 705 | primrec choice :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a"
 | 
| 706 | where \<comment> \<open>An integer-indexed chain of choices\<close> | |
| 707 | choice_0: "choice P r 0 = (SOME x. P x)" | |
| 708 | | choice_Suc: "choice P r (Suc n) = (SOME y. P y \<and> (choice P r n, y) \<in> r)" | |
| 19944 | 709 | |
| 46575 | 710 | lemma choice_n: | 
| 19944 | 711 | assumes P0: "P x0" | 
| 65075 | 712 | and Pstep: "\<And>x. P x \<Longrightarrow> \<exists>y. P y \<and> (x, y) \<in> r" | 
| 19944 | 713 | shows "P (choice P r n)" | 
| 19948 | 714 | proof (induct n) | 
| 65075 | 715 | case 0 | 
| 716 | show ?case by (force intro: someI P0) | |
| 19948 | 717 | next | 
| 65075 | 718 | case Suc | 
| 719 | then show ?case by (auto intro: someI2_ex [OF Pstep]) | |
| 19948 | 720 | qed | 
| 19944 | 721 | |
| 46575 | 722 | lemma dependent_choice: | 
| 19944 | 723 | assumes trans: "trans r" | 
| 65075 | 724 | and P0: "P x0" | 
| 725 | and Pstep: "\<And>x. P x \<Longrightarrow> \<exists>y. P y \<and> (x, y) \<in> r" | |
| 726 | obtains f :: "nat \<Rightarrow> 'a" where "\<And>n. P (f n)" and "\<And>n m. n < m \<Longrightarrow> (f n, f m) \<in> r" | |
| 19954 | 727 | proof | 
| 728 | fix n | |
| 65075 | 729 | show "P (choice P r n)" | 
| 730 | by (blast intro: choice_n [OF P0 Pstep]) | |
| 19944 | 731 | next | 
| 65075 | 732 | fix n m :: nat | 
| 733 | assume "n < m" | |
| 734 | from Pstep [OF choice_n [OF P0 Pstep]] have "(choice P r k, choice P r (Suc k)) \<in> r" for k | |
| 19944 | 735 | by (auto intro: someI2_ex) | 
| 65075 | 736 | then show "(choice P r n, choice P r m) \<in> r" | 
| 737 | by (auto intro: less_Suc_induct [OF \<open>n < m\<close>] transD [OF trans]) | |
| 19954 | 738 | qed | 
| 19944 | 739 | |
| 740 | ||
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 741 | subsubsection \<open>Partition functions\<close> | 
| 19944 | 742 | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 743 | definition part_fn :: "nat \<Rightarrow> nat \<Rightarrow> 'a set \<Rightarrow> ('a set \<Rightarrow> nat) \<Rightarrow> bool"
 | 
| 69593 | 744 | \<comment> \<open>the function \<^term>\<open>f\<close> partitions the \<^term>\<open>r\<close>-subsets of the typically | 
| 745 | infinite set \<^term>\<open>Y\<close> into \<^term>\<open>s\<close> distinct categories.\<close> | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 746 |   where "part_fn r s Y f \<longleftrightarrow> (f \<in> nsets Y r \<rightarrow> {..<s})"
 | 
| 19944 | 747 | |
| 69593 | 748 | text \<open>For induction, we decrease the value of \<^term>\<open>r\<close> in partitions.\<close> | 
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 749 | lemma part_fn_Suc_imp_part_fn: | 
| 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 750 |   "\<lbrakk>infinite Y; part_fn (Suc r) s Y f; y \<in> Y\<rbrakk> \<Longrightarrow> part_fn r s (Y - {y}) (\<lambda>u. f (insert y u))"
 | 
| 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 751 | by (simp add: part_fn_def nsets_def Pi_def subset_Diff_insert) | 
| 19944 | 752 | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 753 | lemma part_fn_subset: "part_fn r s YY f \<Longrightarrow> Y \<subseteq> YY \<Longrightarrow> part_fn r s Y f" | 
| 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 754 | unfolding part_fn_def nsets_def by blast | 
| 46575 | 755 | |
| 19944 | 756 | |
| 60500 | 757 | subsection \<open>Ramsey's Theorem: Infinitary Version\<close> | 
| 19944 | 758 | |
| 46575 | 759 | lemma Ramsey_induction: | 
| 65075 | 760 | fixes s r :: nat | 
| 761 | and YY :: "'a set" | |
| 762 | and f :: "'a set \<Rightarrow> nat" | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 763 | assumes "infinite YY" "part_fn r s YY f" | 
| 65075 | 764 | shows "\<exists>Y' t'. Y' \<subseteq> YY \<and> infinite Y' \<and> t' < s \<and> (\<forall>X. X \<subseteq> Y' \<and> finite X \<and> card X = r \<longrightarrow> f X = t')" | 
| 765 | using assms | |
| 766 | proof (induct r arbitrary: YY f) | |
| 19944 | 767 | case 0 | 
| 65075 | 768 | then show ?case | 
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 769 | by (auto simp add: part_fn_def card_eq_0_iff cong: conj_cong) | 
| 19944 | 770 | next | 
| 46575 | 771 | case (Suc r) | 
| 19944 | 772 | show ?case | 
| 773 | proof - | |
| 65075 | 774 | from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY" | 
| 775 | by blast | |
| 776 |     let ?ramr = "{((y, Y, t), (y', Y', t')). y' \<in> Y \<and> Y' \<subseteq> Y}"
 | |
| 777 | let ?propr = "\<lambda>(y, Y, t). | |
| 778 | y \<in> YY \<and> y \<notin> Y \<and> Y \<subseteq> YY \<and> infinite Y \<and> t < s | |
| 779 | \<and> (\<forall>X. X\<subseteq>Y \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert y) X = t)" | |
| 780 |     from Suc.prems have infYY': "infinite (YY - {yy})" by auto
 | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 781 |     from Suc.prems have partf': "part_fn r s (YY - {yy}) (f \<circ> insert yy)"
 | 
| 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 782 | by (simp add: o_def part_fn_Suc_imp_part_fn yy) | 
| 46575 | 783 | have transr: "trans ?ramr" by (force simp add: trans_def) | 
| 19944 | 784 | from Suc.hyps [OF infYY' partf'] | 
| 65075 | 785 |     obtain Y0 and t0 where "Y0 \<subseteq> YY - {yy}" "infinite Y0" "t0 < s"
 | 
| 786 | "X \<subseteq> Y0 \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yy) X = t0" for X | |
| 787 | by blast | |
| 788 | with yy have propr0: "?propr(yy, Y0, t0)" by blast | |
| 789 | have proprstep: "\<exists>y. ?propr y \<and> (x, y) \<in> ?ramr" if x: "?propr x" for x | |
| 790 | proof (cases x) | |
| 791 | case (fields yx Yx tx) | |
| 792 | with x obtain yx' where yx': "yx' \<in> Yx" | |
| 793 | by (blast dest: infinite_imp_nonempty) | |
| 794 |       from fields x have infYx': "infinite (Yx - {yx'})" by auto
 | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 795 |       with fields x yx' Suc.prems have partfx': "part_fn r s (Yx - {yx'}) (f \<circ> insert yx')"
 | 
| 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 796 | by (simp add: o_def part_fn_Suc_imp_part_fn part_fn_subset [where YY=YY and Y=Yx]) | 
| 65075 | 797 | from Suc.hyps [OF infYx' partfx'] obtain Y' and t' | 
| 798 |         where Y': "Y' \<subseteq> Yx - {yx'}" "infinite Y'" "t' < s"
 | |
| 799 | "X \<subseteq> Y' \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yx') X = t'" for X | |
| 46575 | 800 | by blast | 
| 65075 | 801 | from fields x Y' yx' have "?propr (yx', Y', t') \<and> (x, (yx', Y', t')) \<in> ?ramr" | 
| 802 | by blast | |
| 803 | then show ?thesis .. | |
| 19944 | 804 | qed | 
| 805 | from dependent_choice [OF transr propr0 proprstep] | |
| 65075 | 806 | obtain g where pg: "?propr (g n)" and rg: "n < m \<Longrightarrow> (g n, g m) \<in> ?ramr" for n m :: nat | 
| 63060 | 807 | by blast | 
| 65075 | 808 | let ?gy = "fst \<circ> g" | 
| 809 | let ?gt = "snd \<circ> snd \<circ> g" | |
| 19944 | 810 |     have rangeg: "\<exists>k. range ?gt \<subseteq> {..<k}"
 | 
| 811 | proof (intro exI subsetI) | |
| 812 | fix x | |
| 813 | assume "x \<in> range ?gt" | |
| 814 | then obtain n where "x = ?gt n" .. | |
| 815 |       with pg [of n] show "x \<in> {..<s}" by (cases "g n") auto
 | |
| 816 | qed | |
| 69661 | 817 | from rangeg have "finite (range ?gt)" | 
| 818 | by (simp add: finite_nat_iff_bounded) | |
| 65075 | 819 |     then obtain s' and n' where s': "s' = ?gt n'" and infeqs': "infinite {n. ?gt n = s'}"
 | 
| 54580 | 820 | by (rule inf_img_fin_domE) (auto simp add: vimage_def intro: infinite_UNIV_nat) | 
| 19944 | 821 | with pg [of n'] have less': "s'<s" by (cases "g n'") auto | 
| 822 | have inj_gy: "inj ?gy" | |
| 823 | proof (rule linorder_injI) | |
| 65075 | 824 | fix m m' :: nat | 
| 825 | assume "m < m'" | |
| 826 | from rg [OF this] pg [of m] show "?gy m \<noteq> ?gy m'" | |
| 827 | by (cases "g m", cases "g m'") auto | |
| 19944 | 828 | qed | 
| 829 | show ?thesis | |
| 830 | proof (intro exI conjI) | |
| 65075 | 831 |       from pg show "?gy ` {n. ?gt n = s'} \<subseteq> YY"
 | 
| 46575 | 832 | by (auto simp add: Let_def split_beta) | 
| 65075 | 833 |       from infeqs' show "infinite (?gy ` {n. ?gt n = s'})"
 | 
| 46575 | 834 | by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) | 
| 19944 | 835 | show "s' < s" by (rule less') | 
| 65075 | 836 |       show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} \<and> finite X \<and> card X = Suc r \<longrightarrow> f X = s'"
 | 
| 19944 | 837 | proof - | 
| 65075 | 838 | have "f X = s'" | 
| 839 |           if X: "X \<subseteq> ?gy ` {n. ?gt n = s'}"
 | |
| 840 | and cardX: "finite X" "card X = Suc r" | |
| 841 | for X | |
| 842 | proof - | |
| 843 |           from X obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA"
 | |
| 844 | by (auto simp add: subset_image_iff) | |
| 845 |           with cardX have "AA \<noteq> {}" by auto
 | |
| 846 | then have AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex) | |
| 847 | show ?thesis | |
| 848 | proof (cases "g (LEAST x. x \<in> AA)") | |
| 849 | case (fields ya Ya ta) | |
| 850 | with AAleast Xeq have ya: "ya \<in> X" by (force intro!: rev_image_eqI) | |
| 851 |             then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
 | |
| 852 | also have "\<dots> = ta" | |
| 853 | proof - | |
| 854 |               have *: "X - {ya} \<subseteq> Ya"
 | |
| 855 | proof | |
| 856 |                 fix x assume x: "x \<in> X - {ya}"
 | |
| 857 | then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA" | |
| 858 | by (auto simp add: Xeq) | |
| 859 | with fields x have "a' \<noteq> (LEAST x. x \<in> AA)" by auto | |
| 860 | with Least_le [of "\<lambda>x. x \<in> AA", OF a'] have "(LEAST x. x \<in> AA) < a'" | |
| 861 | by arith | |
| 862 | from xeq fields rg [OF this] show "x \<in> Ya" by auto | |
| 863 | qed | |
| 864 |               have "card (X - {ya}) = r"
 | |
| 865 | by (simp add: cardX ya) | |
| 866 | with pg [of "LEAST x. x \<in> AA"] fields cardX * show ?thesis | |
| 867 | by (auto simp del: insert_Diff_single) | |
| 868 | qed | |
| 869 | also from AA AAleast fields have "\<dots> = s'" by auto | |
| 870 | finally show ?thesis . | |
| 871 | qed | |
| 872 | qed | |
| 46575 | 873 | then show ?thesis by blast | 
| 874 | qed | |
| 875 | qed | |
| 19944 | 876 | qed | 
| 877 | qed | |
| 878 | ||
| 879 | ||
| 880 | theorem Ramsey: | |
| 65075 | 881 | fixes s r :: nat | 
| 882 | and Z :: "'a set" | |
| 883 | and f :: "'a set \<Rightarrow> nat" | |
| 19944 | 884 | shows | 
| 65075 | 885 | "\<lbrakk>infinite Z; | 
| 886 | \<forall>X. X \<subseteq> Z \<and> finite X \<and> card X = r \<longrightarrow> f X < s\<rbrakk> | |
| 887 | \<Longrightarrow> \<exists>Y t. Y \<subseteq> Z \<and> infinite Y \<and> t < s | |
| 888 | \<and> (\<forall>X. X \<subseteq> Y \<and> finite X \<and> card X = r \<longrightarrow> f X = t)" | |
| 71260 
308baf6b450a
corrected some confusing terminology / notation
 paulson <lp15@cam.ac.uk> parents: 
71259diff
changeset | 889 | by (blast intro: Ramsey_induction [unfolded part_fn_def nsets_def]) | 
| 19954 | 890 | |
| 891 | corollary Ramsey2: | |
| 65075 | 892 | fixes s :: nat | 
| 893 | and Z :: "'a set" | |
| 894 | and f :: "'a set \<Rightarrow> nat" | |
| 19954 | 895 | assumes infZ: "infinite Z" | 
| 65075 | 896 |     and part: "\<forall>x\<in>Z. \<forall>y\<in>Z. x \<noteq> y \<longrightarrow> f {x, y} < s"
 | 
| 897 |   shows "\<exists>Y t. Y \<subseteq> Z \<and> infinite Y \<and> t < s \<and> (\<forall>x\<in>Y. \<forall>y\<in>Y. x\<noteq>y \<longrightarrow> f {x, y} = t)"
 | |
| 19954 | 898 | proof - | 
| 65075 | 899 | from part have part2: "\<forall>X. X \<subseteq> Z \<and> finite X \<and> card X = 2 \<longrightarrow> f X < s" | 
| 900 | by (fastforce simp add: eval_nat_numeral card_Suc_eq) | |
| 901 | obtain Y t where *: | |
| 902 | "Y \<subseteq> Z" "infinite Y" "t < s" "(\<forall>X. X \<subseteq> Y \<and> finite X \<and> card X = 2 \<longrightarrow> f X = t)" | |
| 19954 | 903 | by (insert Ramsey [OF infZ part2]) auto | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
46575diff
changeset | 904 |   then have "\<forall>x\<in>Y. \<forall>y\<in>Y. x \<noteq> y \<longrightarrow> f {x, y} = t" by auto
 | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
46575diff
changeset | 905 | with * show ?thesis by iprover | 
| 19954 | 906 | qed | 
| 907 | ||
| 71848 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 908 | corollary Ramsey_nsets: | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 909 | fixes f :: "'a set \<Rightarrow> nat" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 910 |   assumes "infinite Z" "f ` nsets Z r \<subseteq> {..<s}"
 | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 911 |   obtains Y t where "Y \<subseteq> Z" "infinite Y" "t < s" "f ` nsets Y r \<subseteq> {t}"
 | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 912 | using Ramsey [of Z r f s] assms by (auto simp: nsets_def image_subset_iff) | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
71472diff
changeset | 913 | |
| 19954 | 914 | |
| 60500 | 915 | subsection \<open>Disjunctive Well-Foundedness\<close> | 
| 19954 | 916 | |
| 60500 | 917 | text \<open> | 
| 22367 | 918 | An application of Ramsey's theorem to program termination. See | 
| 58622 | 919 |   @{cite "Podelski-Rybalchenko"}.
 | 
| 60500 | 920 | \<close> | 
| 19954 | 921 | |
| 65075 | 922 | definition disj_wf :: "('a \<times> 'a) set \<Rightarrow> bool"
 | 
| 923 | where "disj_wf r \<longleftrightarrow> (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf (T i)) \<and> r = (\<Union>i<n. T i))" | |
| 19954 | 924 | |
| 65075 | 925 | definition transition_idx :: "(nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> ('a \<times> 'a) set) \<Rightarrow> nat set \<Rightarrow> nat"
 | 
| 926 |   where "transition_idx s T A = (LEAST k. \<exists>i j. A = {i, j} \<and> i < j \<and> (s j, s i) \<in> T k)"
 | |
| 19954 | 927 | |
| 928 | ||
| 929 | lemma transition_idx_less: | |
| 65075 | 930 | assumes "i < j" "(s j, s i) \<in> T k" "k < n" | 
| 931 |   shows "transition_idx s T {i, j} < n"
 | |
| 932 | proof - | |
| 933 |   from assms(1,2) have "transition_idx s T {i, j} \<le> k"
 | |
| 934 | by (simp add: transition_idx_def, blast intro: Least_le) | |
| 935 | with assms(3) show ?thesis by simp | |
| 936 | qed | |
| 19954 | 937 | |
| 938 | lemma transition_idx_in: | |
| 65075 | 939 | assumes "i < j" "(s j, s i) \<in> T k" | 
| 940 |   shows "(s j, s i) \<in> T (transition_idx s T {i, j})"
 | |
| 941 | using assms | |
| 942 | by (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR cong: conj_cong) (erule LeastI) | |
| 19954 | 943 | |
| 65075 | 944 | text \<open>To be equal to the union of some well-founded relations is equivalent | 
| 945 | to being the subset of such a union.\<close> | |
| 946 | lemma disj_wf: "disj_wf r \<longleftrightarrow> (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) \<and> r \<subseteq> (\<Union>i<n. T i))" | |
| 71083 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 947 | proof - | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 948 |   have *: "\<And>T n. \<lbrakk>\<forall>i<n. wf (T i); r \<subseteq> \<Union> (T ` {..<n})\<rbrakk>
 | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 949 | \<Longrightarrow> (\<forall>i<n. wf (T i \<inter> r)) \<and> r = (\<Union>i<n. T i \<inter> r)" | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 950 | by (force simp add: wf_Int1) | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 951 | show ?thesis | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 952 | unfolding disj_wf_def by auto (metis "*") | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 953 | qed | 
| 19954 | 954 | |
| 955 | theorem trans_disj_wf_implies_wf: | |
| 65075 | 956 | assumes "trans r" | 
| 957 | and "disj_wf r" | |
| 19954 | 958 | shows "wf r" | 
| 959 | proof (simp only: wf_iff_no_infinite_down_chain, rule notI) | |
| 960 | assume "\<exists>s. \<forall>i. (s (Suc i), s i) \<in> r" | |
| 961 | then obtain s where sSuc: "\<forall>i. (s (Suc i), s i) \<in> r" .. | |
| 65075 | 962 | from \<open>disj_wf r\<close> obtain T and n :: nat where wfT: "\<forall>k<n. wf(T k)" and r: "r = (\<Union>k<n. T k)" | 
| 963 | by (auto simp add: disj_wf_def) | |
| 964 | have s_in_T: "\<exists>k. (s j, s i) \<in> T k \<and> k<n" if "i < j" for i j | |
| 19954 | 965 | proof - | 
| 65075 | 966 | from \<open>i < j\<close> have "(s j, s i) \<in> r" | 
| 967 | proof (induct rule: less_Suc_induct) | |
| 968 | case 1 | |
| 969 | then show ?case by (simp add: sSuc) | |
| 970 | next | |
| 971 | case 2 | |
| 972 | with \<open>trans r\<close> show ?case | |
| 973 | unfolding trans_def by blast | |
| 19954 | 974 | qed | 
| 65075 | 975 | then show ?thesis by (auto simp add: r) | 
| 46575 | 976 | qed | 
| 71083 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 977 |   have "i < j \<Longrightarrow> transition_idx s T {i, j} < n" for i j
 | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 978 | using s_in_T transition_idx_less by blast | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 979 |   then have trless: "i \<noteq> j \<Longrightarrow> transition_idx s T {i, j} < n" for i j
 | 
| 
ce92360f0692
A slight tidying up of messy proof steps
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 980 | by (metis doubleton_eq_iff less_linear) | 
| 65075 | 981 | have "\<exists>K k. K \<subseteq> UNIV \<and> infinite K \<and> k < n \<and> | 
| 982 |       (\<forall>i\<in>K. \<forall>j\<in>K. i \<noteq> j \<longrightarrow> transition_idx s T {i, j} = k)"
 | |
| 54580 | 983 | by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat) | 
| 65075 | 984 | then obtain K and k where infK: "infinite K" and "k < n" | 
| 985 |     and allk: "\<forall>i\<in>K. \<forall>j\<in>K. i \<noteq> j \<longrightarrow> transition_idx s T {i, j} = k"
 | |
| 19954 | 986 | by auto | 
| 65075 | 987 | have "(s (enumerate K (Suc m)), s (enumerate K m)) \<in> T k" for m :: nat | 
| 988 | proof - | |
| 19954 | 989 | let ?j = "enumerate K (Suc m)" | 
| 990 | let ?i = "enumerate K m" | |
| 46575 | 991 | have ij: "?i < ?j" by (simp add: enumerate_step infK) | 
| 65075 | 992 | have "?j \<in> K" "?i \<in> K" by (simp_all add: enumerate_in_set infK) | 
| 993 |     with ij have k: "k = transition_idx s T {?i, ?j}" by (simp add: allk)
 | |
| 994 | from s_in_T [OF ij] obtain k' where "(s ?j, s ?i) \<in> T k'" "k'<n" by blast | |
| 995 | then show "(s ?j, s ?i) \<in> T k" by (simp add: k transition_idx_in ij) | |
| 19954 | 996 | qed | 
| 65075 | 997 | then have "\<not> wf (T k)" | 
| 73709 | 998 | unfolding wf_iff_no_infinite_down_chain by iprover | 
| 65075 | 999 | with wfT \<open>k < n\<close> show False by blast | 
| 19954 | 1000 | qed | 
| 1001 | ||
| 19944 | 1002 | end |