src/HOL/Library/Ramsey.thy
author blanchet
Mon, 26 Sep 2011 14:03:43 +0200
changeset 45085 eb7a797ade0f
parent 44890 22f665a2e91c
child 46575 f1e387195a56
permissions -rw-r--r--
put MiniSat back first -- Torlak's eval seemed to suggest that Crypto and Lingeling were better, but Crypto is slower on "Nitpick_Examples" and Crypto crashes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     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
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     3
*)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     4
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     5
header "Ramsey's Theorem"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
     6
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 24853
diff changeset
     7
theory Ramsey
30738
0842e906300c normalized imports
haftmann
parents: 28741
diff changeset
     8
imports Main Infinite_Set
25594
43c718438f9f switched import from Main to PreList
haftmann
parents: 24853
diff changeset
     9
begin
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
    10
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    11
subsection{* Finite Ramsey theorem(s) *}
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    12
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    13
text{* To distinguish the finite and infinite ones, lower and upper case
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    14
names are used.
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    15
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    16
This is the most basic version in terms of cliques and independent
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    17
sets, i.e. the version for graphs and 2 colours. *}
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    18
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    19
definition "clique V E = (\<forall>v\<in>V. \<forall>w\<in>V. v\<noteq>w \<longrightarrow> {v,w} : E)"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    20
definition "indep V E = (\<forall>v\<in>V. \<forall>w\<in>V. v\<noteq>w \<longrightarrow> \<not> {v,w} : E)"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    21
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    22
lemma ramsey2:
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    23
  "\<exists>r\<ge>1. \<forall> (V::'a set) (E::'a set set). finite V \<and> card V \<ge> r \<longrightarrow>
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    24
  (\<exists> R \<subseteq> V. card R = m \<and> clique R E \<or> card R = n \<and> indep R E)"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    25
  (is "\<exists>r\<ge>1. ?R m n r")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    26
proof(induct k == "m+n" arbitrary: m n)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    27
  case 0
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    28
  show ?case (is "EX r. ?R r")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    29
  proof
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    30
    show "?R 1" using 0
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    31
      by (clarsimp simp: indep_def)(metis card.empty emptyE empty_subsetI)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    32
  qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    33
next
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    34
  case (Suc k)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    35
  { assume "m=0"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    36
    have ?case (is "EX r. ?R r")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    37
    proof
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    38
      show "?R 1" using `m=0`
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    39
        by (simp add:clique_def)(metis card.empty emptyE empty_subsetI)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    40
    qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    41
  } moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    42
  { assume "n=0"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    43
    have ?case (is "EX r. ?R r")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    44
    proof
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    45
      show "?R 1" using `n=0`
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    46
        by (simp add:indep_def)(metis card.empty emptyE empty_subsetI)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    47
    qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    48
  } moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    49
  { assume "m\<noteq>0" "n\<noteq>0"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    50
    hence "k = (m - 1) + n" "k = m + (n - 1)" using `Suc k = m+n` by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    51
    from Suc(1)[OF this(1)] Suc(1)[OF this(2)] 
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    52
    obtain r1 r2 where "r1\<ge>1" "r2\<ge>1" "?R (m - 1) n r1" "?R m (n - 1) r2"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    53
      by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    54
    hence "r1+r2 \<ge> 1" by arith
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    55
    moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    56
    have "?R m n (r1+r2)" (is "ALL V E. _ \<longrightarrow> ?EX V E m n")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    57
    proof clarify
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    58
      fix V :: "'a set" and E :: "'a set set"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    59
      assume "finite V" "r1+r2 \<le> card V"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    60
      with `r1\<ge>1` have "V \<noteq> {}" by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    61
      then obtain v where "v : V" by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    62
      let ?M = "{w : V. w\<noteq>v & {v,w} : E}"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    63
      let ?N = "{w : V. w\<noteq>v & {v,w} ~: E}"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    64
      have "V = insert v (?M \<union> ?N)" using `v : V` by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    65
      hence "card V = card(insert v (?M \<union> ?N))" by metis
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    66
      also have "\<dots> = card ?M + card ?N + 1" using `finite V`
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 40695
diff changeset
    67
        by(fastforce intro: card_Un_disjoint)
40695
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    68
      finally have "card V = card ?M + card ?N + 1" .
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    69
      hence "r1+r2 \<le> card ?M + card ?N + 1" using `r1+r2 \<le> card V` by simp
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    70
      hence "r1 \<le> card ?M \<or> r2 \<le> card ?N" by arith
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    71
      moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    72
      { assume "r1 \<le> card ?M"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    73
        moreover have "finite ?M" using `finite V` by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    74
        ultimately have "?EX ?M E (m - 1) n" using `?R (m - 1) n r1` by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    75
        then obtain R where "R \<subseteq> ?M" "v ~: R" and
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    76
          CI: "card R = m - 1 \<and> clique R E \<or>
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    77
               card R = n \<and> indep R E" (is "?C \<or> ?I")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    78
          by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    79
        have "R <= V" using `R <= ?M` by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    80
        have "finite R" using `finite V` `R \<subseteq> V` by (metis finite_subset)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    81
        { assume "?I"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    82
          with `R <= V` have "?EX V E m n" by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    83
        } moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    84
        { assume "?C"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    85
          hence "clique (insert v R) E" using `R <= ?M`
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    86
           by(auto simp:clique_def insert_commute)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    87
          moreover have "card(insert v R) = m"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    88
            using `?C` `finite R` `v ~: R` `m\<noteq>0` by simp
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    89
          ultimately have "?EX V E m n" using `R <= V` `v : V` by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    90
        } ultimately have "?EX V E m n" using CI by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    91
      } moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    92
      { assume "r2 \<le> card ?N"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    93
        moreover have "finite ?N" using `finite V` by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    94
        ultimately have "?EX ?N E m (n - 1)" using `?R m (n - 1) r2` by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    95
        then obtain R where "R \<subseteq> ?N" "v ~: R" and
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    96
          CI: "card R = m \<and> clique R E \<or>
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    97
               card R = n - 1 \<and> indep R E" (is "?C \<or> ?I")
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    98
          by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
    99
        have "R <= V" using `R <= ?N` by auto
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   100
        have "finite R" using `finite V` `R \<subseteq> V` by (metis finite_subset)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   101
        { assume "?C"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   102
          with `R <= V` have "?EX V E m n" by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   103
        } moreover
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   104
        { assume "?I"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   105
          hence "indep (insert v R) E" using `R <= ?N`
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   106
            by(auto simp:indep_def insert_commute)
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   107
          moreover have "card(insert v R) = n"
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   108
            using `?I` `finite R` `v ~: R` `n\<noteq>0` by simp
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   109
          ultimately have "?EX V E m n" using `R <= V` `v : V` by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   110
        } ultimately have "?EX V E m n" using CI by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   111
      } ultimately show "?EX V E m n" by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   112
    qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   113
    ultimately have ?case by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   114
  } ultimately show ?case by blast
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   115
qed
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   116
1b2573c3b222 Added the simplest finite Ramsey theorem
nipkow
parents: 40077
diff changeset
   117
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22367
diff changeset
   118
subsection {* Preliminaries *}
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   119
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22367
diff changeset
   120
subsubsection {* ``Axiom'' of Dependent Choice *}
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   121
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 32960
diff changeset
   122
primrec choice :: "('a => bool) => ('a * 'a) set => nat => 'a" where
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   123
  --{*An integer-indexed chain of choices*}
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 32960
diff changeset
   124
    choice_0:   "choice P r 0 = (SOME x. P x)"
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 32960
diff changeset
   125
  | choice_Suc: "choice P r (Suc n) = (SOME y. P y & (choice P r n, y) \<in> r)"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   126
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   127
lemma choice_n: 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   128
  assumes P0: "P x0"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   129
      and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   130
  shows "P (choice P r n)"
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   131
proof (induct n)
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   132
  case 0 show ?case by (force intro: someI P0) 
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   133
next
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   134
  case Suc thus ?case by (auto intro: someI2_ex [OF Pstep]) 
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   135
qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   136
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   137
lemma dependent_choice: 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   138
  assumes trans: "trans r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   139
      and P0: "P x0"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   140
      and Pstep: "!!x. P x ==> \<exists>y. P y & (x,y) \<in> r"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   141
  obtains f :: "nat => 'a" where
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   142
    "!!n. P (f n)" and "!!n m. n < m ==> (f n, f m) \<in> r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   143
proof
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   144
  fix n
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   145
  show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep])
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   146
next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   147
  have PSuc: "\<forall>n. (choice P r n, choice P r (Suc n)) \<in> r" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   148
    using Pstep [OF choice_n [OF P0 Pstep]]
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   149
    by (auto intro: someI2_ex)
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   150
  fix n m :: nat
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   151
  assume less: "n < m"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   152
  show "(choice P r n, choice P r m) \<in> r" using PSuc
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   153
    by (auto intro: less_Suc_induct [OF less] transD [OF trans])
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   154
qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   155
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   156
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22367
diff changeset
   157
subsubsection {* Partitions of a Set *}
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   158
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   159
definition
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   160
  part :: "nat => nat => 'a set => ('a set => nat) => bool"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   161
  --{*the function @{term f} partitions the @{term r}-subsets of the typically
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   162
       infinite set @{term Y} into @{term s} distinct categories.*}
21634
369e38e35686 fixed definition syntax
krauss
parents: 20810
diff changeset
   163
where
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   164
  "part r s Y f = (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X < s)"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   165
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   166
text{*For induction, we decrease the value of @{term r} in partitions.*}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   167
lemma part_Suc_imp_part:
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   168
     "[| infinite Y; part (Suc r) s Y f; y \<in> Y |] 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   169
      ==> part r s (Y - {y}) (%u. f (insert y u))"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   170
  apply(simp add: part_def, clarify)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   171
  apply(drule_tac x="insert y X" in spec)
24853
aab5798e5a33 added lemmas
nipkow
parents: 22665
diff changeset
   172
  apply(force)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   173
  done
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   174
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   175
lemma part_subset: "part r s YY f ==> Y \<subseteq> YY ==> part r s Y f" 
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   176
  unfolding part_def by blast
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   177
  
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   178
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22367
diff changeset
   179
subsection {* Ramsey's Theorem: Infinitary Version *}
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   180
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   181
lemma Ramsey_induction: 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   182
  fixes s and r::nat
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   183
  shows
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   184
  "!!(YY::'a set) (f::'a set => nat). 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   185
      [|infinite YY; part r s YY f|]
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   186
      ==> \<exists>Y' t'. Y' \<subseteq> YY & infinite Y' & t' < s & 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   187
                  (\<forall>X. X \<subseteq> Y' & finite X & card X = r --> f X = t')"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   188
proof (induct r)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   189
  case 0
24853
aab5798e5a33 added lemmas
nipkow
parents: 22665
diff changeset
   190
  thus ?case by (auto simp add: part_def card_eq_0_iff cong: conj_cong)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   191
next
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   192
  case (Suc r) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   193
  show ?case
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   194
  proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   195
    from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \<in> YY" by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   196
    let ?ramr = "{((y,Y,t),(y',Y',t')). y' \<in> Y & Y' \<subseteq> Y}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   197
    let ?propr = "%(y,Y,t).     
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   198
                 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
   199
                 & (\<forall>X. X\<subseteq>Y & finite X & card X = r --> (f o insert y) X = t)"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   200
    have infYY': "infinite (YY-{yy})" using Suc.prems by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   201
    have partf': "part r s (YY - {yy}) (f \<circ> insert yy)"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   202
      by (simp add: o_def part_Suc_imp_part yy Suc.prems)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   203
    have transr: "trans ?ramr" by (force simp add: trans_def) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   204
    from Suc.hyps [OF infYY' partf']
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   205
    obtain Y0 and t0
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   206
    where "Y0 \<subseteq> YY - {yy}"  "infinite Y0"  "t0 < s"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   207
          "\<forall>X. X\<subseteq>Y0 \<and> finite X \<and> card X = r \<longrightarrow> (f \<circ> insert yy) X = t0"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   208
        by blast 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   209
    with yy have propr0: "?propr(yy,Y0,t0)" by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   210
    have proprstep: "\<And>x. ?propr x \<Longrightarrow> \<exists>y. ?propr y \<and> (x, y) \<in> ?ramr" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   211
    proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   212
      fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   213
      assume px: "?propr x" thus "?thesis x"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   214
      proof (cases x)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   215
        case (fields yx Yx tx)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   216
        then obtain yx' where yx': "yx' \<in> Yx" using px
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   217
               by (blast dest: infinite_imp_nonempty)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   218
        have infYx': "infinite (Yx-{yx'})" using fields px by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   219
        with fields px yx' Suc.prems
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   220
        have partfx': "part r s (Yx - {yx'}) (f \<circ> insert yx')"
35175
61255c81da01 fix more looping simp rules
huffman
parents: 34941
diff changeset
   221
          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
   222
        from Suc.hyps [OF infYx' partfx']
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   223
        obtain Y' and t'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   224
        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
   225
               "\<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
   226
            by blast 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   227
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   228
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   229
          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
   230
            using fields Y' yx' px by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30738
diff changeset
   231
        qed
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   232
      qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   233
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   234
    from dependent_choice [OF transr propr0 proprstep]
19946
e3ddb0812840 beautification
nipkow
parents: 19944
diff changeset
   235
    obtain g where pg: "!!n::nat.  ?propr (g n)"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   236
      and rg: "!!n m. n<m ==> (g n, g m) \<in> ?ramr" by blast
28741
1b257449f804 simproc for let
haftmann
parents: 27487
diff changeset
   237
    let ?gy = "fst o g"
1b257449f804 simproc for let
haftmann
parents: 27487
diff changeset
   238
    let ?gt = "snd o snd o g"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   239
    have rangeg: "\<exists>k. range ?gt \<subseteq> {..<k}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   240
    proof (intro exI subsetI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   241
      fix x
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   242
      assume "x \<in> range ?gt"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   243
      then obtain n where "x = ?gt n" ..
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   244
      with pg [of n] show "x \<in> {..<s}" by (cases "g n") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   245
    qed
20810
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   246
    have "finite (range ?gt)"
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   247
      by (simp add: finite_nat_iff_bounded rangeg)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   248
    then obtain s' and n'
20810
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   249
      where s': "s' = ?gt n'"
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   250
        and infeqs': "infinite {n. ?gt n = s'}"
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   251
      by (rule inf_img_fin_domE) (auto simp add: vimage_def intro: nat_infinite)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   252
    with pg [of n'] have less': "s'<s" by (cases "g n'") auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   253
    have inj_gy: "inj ?gy"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   254
    proof (rule linorder_injI)
19949
0505dce27b0b fix/fixes: tuned type constraints;
wenzelm
parents: 19948
diff changeset
   255
      fix m m' :: nat assume less: "m < m'" show "?gy m \<noteq> ?gy m'"
19948
1be283f3f1ba minor tuning of definitions/proofs;
wenzelm
parents: 19946
diff changeset
   256
        using rg [OF less] pg [of m] by (cases "g m", cases "g m'") auto
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   257
    qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   258
    show ?thesis
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   259
    proof (intro exI conjI)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   260
      show "?gy ` {n. ?gt n = s'} \<subseteq> YY" using pg
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   261
        by (auto simp add: Let_def split_beta) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   262
      show "infinite (?gy ` {n. ?gt n = s'})" using infeqs'
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   263
        by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   264
      show "s' < s" by (rule less')
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   265
      show "\<forall>X. X \<subseteq> ?gy ` {n. ?gt n = s'} & finite X & card X = Suc r 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   266
          --> f X = s'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   267
      proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   268
        {fix X 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   269
         assume "X \<subseteq> ?gy ` {n. ?gt n = s'}"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   270
            and cardX: "finite X" "card X = Suc r"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   271
         then obtain AA where AA: "AA \<subseteq> {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   272
             by (auto simp add: subset_image_iff) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   273
         with cardX have "AA\<noteq>{}" by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   274
         hence AAleast: "(LEAST x. x \<in> AA) \<in> AA" by (auto intro: LeastI_ex) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   275
         have "f X = s'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   276
         proof (cases "g (LEAST x. x \<in> AA)") 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   277
           case (fields ya Ya ta)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   278
           with AAleast Xeq 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   279
           have ya: "ya \<in> X" by (force intro!: rev_image_eqI) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   280
           hence "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb)
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   281
           also have "... = ta" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   282
           proof -
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   283
             have "X - {ya} \<subseteq> Ya"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   284
             proof 
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   285
               fix x assume x: "x \<in> X - {ya}"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   286
               then obtain a' where xeq: "x = ?gy a'" and a': "a' \<in> AA" 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   287
                 by (auto simp add: Xeq) 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   288
               hence "a' \<noteq> (LEAST x. x \<in> AA)" using x fields by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   289
               hence lessa': "(LEAST x. x \<in> AA) < a'"
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   290
                 using Least_le [of "%x. x \<in> AA", OF a'] by arith
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   291
               show "x \<in> Ya" using xeq fields rg [OF lessa'] by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   292
             qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   293
             moreover
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   294
             have "card (X - {ya}) = r"
24853
aab5798e5a33 added lemmas
nipkow
parents: 22665
diff changeset
   295
               by (simp add: cardX ya)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   296
             ultimately show ?thesis 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   297
               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
   298
               by (clarsimp simp del:insert_Diff_single)
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   299
           qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   300
           also have "... = s'" using AA AAleast fields by auto
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   301
           finally show ?thesis .
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   302
         qed}
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   303
        thus ?thesis by blast
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   304
      qed 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   305
    qed 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   306
  qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   307
qed
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   308
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   309
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   310
theorem Ramsey:
19949
0505dce27b0b fix/fixes: tuned type constraints;
wenzelm
parents: 19948
diff changeset
   311
  fixes s r :: nat and Z::"'a set" and f::"'a set => nat"
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   312
  shows
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   313
   "[|infinite Z;
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   314
      \<forall>X. X \<subseteq> Z & finite X & card X = r --> f X < s|]
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   315
  ==> \<exists>Y t. Y \<subseteq> Z & infinite Y & t < s 
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   316
            & (\<forall>X. X \<subseteq> Y & finite X & card X = r --> f X = t)"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   317
by (blast intro: Ramsey_induction [unfolded part_def])
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   318
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   319
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   320
corollary Ramsey2:
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   321
  fixes s::nat and Z::"'a set" and f::"'a set => nat"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   322
  assumes infZ: "infinite Z"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   323
      and part: "\<forall>x\<in>Z. \<forall>y\<in>Z. x\<noteq>y --> f{x,y} < s"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   324
  shows
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   325
   "\<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)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   326
proof -
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   327
  have part2: "\<forall>X. X \<subseteq> Z & finite X & card X = 2 --> f X < s"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 40695
diff changeset
   328
    using part by (fastforce simp add: eval_nat_numeral card_Suc_eq)
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   329
  obtain Y t 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   330
    where "Y \<subseteq> Z" "infinite Y" "t < s"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   331
          "(\<forall>X. X \<subseteq> Y & finite X & card X = 2 --> f X = t)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   332
    by (insert Ramsey [OF infZ part2]) auto
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   333
  moreover from this have  "\<forall>x\<in>Y. \<forall>y\<in>Y. x \<noteq> y \<longrightarrow> f {x, y} = t" by auto
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   334
  ultimately show ?thesis by iprover
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   335
qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   336
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   337
22665
cf152ff55d16 tuned document (headers, sections, spacing);
wenzelm
parents: 22367
diff changeset
   338
subsection {* Disjunctive Well-Foundedness *}
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   339
22367
6860f09242bf tuned document;
wenzelm
parents: 21634
diff changeset
   340
text {*
6860f09242bf tuned document;
wenzelm
parents: 21634
diff changeset
   341
  An application of Ramsey's theorem to program termination. See
6860f09242bf tuned document;
wenzelm
parents: 21634
diff changeset
   342
  \cite{Podelski-Rybalchenko}.
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   343
*}
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   344
20810
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   345
definition
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   346
  disj_wf         :: "('a * 'a)set => bool"
21634
369e38e35686 fixed definition syntax
krauss
parents: 20810
diff changeset
   347
where
20810
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   348
  "disj_wf r = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r = (\<Union>i<n. T i))"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   349
21634
369e38e35686 fixed definition syntax
krauss
parents: 20810
diff changeset
   350
definition
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   351
  transition_idx :: "[nat => 'a, nat => ('a*'a)set, nat set] => nat"
21634
369e38e35686 fixed definition syntax
krauss
parents: 20810
diff changeset
   352
where
20810
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   353
  "transition_idx s T A =
3377a830b727 moved theory Infinite_Set to Library;
wenzelm
parents: 19954
diff changeset
   354
    (LEAST k. \<exists>i j. A = {i,j} & i<j & (s j, s i) \<in> T k)"
19954
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   355
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   356
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   357
lemma transition_idx_less:
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   358
    "[|i<j; (s j, s i) \<in> T k; k<n|] ==> transition_idx s T {i,j} < n"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   359
apply (subgoal_tac "transition_idx s T {i, j} \<le> k", simp) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   360
apply (simp add: transition_idx_def, blast intro: Least_le) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   361
done
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   362
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   363
lemma transition_idx_in:
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   364
    "[|i<j; (s j, s i) \<in> T k|] ==> (s j, s i) \<in> T (transition_idx s T {i,j})"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   365
apply (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   366
            cong: conj_cong) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   367
apply (erule LeastI) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   368
done
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   369
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   370
text{*To be equal to the union of some well-founded relations is equivalent
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   371
to being the subset of such a union.*}
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   372
lemma disj_wf:
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   373
     "disj_wf(r) = (\<exists>T. \<exists>n::nat. (\<forall>i<n. wf(T i)) & r \<subseteq> (\<Union>i<n. T i))"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   374
apply (auto simp add: disj_wf_def) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   375
apply (rule_tac x="%i. T i Int r" in exI) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   376
apply (rule_tac x=n in exI) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   377
apply (force simp add: wf_Int1) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   378
done
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   379
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   380
theorem trans_disj_wf_implies_wf:
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   381
  assumes transr: "trans r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   382
      and dwf:    "disj_wf(r)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   383
  shows "wf r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   384
proof (simp only: wf_iff_no_infinite_down_chain, rule notI)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   385
  assume "\<exists>s. \<forall>i. (s (Suc i), s i) \<in> r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   386
  then obtain s where sSuc: "\<forall>i. (s (Suc i), s i) \<in> r" ..
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   387
  have s: "!!i j. i < j ==> (s j, s i) \<in> r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   388
  proof -
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   389
    fix i and j::nat
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   390
    assume less: "i<j"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   391
    thus "(s j, s i) \<in> r"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   392
    proof (rule less_Suc_induct)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   393
      show "\<And>i. (s (Suc i), s i) \<in> r" by (simp add: sSuc) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   394
      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"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   395
        using transr by (unfold trans_def, blast) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   396
    qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   397
  qed    
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   398
  from dwf
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   399
  obtain T and n::nat where wfT: "\<forall>k<n. wf(T k)" and r: "r = (\<Union>k<n. T k)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   400
    by (auto simp add: disj_wf_def)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   401
  have s_in_T: "\<And>i j. i<j ==> \<exists>k. (s j, s i) \<in> T k & k<n"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   402
  proof -
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   403
    fix i and j::nat
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   404
    assume less: "i<j"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   405
    hence "(s j, s i) \<in> r" by (rule s [of i j]) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   406
    thus "\<exists>k. (s j, s i) \<in> T k & k<n" by (auto simp add: r)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   407
  qed    
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   408
  have trless: "!!i j. i\<noteq>j ==> transition_idx s T {i,j} < n"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   409
    apply (auto simp add: linorder_neq_iff)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   410
    apply (blast dest: s_in_T transition_idx_less) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   411
    apply (subst insert_commute)   
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   412
    apply (blast dest: s_in_T transition_idx_less) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   413
    done
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   414
  have
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   415
   "\<exists>K k. K \<subseteq> UNIV & infinite K & k < n & 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   416
          (\<forall>i\<in>K. \<forall>j\<in>K. i\<noteq>j --> transition_idx s T {i,j} = k)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   417
    by (rule Ramsey2) (auto intro: trless nat_infinite) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   418
  then obtain K and k 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   419
    where infK: "infinite K" and less: "k < n" and
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   420
          allk: "\<forall>i\<in>K. \<forall>j\<in>K. i\<noteq>j --> transition_idx s T {i,j} = k"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   421
    by auto
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   422
  have "\<forall>m. (s (enumerate K (Suc m)), s(enumerate K m)) \<in> T k"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   423
  proof
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   424
    fix m::nat
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   425
    let ?j = "enumerate K (Suc m)"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   426
    let ?i = "enumerate K m"
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   427
    have jK: "?j \<in> K" by (simp add: enumerate_in_set infK) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   428
    have iK: "?i \<in> K" by (simp add: enumerate_in_set infK) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   429
    have ij: "?i < ?j" by (simp add: enumerate_step infK) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   430
    have ijk: "transition_idx s T {?i,?j} = k" using iK jK ij 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   431
      by (simp add: allk)
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   432
    obtain k' where "(s ?j, s ?i) \<in> T k'" "k'<n" 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   433
      using s_in_T [OF ij] by blast
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   434
    thus "(s ?j, s ?i) \<in> T k" 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   435
      by (simp add: ijk [symmetric] transition_idx_in ij) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   436
  qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   437
  hence "~ wf(T k)" by (force simp add: wf_iff_no_infinite_down_chain) 
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   438
  thus False using wfT less by blast
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   439
qed
e4c9f6946db3 disjunctive wellfoundedness
paulson
parents: 19949
diff changeset
   440
19944
60e0cbeae3d8 Introduction of Ramsey's theorem
paulson
parents:
diff changeset
   441
end