paulson@19944: (* Title: HOL/Library/Ramsey.thy wenzelm@32960: Author: Tom Ridge. Converted to structured Isar by L C Paulson paulson@19944: *) paulson@19944: paulson@19944: header "Ramsey's Theorem" paulson@19944: haftmann@25594: theory Ramsey haftmann@30738: imports Main Infinite_Set haftmann@25594: begin paulson@19944: nipkow@40695: subsection{* Finite Ramsey theorem(s) *} nipkow@40695: nipkow@40695: text{* To distinguish the finite and infinite ones, lower and upper case nipkow@40695: names are used. nipkow@40695: nipkow@40695: This is the most basic version in terms of cliques and independent nipkow@40695: sets, i.e. the version for graphs and 2 colours. *} nipkow@40695: nipkow@40695: definition "clique V E = (\v\V. \w\V. v\w \ {v,w} : E)" nipkow@40695: definition "indep V E = (\v\V. \w\V. v\w \ \ {v,w} : E)" nipkow@40695: nipkow@40695: lemma ramsey2: nipkow@40695: "\r\1. \ (V::'a set) (E::'a set set). finite V \ card V \ r \ nipkow@40695: (\ R \ V. card R = m \ clique R E \ card R = n \ indep R E)" nipkow@40695: (is "\r\1. ?R m n r") nipkow@40695: proof(induct k == "m+n" arbitrary: m n) nipkow@40695: case 0 nipkow@40695: show ?case (is "EX r. ?R r") nipkow@40695: proof nipkow@40695: show "?R 1" using 0 nipkow@40695: by (clarsimp simp: indep_def)(metis card.empty emptyE empty_subsetI) nipkow@40695: qed nipkow@40695: next nipkow@40695: case (Suc k) nipkow@40695: { assume "m=0" nipkow@40695: have ?case (is "EX r. ?R r") nipkow@40695: proof nipkow@40695: show "?R 1" using `m=0` nipkow@40695: by (simp add:clique_def)(metis card.empty emptyE empty_subsetI) nipkow@40695: qed nipkow@40695: } moreover nipkow@40695: { assume "n=0" nipkow@40695: have ?case (is "EX r. ?R r") nipkow@40695: proof nipkow@40695: show "?R 1" using `n=0` nipkow@40695: by (simp add:indep_def)(metis card.empty emptyE empty_subsetI) nipkow@40695: qed nipkow@40695: } moreover nipkow@40695: { assume "m\0" "n\0" wenzelm@46575: then have "k = (m - 1) + n" "k = m + (n - 1)" using `Suc k = m+n` by auto wenzelm@46575: from Suc(1)[OF this(1)] Suc(1)[OF this(2)] nipkow@40695: obtain r1 r2 where "r1\1" "r2\1" "?R (m - 1) n r1" "?R m (n - 1) r2" nipkow@40695: by auto wenzelm@46575: then have "r1+r2 \ 1" by arith nipkow@40695: moreover nipkow@40695: have "?R m n (r1+r2)" (is "ALL V E. _ \ ?EX V E m n") nipkow@40695: proof clarify nipkow@40695: fix V :: "'a set" and E :: "'a set set" nipkow@40695: assume "finite V" "r1+r2 \ card V" nipkow@40695: with `r1\1` have "V \ {}" by auto nipkow@40695: then obtain v where "v : V" by blast nipkow@40695: let ?M = "{w : V. w\v & {v,w} : E}" nipkow@40695: let ?N = "{w : V. w\v & {v,w} ~: E}" nipkow@40695: have "V = insert v (?M \ ?N)" using `v : V` by auto wenzelm@46575: then have "card V = card(insert v (?M \ ?N))" by metis nipkow@40695: also have "\ = card ?M + card ?N + 1" using `finite V` nipkow@44890: by(fastforce intro: card_Un_disjoint) nipkow@40695: finally have "card V = card ?M + card ?N + 1" . wenzelm@46575: then have "r1+r2 \ card ?M + card ?N + 1" using `r1+r2 \ card V` by simp wenzelm@46575: then have "r1 \ card ?M \ r2 \ card ?N" by arith nipkow@40695: moreover nipkow@40695: { assume "r1 \ card ?M" nipkow@40695: moreover have "finite ?M" using `finite V` by auto nipkow@40695: ultimately have "?EX ?M E (m - 1) n" using `?R (m - 1) n r1` by blast nipkow@40695: then obtain R where "R \ ?M" "v ~: R" and nipkow@40695: CI: "card R = m - 1 \ clique R E \ nipkow@40695: card R = n \ indep R E" (is "?C \ ?I") nipkow@40695: by blast nipkow@40695: have "R <= V" using `R <= ?M` by auto nipkow@40695: have "finite R" using `finite V` `R \ V` by (metis finite_subset) nipkow@40695: { assume "?I" nipkow@40695: with `R <= V` have "?EX V E m n" by blast nipkow@40695: } moreover nipkow@40695: { assume "?C" wenzelm@46575: then have "clique (insert v R) E" using `R <= ?M` nipkow@40695: by(auto simp:clique_def insert_commute) nipkow@40695: moreover have "card(insert v R) = m" nipkow@40695: using `?C` `finite R` `v ~: R` `m\0` by simp nipkow@40695: ultimately have "?EX V E m n" using `R <= V` `v : V` by blast nipkow@40695: } ultimately have "?EX V E m n" using CI by blast nipkow@40695: } moreover nipkow@40695: { assume "r2 \ card ?N" nipkow@40695: moreover have "finite ?N" using `finite V` by auto nipkow@40695: ultimately have "?EX ?N E m (n - 1)" using `?R m (n - 1) r2` by blast nipkow@40695: then obtain R where "R \ ?N" "v ~: R" and nipkow@40695: CI: "card R = m \ clique R E \ nipkow@40695: card R = n - 1 \ indep R E" (is "?C \ ?I") nipkow@40695: by blast nipkow@40695: have "R <= V" using `R <= ?N` by auto nipkow@40695: have "finite R" using `finite V` `R \ V` by (metis finite_subset) nipkow@40695: { assume "?C" nipkow@40695: with `R <= V` have "?EX V E m n" by blast nipkow@40695: } moreover nipkow@40695: { assume "?I" wenzelm@46575: then have "indep (insert v R) E" using `R <= ?N` nipkow@40695: by(auto simp:indep_def insert_commute) nipkow@40695: moreover have "card(insert v R) = n" nipkow@40695: using `?I` `finite R` `v ~: R` `n\0` by simp nipkow@40695: ultimately have "?EX V E m n" using `R <= V` `v : V` by blast nipkow@40695: } ultimately have "?EX V E m n" using CI by blast nipkow@40695: } ultimately show "?EX V E m n" by blast nipkow@40695: qed nipkow@40695: ultimately have ?case by blast nipkow@40695: } ultimately show ?case by blast nipkow@40695: qed nipkow@40695: nipkow@40695: wenzelm@22665: subsection {* Preliminaries *} paulson@19944: wenzelm@22665: subsubsection {* ``Axiom'' of Dependent Choice *} paulson@19944: haftmann@34941: primrec choice :: "('a => bool) => ('a * 'a) set => nat => 'a" where paulson@19944: --{*An integer-indexed chain of choices*} haftmann@34941: choice_0: "choice P r 0 = (SOME x. P x)" haftmann@34941: | choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \ r)" paulson@19944: wenzelm@46575: lemma choice_n: paulson@19944: assumes P0: "P x0" paulson@19944: and Pstep: "!!x. P x ==> \y. P y & (x,y) \ r" paulson@19944: shows "P (choice P r n)" wenzelm@19948: proof (induct n) wenzelm@46575: case 0 show ?case by (force intro: someI P0) wenzelm@19948: next wenzelm@46575: case Suc then show ?case by (auto intro: someI2_ex [OF Pstep]) wenzelm@19948: qed paulson@19944: wenzelm@46575: lemma dependent_choice: paulson@19944: assumes trans: "trans r" paulson@19944: and P0: "P x0" paulson@19944: and Pstep: "!!x. P x ==> \y. P y & (x,y) \ r" paulson@19954: obtains f :: "nat => 'a" where paulson@19954: "!!n. P (f n)" and "!!n m. n < m ==> (f n, f m) \ r" paulson@19954: proof paulson@19954: fix n paulson@19954: show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep]) paulson@19944: next wenzelm@46575: have PSuc: "\n. (choice P r n, choice P r (Suc n)) \ r" paulson@19944: using Pstep [OF choice_n [OF P0 Pstep]] paulson@19944: by (auto intro: someI2_ex) paulson@19954: fix n m :: nat paulson@19954: assume less: "n < m" paulson@19954: show "(choice P r n, choice P r m) \ r" using PSuc paulson@19954: by (auto intro: less_Suc_induct [OF less] transD [OF trans]) paulson@19954: qed paulson@19944: paulson@19944: wenzelm@22665: subsubsection {* Partitions of a Set *} paulson@19944: wenzelm@46575: definition part :: "nat => nat => 'a set => ('a set => nat) => bool" paulson@19944: --{*the function @{term f} partitions the @{term r}-subsets of the typically paulson@19944: infinite set @{term Y} into @{term s} distinct categories.*} krauss@21634: where wenzelm@19948: "part r s Y f = (\X. X \ Y & finite X & card X = r --> f X < s)" paulson@19944: paulson@19944: text{*For induction, we decrease the value of @{term r} in partitions.*} paulson@19944: lemma part_Suc_imp_part: wenzelm@46575: "[| infinite Y; part (Suc r) s Y f; y \ Y |] paulson@19944: ==> part r s (Y - {y}) (%u. f (insert y u))" paulson@19944: apply(simp add: part_def, clarify) paulson@19944: apply(drule_tac x="insert y X" in spec) nipkow@24853: apply(force) paulson@19944: done paulson@19944: wenzelm@46575: lemma part_subset: "part r s YY f ==> Y \ YY ==> part r s Y f" wenzelm@19948: unfolding part_def by blast wenzelm@46575: paulson@19944: wenzelm@22665: subsection {* Ramsey's Theorem: Infinitary Version *} paulson@19944: wenzelm@46575: lemma Ramsey_induction: paulson@19954: fixes s and r::nat paulson@19944: shows wenzelm@46575: "!!(YY::'a set) (f::'a set => nat). paulson@19944: [|infinite YY; part r s YY f|] wenzelm@46575: ==> \Y' t'. Y' \ YY & infinite Y' & t' < s & paulson@19944: (\X. X \ Y' & finite X & card X = r --> f X = t')" paulson@19944: proof (induct r) paulson@19944: case 0 wenzelm@46575: then show ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong) paulson@19944: next wenzelm@46575: case (Suc r) paulson@19944: show ?case paulson@19944: proof - paulson@19944: from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \ YY" by blast paulson@19944: let ?ramr = "{((y,Y,t),(y',Y',t')). y' \ Y & Y' \ Y}" wenzelm@46575: let ?propr = "%(y,Y,t). wenzelm@32960: y \ YY & y \ Y & Y \ YY & infinite Y & t < s wenzelm@32960: & (\X. X\Y & finite X & card X = r --> (f o insert y) X = t)" paulson@19944: have infYY': "infinite (YY-{yy})" using Suc.prems by auto paulson@19944: have partf': "part r s (YY - {yy}) (f \ insert yy)" paulson@19944: by (simp add: o_def part_Suc_imp_part yy Suc.prems) wenzelm@46575: have transr: "trans ?ramr" by (force simp add: trans_def) paulson@19944: from Suc.hyps [OF infYY' partf'] paulson@19944: obtain Y0 and t0 paulson@19944: where "Y0 \ YY - {yy}" "infinite Y0" "t0 < s" paulson@19944: "\X. X\Y0 \ finite X \ card X = r \ (f \ insert yy) X = t0" wenzelm@46575: by blast paulson@19944: with yy have propr0: "?propr(yy,Y0,t0)" by blast wenzelm@46575: have proprstep: "\x. ?propr x \ \y. ?propr y \ (x, y) \ ?ramr" paulson@19944: proof - paulson@19944: fix x wenzelm@46575: assume px: "?propr x" then show "?thesis x" paulson@19944: proof (cases x) paulson@19944: case (fields yx Yx tx) paulson@19944: then obtain yx' where yx': "yx' \ Yx" using px paulson@19944: by (blast dest: infinite_imp_nonempty) paulson@19944: have infYx': "infinite (Yx-{yx'})" using fields px by auto paulson@19944: with fields px yx' Suc.prems paulson@19944: have partfx': "part r s (Yx - {yx'}) (f \ insert yx')" huffman@35175: by (simp add: o_def part_Suc_imp_part part_subset [where YY=YY and Y=Yx]) wenzelm@32960: from Suc.hyps [OF infYx' partfx'] wenzelm@32960: obtain Y' and t' wenzelm@32960: where Y': "Y' \ Yx - {yx'}" "infinite Y'" "t' < s" wenzelm@32960: "\X. X\Y' \ finite X \ card X = r \ (f \ insert yx') X = t'" wenzelm@46575: by blast wenzelm@32960: show ?thesis wenzelm@32960: proof wenzelm@32960: show "?propr (yx',Y',t') & (x, (yx',Y',t')) \ ?ramr" wenzelm@32960: using fields Y' yx' px by blast wenzelm@32960: qed paulson@19944: qed paulson@19944: qed paulson@19944: from dependent_choice [OF transr propr0 proprstep] nipkow@19946: obtain g where pg: "!!n::nat. ?propr (g n)" paulson@19954: and rg: "!!n m. n (g n, g m) \ ?ramr" by blast haftmann@28741: let ?gy = "fst o g" haftmann@28741: let ?gt = "snd o snd o g" paulson@19944: have rangeg: "\k. range ?gt \ {.. range ?gt" paulson@19944: then obtain n where "x = ?gt n" .. paulson@19944: with pg [of n] show "x \ {.. ?gy m'" wenzelm@19948: using rg [OF less] pg [of m] by (cases "g m", cases "g m'") auto paulson@19944: qed paulson@19944: show ?thesis paulson@19944: proof (intro exI conjI) paulson@19944: show "?gy ` {n. ?gt n = s'} \ YY" using pg wenzelm@46575: by (auto simp add: Let_def split_beta) paulson@19944: show "infinite (?gy ` {n. ?gt n = s'})" using infeqs' wenzelm@46575: by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) paulson@19944: show "s' < s" by (rule less') wenzelm@46575: show "\X. X \ ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r paulson@19944: --> f X = s'" paulson@19944: proof - wenzelm@46575: {fix X paulson@19944: assume "X \ ?gy ` {n. ?gt n = s'}" paulson@19944: and cardX: "finite X" "card X = Suc r" wenzelm@46575: then obtain AA where AA: "AA \ {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" wenzelm@46575: by (auto simp add: subset_image_iff) paulson@19944: with cardX have "AA\{}" by auto wenzelm@46575: then have AAleast: "(LEAST x. x \ AA) \ AA" by (auto intro: LeastI_ex) paulson@19944: have "f X = s'" wenzelm@46575: proof (cases "g (LEAST x. x \ AA)") paulson@19944: case (fields ya Ya ta) wenzelm@46575: with AAleast Xeq wenzelm@46575: have ya: "ya \ X" by (force intro!: rev_image_eqI) wenzelm@46575: then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb) wenzelm@46575: also have "... = ta" paulson@19944: proof - paulson@19944: have "X - {ya} \ Ya" wenzelm@46575: proof paulson@19954: fix x assume x: "x \ X - {ya}" wenzelm@46575: then obtain a' where xeq: "x = ?gy a'" and a': "a' \ AA" wenzelm@46575: by (auto simp add: Xeq) wenzelm@46575: then have "a' \ (LEAST x. x \ AA)" using x fields by auto wenzelm@46575: then have lessa': "(LEAST x. x \ AA) < a'" paulson@19944: using Least_le [of "%x. x \ AA", OF a'] by arith paulson@19944: show "x \ Ya" using xeq fields rg [OF lessa'] by auto paulson@19944: qed paulson@19944: moreover paulson@19944: have "card (X - {ya}) = r" nipkow@24853: by (simp add: cardX ya) wenzelm@46575: ultimately show ?thesis paulson@19944: using pg [of "LEAST x. x \ AA"] fields cardX wenzelm@32960: by (clarsimp simp del:insert_Diff_single) paulson@19944: qed paulson@19944: also have "... = s'" using AA AAleast fields by auto paulson@19944: finally show ?thesis . paulson@19944: qed} wenzelm@46575: then show ?thesis by blast wenzelm@46575: qed wenzelm@46575: qed paulson@19944: qed paulson@19944: qed paulson@19944: paulson@19944: paulson@19944: theorem Ramsey: wenzelm@19949: fixes s r :: nat and Z::"'a set" and f::"'a set => nat" paulson@19944: shows paulson@19944: "[|infinite Z; paulson@19944: \X. X \ Z & finite X & card X = r --> f X < s|] wenzelm@46575: ==> \Y t. Y \ Z & infinite Y & t < s paulson@19944: & (\X. X \ Y & finite X & card X = r --> f X = t)" paulson@19954: by (blast intro: Ramsey_induction [unfolded part_def]) paulson@19954: paulson@19954: paulson@19954: corollary Ramsey2: paulson@19954: fixes s::nat and Z::"'a set" and f::"'a set => nat" paulson@19954: assumes infZ: "infinite Z" paulson@19954: and part: "\x\Z. \y\Z. x\y --> f{x,y} < s" paulson@19954: shows paulson@19954: "\Y t. Y \ Z & infinite Y & t < s & (\x\Y. \y\Y. x\y --> f{x,y} = t)" paulson@19954: proof - paulson@19954: have part2: "\X. X \ Z & finite X & card X = 2 --> f X < s" nipkow@44890: using part by (fastforce simp add: eval_nat_numeral card_Suc_eq) wenzelm@46575: obtain Y t wenzelm@53374: where *: "Y \ Z" "infinite Y" "t < s" paulson@19954: "(\X. X \ Y & finite X & card X = 2 --> f X = t)" paulson@19954: by (insert Ramsey [OF infZ part2]) auto wenzelm@53374: then have "\x\Y. \y\Y. x \ y \ f {x, y} = t" by auto wenzelm@53374: with * show ?thesis by iprover paulson@19954: qed paulson@19954: paulson@19954: wenzelm@22665: subsection {* Disjunctive Well-Foundedness *} paulson@19954: wenzelm@22367: text {* wenzelm@22367: An application of Ramsey's theorem to program termination. See wenzelm@22367: \cite{Podelski-Rybalchenko}. paulson@19954: *} paulson@19954: wenzelm@46575: definition disj_wf :: "('a * 'a)set => bool" wenzelm@46575: where "disj_wf r = (\T. \n::nat. (\ii 'a, nat => ('a*'a)set, nat set] => nat" wenzelm@46575: where wenzelm@46575: "transition_idx s T A = wenzelm@46575: (LEAST k. \i j. A = {i,j} & i T k)" paulson@19954: paulson@19954: paulson@19954: lemma transition_idx_less: paulson@19954: "[|i T k; k transition_idx s T {i,j} < n" wenzelm@46575: apply (subgoal_tac "transition_idx s T {i, j} \ k", simp) wenzelm@46575: apply (simp add: transition_idx_def, blast intro: Least_le) paulson@19954: done paulson@19954: paulson@19954: lemma transition_idx_in: paulson@19954: "[|i T k|] ==> (s j, s i) \ T (transition_idx s T {i,j})" wenzelm@46575: apply (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR wenzelm@46575: cong: conj_cong) wenzelm@46575: apply (erule LeastI) paulson@19954: done paulson@19954: paulson@19954: text{*To be equal to the union of some well-founded relations is equivalent paulson@19954: to being the subset of such a union.*} paulson@19954: lemma disj_wf: paulson@19954: "disj_wf(r) = (\T. \n::nat. (\i (\is. \i. (s (Suc i), s i) \ r" paulson@19954: then obtain s where sSuc: "\i. (s (Suc i), s i) \ r" .. paulson@19954: have s: "!!i j. i < j ==> (s j, s i) \ r" paulson@19954: proof - paulson@19954: fix i and j::nat paulson@19954: assume less: "i r" paulson@19954: proof (rule less_Suc_induct) wenzelm@46575: show "\i. (s (Suc i), s i) \ r" by (simp add: sSuc) paulson@19954: show "\i j k. \(s j, s i) \ r; (s k, s j) \ r\ \ (s k, s i) \ r" wenzelm@46575: using transr by (unfold trans_def, blast) paulson@19954: qed wenzelm@46575: qed paulson@19954: from dwf paulson@19954: obtain T and n::nat where wfT: "\kki j. i \k. (s j, s i) \ T k & k r" by (rule s [of i j]) wenzelm@46575: then show "\k. (s j, s i) \ T k & kj ==> transition_idx s T {i,j} < n" paulson@19954: apply (auto simp add: linorder_neq_iff) wenzelm@46575: apply (blast dest: s_in_T transition_idx_less) wenzelm@46575: apply (subst insert_commute) wenzelm@46575: apply (blast dest: s_in_T transition_idx_less) paulson@19954: done paulson@19954: have wenzelm@46575: "\K k. K \ UNIV & infinite K & k < n & paulson@19954: (\i\K. \j\K. i\j --> transition_idx s T {i,j} = k)" traytel@54580: by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat) wenzelm@46575: then obtain K and k paulson@19954: where infK: "infinite K" and less: "k < n" and paulson@19954: allk: "\i\K. \j\K. i\j --> transition_idx s T {i,j} = k" paulson@19954: by auto paulson@19954: have "\m. (s (enumerate K (Suc m)), s(enumerate K m)) \ T k" paulson@19954: proof paulson@19954: fix m::nat paulson@19954: let ?j = "enumerate K (Suc m)" paulson@19954: let ?i = "enumerate K m" wenzelm@46575: have jK: "?j \ K" by (simp add: enumerate_in_set infK) wenzelm@46575: have iK: "?i \ K" by (simp add: enumerate_in_set infK) wenzelm@46575: have ij: "?i < ?j" by (simp add: enumerate_step infK) wenzelm@46575: have ijk: "transition_idx s T {?i,?j} = k" using iK jK ij paulson@19954: by (simp add: allk) wenzelm@46575: obtain k' where "(s ?j, s ?i) \ T k'" "k' T k" wenzelm@46575: by (simp add: ijk [symmetric] transition_idx_in ij) paulson@19954: qed wenzelm@46575: then have "~ wf(T k)" by (force simp add: wf_iff_no_infinite_down_chain) wenzelm@46575: then show False using wfT less by blast paulson@19954: qed paulson@19954: paulson@19944: end