src/HOL/Proofs/Extraction/Pigeonhole.thy
author wenzelm
Tue, 07 Oct 2014 23:12:08 +0200
changeset 58622 aa99568f56de
parent 55535 10194808430d
child 58889 5b7a9633cfa8
permissions -rw-r--r--
more antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39157
b98909faaea8 more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
wenzelm
parents: 37678
diff changeset
     1
(*  Title:      HOL/Proofs/Extraction/Pigeonhole.thy
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
     3
*)
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
     4
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
     5
header {* The pigeonhole principle *}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
     6
22737
haftmann
parents: 22507
diff changeset
     7
theory Pigeonhole
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 45169
diff changeset
     8
imports Util "~~/src/HOL/Library/Code_Target_Numeral"
22737
haftmann
parents: 22507
diff changeset
     9
begin
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    10
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    11
text {*
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    12
We formalize two proofs of the pigeonhole principle, which lead
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    13
to extracted programs of quite different complexity. The original
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    14
formalization of these proofs in {\sc Nuprl} is due to
58622
aa99568f56de more antiquotations;
wenzelm
parents: 55535
diff changeset
    15
Aleksey Nogin @{cite "Nogin-ENTCS-2000"}.
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    16
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    17
This proof yields a polynomial program.
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    18
*}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    19
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    20
theorem pigeonhole:
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    21
  "\<And>f. (\<And>i. i \<le> Suc n \<Longrightarrow> f i \<le> n) \<Longrightarrow> \<exists>i j. i \<le> Suc n \<and> j < i \<and> f i = f j"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    22
proof (induct n)
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    23
  case 0
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    24
  hence "Suc 0 \<le> Suc 0 \<and> 0 < Suc 0 \<and> f (Suc 0) = f 0" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    25
  thus ?case by iprover
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    26
next
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    27
  case (Suc n)
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    28
  {
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    29
    fix k
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    30
    have
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    31
      "k \<le> Suc (Suc n) \<Longrightarrow>
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    32
      (\<And>i j. Suc k \<le> i \<Longrightarrow> i \<le> Suc (Suc n) \<Longrightarrow> j < i \<Longrightarrow> f i \<noteq> f j) \<Longrightarrow>
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    33
      (\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j)"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    34
    proof (induct k)
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    35
      case 0
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    36
      let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    37
      have "\<not> (\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j)"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    38
      proof
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    39
        assume "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    40
        then obtain i j where i: "i \<le> Suc n" and j: "j < i"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    41
          and f: "?f i = ?f j" by iprover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    42
        from j have i_nz: "Suc 0 \<le> i" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    43
        from i have iSSn: "i \<le> Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    44
        have S0SSn: "Suc 0 \<le> Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    45
        show False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    46
        proof cases
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    47
          assume fi: "f i = Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    48
          show False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    49
          proof cases
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    50
            assume fj: "f j = Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    51
            from i_nz and iSSn and j have "f i \<noteq> f j" by (rule 0)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    52
            moreover from fi have "f i = f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    53
              by (simp add: fj [symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    54
            ultimately show ?thesis ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    55
          next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    56
            from i and j have "j < Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    57
            with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    58
              by (rule 0)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    59
            moreover assume "f j \<noteq> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    60
            with fi and f have "f (Suc (Suc n)) = f j" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    61
            ultimately show False ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    62
          qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    63
        next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    64
          assume fi: "f i \<noteq> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    65
          show False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    66
          proof cases
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    67
            from i have "i < Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    68
            with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f i"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    69
              by (rule 0)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    70
            moreover assume "f j = Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    71
            with fi and f have "f (Suc (Suc n)) = f i" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    72
            ultimately show False ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    73
          next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    74
            from i_nz and iSSn and j
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    75
            have "f i \<noteq> f j" by (rule 0)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    76
            moreover assume "f j \<noteq> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    77
            with fi and f have "f i = f j" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    78
            ultimately show False ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    79
          qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    80
        qed
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    81
      qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    82
      moreover have "\<And>i. i \<le> Suc n \<Longrightarrow> ?f i \<le> n"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    83
      proof -
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    84
        fix i assume "i \<le> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    85
        hence i: "i < Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    86
        have "f (Suc (Suc n)) \<noteq> f i"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    87
          by (rule 0) (simp_all add: i)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    88
        moreover have "f (Suc (Suc n)) \<le> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    89
          by (rule Suc) simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    90
        moreover from i have "i \<le> Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    91
        hence "f i \<le> Suc n" by (rule Suc)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    92
        ultimately show "?thesis i"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    93
          by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    94
      qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    95
      hence "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
    96
        by (rule Suc)
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    97
      ultimately show ?case ..
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    98
    next
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
    99
      case (Suc k)
25418
d4f80cb18c93 Moved nat_eq_dec and search to Util.thy
berghofe
parents: 24348
diff changeset
   100
      from search [OF nat_eq_dec] show ?case
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   101
      proof
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   102
        assume "\<exists>j<Suc k. f (Suc k) = f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   103
        thus ?case by (iprover intro: le_refl)
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   104
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   105
        assume nex: "\<not> (\<exists>j<Suc k. f (Suc k) = f j)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   106
        have "\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   107
        proof (rule Suc)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   108
          from Suc show "k \<le> Suc (Suc n)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   109
          fix i j assume k: "Suc k \<le> i" and i: "i \<le> Suc (Suc n)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   110
            and j: "j < i"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   111
          show "f i \<noteq> f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   112
          proof cases
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   113
            assume eq: "i = Suc k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   114
            show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   115
            proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   116
              assume "f i = f j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   117
              hence "f (Suc k) = f j" by (simp add: eq)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   118
              with nex and j and eq show False by iprover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   119
            qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   120
          next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   121
            assume "i \<noteq> Suc k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   122
            with k have "Suc (Suc k) \<le> i" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   123
            thus ?thesis using i and j by (rule Suc)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   124
          qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   125
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   126
        thus ?thesis by (iprover intro: le_SucI)
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   127
      qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   128
    qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   129
  }
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   130
  note r = this
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   131
  show ?case by (rule r) simp_all
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   132
qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   133
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   134
text {*
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   135
The following proof, although quite elegant from a mathematical point of view,
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   136
leads to an exponential program:
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   137
*}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   138
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   139
theorem pigeonhole_slow:
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   140
  "\<And>f. (\<And>i. i \<le> Suc n \<Longrightarrow> f i \<le> n) \<Longrightarrow> \<exists>i j. i \<le> Suc n \<and> j < i \<and> f i = f j"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   141
proof (induct n)
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   142
  case 0
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   143
  have "Suc 0 \<le> Suc 0" ..
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   144
  moreover have "0 < Suc 0" ..
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   145
  moreover from 0 have "f (Suc 0) = f 0" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
   146
  ultimately show ?case by iprover
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   147
next
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   148
  case (Suc n)
25418
d4f80cb18c93 Moved nat_eq_dec and search to Util.thy
berghofe
parents: 24348
diff changeset
   149
  from search [OF nat_eq_dec] show ?case
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   150
  proof
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   151
    assume "\<exists>j < Suc (Suc n). f (Suc (Suc n)) = f j"
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
   152
    thus ?case by (iprover intro: le_refl)
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   153
  next
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   154
    assume "\<not> (\<exists>j < Suc (Suc n). f (Suc (Suc n)) = f j)"
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
   155
    hence nex: "\<forall>j < Suc (Suc n). f (Suc (Suc n)) \<noteq> f j" by iprover
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   156
    let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   157
    have "\<And>i. i \<le> Suc n \<Longrightarrow> ?f i \<le> n"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   158
    proof -
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   159
      fix i assume i: "i \<le> Suc n"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   160
      show "?thesis i"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   161
      proof (cases "f i = Suc n")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   162
        case True
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   163
        from i and nex have "f (Suc (Suc n)) \<noteq> f i" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   164
        with True have "f (Suc (Suc n)) \<noteq> Suc n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   165
        moreover from Suc have "f (Suc (Suc n)) \<le> Suc n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   166
        ultimately have "f (Suc (Suc n)) \<le> n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   167
        with True show ?thesis by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   168
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   169
        case False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   170
        from Suc and i have "f i \<le> Suc n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   171
        with False show ?thesis by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   172
      qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   173
    qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   174
    hence "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" by (rule Suc)
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   175
    then obtain i j where i: "i \<le> Suc n" and ji: "j < i" and f: "?f i = ?f j"
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
   176
      by iprover
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   177
    have "f i = f j"
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   178
    proof (cases "f i = Suc n")
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   179
      case True
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   180
      show ?thesis
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   181
      proof (cases "f j = Suc n")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   182
        assume "f j = Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   183
        with True show ?thesis by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   184
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   185
        assume "f j \<noteq> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   186
        moreover from i ji nex have "f (Suc (Suc n)) \<noteq> f j" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   187
        ultimately show ?thesis using True f by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   188
      qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   189
    next
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   190
      case False
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   191
      show ?thesis
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   192
      proof (cases "f j = Suc n")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   193
        assume "f j = Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   194
        moreover from i nex have "f (Suc (Suc n)) \<noteq> f i" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   195
        ultimately show ?thesis using False f by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   196
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   197
        assume "f j \<noteq> Suc n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   198
        with False f show ?thesis by simp
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   199
      qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   200
    qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   201
    moreover from i have "i \<le> Suc (Suc n)" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
   202
    ultimately show ?thesis using ji by iprover
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   203
  qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   204
qed
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   205
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   206
extract pigeonhole pigeonhole_slow
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   207
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   208
text {*
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   209
The programs extracted from the above proofs look as follows:
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   210
@{thm [display] pigeonhole_def}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   211
@{thm [display] pigeonhole_slow_def}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   212
The program for searching for an element in an array is
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   213
@{thm [display,eta_contract=false] search_def}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   214
The correctness statement for @{term "pigeonhole"} is
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   215
@{thm [display] pigeonhole_correctness [no_vars]}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   216
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   217
In order to analyze the speed of the above programs,
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   218
we generate ML code from them.
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   219
*}
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   220
27982
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   221
instantiation nat :: default
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   222
begin
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   223
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   224
definition "default = (0::nat)"
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   225
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   226
instance ..
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   227
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   228
end
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   229
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37287
diff changeset
   230
instantiation prod :: (default, default) default
27982
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   231
begin
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   232
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   233
definition "default = (default, default)"
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   234
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   235
instance ..
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   236
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   237
end
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
   238
20837
099877d83d2b added example for code_gen
haftmann
parents: 20593
diff changeset
   239
definition
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 45169
diff changeset
   240
  "test n u = pigeonhole (nat_of_integer n) (\<lambda>m. m - 1)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21127
diff changeset
   241
definition
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 45169
diff changeset
   242
  "test' n u = pigeonhole_slow (nat_of_integer n) (\<lambda>m. m - 1)"
22507
haftmann
parents: 21545
diff changeset
   243
definition
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 45169
diff changeset
   244
  "test'' u = pigeonhole 8 (List.nth [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])"
20837
099877d83d2b added example for code_gen
haftmann
parents: 20593
diff changeset
   245
51272
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   246
ML_val "timeit (@{code test} 10)" 
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   247
ML_val "timeit (@{code test'} 10)"
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   248
ML_val "timeit (@{code test} 20)"
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   249
ML_val "timeit (@{code test'} 20)"
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   250
ML_val "timeit (@{code test} 25)"
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   251
ML_val "timeit (@{code test'} 25)"
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   252
ML_val "timeit (@{code test} 500)"
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
   253
ML_val "timeit @{code test''}"
37287
9834c21c4ba1 modernized
haftmann
parents: 32960
diff changeset
   254
17024
ae4a8446df16 New case study: pigeonhole principle.
berghofe
parents:
diff changeset
   255
end