src/HOL/ex/Primrec.thy
author wenzelm
Tue, 10 Nov 2020 12:48:56 +0100
changeset 72571 ab4a0b19648a
parent 69880 fe3c12990893
child 75014 0778d233964d
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
     1
(*  Title:      HOL/ex/Primrec.thy
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     3
    Copyright   1997  University of Cambridge
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     4
*)
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     5
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
     6
section \<open>Ackermann's Function and the Primitive Recursive Functions\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 11704
diff changeset
     8
theory Primrec imports Main begin
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
     9
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    10
text \<open>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    11
  Proof adopted from
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    12
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    13
  Nora Szasz, A Machine Checked Proof that Ackermann's Function is not
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    14
  Primitive Recursive, In: Huet \& Plotkin, eds., Logical Environments
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    15
  (CUP, 1993), 317-338.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    16
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    17
  See also E. Mendelson, Introduction to Mathematical Logic.  (Van
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    18
  Nostrand, 1964), page 250, exercise 11.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    19
  \medskip
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    20
\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    21
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    22
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    23
subsection\<open>Ackermann's Function\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    24
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    25
fun ack :: "[nat,nat] \<Rightarrow> nat" where
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    26
  "ack 0 n =  Suc n"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    27
| "ack (Suc m) 0 = ack m 1"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    28
| "ack (Suc m) (Suc n) = ack m (ack (Suc m) n)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    29
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    30
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    31
text \<open>PROPERTY A 4\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    32
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    33
lemma less_ack2 [iff]: "j < ack i j"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    34
  by (induct i j rule: ack.induct) simp_all
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    35
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    36
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    37
text \<open>PROPERTY A 5-, the single-step lemma\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    38
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    39
lemma ack_less_ack_Suc2 [iff]: "ack i j < ack i (Suc j)"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    40
  by (induct i j rule: ack.induct) simp_all
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    41
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    42
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    43
text \<open>PROPERTY A 5, monotonicity for \<open><\<close>\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    44
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    45
lemma ack_less_mono2: "j < k \<Longrightarrow> ack i j < ack i k"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    46
  by (simp add: lift_Suc_mono_less)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    47
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    48
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    49
text \<open>PROPERTY A 5', monotonicity for \<open>\<le>\<close>\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    50
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    51
lemma ack_le_mono2: "j \<le> k \<Longrightarrow> ack i j \<le> ack i k"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    52
  by (simp add: ack_less_mono2 less_mono_imp_le_mono)
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
    53
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    54
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    55
text \<open>PROPERTY A 6\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    56
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    57
lemma ack2_le_ack1 [iff]: "ack i (Suc j) \<le> ack (Suc i) j"
26072
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    58
proof (induct j)
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    59
  case 0 show ?case by simp
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    60
next
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    61
  case (Suc j) show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    62
    by (metis Suc ack.simps(3) ack_le_mono2 le_trans less_ack2 less_eq_Suc_le) 
26072
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    63
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    64
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    65
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    66
text \<open>PROPERTY A 7-, the single-step lemma\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    67
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    68
lemma ack_less_ack_Suc1 [iff]: "ack i j < ack (Suc i) j"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    69
  by (blast intro: ack_less_mono2 less_le_trans)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    70
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    71
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63913
diff changeset
    72
text \<open>PROPERTY A 4'? Extra lemma needed for \<^term>\<open>CONSTANT\<close> case, constant functions\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    73
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    74
lemma less_ack1 [iff]: "i < ack i j"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    75
proof (induct i)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    76
  case 0
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    77
  then show ?case 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    78
    by simp
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    79
next
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    80
  case (Suc i)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    81
  then show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    82
    using less_trans_Suc by blast
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    83
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    84
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    85
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    86
text \<open>PROPERTY A 8\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    87
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    88
lemma ack_1 [simp]: "ack (Suc 0) j = j + 2"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    89
  by (induct j) simp_all
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    90
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    91
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63913
diff changeset
    92
text \<open>PROPERTY A 9.  The unary \<open>1\<close> and \<open>2\<close> in \<^term>\<open>ack\<close> is essential for the rewriting.\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    93
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    94
lemma ack_2 [simp]: "ack (Suc (Suc 0)) j = 2 * j + 3"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    95
  by (induct j) simp_all
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
    96
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
    97
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    98
text \<open>PROPERTY A 7, monotonicity for \<open><\<close> [not clear why
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    99
  @{thm [source] ack_1} is now needed first!]\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   100
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   101
lemma ack_less_mono1_aux: "ack i k < ack (Suc (i +i')) k"
34055
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   102
proof (induct i k rule: ack.induct)
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   103
  case (1 n) show ?case
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   104
    using less_le_trans by auto
34055
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   105
next
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   106
  case (2 m) thus ?case by simp
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   107
next
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   108
  case (3 m n) thus ?case
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   109
    using ack_less_mono2 less_trans by fastforce
34055
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   110
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   111
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   112
lemma ack_less_mono1: "i < j \<Longrightarrow> ack i k < ack j k"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   113
  using ack_less_mono1_aux less_iff_Suc_add by auto
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   114
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   115
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   116
text \<open>PROPERTY A 7', monotonicity for \<open>\<le>\<close>\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   117
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   118
lemma ack_le_mono1: "i \<le> j \<Longrightarrow> ack i k \<le> ack j k"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   119
  using ack_less_mono1 le_eq_less_or_eq by auto
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   120
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   121
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   122
text \<open>PROPERTY A 10\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   123
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   124
lemma ack_nest_bound: "ack i1 (ack i2 j) < ack (2 + (i1 + i2)) j"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   125
proof -
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   126
  have "ack i1 (ack i2 j) < ack (i1 + i2) (ack (Suc (i1 + i2)) j)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   127
    by (meson ack_le_mono1 ack_less_mono1 ack_less_mono2 le_add1 le_trans less_add_Suc2 not_less)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   128
  also have "... = ack (Suc (i1 + i2)) (Suc j)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   129
    by simp
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   130
  also have "... \<le> ack (2 + (i1 + i2)) j"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   131
    using ack2_le_ack1 add_2_eq_Suc by presburger
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   132
  finally show ?thesis .
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   133
qed
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   134
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   135
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   136
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   137
text \<open>PROPERTY A 11\<close>
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   138
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   139
lemma ack_add_bound: "ack i1 j + ack i2 j < ack (4 + (i1 + i2)) j"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   140
proof -
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   141
  have "ack i1 j \<le> ack (i1 + i2) j" "ack i2 j \<le> ack (i1 + i2) j"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   142
    by (simp_all add: ack_le_mono1)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   143
  then have "ack i1 j + ack i2 j < ack (Suc (Suc 0)) (ack (i1 + i2) j)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   144
    by simp
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   145
  also have "... < ack (4 + (i1 + i2)) j"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   146
    by (metis ack_nest_bound add.assoc numeral_2_eq_2 numeral_Bit0)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   147
  finally show ?thesis .
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   148
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   149
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   150
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   151
text \<open>PROPERTY A 12.  Article uses existential quantifier but the ALF proof
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   152
  used \<open>k + 4\<close>.  Quantified version must be nested \<open>\<exists>k'. \<forall>i j. ...\<close>\<close>
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   153
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   154
lemma ack_add_bound2: 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   155
  assumes "i < ack k j" shows "i + j < ack (4 + k) j"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   156
proof -
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   157
  have "i + j < ack k j + ack 0 j"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   158
    using assms by auto
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   159
  also have "... < ack (4 + k) j"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   160
    by (metis ack_add_bound add.right_neutral)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   161
  finally show ?thesis .
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   162
qed
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   163
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   164
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   165
subsection\<open>Primitive Recursive Functions\<close>
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   166
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   167
primrec hd0 :: "nat list \<Rightarrow> nat" where
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   168
  "hd0 [] = 0" 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   169
| "hd0 (m # ms) = m"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   170
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   171
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   172
text \<open>Inductive definition of the set of primitive recursive functions of type \<^typ>\<open>nat list \<Rightarrow> nat\<close>.\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   173
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   174
definition SC :: "nat list \<Rightarrow> nat" 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   175
  where "SC l = Suc (hd0 l)"
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   176
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   177
definition CONSTANT :: "nat \<Rightarrow> nat list \<Rightarrow> nat" 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   178
  where "CONSTANT k l = k"
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   179
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   180
definition PROJ :: "nat \<Rightarrow> nat list \<Rightarrow> nat" 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   181
  where "PROJ i l = hd0 (drop i l)"
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   182
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   183
definition COMP :: "[nat list \<Rightarrow> nat, (nat list \<Rightarrow> nat) list, nat list] \<Rightarrow> nat"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   184
  where "COMP g fs l = g (map (\<lambda>f. f l) fs)"
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   185
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   186
fun PREC :: "[nat list \<Rightarrow> nat, nat list \<Rightarrow> nat, nat list] \<Rightarrow> nat"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   187
  where
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   188
    "PREC f g [] = 0"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   189
  | "PREC f g (x # l) = rec_nat (f l) (\<lambda>y r. g (r # y # l)) x"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   190
    \<comment> \<open>Note that \<^term>\<open>g\<close> is applied first to \<^term>\<open>PREC f g y\<close> and then to \<^term>\<open>y\<close>!\<close>
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   191
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   192
inductive PRIMREC :: "(nat list \<Rightarrow> nat) \<Rightarrow> bool" where
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   193
  SC: "PRIMREC SC"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   194
| CONSTANT: "PRIMREC (CONSTANT k)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   195
| PROJ: "PRIMREC (PROJ i)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   196
| COMP: "PRIMREC g \<Longrightarrow> \<forall>f \<in> set fs. PRIMREC f \<Longrightarrow> PRIMREC (COMP g fs)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   197
| PREC: "PRIMREC f \<Longrightarrow> PRIMREC g \<Longrightarrow> PRIMREC (PREC f g)"
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   198
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   199
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   200
text \<open>Useful special cases of evaluation\<close>
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   201
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   202
lemma SC [simp]: "SC (x # l) = Suc x"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   203
  by (simp add: SC_def)
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   204
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   205
lemma PROJ_0 [simp]: "PROJ 0 (x # l) = x"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   206
  by (simp add: PROJ_def)
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   207
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   208
lemma COMP_1 [simp]: "COMP g [f] l = g [f l]"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   209
  by (simp add: COMP_def)
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   210
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   211
lemma PREC_0: "PREC f g (0 # l) = f l"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   212
  by simp
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   213
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   214
lemma PREC_Suc [simp]: "PREC f g (Suc x # l) = g (PREC f g (x # l) # x # l)"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   215
  by auto
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   216
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   217
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   218
subsection \<open>MAIN RESULT\<close>
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   219
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 61933
diff changeset
   220
lemma SC_case: "SC l < ack 1 (sum_list l)"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   221
  unfolding SC_def
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   222
  by (induct l) (simp_all add: le_add1 le_imp_less_Suc)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   223
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 61933
diff changeset
   224
lemma CONSTANT_case: "CONSTANT k l < ack k (sum_list l)"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   225
  by (simp add: CONSTANT_def)
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   226
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 61933
diff changeset
   227
lemma PROJ_case: "PROJ i l < ack 0 (sum_list l)"
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   228
  unfolding PROJ_def
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   229
proof (induct l arbitrary: i)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   230
  case Nil
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   231
  then show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   232
    by simp
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   233
next
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   234
  case (Cons a l)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   235
  then show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   236
    by (metis ack.simps(1) add.commute drop_Cons' hd0.simps(2) leD leI lessI not_less_eq sum_list.Cons trans_le_add2)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   237
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   238
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   239
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63913
diff changeset
   240
text \<open>\<^term>\<open>COMP\<close> case\<close>
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   241
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 61933
diff changeset
   242
lemma COMP_map_aux: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (sum_list l))
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   243
  \<Longrightarrow> \<exists>k. \<forall>l. sum_list (map (\<lambda>f. f l) fs) < ack k (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   244
proof (induct fs)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   245
  case Nil
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   246
  then show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   247
    by auto
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   248
next
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   249
  case (Cons a fs)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   250
  then show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   251
    by simp (blast intro: add_less_mono ack_add_bound less_trans)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   252
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   253
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   254
lemma COMP_case:
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   255
  assumes 1: "\<forall>l. g l < ack kg (sum_list l)" 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   256
      and 2: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (sum_list l))"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   257
  shows "\<exists>k. \<forall>l. COMP g fs  l < ack k (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   258
  unfolding COMP_def
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   259
  using 1 COMP_map_aux [OF 2] by (meson ack_less_mono2 ack_nest_bound less_trans)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   260
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63913
diff changeset
   261
text \<open>\<^term>\<open>PREC\<close> case\<close>
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   262
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   263
lemma PREC_case_aux:
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   264
  assumes f: "\<And>l. f l + sum_list l < ack kf (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   265
      and g: "\<And>l. g l + sum_list l < ack kg (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   266
  shows "PREC f g l + sum_list l < ack (Suc (kf + kg)) (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   267
proof (cases l)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   268
  case Nil
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   269
  then show ?thesis
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   270
    by (simp add: Suc_lessD)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   271
next
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   272
  case (Cons m l)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   273
  have "rec_nat (f l) (\<lambda>y r. g (r # y # l)) m + (m + sum_list l) < ack (Suc (kf + kg)) (m + sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   274
  proof (induct m)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   275
    case 0
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   276
    then show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   277
      using ack_less_mono1_aux f less_trans by fastforce
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   278
  next
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   279
    case (Suc m)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   280
    let ?r = "rec_nat (f l) (\<lambda>y r. g (r # y # l)) m"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   281
    have "\<not> g (?r # m # l) + sum_list (?r # m # l) < g (?r # m # l) + (m + sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   282
      by force
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   283
    then have "g (?r # m # l) + (m + sum_list l) < ack kg (sum_list (?r # m # l))"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   284
      by (meson assms(2) leI less_le_trans)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   285
    moreover 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   286
    have "... < ack (kf + kg) (ack (Suc (kf + kg)) (m + sum_list l))"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   287
      using Suc.hyps by simp (meson ack_le_mono1 ack_less_mono2 le_add2 le_less_trans)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   288
    ultimately show ?case
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   289
      by auto
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   290
  qed
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   291
  then show ?thesis
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   292
    by (simp add: local.Cons)
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   293
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   294
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   295
proposition PREC_case:
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   296
  "\<lbrakk>\<And>l. f l < ack kf (sum_list l); \<And>l. g l < ack kg (sum_list l)\<rbrakk> 
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   297
  \<Longrightarrow> \<exists>k. \<forall>l. PREC f g l < ack k (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   298
  by (metis le_less_trans [OF le_add1 PREC_case_aux] ack_add_bound2)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   299
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   300
lemma ack_bounds_PRIMREC: "PRIMREC f \<Longrightarrow> \<exists>k. \<forall>l. f l < ack k (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   301
  by (erule PRIMREC.induct) (blast intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)+
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   302
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   303
theorem ack_not_PRIMREC:
69880
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   304
  "\<not> PRIMREC (\<lambda>l. case l of [] \<Rightarrow> 0 | x # l' \<Rightarrow> ack x x)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   305
proof
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   306
  assume *: "PRIMREC (\<lambda>l. case l of [] \<Rightarrow> 0 | x # l' \<Rightarrow> ack x x)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   307
  then obtain m where m: "\<And>l. (case l of [] \<Rightarrow> 0 | x # l' \<Rightarrow> ack x x) < ack m (sum_list l)"
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   308
    using ack_bounds_PRIMREC by metis
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   309
  show False
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   310
    using m [of "[m]"] by simp
fe3c12990893 tidied up HOL/ex/Primrec
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   311
qed
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   312
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   313
end