--- a/src/HOL/Library/Ramsey.thy Mon Dec 04 00:06:59 2006 +0100
+++ b/src/HOL/Library/Ramsey.thy Mon Dec 04 15:15:09 2006 +0100
@@ -56,6 +56,7 @@
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.*}
+where
"part r s Y f = (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X < s)"
text{*For induction, we decrease the value of @{term r} in partitions.*}
@@ -242,9 +243,12 @@
definition
disj_wf :: "('a * 'a)set => bool"
+where
"disj_wf r = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r = (\<Union>i<n. T i))"
+definition
transition_idx :: "[nat => 'a, nat => ('a*'a)set, nat set] => nat"
+where
"transition_idx s T A =
(LEAST k. \<exists>i j. A = {i,j} & i<j & (s j, s i) \<in> T k)"