src/HOL/ex/Primrec.thy
author huffman
Thu, 14 Dec 2006 21:03:39 +0100
changeset 21850 bf253f7075b4
parent 21404 eb85850d3eb7
child 22283 26140713540b
permissions -rw-r--r--
remove usage of ultra tactic
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
    ID:         $Id$
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     4
    Copyright   1997  University of Cambridge
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     5
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
     6
Primitive Recursive Functions.  Demonstrates recursive definitions,
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
     7
the TFL package.
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     8
*)
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
     9
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    10
header {* Primitive Recursive Functions *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    11
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 11704
diff changeset
    12
theory Primrec imports Main begin
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    13
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    14
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    15
  Proof adopted from
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    16
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    17
  Nora Szasz, A Machine Checked Proof that Ackermann's Function is not
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    18
  Primitive Recursive, In: Huet \& Plotkin, eds., Logical Environments
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    19
  (CUP, 1993), 317-338.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    20
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    21
  See also E. Mendelson, Introduction to Mathematical Logic.  (Van
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    22
  Nostrand, 1964), page 250, exercise 11.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    23
  \medskip
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    24
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    25
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    26
consts ack :: "nat * nat => nat"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    27
recdef ack  "less_than <*lex*> less_than"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    28
  "ack (0, n) =  Suc n"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    29
  "ack (Suc m, 0) = ack (m, 1)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    30
  "ack (Suc m, Suc n) = ack (m, ack (Suc m, n))"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    31
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    32
consts list_add :: "nat list => nat"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    33
primrec
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    34
  "list_add [] = 0"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    35
  "list_add (m # ms) = m + list_add ms"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    36
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    37
consts zeroHd :: "nat list => nat"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    38
primrec
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    39
  "zeroHd [] = 0"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    40
  "zeroHd (m # ms) = m"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    41
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    42
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    43
text {* The set of primitive recursive functions of type @{typ "nat list => nat"}. *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    44
19736
wenzelm
parents: 19676
diff changeset
    45
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    46
  SC :: "nat list => nat" where
19736
wenzelm
parents: 19676
diff changeset
    47
  "SC l = Suc (zeroHd l)"
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
    48
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    49
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    50
  CONSTANT :: "nat => nat list => nat" where
19736
wenzelm
parents: 19676
diff changeset
    51
  "CONSTANT k l = k"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    52
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    53
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    54
  PROJ :: "nat => nat list => nat" where
19736
wenzelm
parents: 19676
diff changeset
    55
  "PROJ i l = zeroHd (drop i l)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    56
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    57
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    58
  COMP :: "(nat list => nat) => (nat list => nat) list => nat list => nat" where
19736
wenzelm
parents: 19676
diff changeset
    59
  "COMP g fs l = g (map (\<lambda>f. f l) fs)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    60
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    61
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    62
  PREC :: "(nat list => nat) => (nat list => nat) => nat list => nat" where
19736
wenzelm
parents: 19676
diff changeset
    63
  "PREC f g l =
wenzelm
parents: 19676
diff changeset
    64
    (case l of
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    65
      [] => 0
19736
wenzelm
parents: 19676
diff changeset
    66
    | x # l' => nat_rec (f l') (\<lambda>y r. g (r # y # l')) x)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    67
  -- {* Note that @{term g} is applied first to @{term "PREC f g y"} and then to @{term y}! *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    68
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    69
consts PRIMREC :: "(nat list => nat) set"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    70
inductive PRIMREC
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    71
  intros
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    72
    SC: "SC \<in> PRIMREC"
19676
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
    73
    CONSTANT: "CONSTANT k \<in> PRIMREC"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    74
    PROJ: "PROJ i \<in> PRIMREC"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    75
    COMP: "g \<in> PRIMREC ==> fs \<in> lists PRIMREC ==> COMP g fs \<in> PRIMREC"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    76
    PREC: "f \<in> PRIMREC ==> g \<in> PRIMREC ==> PREC f g \<in> PRIMREC"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    77
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    78
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    79
text {* Useful special cases of evaluation *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    80
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    81
lemma SC [simp]: "SC (x # l) = Suc x"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    82
  apply (simp add: SC_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    83
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    84
19676
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
    85
lemma CONSTANT [simp]: "CONSTANT k l = k"
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
    86
  apply (simp add: CONSTANT_def)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    87
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    88
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    89
lemma PROJ_0 [simp]: "PROJ 0 (x # l) = x"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    90
  apply (simp add: PROJ_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    91
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    92
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    93
lemma COMP_1 [simp]: "COMP g [f] l = g [f l]"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    94
  apply (simp add: COMP_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    95
  done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
    96
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    97
lemma PREC_0 [simp]: "PREC f g (0 # l) = f l"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    98
  apply (simp add: PREC_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    99
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   100
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   101
lemma PREC_Suc [simp]: "PREC f g (Suc x # l) = g (PREC f g (x # l) # x # l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   102
  apply (simp add: PREC_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   103
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   104
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   105
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   106
text {* PROPERTY A 4 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   107
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   108
lemma less_ack2 [iff]: "j < ack (i, j)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   109
  apply (induct i j rule: ack.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   110
    apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   111
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   112
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   113
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   114
text {* PROPERTY A 5-, the single-step lemma *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   115
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   116
lemma ack_less_ack_Suc2 [iff]: "ack(i, j) < ack (i, Suc j)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   117
  apply (induct i j rule: ack.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   118
    apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   119
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   120
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   121
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   122
text {* PROPERTY A 5, monotonicity for @{text "<"} *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   123
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   124
lemma ack_less_mono2: "j < k ==> ack (i, j) < ack (i, k)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   125
  apply (induct i k rule: ack.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   126
    apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   127
  apply (blast elim!: less_SucE intro: less_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   128
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   129
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   130
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   131
text {* PROPERTY A 5', monotonicity for @{text \<le>} *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   132
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   133
lemma ack_le_mono2: "j \<le> k ==> ack (i, j) \<le> ack (i, k)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   134
  apply (simp add: order_le_less)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   135
  apply (blast intro: ack_less_mono2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   136
  done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   137
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   138
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   139
text {* PROPERTY A 6 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   140
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   141
lemma ack2_le_ack1 [iff]: "ack (i, Suc j) \<le> ack (Suc i, j)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   142
  apply (induct j)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   143
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   144
  apply (blast intro: ack_le_mono2 less_ack2 [THEN Suc_leI] le_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   145
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   146
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   147
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   148
text {* PROPERTY A 7-, the single-step lemma *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   149
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   150
lemma ack_less_ack_Suc1 [iff]: "ack (i, j) < ack (Suc i, j)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   151
  apply (blast intro: ack_less_mono2 less_le_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   152
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   153
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   154
19676
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
   155
text {* PROPERTY A 4'? Extra lemma needed for @{term CONSTANT} case, constant functions *}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   156
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   157
lemma less_ack1 [iff]: "i < ack (i, j)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   158
  apply (induct i)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   159
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   160
  apply (blast intro: Suc_leI le_less_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   161
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   162
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   163
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   164
text {* PROPERTY A 8 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   165
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   166
lemma ack_1 [simp]: "ack (Suc 0, j) = j + 2"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   167
  apply (induct j)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   168
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   169
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   170
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   171
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11464
diff changeset
   172
text {* PROPERTY A 9.  The unary @{text 1} and @{text 2} in @{term
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   173
  ack} is essential for the rewriting. *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   174
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   175
lemma ack_2 [simp]: "ack (Suc (Suc 0), j) = 2 * j + 3"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   176
  apply (induct j)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   177
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   178
  done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   179
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   180
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   181
text {* PROPERTY A 7, monotonicity for @{text "<"} [not clear why
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   182
  @{thm [source] ack_1} is now needed first!] *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   183
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   184
lemma ack_less_mono1_aux: "ack (i, k) < ack (Suc (i +i'), k)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   185
  apply (induct i k rule: ack.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   186
    apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   187
   prefer 2
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   188
   apply (blast intro: less_trans ack_less_mono2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   189
  apply (induct_tac i' n rule: ack.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   190
    apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   191
  apply (blast intro: Suc_leI [THEN le_less_trans] ack_less_mono2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   192
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   193
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   194
lemma ack_less_mono1: "i < j ==> ack (i, k) < ack (j, k)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   195
  apply (drule less_imp_Suc_add)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   196
  apply (blast intro!: ack_less_mono1_aux)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   197
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   198
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   199
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   200
text {* PROPERTY A 7', monotonicity for @{text "\<le>"} *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   201
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   202
lemma ack_le_mono1: "i \<le> j ==> ack (i, k) \<le> ack (j, k)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   203
  apply (simp add: order_le_less)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   204
  apply (blast intro: ack_less_mono1)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   205
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   206
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   207
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   208
text {* PROPERTY A 10 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   209
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   210
lemma ack_nest_bound: "ack(i1, ack (i2, j)) < ack (2 + (i1 + i2), j)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   211
  apply (simp add: numerals)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   212
  apply (rule ack2_le_ack1 [THEN [2] less_le_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   213
  apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   214
  apply (rule le_add1 [THEN ack_le_mono1, THEN le_less_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   215
  apply (rule ack_less_mono1 [THEN ack_less_mono2])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   216
  apply (simp add: le_imp_less_Suc le_add2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   217
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   218
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   219
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   220
text {* PROPERTY A 11 *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   221
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   222
lemma ack_add_bound: "ack (i1, j) + ack (i2, j) < ack (4 + (i1 + i2), j)"
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11464
diff changeset
   223
  apply (rule_tac j = "ack (Suc (Suc 0), ack (i1 + i2, j))" in less_trans)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   224
   prefer 2
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   225
   apply (rule ack_nest_bound [THEN less_le_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   226
   apply (simp add: Suc3_eq_add_3)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   227
  apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   228
  apply (cut_tac i = i1 and m1 = i2 and k = j in le_add1 [THEN ack_le_mono1])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   229
  apply (cut_tac i = "i2" and m1 = i1 and k = j in le_add2 [THEN ack_le_mono1])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   230
  apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   231
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   232
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   233
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   234
text {* PROPERTY A 12.  Article uses existential quantifier but the ALF proof
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   235
  used @{text "k + 4"}.  Quantified version must be nested @{text
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   236
  "\<exists>k'. \<forall>i j. ..."} *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   237
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   238
lemma ack_add_bound2: "i < ack (k, j) ==> i + j < ack (4 + k, j)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   239
  apply (rule_tac j = "ack (k, j) + ack (0, j)" in less_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   240
   prefer 2
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   241
   apply (rule ack_add_bound [THEN less_le_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   242
   apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   243
  apply (rule add_less_mono less_ack2 | assumption)+
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   244
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   245
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   246
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   247
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   248
text {* Inductive definition of the @{term PR} functions *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   249
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   250
text {* MAIN RESULT *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   251
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   252
lemma SC_case: "SC l < ack (1, list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   253
  apply (unfold SC_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   254
  apply (induct l)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   255
  apply (simp_all add: le_add1 le_imp_less_Suc)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   256
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   257
19676
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
   258
lemma CONSTANT_case: "CONSTANT k l < ack (k, list_add l)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   259
  apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   260
  done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   261
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   262
lemma PROJ_case [rule_format]: "\<forall>i. PROJ i l < ack (0, list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   263
  apply (simp add: PROJ_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   264
  apply (induct l)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   265
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   266
  apply (rule allI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   267
  apply (case_tac i)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   268
  apply (simp (no_asm_simp) add: le_add1 le_imp_less_Suc)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   269
  apply (simp (no_asm_simp))
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   270
  apply (blast intro: less_le_trans intro!: le_add2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   271
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   272
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   273
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   274
text {* @{term COMP} case *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   275
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   276
lemma COMP_map_aux: "fs \<in> lists (PRIMREC \<inter> {f. \<exists>kf. \<forall>l. f l < ack (kf, list_add l)})
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   277
  ==> \<exists>k. \<forall>l. list_add (map (\<lambda>f. f l) fs) < ack (k, list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   278
  apply (erule lists.induct)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   279
  apply (rule_tac x = 0 in exI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   280
   apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   281
  apply safe
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   282
  apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   283
  apply (rule exI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   284
  apply (blast intro: add_less_mono ack_add_bound less_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   285
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   286
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   287
lemma COMP_case:
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   288
  "\<forall>l. g l < ack (kg, list_add l) ==>
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   289
  fs \<in> lists(PRIMREC Int {f. \<exists>kf. \<forall>l. f l < ack(kf, list_add l)})
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   290
  ==> \<exists>k. \<forall>l. COMP g fs  l < ack(k, list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   291
  apply (unfold COMP_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   292
  apply (frule Int_lower1 [THEN lists_mono, THEN subsetD])
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   293
    --{*Now, if meson tolerated map, we could finish with
16731
124b4782944f fixed antiquotation;
wenzelm
parents: 16588
diff changeset
   294
  @{text "(drule COMP_map_aux, meson ack_less_mono2 ack_nest_bound less_trans)"} *}
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   295
  apply (erule COMP_map_aux [THEN exE])
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   296
  apply (rule exI)
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   297
  apply (rule allI)
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   298
  apply (drule spec)+
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   299
  apply (erule less_trans)
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   300
  apply (blast intro: ack_less_mono2 ack_nest_bound less_trans)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   301
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   302
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   303
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   304
text {* @{term PREC} case *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   305
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   306
lemma PREC_case_aux:
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   307
  "\<forall>l. f l + list_add l < ack (kf, list_add l) ==>
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   308
    \<forall>l. g l + list_add l < ack (kg, list_add l) ==>
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   309
    PREC f g l + list_add l < ack (Suc (kf + kg), list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   310
  apply (unfold PREC_def)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   311
  apply (case_tac l)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   312
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   313
   apply (blast intro: less_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   314
  apply (erule ssubst) -- {* get rid of the needless assumption *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   315
  apply (induct_tac a)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   316
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   317
   txt {* base case *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   318
   apply (blast intro: le_add1 [THEN le_imp_less_Suc, THEN ack_less_mono1] less_trans)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   319
  txt {* induction step *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   320
  apply (rule Suc_leI [THEN le_less_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   321
   apply (rule le_refl [THEN add_le_mono, THEN le_less_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   322
    prefer 2
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   323
    apply (erule spec)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   324
   apply (simp add: le_add2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   325
  txt {* final part of the simplification *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   326
  apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   327
  apply (rule le_add2 [THEN ack_le_mono1, THEN le_less_trans])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   328
  apply (erule ack_less_mono2)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   329
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   330
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   331
lemma PREC_case:
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   332
  "\<forall>l. f l < ack (kf, list_add l) ==>
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   333
    \<forall>l. g l < ack (kg, list_add l) ==>
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   334
    \<exists>k. \<forall>l. PREC f g l < ack (k, list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   335
  apply (rule exI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   336
  apply (rule allI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   337
  apply (rule le_less_trans [OF le_add1 PREC_case_aux])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   338
   apply (blast intro: ack_add_bound2)+
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   339
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   340
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   341
lemma ack_bounds_PRIMREC: "f \<in> PRIMREC ==> \<exists>k. \<forall>l. f l < ack (k, list_add l)"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   342
  apply (erule PRIMREC.induct)
19676
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
   343
      apply (blast intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)+
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   344
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   345
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   346
lemma ack_not_PRIMREC: "(\<lambda>l. case l of [] => 0 | x # l' => ack (x, x)) \<notin> PRIMREC"
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   347
  apply (rule notI)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   348
  apply (erule ack_bounds_PRIMREC [THEN exE])
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   349
  apply (rule less_irrefl)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   350
  apply (drule_tac x = "[x]" in spec)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   351
  apply simp
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   352
  done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   353
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   354
end