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