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