# HG changeset patch # User paulson # Date 1151049301 -7200 # Node ID 60e0cbeae3d84b50234b3e411561fcc334d30769 # Parent 26b37721b357eae89853fd867f06eec3720ee62f Introduction of Ramsey's theorem diff -r 26b37721b357 -r 60e0cbeae3d8 src/HOL/Infinite_Set.thy --- a/src/HOL/Infinite_Set.thy Thu Jun 22 18:48:25 2006 +0200 +++ b/src/HOL/Infinite_Set.thy Fri Jun 23 09:55:01 2006 +0200 @@ -221,7 +221,7 @@ *} lemma linorder_injI: - assumes hyp: "\x y. x < (y::'a::linorder) \ f x \ f y" + assumes hyp: "!!x y. x < (y::'a::linorder) ==> f x \ f y" shows "inj f" proof (rule inj_onI) fix x y @@ -295,8 +295,8 @@ qed have inj: "inj pick" proof (rule linorder_injI) - show "\i j. i<(j::nat) \ pick i \ pick j" - proof (clarify) + show "!!i j. i<(j::nat) ==> pick i \ pick j" + proof fix i j assume ij: "i<(j::nat)" and eq: "pick i = pick j" diff -r 26b37721b357 -r 60e0cbeae3d8 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Jun 22 18:48:25 2006 +0200 +++ b/src/HOL/IsaMakefile Fri Jun 23 09:55:01 2006 +0200 @@ -199,7 +199,7 @@ HOL-Library: HOL $(LOG)/HOL-Library.gz $(LOG)/HOL-Library.gz: $(OUT)/HOL \ - Library/SetsAndFunctions.thy Library/BigO.thy \ + Library/SetsAndFunctions.thy Library/BigO.thy Library/Ramsey.thy \ Library/EfficientNat.thy Library/ExecutableSet.thy Library/ExecutableRat.thy \ Library/FuncSet.thy Library/Library.thy \ Library/List_Prefix.thy Library/Multiset.thy Library/NatPair.thy \ diff -r 26b37721b357 -r 60e0cbeae3d8 src/HOL/Library/Ramsey.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Ramsey.thy Fri Jun 23 09:55:01 2006 +0200 @@ -0,0 +1,221 @@ +(* Title: HOL/Library/Ramsey.thy + ID: $Id$ + Author: Tom Ridge. Converted to structured Isar by L C Paulson +*) + +header "Ramsey's Theorem" + +theory Ramsey imports Main begin + + +subsection{*``Axiom'' of Dependent Choice*} + +consts choice :: "('a => bool) => (('a * 'a) set) => nat => 'a" + --{*An integer-indexed chain of choices*} +primrec + choice_0: "choice P r 0 = (SOME x. P x)" + + choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \ r)" + + +lemma choice_n: + assumes P0: "P x0" + and Pstep: "!!x. P x ==> \y. P y & (x,y) \ r" + shows "P (choice P r n)" + proof (induct n) + case 0 show ?case by (force intro: someI P0) + next + case (Suc n) thus ?case by (auto intro: someI2_ex [OF Pstep]) + qed + +lemma dependent_choice: + assumes trans: "trans r" + and P0: "P x0" + and Pstep: "!!x. P x ==> \y. P y & (x,y) \ r" + shows "\f::nat=>'a. (\n. P (f n)) & (\n m. n (f n, f m) \ r)" +proof (intro exI conjI) + show "\n. P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep]) +next + have PSuc: "\n. (choice P r n, choice P r (Suc n)) \ r" + using Pstep [OF choice_n [OF P0 Pstep]] + by (auto intro: someI2_ex) + show "\n m. n (choice P r n, choice P r m) \ r" + proof (intro strip) + fix n and m::nat + assume less: "n r" using PSuc + by (auto intro: less_Suc_induct [OF less] transD [OF trans]) + qed +qed + + +subsection {*Partitions of a Set*} + +constdefs part :: "nat => nat => 'a set => ('a set => nat) => bool" + --{*the function @{term f} partitions the @{term r}-subsets of the typically + infinite set @{term Y} into @{term s} distinct categories.*} + "part r s Y f == \X. X \ Y & finite X & card X = r --> f X < s" + +text{*For induction, we decrease the value of @{term r} in partitions.*} +lemma part_Suc_imp_part: + "[| infinite Y; part (Suc r) s Y f; y \ Y |] + ==> part r s (Y - {y}) (%u. f (insert y u))" + apply(simp add: part_def, clarify) + apply(drule_tac x="insert y X" in spec) + apply(force simp:card_Diff_singleton_if) + done + +lemma part_subset: "part r s YY f ==> Y \ YY ==> part r s Y f" + by (simp add: part_def, blast) + + +subsection {*Ramsey's Theorem: Infinitary Version*} + +lemma ramsey_induction: + fixes s::nat and r::nat + shows + "!!(YY::'a set) (f::'a set => nat). + [|infinite YY; part r s YY f|] + ==> \Y' t'. Y' \ YY & infinite Y' & t' < s & + (\X. X \ Y' & finite X & card X = r --> f X = t')" +proof (induct r) + case 0 + thus ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong) +next + case (Suc r) + show ?case + proof - + from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \ YY" by blast + let ?ramr = "{((y,Y,t),(y',Y',t')). y' \ Y & Y' \ Y}" + let ?propr = "%(y,Y,t). + y \ YY & y \ Y & Y \ YY & infinite Y & t < s + & (\X. X\Y & finite X & card X = r --> (f o insert y) X = t)" + have infYY': "infinite (YY-{yy})" using Suc.prems by auto + have partf': "part r s (YY - {yy}) (f \ insert yy)" + by (simp add: o_def part_Suc_imp_part yy Suc.prems) + have transr: "trans ?ramr" by (force simp add: trans_def) + from Suc.hyps [OF infYY' partf'] + obtain Y0 and t0 + where "Y0 \ YY - {yy}" "infinite Y0" "t0 < s" + "\X. X\Y0 \ finite X \ card X = r \ (f \ insert yy) X = t0" + by blast + with yy have propr0: "?propr(yy,Y0,t0)" by blast + have proprstep: "\x. ?propr x \ \y. ?propr y \ (x, y) \ ?ramr" + proof - + fix x + assume px: "?propr x" thus "?thesis x" + proof (cases x) + case (fields yx Yx tx) + then obtain yx' where yx': "yx' \ Yx" using px + by (blast dest: infinite_imp_nonempty) + have infYx': "infinite (Yx-{yx'})" using fields px by auto + with fields px yx' Suc.prems + have partfx': "part r s (Yx - {yx'}) (f \ insert yx')" + by (simp add: o_def part_Suc_imp_part part_subset [where ?YY=YY]) + from Suc.hyps [OF infYx' partfx'] + obtain Y' and t' + where Y': "Y' \ Yx - {yx'}" "infinite Y'" "t' < s" + "\X. X\Y' \ finite X \ card X = r \ (f \ insert yx') X = t'" + by blast + show ?thesis + proof + show "?propr (yx',Y',t') & (x, (yx',Y',t')) \ ?ramr" + using fields Y' yx' px by blast + qed + qed + qed + from dependent_choice [OF transr propr0 proprstep] + obtain g where "(\n::nat. ?propr(g n)) & (\n m. n(g n, g m) \ ?ramr)" + .. --{*for some reason, can't derive the following directly from dc*} + hence pg: "!!n. ?propr (g n)" + and rg: "!!n m. n (g n, g m) \ ?ramr" by auto + let ?gy = "(\n. let (y,Y,t) = g n in y)" + let ?gt = "(\n. let (y,Y,t) = g n in t)" + have rangeg: "\k. range ?gt \ {.. range ?gt" + then obtain n where "x = ?gt n" .. + with pg [of n] show "x \ {..s' \ range ?gt. infinite (?gt -` {s'})" + by (rule inf_img_fin_dom [OF _ nat_infinite]) + (simp add: finite_nat_iff_bounded rangeg) + then obtain s' and n' + where s': "s' = ?gt n'" + and infeqs': "infinite {n. ?gt n = s'}" + by (auto simp add: vimage_def) + with pg [of n'] have less': "s' ?gy m'" + using rg [OF less] pg [of m] by (cases "g m", cases "g m'", auto) + qed + show ?thesis + proof (intro exI conjI) + show "?gy ` {n. ?gt n = s'} \ YY" using pg + by (auto simp add: Let_def split_beta) + next + show "infinite (?gy ` {n. ?gt n = s'})" using infeqs' + by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) + next + show "s' < s" by (rule less') + next + show "\X. X \ ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r + --> f X = s'" + proof - + {fix X + assume "X \ ?gy ` {n. ?gt n = s'}" + and cardX: "finite X" "card X = Suc r" + then obtain AA where AA: "AA \ {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" + by (auto simp add: subset_image_iff) + with cardX have "AA\{}" by auto + hence AAleast: "(LEAST x. x \ AA) \ AA" by (auto intro: LeastI_ex) + have "f X = s'" + proof (cases "g (LEAST x. x \ AA)") + case (fields ya Ya ta) + with AAleast Xeq + have ya: "ya \ X" by (force intro!: rev_image_eqI) + hence "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb) + also have "... = ta" + proof - + have "X - {ya} \ Ya" + proof + fix x + assume x: "x \ X - {ya}" + then obtain a' where xeq: "x = ?gy a'" and a': "a' \ AA" + by (auto simp add: Xeq) + hence "a' \ (LEAST x. x \ AA)" using x fields by auto + hence lessa': "(LEAST x. x \ AA) < a'" + using Least_le [of "%x. x \ AA", OF a'] by arith + show "x \ Ya" using xeq fields rg [OF lessa'] by auto + qed + moreover + have "card (X - {ya}) = r" + by (simp add: card_Diff_singleton_if cardX ya) + ultimately show ?thesis + using pg [of "LEAST x. x \ AA"] fields cardX + by (clarify, drule_tac x="X-{ya}" in spec, simp) + qed + also have "... = s'" using AA AAleast fields by auto + finally show ?thesis . + qed} + thus ?thesis by blast + qed + qed + qed +qed + + +text{*Repackaging of Tom Ridge's final result*} +theorem Ramsey: + fixes s::nat and r::nat and Z::"'a set" and f::"'a set => nat" + shows + "[|infinite Z; + \X. X \ Z & finite X & card X = r --> f X < s|] + ==> \Y t. Y \ Z & infinite Y & t < s + & (\X. X \ Y & finite X & card X = r --> f X = t)" +by (blast intro: ramsey_induction [unfolded part_def, rule_format]) + +end +