| author | haftmann | 
| Thu, 30 Sep 2010 08:50:45 +0200 | |
| changeset 39794 | 51451d73c3d4 | 
| parent 35175 | 61255c81da01 | 
| child 40077 | c8a9eaaa2f59 | 
| permissions | -rw-r--r-- | 
| 19944 | 1  | 
(* Title: HOL/Library/Ramsey.thy  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
2  | 
Author: Tom Ridge. Converted to structured Isar by L C Paulson  | 
| 19944 | 3  | 
*)  | 
4  | 
||
5  | 
header "Ramsey's Theorem"  | 
|
6  | 
||
| 25594 | 7  | 
theory Ramsey  | 
| 30738 | 8  | 
imports Main Infinite_Set  | 
| 25594 | 9  | 
begin  | 
| 19944 | 10  | 
|
| 22665 | 11  | 
subsection {* Preliminaries *}
 | 
| 19944 | 12  | 
|
| 22665 | 13  | 
subsubsection {* ``Axiom'' of Dependent Choice *}
 | 
| 19944 | 14  | 
|
| 34941 | 15  | 
primrec choice :: "('a => bool) => ('a * 'a) set => nat => 'a" where
 | 
| 19944 | 16  | 
  --{*An integer-indexed chain of choices*}
 | 
| 34941 | 17  | 
choice_0: "choice P r 0 = (SOME x. P x)"  | 
18  | 
| choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \<in> r)"  | 
|
| 19944 | 19  | 
|
20  | 
lemma choice_n:  | 
|
21  | 
assumes P0: "P x0"  | 
|
22  | 
and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"  | 
|
23  | 
shows "P (choice P r n)"  | 
|
| 19948 | 24  | 
proof (induct n)  | 
25  | 
case 0 show ?case by (force intro: someI P0)  | 
|
26  | 
next  | 
|
27  | 
case Suc thus ?case by (auto intro: someI2_ex [OF Pstep])  | 
|
28  | 
qed  | 
|
| 19944 | 29  | 
|
30  | 
lemma dependent_choice:  | 
|
31  | 
assumes trans: "trans r"  | 
|
32  | 
and P0: "P x0"  | 
|
33  | 
and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"  | 
|
| 19954 | 34  | 
obtains f :: "nat => 'a" where  | 
35  | 
"!!n. P (f n)" and "!!n m. n < m ==> (f n, f m) \<in> r"  | 
|
36  | 
proof  | 
|
37  | 
fix n  | 
|
38  | 
show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep])  | 
|
| 19944 | 39  | 
next  | 
40  | 
have PSuc: "\<forall>n. (choice P r n, choice P r (Suc n)) \<in> r"  | 
|
41  | 
using Pstep [OF choice_n [OF P0 Pstep]]  | 
|
42  | 
by (auto intro: someI2_ex)  | 
|
| 19954 | 43  | 
fix n m :: nat  | 
44  | 
assume less: "n < m"  | 
|
45  | 
show "(choice P r n, choice P r m) \<in> r" using PSuc  | 
|
46  | 
by (auto intro: less_Suc_induct [OF less] transD [OF trans])  | 
|
47  | 
qed  | 
|
| 19944 | 48  | 
|
49  | 
||
| 22665 | 50  | 
subsubsection {* Partitions of a Set *}
 | 
| 19944 | 51  | 
|
| 19948 | 52  | 
definition  | 
53  | 
  part :: "nat => nat => 'a set => ('a set => nat) => bool"
 | 
|
| 19944 | 54  | 
  --{*the function @{term f} partitions the @{term r}-subsets of the typically
 | 
55  | 
       infinite set @{term Y} into @{term s} distinct categories.*}
 | 
|
| 21634 | 56  | 
where  | 
| 19948 | 57  | 
"part r s Y f = (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X < s)"  | 
| 19944 | 58  | 
|
59  | 
text{*For induction, we decrease the value of @{term r} in partitions.*}
 | 
|
60  | 
lemma part_Suc_imp_part:  | 
|
61  | 
"[| infinite Y; part (Suc r) s Y f; y \<in> Y |]  | 
|
62  | 
      ==> part r s (Y - {y}) (%u. f (insert y u))"
 | 
|
63  | 
apply(simp add: part_def, clarify)  | 
|
64  | 
apply(drule_tac x="insert y X" in spec)  | 
|
| 24853 | 65  | 
apply(force)  | 
| 19944 | 66  | 
done  | 
67  | 
||
68  | 
lemma part_subset: "part r s YY f ==> Y \<subseteq> YY ==> part r s Y f"  | 
|
| 19948 | 69  | 
unfolding part_def by blast  | 
| 19944 | 70  | 
|
71  | 
||
| 22665 | 72  | 
subsection {* Ramsey's Theorem: Infinitary Version *}
 | 
| 19944 | 73  | 
|
| 19954 | 74  | 
lemma Ramsey_induction:  | 
75  | 
fixes s and r::nat  | 
|
| 19944 | 76  | 
shows  | 
77  | 
"!!(YY::'a set) (f::'a set => nat).  | 
|
78  | 
[|infinite YY; part r s YY f|]  | 
|
79  | 
==> \<exists>Y' t'. Y' \<subseteq> YY & infinite Y' & t' < s &  | 
|
80  | 
(\<forall>X. X \<subseteq> Y' & finite X & card X = r --> f X = t')"  | 
|
81  | 
proof (induct r)  | 
|
82  | 
case 0  | 
|
| 24853 | 83  | 
thus ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong)  | 
| 19944 | 84  | 
next  | 
85  | 
case (Suc r)  | 
|
86  | 
show ?case  | 
|
87  | 
proof -  | 
|
88  | 
from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY" by blast  | 
|
89  | 
    let ?ramr = "{((y,Y,t),(y',Y',t')). y' \<in> Y & Y' \<subseteq> Y}"
 | 
|
90  | 
let ?propr = "%(y,Y,t).  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
91  | 
y \<in> YY & y \<notin> Y & Y \<subseteq> YY & infinite Y & t < s  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
92  | 
& (\<forall>X. X\<subseteq>Y & finite X & card X = r --> (f o insert y) X = t)"  | 
| 19944 | 93  | 
    have infYY': "infinite (YY-{yy})" using Suc.prems by auto
 | 
94  | 
    have partf': "part r s (YY - {yy}) (f \<circ> insert yy)"
 | 
|
95  | 
by (simp add: o_def part_Suc_imp_part yy Suc.prems)  | 
|
96  | 
have transr: "trans ?ramr" by (force simp add: trans_def)  | 
|
97  | 
from Suc.hyps [OF infYY' partf']  | 
|
98  | 
obtain Y0 and t0  | 
|
99  | 
    where "Y0 \<subseteq> YY - {yy}"  "infinite Y0"  "t0 < s"
 | 
|
100  | 
"\<forall>X. X\<subseteq>Y0 \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yy) X = t0"  | 
|
101  | 
by blast  | 
|
102  | 
with yy have propr0: "?propr(yy,Y0,t0)" by blast  | 
|
103  | 
have proprstep: "\<And>x. ?propr x \<Longrightarrow> \<exists>y. ?propr y \<and> (x, y) \<in> ?ramr"  | 
|
104  | 
proof -  | 
|
105  | 
fix x  | 
|
106  | 
assume px: "?propr x" thus "?thesis x"  | 
|
107  | 
proof (cases x)  | 
|
108  | 
case (fields yx Yx tx)  | 
|
109  | 
then obtain yx' where yx': "yx' \<in> Yx" using px  | 
|
110  | 
by (blast dest: infinite_imp_nonempty)  | 
|
111  | 
        have infYx': "infinite (Yx-{yx'})" using fields px by auto
 | 
|
112  | 
with fields px yx' Suc.prems  | 
|
113  | 
        have partfx': "part r s (Yx - {yx'}) (f \<circ> insert yx')"
 | 
|
| 35175 | 114  | 
by (simp add: o_def part_Suc_imp_part part_subset [where YY=YY and Y=Yx])  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
115  | 
from Suc.hyps [OF infYx' partfx']  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
116  | 
obtain Y' and t'  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
117  | 
        where Y': "Y' \<subseteq> Yx - {yx'}"  "infinite Y'"  "t' < s"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
118  | 
"\<forall>X. X\<subseteq>Y' \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yx') X = t'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
119  | 
by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
120  | 
show ?thesis  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
121  | 
proof  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
122  | 
show "?propr (yx',Y',t') & (x, (yx',Y',t')) \<in> ?ramr"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
123  | 
using fields Y' yx' px by blast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
124  | 
qed  | 
| 19944 | 125  | 
qed  | 
126  | 
qed  | 
|
127  | 
from dependent_choice [OF transr propr0 proprstep]  | 
|
| 19946 | 128  | 
obtain g where pg: "!!n::nat. ?propr (g n)"  | 
| 19954 | 129  | 
and rg: "!!n m. n<m ==> (g n, g m) \<in> ?ramr" by blast  | 
| 28741 | 130  | 
let ?gy = "fst o g"  | 
131  | 
let ?gt = "snd o snd o g"  | 
|
| 19944 | 132  | 
    have rangeg: "\<exists>k. range ?gt \<subseteq> {..<k}"
 | 
133  | 
proof (intro exI subsetI)  | 
|
134  | 
fix x  | 
|
135  | 
assume "x \<in> range ?gt"  | 
|
136  | 
then obtain n where "x = ?gt n" ..  | 
|
137  | 
      with pg [of n] show "x \<in> {..<s}" by (cases "g n") auto
 | 
|
138  | 
qed  | 
|
| 20810 | 139  | 
have "finite (range ?gt)"  | 
140  | 
by (simp add: finite_nat_iff_bounded rangeg)  | 
|
| 19944 | 141  | 
then obtain s' and n'  | 
| 20810 | 142  | 
where s': "s' = ?gt n'"  | 
143  | 
        and infeqs': "infinite {n. ?gt n = s'}"
 | 
|
144  | 
by (rule inf_img_fin_domE) (auto simp add: vimage_def intro: nat_infinite)  | 
|
| 19944 | 145  | 
with pg [of n'] have less': "s'<s" by (cases "g n'") auto  | 
146  | 
have inj_gy: "inj ?gy"  | 
|
147  | 
proof (rule linorder_injI)  | 
|
| 19949 | 148  | 
fix m m' :: nat assume less: "m < m'" show "?gy m \<noteq> ?gy m'"  | 
| 19948 | 149  | 
using rg [OF less] pg [of m] by (cases "g m", cases "g m'") auto  | 
| 19944 | 150  | 
qed  | 
151  | 
show ?thesis  | 
|
152  | 
proof (intro exI conjI)  | 
|
153  | 
      show "?gy ` {n. ?gt n = s'} \<subseteq> YY" using pg
 | 
|
154  | 
by (auto simp add: Let_def split_beta)  | 
|
155  | 
      show "infinite (?gy ` {n. ?gt n = s'})" using infeqs'
 | 
|
156  | 
by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD)  | 
|
157  | 
show "s' < s" by (rule less')  | 
|
158  | 
      show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r 
 | 
|
159  | 
--> f X = s'"  | 
|
160  | 
proof -  | 
|
161  | 
        {fix X 
 | 
|
162  | 
         assume "X \<subseteq> ?gy ` {n. ?gt n = s'}"
 | 
|
163  | 
and cardX: "finite X" "card X = Suc r"  | 
|
164  | 
         then obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" 
 | 
|
165  | 
by (auto simp add: subset_image_iff)  | 
|
166  | 
         with cardX have "AA\<noteq>{}" by auto
 | 
|
167  | 
hence AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex)  | 
|
168  | 
have "f X = s'"  | 
|
169  | 
proof (cases "g (LEAST x. x \<in> AA)")  | 
|
170  | 
case (fields ya Ya ta)  | 
|
171  | 
with AAleast Xeq  | 
|
172  | 
have ya: "ya \<in> X" by (force intro!: rev_image_eqI)  | 
|
173  | 
           hence "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
 | 
|
174  | 
also have "... = ta"  | 
|
175  | 
proof -  | 
|
176  | 
             have "X - {ya} \<subseteq> Ya"
 | 
|
177  | 
proof  | 
|
| 19954 | 178  | 
               fix x assume x: "x \<in> X - {ya}"
 | 
| 19944 | 179  | 
then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA"  | 
180  | 
by (auto simp add: Xeq)  | 
|
181  | 
hence "a' \<noteq> (LEAST x. x \<in> AA)" using x fields by auto  | 
|
182  | 
hence lessa': "(LEAST x. x \<in> AA) < a'"  | 
|
183  | 
using Least_le [of "%x. x \<in> AA", OF a'] by arith  | 
|
184  | 
show "x \<in> Ya" using xeq fields rg [OF lessa'] by auto  | 
|
185  | 
qed  | 
|
186  | 
moreover  | 
|
187  | 
             have "card (X - {ya}) = r"
 | 
|
| 24853 | 188  | 
by (simp add: cardX ya)  | 
| 19944 | 189  | 
ultimately show ?thesis  | 
190  | 
using pg [of "LEAST x. x \<in> AA"] fields cardX  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
30738 
diff
changeset
 | 
191  | 
by (clarsimp simp del:insert_Diff_single)  | 
| 19944 | 192  | 
qed  | 
193  | 
also have "... = s'" using AA AAleast fields by auto  | 
|
194  | 
finally show ?thesis .  | 
|
195  | 
qed}  | 
|
196  | 
thus ?thesis by blast  | 
|
197  | 
qed  | 
|
198  | 
qed  | 
|
199  | 
qed  | 
|
200  | 
qed  | 
|
201  | 
||
202  | 
||
203  | 
theorem Ramsey:  | 
|
| 19949 | 204  | 
fixes s r :: nat and Z::"'a set" and f::"'a set => nat"  | 
| 19944 | 205  | 
shows  | 
206  | 
"[|infinite Z;  | 
|
207  | 
\<forall>X. X \<subseteq> Z & finite X & card X = r --> f X < s|]  | 
|
208  | 
==> \<exists>Y t. Y \<subseteq> Z & infinite Y & t < s  | 
|
209  | 
& (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X = t)"  | 
|
| 19954 | 210  | 
by (blast intro: Ramsey_induction [unfolded part_def])  | 
211  | 
||
212  | 
||
213  | 
corollary Ramsey2:  | 
|
214  | 
fixes s::nat and Z::"'a set" and f::"'a set => nat"  | 
|
215  | 
assumes infZ: "infinite Z"  | 
|
216  | 
      and part: "\<forall>x\<in>Z. \<forall>y\<in>Z. x\<noteq>y --> f{x,y} < s"
 | 
|
217  | 
shows  | 
|
218  | 
   "\<exists>Y t. Y \<subseteq> Z & infinite Y & t < s & (\<forall>x\<in>Y. \<forall>y\<in>Y. x\<noteq>y --> f{x,y} = t)"
 | 
|
219  | 
proof -  | 
|
220  | 
have part2: "\<forall>X. X \<subseteq> Z & finite X & card X = 2 --> f X < s"  | 
|
| 24853 | 221  | 
using part by (fastsimp simp add: nat_number card_Suc_eq)  | 
| 19954 | 222  | 
obtain Y t  | 
223  | 
where "Y \<subseteq> Z" "infinite Y" "t < s"  | 
|
224  | 
"(\<forall>X. X \<subseteq> Y & finite X & card X = 2 --> f X = t)"  | 
|
225  | 
by (insert Ramsey [OF infZ part2]) auto  | 
|
226  | 
  moreover from this have  "\<forall>x\<in>Y. \<forall>y\<in>Y. x \<noteq> y \<longrightarrow> f {x, y} = t" by auto
 | 
|
227  | 
ultimately show ?thesis by iprover  | 
|
228  | 
qed  | 
|
229  | 
||
230  | 
||
| 22665 | 231  | 
subsection {* Disjunctive Well-Foundedness *}
 | 
| 19954 | 232  | 
|
| 22367 | 233  | 
text {*
 | 
234  | 
An application of Ramsey's theorem to program termination. See  | 
|
235  | 
  \cite{Podelski-Rybalchenko}.
 | 
|
| 19954 | 236  | 
*}  | 
237  | 
||
| 20810 | 238  | 
definition  | 
| 19954 | 239  | 
  disj_wf         :: "('a * 'a)set => bool"
 | 
| 21634 | 240  | 
where  | 
| 20810 | 241  | 
"disj_wf r = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r = (\<Union>i<n. T i))"  | 
| 19954 | 242  | 
|
| 21634 | 243  | 
definition  | 
| 19954 | 244  | 
  transition_idx :: "[nat => 'a, nat => ('a*'a)set, nat set] => nat"
 | 
| 21634 | 245  | 
where  | 
| 20810 | 246  | 
"transition_idx s T A =  | 
247  | 
    (LEAST k. \<exists>i j. A = {i,j} & i<j & (s j, s i) \<in> T k)"
 | 
|
| 19954 | 248  | 
|
249  | 
||
250  | 
lemma transition_idx_less:  | 
|
251  | 
    "[|i<j; (s j, s i) \<in> T k; k<n|] ==> transition_idx s T {i,j} < n"
 | 
|
252  | 
apply (subgoal_tac "transition_idx s T {i, j} \<le> k", simp) 
 | 
|
253  | 
apply (simp add: transition_idx_def, blast intro: Least_le)  | 
|
254  | 
done  | 
|
255  | 
||
256  | 
lemma transition_idx_in:  | 
|
257  | 
    "[|i<j; (s j, s i) \<in> T k|] ==> (s j, s i) \<in> T (transition_idx s T {i,j})"
 | 
|
258  | 
apply (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR  | 
|
259  | 
cong: conj_cong)  | 
|
260  | 
apply (erule LeastI)  | 
|
261  | 
done  | 
|
262  | 
||
263  | 
text{*To be equal to the union of some well-founded relations is equivalent
 | 
|
264  | 
to being the subset of such a union.*}  | 
|
265  | 
lemma disj_wf:  | 
|
266  | 
"disj_wf(r) = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r \<subseteq> (\<Union>i<n. T i))"  | 
|
267  | 
apply (auto simp add: disj_wf_def)  | 
|
268  | 
apply (rule_tac x="%i. T i Int r" in exI)  | 
|
269  | 
apply (rule_tac x=n in exI)  | 
|
270  | 
apply (force simp add: wf_Int1)  | 
|
271  | 
done  | 
|
272  | 
||
273  | 
theorem trans_disj_wf_implies_wf:  | 
|
274  | 
assumes transr: "trans r"  | 
|
275  | 
and dwf: "disj_wf(r)"  | 
|
276  | 
shows "wf r"  | 
|
277  | 
proof (simp only: wf_iff_no_infinite_down_chain, rule notI)  | 
|
278  | 
assume "\<exists>s. \<forall>i. (s (Suc i), s i) \<in> r"  | 
|
279  | 
then obtain s where sSuc: "\<forall>i. (s (Suc i), s i) \<in> r" ..  | 
|
280  | 
have s: "!!i j. i < j ==> (s j, s i) \<in> r"  | 
|
281  | 
proof -  | 
|
282  | 
fix i and j::nat  | 
|
283  | 
assume less: "i<j"  | 
|
284  | 
thus "(s j, s i) \<in> r"  | 
|
285  | 
proof (rule less_Suc_induct)  | 
|
286  | 
show "\<And>i. (s (Suc i), s i) \<in> r" by (simp add: sSuc)  | 
|
287  | 
show "\<And>i j k. \<lbrakk>(s j, s i) \<in> r; (s k, s j) \<in> r\<rbrakk> \<Longrightarrow> (s k, s i) \<in> r"  | 
|
288  | 
using transr by (unfold trans_def, blast)  | 
|
289  | 
qed  | 
|
290  | 
qed  | 
|
291  | 
from dwf  | 
|
292  | 
obtain T and n::nat where wfT: "\<forall>k<n. wf(T k)" and r: "r = (\<Union>k<n. T k)"  | 
|
293  | 
by (auto simp add: disj_wf_def)  | 
|
294  | 
have s_in_T: "\<And>i j. i<j ==> \<exists>k. (s j, s i) \<in> T k & k<n"  | 
|
295  | 
proof -  | 
|
296  | 
fix i and j::nat  | 
|
297  | 
assume less: "i<j"  | 
|
298  | 
hence "(s j, s i) \<in> r" by (rule s [of i j])  | 
|
299  | 
thus "\<exists>k. (s j, s i) \<in> T k & k<n" by (auto simp add: r)  | 
|
300  | 
qed  | 
|
301  | 
  have trless: "!!i j. i\<noteq>j ==> transition_idx s T {i,j} < n"
 | 
|
302  | 
apply (auto simp add: linorder_neq_iff)  | 
|
303  | 
apply (blast dest: s_in_T transition_idx_less)  | 
|
304  | 
apply (subst insert_commute)  | 
|
305  | 
apply (blast dest: s_in_T transition_idx_less)  | 
|
306  | 
done  | 
|
307  | 
have  | 
|
308  | 
"\<exists>K k. K \<subseteq> UNIV & infinite K & k < n &  | 
|
309  | 
          (\<forall>i\<in>K. \<forall>j\<in>K. i\<noteq>j --> transition_idx s T {i,j} = k)"
 | 
|
310  | 
by (rule Ramsey2) (auto intro: trless nat_infinite)  | 
|
311  | 
then obtain K and k  | 
|
312  | 
where infK: "infinite K" and less: "k < n" and  | 
|
313  | 
          allk: "\<forall>i\<in>K. \<forall>j\<in>K. i\<noteq>j --> transition_idx s T {i,j} = k"
 | 
|
314  | 
by auto  | 
|
315  | 
have "\<forall>m. (s (enumerate K (Suc m)), s(enumerate K m)) \<in> T k"  | 
|
316  | 
proof  | 
|
317  | 
fix m::nat  | 
|
318  | 
let ?j = "enumerate K (Suc m)"  | 
|
319  | 
let ?i = "enumerate K m"  | 
|
320  | 
have jK: "?j \<in> K" by (simp add: enumerate_in_set infK)  | 
|
321  | 
have iK: "?i \<in> K" by (simp add: enumerate_in_set infK)  | 
|
322  | 
have ij: "?i < ?j" by (simp add: enumerate_step infK)  | 
|
323  | 
    have ijk: "transition_idx s T {?i,?j} = k" using iK jK ij 
 | 
|
324  | 
by (simp add: allk)  | 
|
325  | 
obtain k' where "(s ?j, s ?i) \<in> T k'" "k'<n"  | 
|
326  | 
using s_in_T [OF ij] by blast  | 
|
327  | 
thus "(s ?j, s ?i) \<in> T k"  | 
|
328  | 
by (simp add: ijk [symmetric] transition_idx_in ij)  | 
|
329  | 
qed  | 
|
330  | 
hence "~ wf(T k)" by (force simp add: wf_iff_no_infinite_down_chain)  | 
|
331  | 
thus False using wfT less by blast  | 
|
332  | 
qed  | 
|
333  | 
||
| 19944 | 334  | 
end  |