src/HOL/Library/Ramsey.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81871 e8ecc32d18c1
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Library/Ramsey.thy
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
     2
    Author:     Tom Ridge. Full finite version by L C Paulson.
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     3
*)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     4
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
     5
section \<open>Ramsey's Theorem\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     6
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 24853
diff changeset
     7
theory Ramsey
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
     8
  imports Infinite_Set Equipollence FuncSet
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 24853
diff changeset
     9
begin
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    10
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
    11
subsection \<open>Preliminary definitions\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
    12
73709
58bd53caf800 things need to be ugly
paulson <lp15@cam.ac.uk>
parents: 73655
diff changeset
    13
abbreviation strict_sorted :: "'a::linorder list \<Rightarrow> bool" where
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    14
  "strict_sorted \<equiv> sorted_wrt (<)"
73709
58bd53caf800 things need to be ugly
paulson <lp15@cam.ac.uk>
parents: 73655
diff changeset
    15
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff 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: 71083
diff changeset
    17
81142
6ad2c917dd2e more inner-syntax markup;
wenzelm
parents: 80914
diff changeset
    18
definition nsets :: "['a set, nat] \<Rightarrow> 'a set set" (\<open>(\<open>notation=\<open>mixfix nsets\<close>\<close>[_]\<^bsup>_\<^esup>)\<close> [0,999] 999)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff 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: 71083
diff changeset
    20
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    21
lemma finite_imp_finite_nsets: "finite A \<Longrightarrow> finite ([A]\<^bsup>k\<^esup>)"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    22
  by (simp add: nsets_def)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    23
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
    24
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: 71083
diff changeset
    25
  by (auto simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
    26
72231
6b620d91e8cc a new lemma
paulson <lp15@cam.ac.uk>
parents: 71848
diff changeset
    27
lemma nsets_Pi_contra: "A' \<subseteq> A \<Longrightarrow> Pi ([A]\<^bsup>n\<^esup>) B \<subseteq> Pi ([A']\<^bsup>n\<^esup>) B"
6b620d91e8cc a new lemma
paulson <lp15@cam.ac.uk>
parents: 71848
diff changeset
    28
  by (auto simp: nsets_def)
6b620d91e8cc a new lemma
paulson <lp15@cam.ac.uk>
parents: 71848
diff changeset
    29
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    30
lemma nsets_2_eq: "[A]\<^bsup>2\<^esup> = (\<Union>x\<in>A. \<Union>y\<in>A - {x}. {{x, y}})"
72231
6b620d91e8cc a new lemma
paulson <lp15@cam.ac.uk>
parents: 71848
diff changeset
    31
  by (auto simp: nsets_def card_2_iff)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
    32
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    33
lemma nsets2_E:
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    34
  assumes "e \<in> [A]\<^bsup>2\<^esup>"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    35
  obtains x y where "e = {x,y}" "x \<in> A" "y \<in> A" "x\<noteq>y"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    36
  using assms by (auto simp: nsets_def card_2_iff)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
    37
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
    38
lemma nsets_doubleton_2_eq [simp]: "[{x, y}]\<^bsup>2\<^esup> = (if x=y then {} else {{x, y}})"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
    39
  by (auto simp: nsets_2_eq)
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
    40
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
    41
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"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
    42
  by (auto simp: nsets_2_eq Set.doubleton_eq_iff)
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
    43
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    44
lemma nsets_3_eq: "[A]\<^bsup>3\<^esup> = (\<Union>x\<in>A. \<Union>y\<in>A - {x}. \<Union>z\<in>A - {x,y}. {{x,y,z}})"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    45
  by (simp add: eval_nat_numeral nsets_def card_Suc_eq) blast
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    46
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    47
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}})"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    48
     (is "_ = ?rhs")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    49
proof
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    50
  show "[A]\<^bsup>4\<^esup> \<subseteq> ?rhs"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    51
    by (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) blast
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    52
  have "\<And>X. X \<in> ?rhs \<Longrightarrow> card X = 4"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    53
    by (force simp: card_2_iff)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    54
  then show "?rhs \<subseteq> [A]\<^bsup>4\<^esup>"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    55
    by (auto simp: nsets_def)
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    56
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    57
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    58
lemma nsets_disjoint_2:
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    59
  "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}})"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    60
  by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    61
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    62
lemma ordered_nsets_2_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    63
  fixes A :: "'a::linorder set"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    64
  shows "[A]\<^bsup>2\<^esup> = {{x,y} | x y. x \<in> A \<and> y \<in> A \<and> x<y}"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    65
     (is "_ = ?rhs")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    66
proof
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    67
  show "[A]\<^bsup>2\<^esup> \<subseteq> ?rhs"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    68
    by (auto simp: nsets_def card_2_iff doubleton_eq_iff neq_iff)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    69
  show "?rhs \<subseteq> [A]\<^bsup>2\<^esup>"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    70
    unfolding numeral_nat by (auto simp: nsets_def card_Suc_eq)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    71
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    72
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    73
lemma ordered_nsets_3_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    74
  fixes A :: "'a::linorder set"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    75
  shows "[A]\<^bsup>3\<^esup> = {{x,y,z} | x y z. x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> x<y \<and> y<z}"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    76
     (is "_ = ?rhs")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    77
proof
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    78
  show "[A]\<^bsup>3\<^esup> \<subseteq> ?rhs"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    79
    unfolding nsets_def card_3_iff
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    80
    by (smt (verit, del_insts) Collect_mono_iff insert_commute insert_subset
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    81
        linorder_less_linear)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    82
  have "\<And>X. X \<in> ?rhs \<Longrightarrow> card X = 3"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    83
    by (force simp: card_3_iff)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    84
  then show "?rhs \<subseteq> [A]\<^bsup>3\<^esup>"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    85
    by (auto simp: nsets_def)
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    86
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    87
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    88
lemma ordered_nsets_4_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    89
  fixes A :: "'a::linorder set"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    90
  defines "rhs \<equiv> \<lambda>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"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    91
  shows "[A]\<^bsup>4\<^esup> = Collect rhs"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    92
proof -
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    93
  have "rhs U" if "U \<in> [A]\<^bsup>4\<^esup>" for U
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    94
  proof -
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    95
    from that obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \<subseteq> A"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    96
      by (simp add: nsets_def) (metis finite_set_strict_sorted)
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    97
    then show ?thesis
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
    98
      unfolding numeral_nat length_Suc_conv rhs_def by auto blast
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    99
  qed
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   100
  moreover
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   101
  have "\<And>X. X \<in> Collect rhs \<Longrightarrow> card X = 4 \<and> finite X \<and> X \<subseteq> A"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   102
    by (auto simp: rhs_def card_insert_if)
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   103
  ultimately show ?thesis
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   104
    unfolding nsets_def by blast
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   105
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   106
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   107
lemma ordered_nsets_5_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   108
  fixes A :: "'a::linorder set"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   109
  defines "rhs \<equiv> \<lambda>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"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   110
  shows "[A]\<^bsup>5\<^esup> = Collect rhs"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   111
proof -
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   112
  have "rhs U" if "U \<in> [A]\<^bsup>5\<^esup>" for U
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   113
  proof -
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   114
    from that obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \<subseteq> A"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   115
      by (simp add: nsets_def) (metis finite_set_strict_sorted)
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   116
    then show ?thesis
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   117
      unfolding numeral_nat length_Suc_conv rhs_def by auto blast
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   118
  qed
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   119
  moreover
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   120
  have "\<And>X. X \<in> Collect rhs \<Longrightarrow> card X = 5 \<and> finite X \<and> X \<subseteq> A"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   121
    by (auto simp: rhs_def card_insert_if)
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   122
  ultimately show ?thesis
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   123
    unfolding nsets_def by blast
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   124
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   125
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   126
lemma binomial_eq_nsets: "n choose k = card (nsets {0..<n} k)"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   127
proof -
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   128
  have "{K. K \<subseteq> {0..<n} \<and> card K = k} = {N. N \<subseteq> {0..<n} \<and> finite N \<and> card N = k}"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   129
    using infinite_super by blast
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   130
  then show ?thesis
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   131
    by (simp add: binomial_def nsets_def)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   132
qed
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   133
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   134
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: 71083
diff changeset
   135
  unfolding nsets_def
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   136
proof (intro iffI conjI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   137
  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: 71083
diff changeset
   138
  show "finite A"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   139
    using infinite_arbitrarily_large that by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   140
  then have "\<not> r \<le> card A"
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   141
    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: 71083
diff changeset
   142
  then show "card A < r"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   143
    using not_less by blast
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   144
next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   145
  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: 71083
diff changeset
   146
    if "finite A \<and> card A < r"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   147
    using that card_mono leD by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   148
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   149
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   150
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: 71083
diff changeset
   151
  by (simp add: nsets_eq_empty_iff)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   152
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   153
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: 71083
diff changeset
   154
  by (auto simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   155
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   156
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: 71083
diff changeset
   157
  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: 71083
diff changeset
   158
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   159
lemma nsets_self [simp]: "nsets {..<m} m = {{..<m}}"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   160
proof
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   161
  show "[{..<m}]\<^bsup>m\<^esup> \<subseteq> {{..<m}}"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   162
    by (force simp add: card_subset_eq nsets_def)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   163
qed (simp add: nsets_def)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   164
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   165
lemma nsets_zero [simp]: "nsets A 0 = {{}}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   166
  by (auto simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   167
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   168
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: 71083
diff changeset
   169
  using card_eq_SucD by (force simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   170
71405
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   171
lemma inj_on_nsets:
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   172
  assumes "inj_on f A"
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   173
  shows "inj_on (\<lambda>X. f ` X) ([A]\<^bsup>n\<^esup>)"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   174
  using assms by (simp add: nsets_def inj_on_def inj_on_image_eq_iff)
71405
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   175
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   176
lemma bij_betw_nsets:
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   177
  assumes "bij_betw f A B"
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   178
  shows "bij_betw (\<lambda>X. f ` X) ([A]\<^bsup>n\<^esup>) ([B]\<^bsup>n\<^esup>)"
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   179
proof -
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   180
  have "\<And>X. \<lbrakk>X \<subseteq> f ` A; finite X\<rbrakk> \<Longrightarrow> \<exists>Y\<subseteq>A. finite Y \<and> card Y = card X \<and> X = f ` Y"
71405
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   181
    by (metis card_image inj_on_finite order_refl subset_image_inj)
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   182
  then have "(`) f ` [A]\<^bsup>n\<^esup> = [f ` A]\<^bsup>n\<^esup>"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   183
    using assms by (auto simp: nsets_def bij_betw_def image_iff card_image inj_on_subset)
71405
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   184
  with assms show ?thesis
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   185
    by (auto simp: bij_betw_def inj_on_nsets)
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   186
qed
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   187
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   188
lemma nset_image_obtains:
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   189
  assumes "X \<in> [f`A]\<^bsup>k\<^esup>" "inj_on f A"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   190
  obtains Y where "Y \<in> [A]\<^bsup>k\<^esup>" "X = f ` Y"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   191
proof
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   192
  show "X = f ` (A \<inter> f -` X)"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   193
    using assms by (auto simp: nsets_def)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   194
  then show "A Int (f -` X) \<in> [A]\<^bsup>k\<^esup>"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   195
    using assms
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   196
    unfolding nsets_def mem_Collect_eq
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   197
    by (metis card_image finite_image_iff inf_le1 subset_inj_on)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   198
qed
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   199
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   200
lemma nsets_image_funcset:
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   201
  assumes "g \<in> S \<rightarrow> T" and "inj_on g S"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   202
  shows "(\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> [T]\<^bsup>k\<^esup>"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   203
    using assms
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   204
    by (fastforce simp: nsets_def card_image inj_on_subset subset_iff simp flip: image_subset_iff_funcset)
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   205
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   206
lemma nsets_compose_image_funcset:
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   207
  assumes f: "f \<in> [T]\<^bsup>k\<^esup> \<rightarrow> D" and "g \<in> S \<rightarrow> T" and "inj_on g S"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   208
  shows "f \<circ> (\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> D"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   209
proof -
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   210
  have "(\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> [T]\<^bsup>k\<^esup>"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   211
    using assms by (simp add: nsets_image_funcset)
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   212
  then show ?thesis
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   213
    using f by fastforce
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   214
qed
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   215
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   216
subsubsection \<open>Further properties, involving equipollence\<close>
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   217
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   218
lemma nsets_lepoll_cong:
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   219
  assumes "A \<lesssim> B"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   220
  shows "[A]\<^bsup>k\<^esup> \<lesssim> [B]\<^bsup>k\<^esup>"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   221
proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   222
  obtain f where f: "inj_on f A" "f ` A \<subseteq> B"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   223
    by (meson assms lepoll_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   224
  define F where "F \<equiv> \<lambda>N. f ` N"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   225
  have "inj_on F ([A]\<^bsup>k\<^esup>)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   226
    using F_def f inj_on_nsets by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   227
  moreover
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   228
  have "F ` ([A]\<^bsup>k\<^esup>) \<subseteq> [B]\<^bsup>k\<^esup>"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   229
    by (metis F_def bij_betw_def bij_betw_nsets f nsets_mono)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   230
  ultimately show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   231
    by (meson lepoll_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   232
qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   233
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   234
lemma nsets_eqpoll_cong:
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   235
  assumes "A\<approx>B"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   236
  shows "[A]\<^bsup>k\<^esup> \<approx> [B]\<^bsup>k\<^esup>"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   237
  by (meson assms eqpoll_imp_lepoll eqpoll_sym lepoll_antisym nsets_lepoll_cong)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   238
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   239
lemma infinite_imp_infinite_nsets:
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   240
  assumes inf: "infinite A" and "k>0"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   241
  shows "infinite ([A]\<^bsup>k\<^esup>)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   242
proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   243
  obtain B where "B \<subset> A" "A\<approx>B"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   244
    by (meson inf infinite_iff_psubset)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   245
  then obtain a where a: "a \<in> A" "a \<notin> B"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   246
    by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   247
  then obtain N where "N \<subseteq> B" "finite N" "card N = k-1" "a \<notin> N"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   248
    by (metis \<open>A \<approx> B\<close> inf eqpoll_finite_iff infinite_arbitrarily_large subset_eq)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   249
  with a \<open>k>0\<close> \<open>B \<subset> A\<close> have "insert a N \<in> [A]\<^bsup>k\<^esup>"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   250
    by (simp add: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   251
  with a have "nsets B k \<noteq> nsets A k"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   252
    by (metis (no_types, lifting) in_mono insertI1 mem_Collect_eq nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   253
  moreover have "nsets B k \<subseteq> nsets A k"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   254
    using \<open>B \<subset> A\<close> nsets_mono by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   255
  ultimately show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   256
    unfolding infinite_iff_psubset_le
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   257
    by (meson \<open>A \<approx> B\<close> eqpoll_imp_lepoll nsets_eqpoll_cong psubsetI)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   258
qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   259
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   260
lemma finite_nsets_iff:
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   261
  assumes "k>0"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   262
  shows "finite ([A]\<^bsup>k\<^esup>) \<longleftrightarrow> finite A"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   263
  using assms finite_imp_finite_nsets infinite_imp_infinite_nsets by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   264
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   265
lemma card_nsets [simp]: "card (nsets A k) = card A choose k"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   266
proof (cases "finite A")
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   267
  case True
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   268
  then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   269
    by (metis bij_betw_nsets bij_betw_same_card binomial_eq_nsets ex_bij_betw_nat_finite)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   270
next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   271
  case False
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   272
  then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   273
    by (cases "k=0"; simp add: finite_nsets_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   274
qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   275
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   276
subsubsection \<open>Partition predicates\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   277
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   278
definition "monochromatic \<equiv> \<lambda>\<beta> \<alpha> \<gamma> f i. \<exists>H \<in> nsets \<beta> \<alpha>. f ` (nsets H \<gamma>) \<subseteq> {i}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   279
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   280
text \<open>uniform partition sizes\<close>
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   281
definition partn :: "'a set \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'b set \<Rightarrow> bool"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   282
  where "partn \<beta> \<alpha> \<gamma> \<delta> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  \<delta>. \<exists>\<xi>\<in>\<delta>. monochromatic \<beta> \<alpha> \<gamma> f \<xi>"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   283
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   284
text \<open>partition sizes enumerated in a list\<close>
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   285
definition partn_lst :: "'a set \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   286
  where "partn_lst \<beta> \<alpha> \<gamma> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  {..<length \<alpha>}. \<exists>i < length \<alpha>. monochromatic \<beta> (\<alpha>!i) \<gamma> f i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   287
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   288
text \<open>There's always a 0-clique\<close>
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   289
lemma partn_lst_0: "\<gamma> > 0 \<Longrightarrow> partn_lst \<beta> (0#\<alpha>) \<gamma>"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   290
  by (force simp: partn_lst_def monochromatic_def nsets_empty_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   291
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   292
lemma partn_lst_0': "\<gamma> > 0 \<Longrightarrow> partn_lst \<beta> (a#0#\<alpha>) \<gamma>"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   293
  by (force simp: partn_lst_def monochromatic_def nsets_empty_iff)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   294
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   295
lemma partn_lst_greater_resource:
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   296
  fixes M::nat
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   297
  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: 71083
diff changeset
   298
  shows "partn_lst {..<N} \<alpha> \<gamma>"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   299
proof (clarsimp simp: partn_lst_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   300
  fix f
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   301
  assume "f \<in> nsets {..<N} \<gamma> \<rightarrow> {..<length \<alpha>}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   302
  then have "f \<in> nsets {..<M} \<gamma> \<rightarrow> {..<length \<alpha>}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   303
    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: 71083
diff changeset
   304
  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}"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   305
    using M unfolding partn_lst_def monochromatic_def by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   306
  have "H \<in> nsets {..<N} (\<alpha> ! i)"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   307
    using \<open>M \<le> N\<close> H by (auto simp: nsets_def subset_iff)
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   308
  then show "\<exists>i<length \<alpha>. monochromatic {..<N} (\<alpha>!i) \<gamma> f i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   309
    using i subi unfolding monochromatic_def by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   310
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   311
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   312
lemma partn_lst_fewer_colours:
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   313
  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: 71083
diff changeset
   314
  shows "partn_lst \<beta> \<alpha> \<gamma>"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   315
proof (clarsimp simp: partn_lst_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   316
  fix f :: "'a set \<Rightarrow> nat"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   317
  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: 71083
diff changeset
   318
  then obtain i H where i: "i < Suc (length \<alpha>)"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   319
      and H: "H \<in> [\<beta>]\<^bsup>((n # \<alpha>) ! i)\<^esup>"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   320
      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: 71083
diff changeset
   321
    using \<open>n \<ge> \<gamma>\<close> major [unfolded partn_lst_def, rule_format, of "Suc o f"]
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   322
    by (fastforce simp: image_subset_iff nsets_eq_empty_iff monochromatic_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   323
  show "\<exists>i<length \<alpha>. monochromatic \<beta> (\<alpha>!i) \<gamma> f i"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   324
  proof (cases i)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   325
    case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   326
    then have "[H]\<^bsup>\<gamma>\<^esup> = {}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   327
      using hom by blast
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   328
    then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   329
      using 0 H \<open>n \<ge> \<gamma>\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   330
      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: 71083
diff changeset
   331
  next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   332
    case (Suc i')
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   333
    then show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   334
      unfolding monochromatic_def using i H hom by auto
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   335
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   336
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   337
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   338
lemma partn_lst_eq_partn: "partn_lst {..<n} [m,m] 2 = partn {..<n} m 2 {..<2::nat}"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   339
proof -
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   340
  have "\<And>i. i < 2 \<Longrightarrow> [m, m] ! i = m"
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   341
    using less_2_cases_iff by force
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   342
  then show ?thesis
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   343
    by (auto simp: partn_lst_def partn_def numeral_2_eq_2 cong: conj_cong)
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   344
qed
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   345
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   346
lemma partn_lstE:
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   347
  assumes "partn_lst \<beta> \<alpha> \<gamma>" "f \<in> nsets \<beta> \<gamma>  \<rightarrow>  {..<l}" "length \<alpha> = l"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   348
  obtains i H where "i < length \<alpha>" "H \<in> nsets \<beta> (\<alpha>!i)" "f ` (nsets H \<gamma>) \<subseteq> {i}"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   349
  using partn_lst_def monochromatic_def assms by metis
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   350
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   351
lemma partn_lst_less:
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   352
  assumes M: "partn_lst \<beta> \<alpha> n" and eq: "length \<alpha>' = length \<alpha>"
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   353
    and le: "\<And>i. i < length \<alpha> \<Longrightarrow> \<alpha>'!i \<le> \<alpha>!i "
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   354
  shows "partn_lst \<beta> \<alpha>' n"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   355
proof (clarsimp simp: partn_lst_def)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   356
  fix f
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   357
  assume "f \<in> [\<beta>]\<^bsup>n\<^esup> \<rightarrow> {..<length \<alpha>'}"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   358
  then obtain i H where i: "i < length \<alpha>"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   359
                   and "H \<subseteq> \<beta>"  and H: "card H = (\<alpha>!i)" and "finite H"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   360
                   and fi: "f ` nsets H n \<subseteq> {i}"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   361
    using assms by (auto simp: partn_lst_def monochromatic_def nsets_def)
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   362
  then obtain bij where bij: "bij_betw bij H {0..<\<alpha>!i}"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   363
    by (metis ex_bij_betw_finite_nat)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   364
  then have inj: "inj_on (inv_into H bij) {0..<\<alpha>' ! i}"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   365
    by (metis bij_betw_def dual_order.refl i inj_on_inv_into ivl_subset le)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   366
  define H' where "H' = inv_into H bij ` {0..<\<alpha>'!i}"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   367
  show "\<exists>i<length \<alpha>'. monochromatic \<beta> (\<alpha>'!i) n f i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   368
    unfolding monochromatic_def
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   369
  proof (intro exI bexI conjI)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   370
    show "i < length \<alpha>'"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   371
      by (simp add: assms(2) i)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   372
    have "H' \<subseteq> H"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   373
      using bij \<open>i < length \<alpha>\<close> bij_betw_imp_surj_on le
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   374
      by (force simp: H'_def image_subset_iff intro: inv_into_into)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   375
    then have "finite H'"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   376
      by (simp add: \<open>finite H\<close> finite_subset)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   377
    with \<open>H' \<subseteq> H\<close> have cardH': "card H' = (\<alpha>'!i)"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   378
      unfolding H'_def by (simp add: inj card_image)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   379
    show "f ` [H']\<^bsup>n\<^esup> \<subseteq> {i}"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   380
      by (meson \<open>H' \<subseteq> H\<close> dual_order.trans fi image_mono nsets_mono)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   381
    show "H' \<in> [\<beta>]\<^bsup>(\<alpha>'! i)\<^esup>"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   382
      using \<open>H \<subseteq> \<beta>\<close> \<open>H' \<subseteq> H\<close> \<open>finite H'\<close> cardH' nsets_def by fastforce
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   383
  qed
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   384
qed
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   385
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   386
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   387
subsection \<open>Finite versions of Ramsey's theorem\<close>
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   388
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   389
text \<open>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   390
  To distinguish the finite and infinite ones, lower and upper case
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   391
  names are used (ramsey vs Ramsey).
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   392
\<close>
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   393
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   394
subsubsection \<open>The Erdős--Szekeres theorem exhibits an upper bound for Ramsey numbers\<close>
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   395
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   396
text \<open>The Erdős--Szekeres bound, essentially extracted from the proof\<close>
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   397
fun ES :: "[nat,nat,nat] \<Rightarrow> nat"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   398
  where "ES 0 k l = max k l"
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   399
  |     "ES (Suc r) k l =
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   400
            (if r=0 then k+l-1
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   401
             else if k=0 \<or> l=0 then 1 else Suc (ES r (ES (Suc r) (k-1) l) (ES (Suc r) k (l-1))))"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   402
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   403
declare ES.simps [simp del]
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   404
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   405
lemma ES_0 [simp]: "ES 0 k l = max k l"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   406
  using ES.simps(1) by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   407
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   408
lemma ES_1 [simp]: "ES 1 k l = k+l-1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   409
  using ES.simps(2) [of 0 k l] by simp
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   410
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   411
lemma ES_2: "ES 2 k l = (if k=0 \<or> l=0 then 1 else ES 2 (k-1) l + ES 2 k (l-1))"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   412
  unfolding numeral_2_eq_2
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   413
  by (smt (verit) ES.elims One_nat_def Suc_pred add_gr_0 neq0_conv nat.inject zero_less_Suc)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   414
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   415
text \<open>The Erdős--Szekeres upper bound\<close>
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   416
lemma ES2_choose: "ES 2 k l = (k+l) choose k"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   417
proof (induct n \<equiv> "k+l" arbitrary: k l)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   418
  case 0
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   419
  then show ?case
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   420
    by (auto simp: ES_2)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   421
next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   422
  case (Suc n)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   423
  then have "k>0 \<Longrightarrow> l>0 \<Longrightarrow> ES 2 (k - 1) l + ES 2 k (l - 1) = k + l choose k"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   424
    using choose_reduce_nat by force
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   425
  then show ?case
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   426
    by (metis ES_2 Nat.add_0_right binomial_n_0 binomial_n_n gr0I)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   427
qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   428
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   429
subsubsection \<open>Trivial cases\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   430
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   431
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: 71083
diff changeset
   432
lemma ramsey0: "\<exists>N::nat. partn_lst {..<N} [q1,q2] 0"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   433
  by (force simp: partn_lst_def monochromatic_def ex_in_conv less_Suc_eq nsets_eq_empty_iff)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   434
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   435
text \<open>Just the pigeon hole principle, since we are dealing with 1-sets\<close>
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   436
lemma ramsey1_explicit: "partn_lst {..<q0 + q1 - Suc 0} [q0,q1] 1"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   437
proof -
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   438
  have "\<exists>i<Suc (Suc 0). \<exists>H\<in>nsets {..<q0 + q1 - 1} ([q0, q1] ! i). f ` nsets H 1 \<subseteq> {i}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   439
    if "f \<in> nsets {..<q0 + q1 - 1} (Suc 0) \<rightarrow> {..<Suc (Suc 0)}" for f
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   440
  proof -
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   441
    define A where "A \<equiv> \<lambda>i. {q. q < q0+q1-1 \<and> f {q} = i}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   442
    have "A 0 \<union> A 1 = {..<q0 + q1-1}"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   443
      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: 71083
diff changeset
   444
    moreover have "A 0 \<inter> A 1 = {}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   445
      by (auto simp: A_def)
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   446
    ultimately have "q0 + q1 \<le> card (A 0) + card (A 1) + 1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   447
      by (metis card_Un_le card_lessThan le_diff_conv)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   448
    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: 71083
diff changeset
   449
      by linarith
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   450
    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: 71083
diff changeset
   451
      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: 71083
diff changeset
   452
    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: 71464
diff changeset
   453
      by (meson obtain_subset_with_card_n)
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   454
    then have "B \<in> nsets {..<q0 + q1 - 1} ([q0, q1] ! i) \<and> f ` nsets B (Suc 0) \<subseteq> {i}"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   455
      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: 71083
diff changeset
   456
    then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   457
      using \<open>i < Suc (Suc 0)\<close> by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   458
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   459
  then show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   460
    by (simp add: partn_lst_def monochromatic_def)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   461
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   462
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   463
lemma ramsey1: "\<exists>N::nat. partn_lst {..<N} [q0,q1] 1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   464
  using ramsey1_explicit by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   465
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   466
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   467
subsubsection \<open>Ramsey's theorem with TWO colours and arbitrary exponents (hypergraph version)\<close>
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   468
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   469
lemma ramsey_induction_step:
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   470
  fixes p::nat
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   471
  assumes p1: "partn_lst {..<p1} [q1-1,q2] (Suc r)" and p2: "partn_lst {..<p2} [q1,q2-1] (Suc r)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   472
    and p: "partn_lst {..<p} [p1,p2] r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   473
    and "q1>0" "q2>0"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   474
  shows "partn_lst {..<Suc p} [q1, q2] (Suc r)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   475
proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   476
  have "\<exists>i<Suc (Suc 0). \<exists>H\<in>nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \<subseteq> {i}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   477
    if f: "f \<in> nsets {..p} (Suc r) \<rightarrow> {..<Suc (Suc 0)}" for f
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   478
  proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   479
    define g where "g \<equiv> \<lambda>R. f (insert p R)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   480
    have "f (insert p i) \<in> {..<Suc (Suc 0)}" if "i \<in> nsets {..<p} r" for i
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   481
      using that card_insert_if by (fastforce simp: nsets_def intro!: Pi_mem [OF f])
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   482
    then have g: "g \<in> nsets {..<p} r \<rightarrow> {..<Suc (Suc 0)}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   483
      by (force simp: g_def PiE_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   484
    then obtain i U where i: "i < Suc (Suc 0)" and gi: "g ` nsets U r \<subseteq> {i}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   485
      and U: "U \<in> nsets {..<p} ([p1, p2] ! i)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   486
      using p by (auto simp: partn_lst_def monochromatic_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   487
    then have Usub: "U \<subseteq> {..<p}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   488
      by (auto simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   489
    consider (izero) "i = 0" | (ione) "i = Suc 0"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   490
      using i by linarith
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   491
    then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   492
    proof cases
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   493
      case izero
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   494
      then have "U \<in> nsets {..<p} p1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   495
        using U by simp
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   496
      then obtain u where u: "bij_betw u {..<p1} U"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   497
        using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   498
      have u_nsets: "u ` X \<in> nsets {..p} n" if "X \<in> nsets {..<p1} n" for X n
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   499
      proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   500
        have "inj_on u X"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   501
          using u that bij_betw_imp_inj_on inj_on_subset by (force simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   502
        then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   503
          using Usub u that bij_betwE
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   504
          by (fastforce simp: nsets_def card_image)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   505
      qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   506
      define h where "h \<equiv> \<lambda>R. f (u ` R)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   507
      have "h \<in> nsets {..<p1} (Suc r) \<rightarrow> {..<Suc (Suc 0)}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   508
        unfolding h_def using f u_nsets by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   509
      then obtain j V where j: "j <Suc (Suc 0)" and hj: "h ` nsets V (Suc r) \<subseteq> {j}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   510
        and V: "V \<in> nsets {..<p1} ([q1 - Suc 0, q2] ! j)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   511
        using p1 by (auto simp: partn_lst_def monochromatic_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   512
      then have Vsub: "V \<subseteq> {..<p1}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   513
        by (auto simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   514
      have invinv_eq: "u ` inv_into {..<p1} u ` X = X" if "X \<subseteq> u ` {..<p1}" for X
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   515
        by (simp add: image_inv_into_cancel that)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   516
      let ?W = "insert p (u ` V)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   517
      consider (jzero) "j = 0" | (jone) "j = Suc 0"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   518
        using j by linarith
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   519
      then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   520
      proof cases
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   521
        case jzero
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   522
        then have "V \<in> nsets {..<p1} (q1 - Suc 0)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   523
          using V by simp
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   524
        then have "u ` V \<in> nsets {..<p} (q1 - Suc 0)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   525
          using u_nsets [of _ "q1 - Suc 0"] nsets_mono [OF Vsub] Usub u
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   526
          unfolding bij_betw_def nsets_def
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   527
          by (fastforce elim!: subsetD)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   528
        then have inq1: "?W \<in> nsets {..p} q1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   529
          unfolding nsets_def using \<open>q1 > 0\<close> card_insert_if by fastforce
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   530
        have invu_nsets: "inv_into {..<p1} u ` X \<in> nsets V r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   531
          if "X \<in> nsets (u ` V) r" for X r
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   532
        proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   533
          have "X \<subseteq> u ` V \<and> finite X \<and> card X = r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   534
            using nsets_def that by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   535
          then have [simp]: "card (inv_into {..<p1} u ` X) = card X"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   536
            by (meson Vsub bij_betw_def bij_betw_inv_into card_image image_mono inj_on_subset u)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   537
          show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   538
            using that u Vsub by (fastforce simp: nsets_def bij_betw_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   539
        qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   540
        have "f X = i" if X: "X \<in> nsets ?W (Suc r)" for X
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   541
        proof (cases "p \<in> X")
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   542
          case True
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   543
          then have Xp: "X - {p} \<in> nsets (u ` V) r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   544
            using X by (auto simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   545
          moreover have "u ` V \<subseteq> U"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   546
            using Vsub bij_betwE u by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   547
          ultimately have "X - {p} \<in> nsets U r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   548
            by (meson in_mono nsets_mono)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   549
          then have "g (X - {p}) = i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   550
            using gi by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   551
          have "f X = i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   552
            using gi True \<open>X - {p} \<in> nsets U r\<close> insert_Diff
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   553
            by (fastforce simp: g_def image_subset_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   554
          then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   555
            by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   556
        next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   557
          case False
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   558
          then have Xim: "X \<in> nsets (u ` V) (Suc r)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   559
            using X by (auto simp: nsets_def subset_insert)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   560
          then have "u ` inv_into {..<p1} u ` X = X"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   561
            using Vsub bij_betw_imp_inj_on u
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   562
            by (fastforce simp: nsets_def image_mono invinv_eq subset_trans)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   563
          then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   564
            using izero jzero hj Xim invu_nsets unfolding h_def
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   565
            by (fastforce simp: image_subset_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   566
        qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   567
        moreover have "insert p (u ` V) \<in> nsets {..p} q1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   568
          by (simp add: izero inq1)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   569
        ultimately show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   570
          by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   571
      next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   572
        case jone
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   573
        then have "u ` V \<in> nsets {..p} q2"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   574
          using V u_nsets by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   575
        moreover have "f ` nsets (u ` V) (Suc r) \<subseteq> {j}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   576
          using hj
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   577
          by (force simp: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   578
        ultimately show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   579
          using jone not_less_eq by fastforce
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   580
      qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   581
    next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   582
      case ione
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   583
      then have "U \<in> nsets {..<p} p2"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   584
        using U by simp
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   585
      then obtain u where u: "bij_betw u {..<p2} U"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   586
        using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   587
      have u_nsets: "u ` X \<in> nsets {..p} n" if "X \<in> nsets {..<p2} n" for X n
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   588
      proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   589
        have "inj_on u X"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   590
          using u that bij_betw_imp_inj_on inj_on_subset by (force simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   591
        then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   592
          using Usub u that bij_betwE
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   593
          by (fastforce simp: nsets_def card_image)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   594
      qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   595
      define h where "h \<equiv> \<lambda>R. f (u ` R)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   596
      have "h \<in> nsets {..<p2} (Suc r) \<rightarrow> {..<Suc (Suc 0)}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   597
        unfolding h_def using f u_nsets by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   598
      then obtain j V where j: "j <Suc (Suc 0)" and hj: "h ` nsets V (Suc r) \<subseteq> {j}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   599
        and V: "V \<in> nsets {..<p2} ([q1, q2 - Suc 0] ! j)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   600
        using p2 by (auto simp: partn_lst_def monochromatic_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   601
      then have Vsub: "V \<subseteq> {..<p2}"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   602
        by (auto simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   603
      have invinv_eq: "u ` inv_into {..<p2} u ` X = X" if "X \<subseteq> u ` {..<p2}" for X
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   604
        by (simp add: image_inv_into_cancel that)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   605
      let ?W = "insert p (u ` V)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   606
      consider (jzero) "j = 0" | (jone) "j = Suc 0"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   607
        using j by linarith
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   608
      then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   609
      proof cases
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   610
        case jone
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   611
        then have "V \<in> nsets {..<p2} (q2 - Suc 0)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   612
          using V by simp
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   613
        then have "u ` V \<in> nsets {..<p} (q2 - Suc 0)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   614
          using u_nsets [of _ "q2 - Suc 0"] nsets_mono [OF Vsub] Usub u
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   615
          unfolding bij_betw_def nsets_def by blast
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   616
        then have inq1: "?W \<in> nsets {..p} q2"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   617
          unfolding nsets_def using \<open>q2 > 0\<close> card_insert_if by fastforce
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   618
        have invu_nsets: "inv_into {..<p2} u ` X \<in> nsets V r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   619
          if "X \<in> nsets (u ` V) r" for X r
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   620
        proof -
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   621
          have "X \<subseteq> u ` V \<and> finite X \<and> card X = r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   622
            using nsets_def that by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   623
          then have [simp]: "card (inv_into {..<p2} u ` X) = card X"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   624
            by (meson Vsub bij_betw_def bij_betw_inv_into card_image image_mono inj_on_subset u)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   625
          show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   626
            using that u Vsub by (fastforce simp: nsets_def bij_betw_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   627
        qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   628
        have "f X = i" if X: "X \<in> nsets ?W (Suc r)" for X
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   629
        proof (cases "p \<in> X")
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   630
          case True
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   631
          then have Xp: "X - {p} \<in> nsets (u ` V) r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   632
            using X by (auto simp: nsets_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   633
          moreover have "u ` V \<subseteq> U"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   634
            using Vsub bij_betwE u by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   635
          ultimately have "X - {p} \<in> nsets U r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   636
            by (meson in_mono nsets_mono)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   637
          then have "g (X - {p}) = i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   638
            using gi by blast
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   639
          have "f X = i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   640
            using gi True \<open>X - {p} \<in> nsets U r\<close> insert_Diff
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   641
            by (fastforce simp: g_def image_subset_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   642
          then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   643
            by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   644
        next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   645
          case False
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   646
          then have Xim: "X \<in> nsets (u ` V) (Suc r)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   647
            using X by (auto simp: nsets_def subset_insert)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   648
          then have "u ` inv_into {..<p2} u ` X = X"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   649
            using Vsub bij_betw_imp_inj_on u
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   650
            by (fastforce simp: nsets_def image_mono invinv_eq subset_trans)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   651
          then show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   652
            using ione jone hj Xim invu_nsets unfolding h_def
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   653
            by (fastforce simp: image_subset_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   654
        qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   655
        moreover have "insert p (u ` V) \<in> nsets {..p} q2"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   656
          by (simp add: ione inq1)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   657
        ultimately show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   658
          by (metis ione image_subsetI insertI1 lessI nth_Cons_0 nth_Cons_Suc)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   659
      next
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   660
        case jzero
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   661
        then have "u ` V \<in> nsets {..p} q1"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   662
          using V u_nsets by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   663
        moreover have "f ` nsets (u ` V) (Suc r) \<subseteq> {j}"
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   664
          using hj unfolding h_def image_subset_iff nsets_def
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   665
          apply (clarsimp simp add: h_def image_subset_iff nsets_def)
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   666
          by (metis card_image finite_imageD subset_image_inj)
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   667
        ultimately show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   668
          using jzero not_less_eq by fastforce
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   669
      qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   670
    qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   671
  qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   672
  then show "?thesis"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   673
    using lessThan_Suc lessThan_Suc_atMost
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   674
    by (auto simp: partn_lst_def monochromatic_def insert_commute)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   675
qed
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   676
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   677
proposition ramsey2_full: "partn_lst {..<ES r q1 q2} [q1,q2] r"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   678
proof (induction r arbitrary: q1 q2)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   679
  case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   680
  then show ?case
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   681
    by (auto simp: partn_lst_def monochromatic_def less_Suc_eq ex_in_conv nsets_eq_empty_iff)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   682
next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   683
  case (Suc r)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   684
  note outer = this
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   685
  show ?case
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   686
  proof (cases "r = 0")
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   687
    case True
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   688
    then show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   689
      using ramsey1_explicit by (force simp: ES.simps)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   690
  next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   691
    case False
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   692
    then have "r > 0"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   693
      by simp
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   694
    show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   695
      using Suc.prems
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   696
    proof (induct k \<equiv> "q1 + q2" arbitrary: q1 q2)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   697
      case 0
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   698
      with partn_lst_0 show ?case by auto
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   699
    next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   700
      case (Suc k)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   701
      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: 71464
diff changeset
   702
      then show ?case
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   703
      proof cases
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   704
        case 1
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   705
        with False partn_lst_0 partn_lst_0' show ?thesis
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   706
          by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   707
      next
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   708
        define p1 where "p1 \<equiv> ES (Suc r) (q1-1) q2"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   709
        define p2 where "p2 \<equiv> ES (Suc r) q1 (q2-1)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   710
        define p where "p \<equiv> ES r p1 p2"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   711
        case 2
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   712
        with Suc have "k = (q1-1) + q2" "k = q1 + (q2 - 1)" by auto
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   713
        then have p1: "partn_lst {..<p1} [q1-1,q2] (Suc r)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   714
              and p2: "partn_lst {..<p2} [q1,q2-1] (Suc r)"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   715
          using Suc.hyps unfolding p1_def p2_def by blast+
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   716
        then have p: "partn_lst {..<p} [p1,p2] r"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   717
          using outer Suc.prems unfolding p_def by auto
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   718
        show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   719
          using ramsey_induction_step [OF p1 p2 p] "2" ES.simps(2) False p1_def p2_def p_def by auto
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   720
      qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   721
    qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   722
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   723
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   724
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   725
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: 71083
diff changeset
   726
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   727
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: 71083
diff changeset
   728
proof (induction k \<equiv> "length qs" arbitrary: qs)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   729
  case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   730
  then show ?case
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   731
    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: 71083
diff changeset
   732
next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   733
  case (Suc k)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   734
  note IH = this
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   735
  show ?case
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   736
  proof (cases k)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   737
    case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   738
    with Suc obtain q where "qs = [q]"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   739
      by (metis length_0_conv length_Suc_conv)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   740
    then show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   741
      by (rule_tac x=q in exI) (auto simp: partn_lst_def monochromatic_def funcset_to_empty_iff)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   742
  next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   743
    case (Suc k')
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   744
    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: 71083
diff changeset
   745
      by (metis Suc.hyps(2) length_Suc_conv)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   746
    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: 71464
diff changeset
   747
      using ramsey2_full by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   748
    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: 71464
diff changeset
   749
      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: 71083
diff changeset
   750
    have keq: "Suc (length l) = k"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   751
      using IH qs by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   752
    show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   753
    proof (intro exI conjI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   754
      show "partn_lst {..<p} qs r"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   755
      proof (auto simp: partn_lst_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   756
        fix f
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   757
        assume f: "f \<in> nsets {..<p} r \<rightarrow> {..<length qs}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   758
        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: 71083
diff changeset
   759
        have "g \<in> nsets {..<p} r \<rightarrow> {..<k}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   760
          unfolding g_def using f Suc IH
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   761
          by (auto simp: Pi_def not_less)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   762
        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: 71464
diff changeset
   763
                and U: "U \<in> nsets {..<p} ((q#l) ! i)"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   764
          using p keq by (auto simp: partn_lst_def monochromatic_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   765
        show "\<exists>i<length qs. monochromatic {..<p} (qs!i) r f i"
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   766
        proof (cases "i = 0")
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   767
          case True
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   768
          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: 71083
diff changeset
   769
            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: 71464
diff changeset
   770
          then obtain u where u: "bij_betw u {..<q} U"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   771
            using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp: nsets_def)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   772
          then have Usub: "U \<subseteq> {..<p}"
73655
26a1d66b9077 tuned proofs --- avoid z3, which is absent on arm64-linux;
wenzelm
parents: 72231
diff changeset
   773
            by (smt (verit) U mem_Collect_eq nsets_def)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   774
          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: 71083
diff changeset
   775
          proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   776
            have "inj_on u X"
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   777
              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: 71083
diff changeset
   778
              by (force simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   779
            then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   780
              using Usub u that bij_betwE
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   781
              by (fastforce simp: nsets_def card_image)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   782
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   783
          define h where "h \<equiv> \<lambda>X. f (u ` X)"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   784
          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: 71083
diff changeset
   785
          proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   786
            have "u ` X \<in> nsets U r"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   787
              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: 71083
diff changeset
   788
            then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   789
              using f01 by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   790
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   791
          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: 71464
diff changeset
   792
            unfolding h_def by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   793
          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: 71464
diff changeset
   794
            and V: "V \<in> nsets {..<q} ([q1,q2] ! j)"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   795
            using q by (auto simp: partn_lst_def monochromatic_def)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   796
          show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   797
            unfolding monochromatic_def
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   798
          proof (intro exI conjI bexI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   799
            show "j < length qs"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   800
              using Suc Suc.hyps(2) j by linarith
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   801
            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: 71083
diff changeset
   802
              apply (clarsimp simp add: nsets_def image_iff)
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   803
              by (metis card_image finite_imageD subset_image_inj)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   804
            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: 71083
diff changeset
   805
              by (auto simp: h_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   806
            then show "f ` nsets (u ` V) r \<subseteq> {j}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   807
              using hj by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   808
            show "(u ` V) \<in> nsets {..<p} (qs ! j)"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   809
              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: 71083
diff changeset
   810
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   811
        next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   812
          case False
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   813
          then have eq: "\<And>A. \<lbrakk>A \<in> [U]\<^bsup>r\<^esup>\<rbrakk> \<Longrightarrow> f A = Suc i"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   814
            by (metis Suc_pred diff_0_eq_0 g_def gi image_subset_iff not_gr0 singletonD)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   815
          show ?thesis
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   816
            unfolding monochromatic_def
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   817
          proof (intro exI conjI bexI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   818
            show "Suc i < length qs"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   819
              using Suc.hyps(2) i by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   820
            show "f ` nsets U r \<subseteq> {Suc i}"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   821
              using False by (auto simp: eq)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   822
            show "U \<in> nsets {..<p} (qs ! (Suc i))"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   823
              using False U qs by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   824
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   825
        qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   826
      qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   827
    qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   828
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   829
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   830
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   831
subsubsection \<open>Simple graph version\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   832
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   833
text \<open>This is the most basic version in terms of cliques and independent
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   834
  sets, i.e. the version for graphs and 2 colours.
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   835
\<close>
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   836
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   837
definition "clique V E \<longleftrightarrow> (\<forall>v\<in>V. \<forall>w\<in>V. v \<noteq> w \<longrightarrow> {v, w} \<in> E)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   838
definition "indep V E \<longleftrightarrow> (\<forall>v\<in>V. \<forall>w\<in>V. v \<noteq> w \<longrightarrow> {v, w} \<notin> E)"
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   839
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   840
lemma clique_Un: "\<lbrakk>clique K F; clique L F; \<forall>v\<in>K. \<forall>w\<in>L. v\<noteq>w \<longrightarrow> {v,w} \<in> F\<rbrakk> \<Longrightarrow> clique (K \<union> L) F"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   841
  by (metis UnE clique_def doubleton_eq_iff)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   842
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   843
lemma null_clique[simp]: "clique {} E" and null_indep[simp]: "indep {} E"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   844
  by (auto simp: clique_def indep_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   845
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   846
lemma smaller_clique: "\<lbrakk>clique R E; R' \<subseteq> R\<rbrakk> \<Longrightarrow> clique R' E"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   847
  by (auto simp: clique_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   848
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   849
lemma smaller_indep: "\<lbrakk>indep R E; R' \<subseteq> R\<rbrakk> \<Longrightarrow> indep R' E"
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   850
  by (auto simp: indep_def)
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   851
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   852
lemma ramsey2:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   853
  "\<exists>r\<ge>1. \<forall>(V::'a set) (E::'a set set). finite V \<and> card V \<ge> r \<longrightarrow>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   854
    (\<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: 71083
diff changeset
   855
proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   856
  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: 71083
diff changeset
   857
    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: 71464
diff changeset
   858
  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: 71083
diff changeset
   859
    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: 71083
diff changeset
   860
  proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   861
    from that
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   862
    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: 71083
diff changeset
   863
      by (metis card_le_inj card_lessThan finite_lessThan)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   864
    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: 71083
diff changeset
   865
    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: 71083
diff changeset
   866
      by (simp add: f_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   867
    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: 71464
diff changeset
   868
      and U: "U \<in> nsets {..<N} ([m,n] ! i)"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   869
      using N numeral_2_eq_2 by (auto simp: partn_lst_def monochromatic_def)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   870
    show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   871
    proof (intro exI conjI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   872
      show "v ` U \<subseteq> V"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   873
        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: 71083
diff changeset
   874
      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: 71083
diff changeset
   875
        using i unfolding numeral_2_eq_2
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   876
          using gi U u
81871
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   877
          unfolding image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq
e8ecc32d18c1 A variety of tweaks
paulson <lp15@cam.ac.uk>
parents: 81332
diff changeset
   878
          by (auto simp: f_def nsets_def card_image inj_on_subset split: if_splits)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   879
    qed
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   880
  qed
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   881
  then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   882
    using \<open>Suc 0 \<le> N\<close> by auto
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   883
qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   884
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   885
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   886
subsection \<open>Preliminaries for the infinitary version\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   887
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   888
subsubsection \<open>``Axiom'' of Dependent Choice\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   889
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   890
primrec choice :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   891
  where \<comment> \<open>An integer-indexed chain of choices\<close>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   892
    choice_0: "choice P r 0 = (SOME x. P x)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   893
  | choice_Suc: "choice P r (Suc n) = (SOME y. P y \<and> (choice P r n, y) \<in> r)"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   894
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   895
lemma choice_n:
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   896
  assumes P0: "P x0"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   897
    and Pstep: "\<And>x. P x \<Longrightarrow> \<exists>y. P y \<and> (x, y) \<in> r"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   898
  shows "P (choice P r n)"
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   899
proof (induct n)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   900
  case 0
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   901
  show ?case by (force intro: someI P0)
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   902
next
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   903
  case Suc
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   904
  then show ?case by (auto intro: someI2_ex [OF Pstep])
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   905
qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   906
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   907
lemma dependent_choice:
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   908
  assumes trans: "trans r"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   909
    and P0: "P x0"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   910
    and Pstep: "\<And>x. P x \<Longrightarrow> \<exists>y. P y \<and> (x, y) \<in> r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   911
  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
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   912
proof
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   913
  fix n
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   914
  show "P (choice P r n)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   915
    by (blast intro: choice_n [OF P0 Pstep])
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   916
next
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   917
  fix n m :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   918
  assume "n < m"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   919
  from Pstep [OF choice_n [OF P0 Pstep]] have "(choice P r k, choice P r (Suc k)) \<in> r" for k
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   920
    by (auto intro: someI2_ex)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   921
  then show "(choice P r n, choice P r m) \<in> r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   922
    by (auto intro: less_Suc_induct [OF \<open>n < m\<close>] transD [OF trans])
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   923
qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   924
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   925
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   926
subsubsection \<open>Partition functions\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   927
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   928
definition part_fn :: "nat \<Rightarrow> nat \<Rightarrow> 'a set \<Rightarrow> ('a set \<Rightarrow> nat) \<Rightarrow> bool"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67091
diff changeset
   929
  \<comment> \<open>the function \<^term>\<open>f\<close> partitions the \<^term>\<open>r\<close>-subsets of the typically
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67091
diff changeset
   930
      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: 71259
diff changeset
   931
  where "part_fn r s Y f \<longleftrightarrow> (f \<in> nsets Y r \<rightarrow> {..<s})"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   932
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67091
diff changeset
   933
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: 71259
diff changeset
   934
lemma part_fn_Suc_imp_part_fn:
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   935
  "\<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: 71259
diff changeset
   936
  by (simp add: part_fn_def nsets_def Pi_def subset_Diff_insert)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   937
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   938
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: 71259
diff changeset
   939
  unfolding part_fn_def nsets_def by blast
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   940
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   941
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   942
subsection \<open>Ramsey's Theorem: Infinitary Version\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   943
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   944
lemma Ramsey_induction:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   945
  fixes s r :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   946
    and YY :: "'a set"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   947
    and f :: "'a set \<Rightarrow> nat"
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   948
  assumes "infinite YY" "part_fn r s YY f"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   949
  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')"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   950
  using assms
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   951
proof (induct r arbitrary: YY f)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   952
  case 0
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   953
  then show ?case
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   954
    by (auto simp add: part_fn_def card_eq_0_iff cong: conj_cong)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   955
next
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   956
  case (Suc r)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   957
  show ?case
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   958
  proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   959
    from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   960
      by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   961
    let ?ramr = "{((y, Y, t), (y', Y', t')). y' \<in> Y \<and> Y' \<subseteq> Y}"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   962
    let ?propr = "\<lambda>(y, Y, t).
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   963
                 y \<in> YY \<and> y \<notin> Y \<and> Y \<subseteq> YY \<and> infinite Y \<and> t < s
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   964
                 \<and> (\<forall>X. X\<subseteq>Y \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert y) X = t)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   965
    from Suc.prems have infYY': "infinite (YY - {yy})" by auto
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   966
    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: 71259
diff changeset
   967
      by (simp add: o_def part_fn_Suc_imp_part_fn yy)
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   968
    have transr: "trans ?ramr" by (force simp: trans_def)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   969
    from Suc.hyps [OF infYY' partf']
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   970
    obtain Y0 and t0 where "Y0 \<subseteq> YY - {yy}" "infinite Y0" "t0 < s"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   971
      "X \<subseteq> Y0 \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yy) X = t0" for X
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   972
      by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   973
    with yy have propr0: "?propr(yy, Y0, t0)" by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   974
    have proprstep: "\<exists>y. ?propr y \<and> (x, y) \<in> ?ramr" if x: "?propr x" for x
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   975
    proof (cases x)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   976
      case (fields yx Yx tx)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   977
      with x obtain yx' where yx': "yx' \<in> Yx"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   978
        by (blast dest: infinite_imp_nonempty)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   979
      from fields x have infYx': "infinite (Yx - {yx'})" by auto
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   980
      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: 71259
diff changeset
   981
        by (simp add: o_def part_fn_Suc_imp_part_fn part_fn_subset [where YY=YY and Y=Yx])
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   982
      from Suc.hyps [OF infYx' partfx'] obtain Y' and t'
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   983
        where Y': "Y' \<subseteq> Yx - {yx'}" "infinite Y'" "t' < s"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   984
          "X \<subseteq> Y' \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yx') X = t'" for X
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   985
        by blast
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   986
      from fields x Y' yx' have "?propr (yx', Y', t') \<and> (x, (yx', Y', t')) \<in> ?ramr"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   987
        by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   988
      then show ?thesis ..
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   989
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   990
    from dependent_choice [OF transr propr0 proprstep]
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   991
    obtain g where pg: "?propr (g n)" and rg: "n < m \<Longrightarrow> (g n, g m) \<in> ?ramr" for n m :: nat
63060
293ede07b775 some uses of 'obtain' with structure statement;
wenzelm
parents: 61585
diff changeset
   992
      by blast
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   993
    let ?gy = "fst \<circ> g"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   994
    let ?gt = "snd \<circ> snd \<circ> g"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   995
    have rangeg: "\<exists>k. range ?gt \<subseteq> {..<k}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   996
    proof (intro exI subsetI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   997
      fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   998
      assume "x \<in> range ?gt"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   999
      then obtain n where "x = ?gt n" ..
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1000
      with pg [of n] show "x \<in> {..<s}" by (cases "g n") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1001
    qed
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69593
diff changeset
  1002
    from rangeg have "finite (range ?gt)"
a03a63b81f44 tuned proofs
haftmann
parents: 69593
diff changeset
  1003
      by (simp add: finite_nat_iff_bounded)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1004
    then obtain s' and n' where s': "s' = ?gt n'" and infeqs': "infinite {n. ?gt n = s'}"
54580
7b9336176a1c adapt to 9733ab5c1df6
traytel
parents: 53374
diff changeset
  1005
      by (rule inf_img_fin_domE) (auto simp add: vimage_def intro: infinite_UNIV_nat)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1006
    with pg [of n'] have less': "s'<s" by (cases "g n'") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1007
    have inj_gy: "inj ?gy"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1008
    proof (rule linorder_injI)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1009
      fix m m' :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1010
      assume "m < m'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1011
      from rg [OF this] pg [of m] show "?gy m \<noteq> ?gy m'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1012
        by (cases "g m", cases "g m'") auto
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1013
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1014
    show ?thesis
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1015
    proof (intro exI conjI)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1016
      from pg show "?gy ` {n. ?gt n = s'} \<subseteq> YY"
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1017
        by (auto simp add: Let_def split_beta)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1018
      from infeqs' show "infinite (?gy ` {n. ?gt n = s'})"
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1019
        by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1020
      show "s' < s" by (rule less')
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1021
      show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} \<and> finite X \<and> card X = Suc r \<longrightarrow> f X = s'"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1022
      proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1023
        have "f X = s'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1024
          if X: "X \<subseteq> ?gy ` {n. ?gt n = s'}"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1025
          and cardX: "finite X" "card X = Suc r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1026
          for X
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1027
        proof -
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1028
          from X obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1029
            by (auto simp add: subset_image_iff)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1030
          with cardX have "AA \<noteq> {}" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1031
          then have AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1032
          show ?thesis
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1033
          proof (cases "g (LEAST x. x \<in> AA)")
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1034
            case (fields ya Ya ta)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1035
            with AAleast Xeq have ya: "ya \<in> X" by (force intro!: rev_image_eqI)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1036
            then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1037
            also have "\<dots> = ta"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1038
            proof -
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1039
              have *: "X - {ya} \<subseteq> Ya"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1040
              proof
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1041
                fix x assume x: "x \<in> X - {ya}"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1042
                then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1043
                  by (auto simp add: Xeq)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1044
                with fields x have "a' \<noteq> (LEAST x. x \<in> AA)" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1045
                with Least_le [of "\<lambda>x. x \<in> AA", OF a'] have "(LEAST x. x \<in> AA) < a'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1046
                  by arith
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1047
                from xeq fields rg [OF this] show "x \<in> Ya" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1048
              qed
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1049
              have "card (X - {ya}) = r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1050
                by (simp add: cardX ya)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1051
              with pg [of "LEAST x. x \<in> AA"] fields cardX * show ?thesis
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1052
                by (auto simp del: insert_Diff_single)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1053
            qed
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1054
            also from AA AAleast fields have "\<dots> = s'" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1055
            finally show ?thesis .
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1056
          qed
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1057
        qed
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1058
        then show ?thesis by blast
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1059
      qed
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1060
    qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1061
  qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1062
qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1063
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1064
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1065
theorem Ramsey:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1066
  fixes s r :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1067
    and Z :: "'a set"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1068
    and f :: "'a set \<Rightarrow> nat"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1069
  shows
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1070
   "\<lbrakk>infinite Z;
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1071
      \<forall>X. X \<subseteq> Z \<and> finite X \<and> card X = r \<longrightarrow> f X < s\<rbrakk>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1072
    \<Longrightarrow> \<exists>Y t. Y \<subseteq> Z \<and> infinite Y \<and> t < s
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1073
            \<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: 71259
diff changeset
  1074
  by (blast intro: Ramsey_induction [unfolded part_fn_def nsets_def])
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1075
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1076
corollary Ramsey2:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1077
  fixes s :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1078
    and Z :: "'a set"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1079
    and f :: "'a set \<Rightarrow> nat"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1080
  assumes infZ: "infinite Z"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1081
    and part: "\<forall>x\<in>Z. \<forall>y\<in>Z. x \<noteq> y \<longrightarrow> f {x, y} < s"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1082
  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
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1083
proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1084
  from part have part2: "\<forall>X. X \<subseteq> Z \<and> finite X \<and> card X = 2 \<longrightarrow> f X < s"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
  1085
    by (fastforce simp: eval_nat_numeral card_Suc_eq)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1086
  obtain Y t where *:
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1087
    "Y \<subseteq> Z" "infinite Y" "t < s" "(\<forall>X. X \<subseteq> Y \<and> finite X \<and> card X = 2 \<longrightarrow> f X = t)"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1088
    by (insert Ramsey [OF infZ part2]) auto
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 46575
diff changeset
  1089
  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: 46575
diff changeset
  1090
  with * show ?thesis by iprover
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1091
qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1092
71848
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1093
corollary Ramsey_nsets:
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1094
  fixes f :: "'a set \<Rightarrow> nat"
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1095
  assumes "infinite Z" "f ` nsets Z r \<subseteq> {..<s}"
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1096
  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: 71472
diff changeset
  1097
  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: 71472
diff changeset
  1098
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1099
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
  1100
subsection \<open>Disjunctive Well-Foundedness\<close>
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1101
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
  1102
text \<open>
22367
6860f09242bf tuned document;
wenzelm
parents: 21634
diff changeset
  1103
  An application of Ramsey's theorem to program termination. See
76987
4c275405faae isabelle update -u cite;
wenzelm
parents: 73709
diff changeset
  1104
  \<^cite>\<open>"Podelski-Rybalchenko"\<close>.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
  1105
\<close>
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1106
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1107
definition disj_wf :: "('a \<times> 'a) set \<Rightarrow> bool"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1108
  where "disj_wf r \<longleftrightarrow> (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf (T i)) \<and> r = (\<Union>i<n. T i))"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1109
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1110
definition transition_idx :: "(nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> ('a \<times> 'a) set) \<Rightarrow> nat set \<Rightarrow> nat"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1111
  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
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1112
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1113
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1114
lemma transition_idx_less:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1115
  assumes "i < j" "(s j, s i) \<in> T k" "k < n"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1116
  shows "transition_idx s T {i, j} < n"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1117
proof -
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1118
  from assms(1,2) have "transition_idx s T {i, j} \<le> k"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1119
    by (simp add: transition_idx_def, blast intro: Least_le)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1120
  with assms(3) show ?thesis by simp
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1121
qed
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1122
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1123
lemma transition_idx_in:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1124
  assumes "i < j" "(s j, s i) \<in> T k"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1125
  shows "(s j, s i) \<in> T (transition_idx s T {i, j})"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1126
  using assms
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1127
  by (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR cong: conj_cong) (erule LeastI)
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1128
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1129
text \<open>To be equal to the union of some well-founded relations is equivalent
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1130
  to being the subset of such a union.\<close>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1131
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: 69661
diff changeset
  1132
proof -
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1133
  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: 69661
diff changeset
  1134
           \<Longrightarrow> (\<forall>i<n. wf (T i \<inter> r)) \<and> r = (\<Union>i<n. T i \<inter> r)"
79735
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
  1135
    by (force simp: wf_Int1)
71083
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1136
  show ?thesis
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1137
    unfolding disj_wf_def by auto (metis "*")
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1138
qed
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1139
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1140
theorem trans_disj_wf_implies_wf:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1141
  assumes "trans r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1142
    and "disj_wf r"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1143
  shows "wf r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1144
proof (simp only: wf_iff_no_infinite_down_chain, rule notI)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1145
  assume "\<exists>s. \<forall>i. (s (Suc i), s i) \<in> r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1146
  then obtain s where sSuc: "\<forall>i. (s (Suc i), s i) \<in> r" ..
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1147
  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)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1148
    by (auto simp add: disj_wf_def)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1149
  have s_in_T: "\<exists>k. (s j, s i) \<in> T k \<and> k<n" if "i < j" for i j
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1150
  proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1151
    from \<open>i < j\<close> have "(s j, s i) \<in> r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1152
    proof (induct rule: less_Suc_induct)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1153
      case 1
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1154
      then show ?case by (simp add: sSuc)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1155
    next
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1156
      case 2
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1157
      with \<open>trans r\<close> show ?case
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1158
        unfolding trans_def by blast
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1159
    qed
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1160
    then show ?thesis by (auto simp add: r)
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1161
  qed
71083
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1162
  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: 69661
diff changeset
  1163
    using s_in_T transition_idx_less by blast
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1164
  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: 69661
diff changeset
  1165
    by (metis doubleton_eq_iff less_linear)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1166
  have "\<exists>K k. K \<subseteq> UNIV \<and> infinite K \<and> k < n \<and>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1167
      (\<forall>i\<in>K. \<forall>j\<in>K. i \<noteq> j \<longrightarrow> transition_idx s T {i, j} = k)"
54580
7b9336176a1c adapt to 9733ab5c1df6
traytel
parents: 53374
diff changeset
  1168
    by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1169
  then obtain K and k where infK: "infinite K" and "k < n"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1170
    and allk: "\<forall>i\<in>K. \<forall>j\<in>K. i \<noteq> j \<longrightarrow> transition_idx s T {i, j} = k"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1171
    by auto
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1172
  have "(s (enumerate K (Suc m)), s (enumerate K m)) \<in> T k" for m :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1173
  proof -
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1174
    let ?j = "enumerate K (Suc m)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1175
    let ?i = "enumerate K m"
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1176
    have ij: "?i < ?j" by (simp add: enumerate_step infK)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1177
    have "?j \<in> K" "?i \<in> K" by (simp_all add: enumerate_in_set infK)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1178
    with ij have k: "k = transition_idx s T {?i, ?j}" by (simp add: allk)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1179
    from s_in_T [OF ij] obtain k' where "(s ?j, s ?i) \<in> T k'" "k'<n" by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1180
    then show "(s ?j, s ?i) \<in> T k" by (simp add: k transition_idx_in ij)
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1181
  qed
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1182
  then have "\<not> wf (T k)"
73709
58bd53caf800 things need to be ugly
paulson <lp15@cam.ac.uk>
parents: 73655
diff changeset
  1183
    unfolding wf_iff_no_infinite_down_chain by iprover
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1184
  with wfT \<open>k < n\<close> show False by blast
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1185
qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1186
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1187
end