src/HOL/Library/Ramsey.thy
author nipkow
Fri, 23 Jun 2006 13:42:19 +0200
changeset 19946 e3ddb0812840
parent 19944 60e0cbeae3d8
child 19948 1be283f3f1ba
permissions -rwxr-xr-x
beautification
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
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     2
    ID:         $Id$
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     3
    Author:     Tom Ridge. Converted to structured Isar by L C Paulson
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     4
*)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     5
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     6
header "Ramsey's Theorem"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     7
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     8
theory Ramsey imports Main begin
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     9
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    10
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    11
subsection{*``Axiom'' of Dependent Choice*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    12
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    13
consts choice :: "('a => bool) => (('a * 'a) set) => nat => 'a"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    14
  --{*An integer-indexed chain of choices*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    15
primrec
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    16
  choice_0:   "choice P r 0 = (SOME x. P x)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    17
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    18
  choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \<in> r)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    19
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    20
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    21
lemma choice_n: 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    22
  assumes P0: "P x0"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    23
      and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    24
  shows "P (choice P r n)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    25
 proof (induct n)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    26
   case 0 show ?case by (force intro: someI P0) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    27
 next
19946
e3ddb0812840 beautification
nipkow
parents: 19944
diff changeset
    28
   case Suc thus ?case by (auto intro: someI2_ex [OF Pstep]) 
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    29
 qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    30
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    31
lemma dependent_choice: 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    32
  assumes trans: "trans r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    33
      and P0: "P x0"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    34
      and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    35
  shows "\<exists>f::nat=>'a. (\<forall>n. P (f n)) & (\<forall>n m. n<m --> (f n, f m) \<in> r)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    36
proof (intro exI conjI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    37
  show "\<forall>n. P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep]) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    38
next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    39
  have PSuc: "\<forall>n. (choice P r n, choice P r (Suc n)) \<in> r" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    40
    using Pstep [OF choice_n [OF P0 Pstep]]
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    41
    by (auto intro: someI2_ex)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    42
  show "\<forall>n m. n<m --> (choice P r n, choice P r m) \<in> r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    43
  proof (intro strip)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    44
    fix n and m::nat
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    45
    assume less: "n<m"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    46
    show "(choice P r n, choice P r m) \<in> r" using PSuc
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    47
      by (auto intro: less_Suc_induct [OF less] transD [OF trans])
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    48
  qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    49
qed 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    50
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    51
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    52
subsection {*Partitions of a Set*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    53
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    54
constdefs part :: "nat => nat => 'a set => ('a set => nat) => bool"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    55
  --{*the function @{term f} partitions the @{term r}-subsets of the typically
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    56
       infinite set @{term Y} into @{term s} distinct categories.*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    57
  "part r s Y f == \<forall>X. X \<subseteq> Y & finite X & card X = r --> f X < s"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    58
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    59
text{*For induction, we decrease the value of @{term r} in partitions.*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    60
lemma part_Suc_imp_part:
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    61
     "[| infinite Y; part (Suc r) s Y f; y \<in> Y |] 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    62
      ==> part r s (Y - {y}) (%u. f (insert y u))"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    63
  apply(simp add: part_def, clarify)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    64
  apply(drule_tac x="insert y X" in spec)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    65
  apply(force simp:card_Diff_singleton_if)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    66
  done
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    67
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    68
lemma part_subset: "part r s YY f ==> Y \<subseteq> YY ==> part r s Y f" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    69
  by (simp add: part_def, blast)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    70
  
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    71
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    72
subsection {*Ramsey's Theorem: Infinitary Version*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    73
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    74
lemma ramsey_induction: 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    75
  fixes s::nat and r::nat
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    76
  shows
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    77
  "!!(YY::'a set) (f::'a set => nat). 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    78
      [|infinite YY; part r s YY f|]
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    79
      ==> \<exists>Y' t'. Y' \<subseteq> YY & infinite Y' & t' < s & 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    80
                  (\<forall>X. X \<subseteq> Y' & finite X & card X = r --> f X = t')"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    81
proof (induct r)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    82
  case 0
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    83
  thus ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    84
next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    85
  case (Suc r) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    86
  show ?case
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    87
  proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    88
    from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY" by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    89
    let ?ramr = "{((y,Y,t),(y',Y',t')). y' \<in> Y & Y' \<subseteq> Y}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    90
    let ?propr = "%(y,Y,t).     
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    91
		 y \<in> YY & y \<notin> Y & Y \<subseteq> YY & infinite Y & t < s
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    92
		 & (\<forall>X. X\<subseteq>Y & finite X & card X = r --> (f o insert y) X = t)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    93
    have infYY': "infinite (YY-{yy})" using Suc.prems by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    94
    have partf': "part r s (YY - {yy}) (f \<circ> insert yy)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    95
      by (simp add: o_def part_Suc_imp_part yy Suc.prems)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    96
    have transr: "trans ?ramr" by (force simp add: trans_def) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    97
    from Suc.hyps [OF infYY' partf']
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    98
    obtain Y0 and t0
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    99
    where "Y0 \<subseteq> YY - {yy}"  "infinite Y0"  "t0 < s"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   100
          "\<forall>X. X\<subseteq>Y0 \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yy) X = t0"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   101
        by blast 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   102
    with yy have propr0: "?propr(yy,Y0,t0)" by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   103
    have proprstep: "\<And>x. ?propr x \<Longrightarrow> \<exists>y. ?propr y \<and> (x, y) \<in> ?ramr" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   104
    proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   105
      fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   106
      assume px: "?propr x" thus "?thesis x"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   107
      proof (cases x)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   108
        case (fields yx Yx tx)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   109
        then obtain yx' where yx': "yx' \<in> Yx" using px
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   110
               by (blast dest: infinite_imp_nonempty)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   111
        have infYx': "infinite (Yx-{yx'})" using fields px by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   112
        with fields px yx' Suc.prems
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   113
        have partfx': "part r s (Yx - {yx'}) (f \<circ> insert yx')"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   114
          by (simp add: o_def part_Suc_imp_part part_subset [where ?YY=YY]) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   115
	from Suc.hyps [OF infYx' partfx']
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   116
	obtain Y' and t'
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   117
	where Y': "Y' \<subseteq> Yx - {yx'}"  "infinite Y'"  "t' < s"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   118
	       "\<forall>X. X\<subseteq>Y' \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yx') X = t'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   119
	    by blast 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   120
	show ?thesis
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   121
	proof
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   122
	  show "?propr (yx',Y',t') & (x, (yx',Y',t')) \<in> ?ramr"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   123
  	    using fields Y' yx' px by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   124
	qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   125
      qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   126
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   127
    from dependent_choice [OF transr propr0 proprstep]
19946
e3ddb0812840 beautification
nipkow
parents: 19944
diff changeset
   128
    obtain g where pg: "!!n::nat.  ?propr (g n)"
e3ddb0812840 beautification
nipkow
parents: 19944
diff changeset
   129
      and rg: "!!n m. n<m ==> (g n, g m) \<in> ?ramr" by force
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   130
    let ?gy = "(\<lambda>n. let (y,Y,t) = g n in y)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   131
    let ?gt = "(\<lambda>n. let (y,Y,t) = g n in t)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   132
    have rangeg: "\<exists>k. range ?gt \<subseteq> {..<k}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   133
    proof (intro exI subsetI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   134
      fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   135
      assume "x \<in> range ?gt"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   136
      then obtain n where "x = ?gt n" ..
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   137
      with pg [of n] show "x \<in> {..<s}" by (cases "g n") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   138
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   139
    have "\<exists>s' \<in> range ?gt. infinite (?gt -` {s'})" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   140
     by (rule inf_img_fin_dom [OF _ nat_infinite]) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   141
        (simp add: finite_nat_iff_bounded rangeg)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   142
    then obtain s' and n'
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   143
            where s':      "s' = ?gt n'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   144
              and infeqs': "infinite {n. ?gt n = s'}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   145
       by (auto simp add: vimage_def)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   146
    with pg [of n'] have less': "s'<s" by (cases "g n'") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   147
    have inj_gy: "inj ?gy"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   148
    proof (rule linorder_injI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   149
      fix m and m'::nat assume less: "m < m'" show "?gy m \<noteq> ?gy m'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   150
        using rg [OF less] pg [of m] by (cases "g m", cases "g m'", auto) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   151
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   152
    show ?thesis
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   153
    proof (intro exI conjI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   154
      show "?gy ` {n. ?gt n = s'} \<subseteq> YY" using pg
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   155
        by (auto simp add: Let_def split_beta) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   156
    next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   157
      show "infinite (?gy ` {n. ?gt n = s'})" using infeqs'
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   158
        by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   159
    next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   160
      show "s' < s" by (rule less')
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   161
    next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   162
      show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   163
          --> f X = s'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   164
      proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   165
        {fix X 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   166
         assume "X \<subseteq> ?gy ` {n. ?gt n = s'}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   167
            and cardX: "finite X" "card X = Suc r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   168
         then obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   169
             by (auto simp add: subset_image_iff) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   170
         with cardX have "AA\<noteq>{}" by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   171
         hence AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   172
         have "f X = s'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   173
         proof (cases "g (LEAST x. x \<in> AA)") 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   174
           case (fields ya Ya ta)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   175
           with AAleast Xeq 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   176
           have ya: "ya \<in> X" by (force intro!: rev_image_eqI) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   177
           hence "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   178
           also have "... = ta" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   179
           proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   180
             have "X - {ya} \<subseteq> Ya"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   181
             proof 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   182
               fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   183
               assume x: "x \<in> X - {ya}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   184
               then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   185
                 by (auto simp add: Xeq) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   186
               hence "a' \<noteq> (LEAST x. x \<in> AA)" using x fields by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   187
               hence lessa': "(LEAST x. x \<in> AA) < a'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   188
                 using Least_le [of "%x. x \<in> AA", OF a'] by arith
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   189
               show "x \<in> Ya" using xeq fields rg [OF lessa'] by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   190
             qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   191
             moreover
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   192
             have "card (X - {ya}) = r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   193
               by (simp add: card_Diff_singleton_if cardX ya)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   194
             ultimately show ?thesis 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   195
               using pg [of "LEAST x. x \<in> AA"] fields cardX
19946
e3ddb0812840 beautification
nipkow
parents: 19944
diff changeset
   196
	       by (clarsimp simp del:insert_Diff_single)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   197
           qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   198
           also have "... = s'" using AA AAleast fields by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   199
           finally show ?thesis .
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   200
         qed}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   201
        thus ?thesis by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   202
      qed 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   203
    qed 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   204
  qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   205
qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   206
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   207
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   208
text{*Repackaging of Tom Ridge's final result*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   209
theorem Ramsey:
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   210
  fixes s::nat and r::nat and Z::"'a set" and f::"'a set => nat"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   211
  shows
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   212
   "[|infinite Z;
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   213
      \<forall>X. X \<subseteq> Z & finite X & card X = r --> f X < s|]
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   214
  ==> \<exists>Y t. Y \<subseteq> Z & infinite Y & t < s 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   215
            & (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X = t)"
19946
e3ddb0812840 beautification
nipkow
parents: 19944
diff changeset
   216
by (blast intro: ramsey_induction [unfolded part_def]) 
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   217
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   218
end
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   219