src/HOL/Library/Ramsey.thy
author wenzelm
Sat, 28 Dec 2024 23:44:56 +0100
changeset 81683 b31d09029b94
parent 81332 f94b30fa2b6c
child 81871 e8ecc32d18c1
permissions -rw-r--r--
clarified signature: more direct indent_markup; minor performance tuning: bypass "ind" buffer when unused;
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
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
    30
lemma nsets_2_eq: "nsets A 2 = (\<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
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    44
lemma nsets_3_eq: "nsets A 3 = (\<Union>x\<in>A. \<Union>y\<in>A - {x}. \<Union>z\<in>A - {x,y}. {{x,y,z}})"
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
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    52
  show "?rhs \<subseteq> [A]\<^bsup>4\<^esup>"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    53
    apply (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    54
    by (metis insert_iff singletonD)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    55
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    56
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    57
lemma nsets_disjoint_2:
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    58
  "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
    59
  by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    60
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    61
lemma ordered_nsets_2_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    62
  fixes A :: "'a::linorder set"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    63
  shows "nsets A 2 = {{x,y} | x y. x \<in> A \<and> y \<in> A \<and> x<y}"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    64
     (is "_ = ?rhs")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    65
proof
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    66
  show "nsets A 2 \<subseteq> ?rhs"
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    67
    unfolding numeral_nat
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    68
    apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less)
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    69
    by (metis antisym)
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    70
  show "?rhs \<subseteq> nsets A 2"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    71
    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
    72
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    73
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    74
lemma ordered_nsets_3_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    75
  fixes A :: "'a::linorder set"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    76
  shows "nsets A 3 = {{x,y,z} | x y z. x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> x<y \<and> y<z}"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    77
     (is "_ = ?rhs")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    78
proof
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    79
  show "nsets A 3 \<subseteq> ?rhs"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    80
    apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    81
    by (metis insert_commute linorder_cases)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    82
  show "?rhs \<subseteq> nsets A 3"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    83
    apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    84
  by (metis empty_iff insert_iff not_less_iff_gr_or_eq)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    85
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    86
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    87
lemma ordered_nsets_4_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
    88
  fixes A :: "'a::linorder set"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    89
  shows "[A]\<^bsup>4\<^esup> = {U. \<exists>u x y z. U = {u,x,y,z} \<and> u \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < x \<and> x < y \<and> y < z}"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    90
    (is "_ = Collect ?RHS")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    91
proof -
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    92
  have "?RHS U" if "U \<in> [A]\<^bsup>4\<^esup>" for U
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    93
  proof -
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    94
    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
    95
      by (simp add: nsets_def) (metis finite_set_strict_sorted)
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    96
    then show ?thesis
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    97
      unfolding numeral_nat length_Suc_conv by auto blast
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
    98
  qed
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
    99
  moreover
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   100
  have "Collect ?RHS \<subseteq> [A]\<^bsup>4\<^esup>"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   101
    apply (clarsimp simp add: nsets_def eval_nat_numeral)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   102
    apply (subst card_insert_disjoint, auto)+
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   103
    done
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   104
  ultimately show ?thesis
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   105
    by auto
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   106
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   107
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   108
lemma ordered_nsets_5_eq:
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   109
  fixes A :: "'a::linorder set"
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   110
  shows "[A]\<^bsup>5\<^esup> = {U. \<exists>u v x y z. U = {u,v,x,y,z} \<and> u \<in> A \<and> v \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < v \<and> v < x \<and> x < y \<and> y < z}"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   111
    (is "_ = Collect ?RHS")
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   112
proof -
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   113
  have "?RHS U" if "U \<in> [A]\<^bsup>5\<^esup>" for U
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   114
  proof -
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   115
    from that obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \<subseteq> A"
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   116
      apply (simp add: nsets_def)
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   117
      by (metis finite_set_strict_sorted)
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   118
    then show ?thesis
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   119
      unfolding numeral_nat length_Suc_conv by auto blast
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   120
  qed
71464
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   121
  moreover
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   122
  have "Collect ?RHS \<subseteq> [A]\<^bsup>5\<^esup>"
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   123
    apply (clarsimp simp add: nsets_def eval_nat_numeral)
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   124
    apply (subst card_insert_disjoint, auto)+
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   125
    done
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   126
  ultimately show ?thesis
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   127
    by auto
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   128
qed
4a04b6bd628b a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71453
diff changeset
   129
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   130
lemma binomial_eq_nsets: "n choose k = card (nsets {0..<n} k)"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   131
  apply (simp add: binomial_def nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   132
  by (meson subset_eq_atLeast0_lessThan_finite)
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}}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   160
  unfolding nsets_def
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   161
  apply auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   162
  by (metis add.left_neutral lessThan_atLeast0 lessThan_iff subset_card_intvl_is_intvl)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   163
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   164
lemma nsets_zero [simp]: "nsets A 0 = {{}}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   165
  by (auto simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   166
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   167
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
   168
  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
   169
71405
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   170
lemma inj_on_nsets:
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   171
  assumes "inj_on f A"
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   172
  shows "inj_on (\<lambda>X. f ` X) ([A]\<^bsup>n\<^esup>)"
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   173
  using assms unfolding nsets_def
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   174
  by (metis (no_types, lifting) inj_on_inverseI inv_into_image_cancel mem_Collect_eq)
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 -
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   180
  have "(`) f ` [A]\<^bsup>n\<^esup> = [f ` A]\<^bsup>n\<^esup>"
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   181
    using assms
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   182
    apply (auto simp: nsets_def bij_betw_def image_iff card_image inj_on_subset)
3ab52e4a8b45 Two lemmas about nsets
paulson <lp15@cam.ac.uk>
parents: 71260
diff changeset
   183
    by (metis card_image inj_on_finite order_refl subset_image_inj)
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"
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   191
  using assms
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   192
  apply (clarsimp simp add: nsets_def subset_image_iff)
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   193
  by (metis card_image finite_imageD inj_on_subset)
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   194
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   195
lemma nsets_image_funcset:
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   196
  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
   197
  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
   198
    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
   199
    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
   200
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   201
lemma nsets_compose_image_funcset:
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   202
  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
   203
  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
   204
proof -
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   205
  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
   206
    using assms by (simp add: nsets_image_funcset)
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   207
  then show ?thesis
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   208
    using f by fastforce
71452
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   209
qed
9edb7fb69bc2 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 71405
diff changeset
   210
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
   211
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
   212
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   213
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
   214
  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
   215
  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
   216
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
   217
  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
   218
    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
   219
  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
   220
  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
   221
    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
   222
  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
   223
  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
   224
    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
   225
  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
   226
    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
   227
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
   228
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   230
  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
   231
  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
   232
  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
   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 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
   235
  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
   236
  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
   237
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
   238
  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
   239
    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
   240
  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
   241
    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
   242
  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
   243
    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
   244
  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
   245
    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
   246
  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
   247
    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
   248
  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
   249
    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
   250
  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
   251
    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
   252
    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
   253
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
   254
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   256
  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
   257
  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
   258
  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
   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 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
   261
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
   262
  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
   263
  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
   264
    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
   265
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
   266
  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
   267
  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
   268
    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
   269
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
   270
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   271
subsubsection \<open>Partition predicates\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   272
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
   273
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
   274
d11cee9c3a7c Some new material about Ramsey's theorem, 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
text \<open>uniform partition sizes\<close>
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   276
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
   277
  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
   278
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
   279
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
   280
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
   281
  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
   282
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   283
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
   284
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
   285
  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
   286
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   288
  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
   289
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   290
lemma partn_lst_greater_resource:
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   291
  fixes M::nat
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   292
  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
   293
  shows "partn_lst {..<N} \<alpha> \<gamma>"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   294
proof (clarsimp simp: partn_lst_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   295
  fix f
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   296
  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
   297
  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
   298
    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
   299
  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
   300
    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
   301
  have "H \<in> nsets {..<N} (\<alpha> ! i)"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   302
    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
   303
  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
   304
    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
   305
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   306
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   307
lemma partn_lst_fewer_colours:
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   308
  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
   309
  shows "partn_lst \<beta> \<alpha> \<gamma>"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   310
proof (clarsimp simp: partn_lst_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   311
  fix f :: "'a set \<Rightarrow> nat"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   312
  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
   313
  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
   314
      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
   315
      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
   316
    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
   317
    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
   318
  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
   319
  proof (cases i)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   320
    case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   321
    then have "[H]\<^bsup>\<gamma>\<^esup> = {}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   322
      using hom by blast
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   323
    then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   324
      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
   325
      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
   326
  next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   327
    case (Suc i')
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   328
    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
   329
      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
   330
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   331
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   332
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   333
lemma partn_lst_eq_partn: "partn_lst {..<n} [m,m] 2 = partn {..<n} m 2 {..<2::nat}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   334
  apply (simp add: partn_lst_def partn_def numeral_2_eq_2)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   335
  by (metis less_2_cases numeral_2_eq_2 lessThan_iff nth_Cons_0 nth_Cons_Suc)
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   336
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   337
lemma partn_lstE:
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   338
  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
   339
  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
   340
  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
   341
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   342
lemma partn_lst_less:
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   343
  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
   344
    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
   345
  shows "partn_lst \<beta> \<alpha>' n"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   346
proof (clarsimp simp: partn_lst_def)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   347
  fix f
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   348
  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
   349
  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
   350
                   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
   351
                   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
   352
    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
   353
  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
   354
    by (metis ex_bij_betw_finite_nat)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   355
  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
   356
    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
   357
  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
   358
  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
   359
    unfolding monochromatic_def
79589
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   360
  proof (intro exI bexI conjI)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   361
    show "i < length \<alpha>'"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   362
      by (simp add: assms(2) i)
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   363
    have "H' \<subseteq> H"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   364
      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
   365
      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
   366
    then have "finite H'"
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   367
      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
   368
    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
   369
      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
   370
    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
   371
      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
   372
    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
   373
      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
   374
  qed
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   375
qed
9dee3b4fdb06 new lemmas involving Ramsey numbers, infinite sets
paulson <lp15@cam.ac.uk>
parents: 76987
diff changeset
   376
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   377
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   378
subsection \<open>Finite versions of Ramsey's theorem\<close>
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   379
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   380
text \<open>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   381
  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
   382
  names are used (ramsey vs Ramsey).
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   383
\<close>
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   384
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
   385
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
   386
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   387
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
   388
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
   389
  where "ES 0 k l = max k l"
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   390
  |     "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
   391
            (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
   392
             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
   393
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   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
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
   397
  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
   398
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   399
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
   400
  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
   401
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   403
  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
   404
  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
   405
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   407
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
   408
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
   409
  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
   410
  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
   411
    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
   412
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
   413
  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
   414
  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
   415
    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
   416
  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
   417
    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
   418
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
   419
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   420
subsubsection \<open>Trivial cases\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   421
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   422
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
   423
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
   424
  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
   425
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   426
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
   427
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
   428
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
   429
  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
   430
    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
   431
  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
   432
    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
   433
    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
   434
      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
   435
    moreover have "A 0 \<inter> A 1 = {}"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   436
      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
   437
    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
   438
      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
   439
    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
   440
      by linarith
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   441
    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
   442
      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
   443
    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
   444
      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
   445
    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
   446
      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
   447
    then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   448
      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
   449
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   450
  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
   451
    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
   452
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   453
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
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
   455
  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
   456
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   457
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
   458
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
   459
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
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
   461
  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
   462
  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
   463
    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
   464
    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
   465
  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
   466
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
   467
  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
   468
    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
   469
  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
   470
    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
   471
    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
   472
      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
   473
    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
   474
      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
   475
    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
   476
      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
   477
      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
   478
    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
   479
      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
   480
    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
   481
      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
   482
    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
   483
    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
   484
      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
   485
      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
   486
        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
   487
      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
   488
        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
   489
      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
   490
      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
   491
        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
   492
          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
   493
        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
   494
          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
   495
          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
   496
      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
   497
      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
   498
      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
   499
        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
   500
      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
   501
        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
   502
        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
   503
      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
   504
        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
   505
      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
   506
        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
   507
      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
   508
      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
   509
        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
   510
      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
   511
      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
   512
        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
   513
        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
   514
          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
   515
        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
   516
          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
   517
          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
   518
          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
   519
        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
   520
          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
   521
        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
   522
          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
   523
        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
   524
          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
   525
            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
   526
          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
   527
            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
   528
          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
   529
            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
   530
        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
   531
        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
   532
        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
   533
          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
   534
          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
   535
            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
   536
          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
   537
            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
   538
          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
   539
            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
   540
          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
   541
            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
   542
          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
   543
            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
   544
            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
   545
          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
   546
            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
   547
        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
   548
          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
   549
          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
   550
            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
   551
          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
   552
            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
   553
            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
   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
            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
   556
            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
   557
        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
   558
        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
   559
          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
   560
        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
   561
          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
   562
      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
   563
        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
   564
        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
   565
          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
   566
        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
   567
          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
   568
          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
   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
          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
   571
      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
   572
    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
   573
      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
   574
      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
   575
        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
   576
      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
   577
        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
   578
      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
   579
      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
   580
        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
   581
          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
   582
        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
   583
          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
   584
          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
   585
      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
   586
      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
   587
      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
   588
        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
   589
      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
   590
        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
   591
        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
   592
      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
   593
        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
   594
      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
   595
        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
   596
      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
   597
      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
   598
        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
   599
      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
   600
      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
   601
        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
   602
        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
   603
          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
   604
        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
   605
          using u_nsets [of _ "q2 - 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
   606
          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
   607
          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
   608
        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
   609
          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
   610
        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
   611
          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
   612
        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
   613
          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
   614
            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
   615
          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
   616
            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
   617
          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
   618
            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
   619
        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
   620
        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
   621
        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
   622
          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
   623
          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
   624
            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
   625
          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
   626
            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
   627
          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
   628
            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
   629
          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
   630
            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
   631
          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
   632
            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
   633
            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
   634
          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
   635
            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
   636
        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
   637
          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
   638
          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
   639
            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
   640
          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
   641
            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
   642
            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
   643
          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
   644
            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
   645
            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
   646
        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
   647
        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
   648
          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
   649
        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
   650
          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
   651
      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
   652
        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
   653
        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
   654
          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
   655
        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
   656
          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
   657
          apply (clarsimp simp add: h_def image_subset_iff 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
   658
          by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj)
d11cee9c3a7c Some new material about Ramsey's theorem, 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
        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
   660
          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
   661
      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
   662
    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
   663
  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
   664
  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
   665
    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
   666
    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
   667
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
   668
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   670
proof (induction r arbitrary: q1 q2)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   671
  case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   672
  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
   673
    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
   674
next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   675
  case (Suc r)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   676
  note outer = this
71472
c213d067e60f Moved a number of general-purpose lemmas into HOL
paulson <lp15@cam.ac.uk>
parents: 71464
diff changeset
   677
  show ?case
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   678
  proof (cases "r = 0")
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   679
    case True
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   680
    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
   681
      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
   682
  next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   683
    case False
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   684
    then have "r > 0"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   685
      by simp
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   686
    show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   687
      using Suc.prems
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   688
    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
   689
      case 0
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 81142
diff changeset
   690
      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
   691
    next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   692
      case (Suc k)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   693
      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
   694
      then show ?case
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   695
      proof cases
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   696
        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
   697
        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
   698
          by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   699
      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
   700
        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
   701
        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
   702
        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
   703
        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
   704
        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
   705
        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
   706
              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
   707
          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
   708
        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
   709
          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
   710
        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
   711
          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
   712
      qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   713
    qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   714
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   715
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   716
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   717
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
   718
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   719
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
   720
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
   721
  case 0
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   722
  then show ?case
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   723
    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
   724
next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   725
  case (Suc k)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   726
  note IH = this
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   727
  show ?case
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   728
  proof (cases k)
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
    with Suc obtain q where "qs = [q]"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   731
      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
   732
    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
   733
      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
   734
  next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   735
    case (Suc k')
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   736
    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
   737
      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
   738
    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
   739
      using ramsey2_full by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   740
    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
   741
      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
   742
    have keq: "Suc (length l) = k"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   743
      using IH qs by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   744
    show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   745
    proof (intro exI conjI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   746
      show "partn_lst {..<p} qs r"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   747
      proof (auto simp: partn_lst_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   748
        fix f
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   749
        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
   750
        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
   751
        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
   752
          unfolding g_def using f Suc IH
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   753
          by (auto simp: Pi_def not_less)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   754
        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
   755
                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
   756
          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
   757
        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
   758
        proof (cases "i = 0")
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   759
          case True
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   760
          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
   761
            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
   762
          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
   763
            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
   764
          then have Usub: "U \<subseteq> {..<p}"
73655
26a1d66b9077 tuned proofs --- avoid z3, which is absent on arm64-linux;
wenzelm
parents: 72231
diff changeset
   765
            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
   766
          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
   767
          proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   768
            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
   769
              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
   770
              by (force simp: nsets_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   771
            then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   772
              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
   773
              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
   774
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   775
          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
   776
          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
   777
          proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   778
            have "u ` X \<in> nsets U r"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   779
              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
   780
            then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   781
              using f01 by auto
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
          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
   784
            unfolding h_def by blast
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   785
          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
   786
            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
   787
            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
   788
          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
   789
            unfolding monochromatic_def
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   790
          proof (intro exI conjI bexI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   791
            show "j < length qs"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   792
              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
   793
            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
   794
              apply (clarsimp simp add: nsets_def image_iff)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   795
              by (metis card_eq_0_iff card_image image_is_empty subset_image_inj)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   796
            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
   797
              by (auto simp: h_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   798
            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
   799
              using hj by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   800
            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
   801
              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
   802
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   803
        next
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   804
          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
   805
          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
   806
            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
   807
          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
   808
            unfolding monochromatic_def
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   809
          proof (intro exI conjI bexI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   810
            show "Suc i < length qs"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   811
              using Suc.hyps(2) i by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   812
            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
   813
              using False by (auto simp: eq)
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   814
            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
   815
              using False U qs by auto
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   816
          qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   817
        qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   818
      qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   819
    qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   820
  qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   821
qed
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   822
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   823
subsubsection \<open>Simple graph version\<close>
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   824
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   825
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
   826
  sets, i.e. the version for graphs and 2 colours.
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   827
\<close>
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   828
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   829
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
   830
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
   831
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
   832
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
   833
  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
   834
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   835
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
   836
  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
   837
d11cee9c3a7c Some new material about Ramsey's theorem, also sharpening the proof to deliver the Erdős–Szekeres upper bound on Ramsey numbers
paulson <lp15@cam.ac.uk>
parents: 79589
diff changeset
   838
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
   839
  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
   840
d11cee9c3a7c Some new material about Ramsey's theorem, 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
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
   842
  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
   843
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   844
lemma ramsey2:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   845
  "\<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
   846
    (\<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
   847
proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   848
  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
   849
    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
   850
  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
   851
    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
   852
  proof -
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   853
    from that
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   854
    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
   855
      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
   856
    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
   857
    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
   858
      by (simp add: f_def)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   859
    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
   860
      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
   861
      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
   862
    show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   863
    proof (intro exI conjI)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   864
      show "v ` U \<subseteq> V"
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   865
        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
   866
      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
   867
        using i unfolding numeral_2_eq_2
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   868
          using gi U u
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   869
          apply (simp add: image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   870
          apply (auto simp: f_def nsets_def card_image inj_on_subset split: if_split_asm)
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   871
          done
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   872
    qed
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   873
  qed
71259
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   874
  then show ?thesis
09aee7f5b447 Ramsey with multiple colours and arbitrary exponents
paulson <lp15@cam.ac.uk>
parents: 71083
diff changeset
   875
    using \<open>Suc 0 \<le> N\<close> by auto
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   876
qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   877
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   878
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
   879
subsection \<open>Preliminaries for the infinitary version\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   880
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   881
subsubsection \<open>``Axiom'' of Dependent Choice\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   882
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   883
primrec choice :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   884
  where \<comment> \<open>An integer-indexed chain of choices\<close>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   885
    choice_0: "choice P r 0 = (SOME x. P x)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   886
  | 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
   887
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   888
lemma choice_n:
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   889
  assumes P0: "P x0"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   890
    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
   891
  shows "P (choice P r n)"
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   892
proof (induct n)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   893
  case 0
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   894
  show ?case by (force intro: someI P0)
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   895
next
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   896
  case Suc
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   897
  then show ?case by (auto intro: someI2_ex [OF Pstep])
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   898
qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   899
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   900
lemma dependent_choice:
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   901
  assumes trans: "trans r"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   902
    and P0: "P x0"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   903
    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
   904
  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
   905
proof
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   906
  fix n
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   907
  show "P (choice P r n)"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   908
    by (blast intro: choice_n [OF P0 Pstep])
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   909
next
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   910
  fix n m :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   911
  assume "n < m"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   912
  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
   913
    by (auto intro: someI2_ex)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   914
  then show "(choice P r n, choice P r m) \<in> r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   915
    by (auto intro: less_Suc_induct [OF \<open>n < m\<close>] transD [OF trans])
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   916
qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   917
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   918
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   919
subsubsection \<open>Partition functions\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   920
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   921
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
   922
  \<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
   923
      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
   924
  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
   925
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67091
diff changeset
   926
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
   927
lemma part_fn_Suc_imp_part_fn:
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   928
  "\<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
   929
  by (simp add: part_fn_def nsets_def Pi_def subset_Diff_insert)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   930
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   931
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
   932
  unfolding part_fn_def nsets_def by blast
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   933
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   934
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   935
subsection \<open>Ramsey's Theorem: Infinitary Version\<close>
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   936
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   937
lemma Ramsey_induction:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   938
  fixes s r :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   939
    and YY :: "'a set"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   940
    and f :: "'a set \<Rightarrow> nat"
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   941
  assumes "infinite YY" "part_fn r s YY f"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   942
  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
   943
  using assms
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   944
proof (induct r arbitrary: YY f)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   945
  case 0
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   946
  then show ?case
71260
308baf6b450a corrected some confusing terminology / notation
paulson <lp15@cam.ac.uk>
parents: 71259
diff changeset
   947
    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
   948
next
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
   949
  case (Suc r)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   950
  show ?case
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   951
  proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   952
    from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   953
      by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   954
    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
   955
    let ?propr = "\<lambda>(y, Y, t).
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   956
                 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
   957
                 \<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
   958
    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
   959
    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
   960
      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
   961
    have transr: "trans ?ramr" by (force simp: trans_def)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   962
    from Suc.hyps [OF infYY' partf']
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   963
    obtain Y0 and t0 where "Y0 \<subseteq> YY - {yy}" "infinite Y0" "t0 < s"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   964
      "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
   965
      by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   966
    with yy have propr0: "?propr(yy, Y0, t0)" by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   967
    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
   968
    proof (cases x)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   969
      case (fields yx Yx tx)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   970
      with x obtain yx' where yx': "yx' \<in> Yx"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   971
        by (blast dest: infinite_imp_nonempty)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   972
      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
   973
      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
   974
        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
   975
      from Suc.hyps [OF infYx' partfx'] obtain Y' and t'
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   976
        where Y': "Y' \<subseteq> Yx - {yx'}" "infinite Y'" "t' < s"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   977
          "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
   978
        by blast
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   979
      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
   980
        by blast
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   981
      then show ?thesis ..
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   982
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   983
    from dependent_choice [OF transr propr0 proprstep]
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   984
    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
   985
      by blast
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   986
    let ?gy = "fst \<circ> g"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   987
    let ?gt = "snd \<circ> snd \<circ> g"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   988
    have rangeg: "\<exists>k. range ?gt \<subseteq> {..<k}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   989
    proof (intro exI subsetI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   990
      fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   991
      assume "x \<in> range ?gt"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   992
      then obtain n where "x = ?gt n" ..
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   993
      with pg [of n] show "x \<in> {..<s}" by (cases "g n") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   994
    qed
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69593
diff changeset
   995
    from rangeg have "finite (range ?gt)"
a03a63b81f44 tuned proofs
haftmann
parents: 69593
diff changeset
   996
      by (simp add: finite_nat_iff_bounded)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
   997
    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
   998
      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
   999
    with pg [of n'] have less': "s'<s" by (cases "g n'") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1000
    have inj_gy: "inj ?gy"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1001
    proof (rule linorder_injI)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1002
      fix m m' :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1003
      assume "m < m'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1004
      from rg [OF this] pg [of m] show "?gy m \<noteq> ?gy m'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1005
        by (cases "g m", cases "g m'") auto
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1006
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1007
    show ?thesis
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1008
    proof (intro exI conjI)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1009
      from pg show "?gy ` {n. ?gt n = s'} \<subseteq> YY"
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1010
        by (auto simp add: Let_def split_beta)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1011
      from infeqs' show "infinite (?gy ` {n. ?gt n = s'})"
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1012
        by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1013
      show "s' < s" by (rule less')
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1014
      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
  1015
      proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1016
        have "f X = s'"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1017
          if X: "X \<subseteq> ?gy ` {n. ?gt n = s'}"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1018
          and cardX: "finite X" "card X = Suc r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1019
          for X
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1020
        proof -
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1021
          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
  1022
            by (auto simp add: subset_image_iff)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1023
          with cardX have "AA \<noteq> {}" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1024
          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
  1025
          show ?thesis
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1026
          proof (cases "g (LEAST x. x \<in> AA)")
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1027
            case (fields ya Ya ta)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1028
            with AAleast Xeq have ya: "ya \<in> X" by (force intro!: rev_image_eqI)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1029
            then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1030
            also have "\<dots> = ta"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1031
            proof -
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1032
              have *: "X - {ya} \<subseteq> Ya"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1033
              proof
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1034
                fix x assume x: "x \<in> X - {ya}"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1035
                then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1036
                  by (auto simp add: Xeq)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1037
                with fields x have "a' \<noteq> (LEAST x. x \<in> AA)" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1038
                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
  1039
                  by arith
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1040
                from xeq fields rg [OF this] show "x \<in> Ya" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1041
              qed
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1042
              have "card (X - {ya}) = r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1043
                by (simp add: cardX ya)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1044
              with pg [of "LEAST x. x \<in> AA"] fields cardX * show ?thesis
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1045
                by (auto simp del: insert_Diff_single)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1046
            qed
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1047
            also from AA AAleast fields have "\<dots> = s'" by auto
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1048
            finally show ?thesis .
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1049
          qed
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1050
        qed
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1051
        then show ?thesis by blast
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1052
      qed
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1053
    qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1054
  qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1055
qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1056
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1057
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1058
theorem Ramsey:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1059
  fixes s r :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1060
    and Z :: "'a set"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1061
    and f :: "'a set \<Rightarrow> nat"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1062
  shows
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1063
   "\<lbrakk>infinite Z;
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1064
      \<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
  1065
    \<Longrightarrow> \<exists>Y t. Y \<subseteq> Z \<and> infinite Y \<and> t < s
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1066
            \<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
  1067
  by (blast intro: Ramsey_induction [unfolded part_fn_def nsets_def])
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1068
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1069
corollary Ramsey2:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1070
  fixes s :: nat
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1071
    and Z :: "'a set"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1072
    and f :: "'a set \<Rightarrow> nat"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1073
  assumes infZ: "infinite Z"
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1074
    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
  1075
  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
  1076
proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1077
  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
  1078
    by (fastforce simp: eval_nat_numeral card_Suc_eq)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1079
  obtain Y t where *:
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1080
    "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
  1081
    by (insert Ramsey [OF infZ part2]) auto
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 46575
diff changeset
  1082
  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
  1083
  with * show ?thesis by iprover
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1084
qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1085
71848
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1086
corollary Ramsey_nsets:
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1087
  fixes f :: "'a set \<Rightarrow> nat"
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 71472
diff changeset
  1088
  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
  1089
  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
  1090
  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
  1091
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1092
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
  1093
subsection \<open>Disjunctive Well-Foundedness\<close>
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1094
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
  1095
text \<open>
22367
6860f09242bf tuned document;
wenzelm
parents: 21634
diff changeset
  1096
  An application of Ramsey's theorem to program termination. See
76987
4c275405faae isabelle update -u cite;
wenzelm
parents: 73709
diff changeset
  1097
  \<^cite>\<open>"Podelski-Rybalchenko"\<close>.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
  1098
\<close>
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1099
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1100
definition disj_wf :: "('a \<times> 'a) set \<Rightarrow> bool"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1101
  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
  1102
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1103
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
  1104
  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
  1105
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1106
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1107
lemma transition_idx_less:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1108
  assumes "i < j" "(s j, s i) \<in> T k" "k < n"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1109
  shows "transition_idx s T {i, j} < n"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1110
proof -
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1111
  from assms(1,2) have "transition_idx s T {i, j} \<le> k"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1112
    by (simp add: transition_idx_def, blast intro: Least_le)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1113
  with assms(3) show ?thesis by simp
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1114
qed
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1115
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1116
lemma transition_idx_in:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1117
  assumes "i < j" "(s j, s i) \<in> T k"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1118
  shows "(s j, s i) \<in> T (transition_idx s T {i, j})"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1119
  using assms
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1120
  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
  1121
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1122
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
  1123
  to being the subset of such a union.\<close>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1124
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
  1125
proof -
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1126
  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
  1127
           \<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
  1128
    by (force simp: wf_Int1)
71083
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1129
  show ?thesis
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1130
    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
  1131
qed
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1132
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1133
theorem trans_disj_wf_implies_wf:
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1134
  assumes "trans r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1135
    and "disj_wf r"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1136
  shows "wf r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1137
proof (simp only: wf_iff_no_infinite_down_chain, rule notI)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1138
  assume "\<exists>s. \<forall>i. (s (Suc i), s i) \<in> r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1139
  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
  1140
  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
  1141
    by (auto simp add: disj_wf_def)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1142
  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
  1143
  proof -
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1144
    from \<open>i < j\<close> have "(s j, s i) \<in> r"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1145
    proof (induct rule: less_Suc_induct)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1146
      case 1
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1147
      then show ?case by (simp add: sSuc)
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1148
    next
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1149
      case 2
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1150
      with \<open>trans r\<close> show ?case
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1151
        unfolding trans_def by blast
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1152
    qed
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1153
    then show ?thesis by (auto simp add: r)
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1154
  qed
71083
ce92360f0692 A slight tidying up of messy proof steps
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  1155
  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
  1156
    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
  1157
  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
  1158
    by (metis doubleton_eq_iff less_linear)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1159
  have "\<exists>K k. K \<subseteq> UNIV \<and> infinite K \<and> k < n \<and>
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1160
      (\<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
  1161
    by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1162
  then obtain K and k where infK: "infinite K" and "k < n"
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1163
    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
  1164
    by auto
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1165
  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
  1166
  proof -
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1167
    let ?j = "enumerate K (Suc m)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1168
    let ?i = "enumerate K m"
46575
f1e387195a56 misc tuning;
wenzelm
parents: 44890
diff changeset
  1169
    have ij: "?i < ?j" by (simp add: enumerate_step infK)
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1170
    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
  1171
    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
  1172
    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
  1173
    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
  1174
  qed
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1175
  then have "\<not> wf (T k)"
73709
58bd53caf800 things need to be ugly
paulson <lp15@cam.ac.uk>
parents: 73655
diff changeset
  1176
    unfolding wf_iff_no_infinite_down_chain by iprover
65075
03e6aa683c4d misc tuning and modernization;
wenzelm
parents: 63060
diff changeset
  1177
  with wfT \<open>k < n\<close> show False by blast
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1178
qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
  1179
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
  1180
end