src/HOL/ex/Primrec.thy
author haftmann
Fri, 11 Jun 2010 17:14:02 +0200
changeset 37407 61dd8c145da7
parent 34055 fdf294ee08b2
child 41460 ea56b98aee83
permissions -rw-r--r--
declare lex_prod_def [code del]
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
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
     6
Ackermann's Function and the
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
     7
Primitive Recursive Functions.
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
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    27
subsection{* Ackermann's Function *}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    28
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    29
fun ack :: "nat => nat => nat" where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    30
"ack 0 n =  Suc n" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    31
"ack (Suc m) 0 = ack m 1" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    32
"ack (Suc m) (Suc n) = ack m (ack (Suc m) n)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    33
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    34
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    35
text {* PROPERTY A 4 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    36
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    37
lemma less_ack2 [iff]: "j < ack i j"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    38
by (induct i j rule: ack.induct) simp_all
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    39
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    40
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    41
text {* PROPERTY A 5-, the single-step lemma *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    42
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    43
lemma ack_less_ack_Suc2 [iff]: "ack i j < ack i (Suc j)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    44
by (induct i j rule: ack.induct) simp_all
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    45
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    46
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    47
text {* PROPERTY A 5, monotonicity for @{text "<"} *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    48
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    49
lemma ack_less_mono2: "j < k ==> ack i j < ack i k"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    50
using lift_Suc_mono_less[where f = "ack i"]
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    51
by (metis ack_less_ack_Suc2)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    52
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    53
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    54
text {* PROPERTY A 5', monotonicity for @{text \<le>} *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    55
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    56
lemma ack_le_mono2: "j \<le> k ==> ack i j \<le> ack i k"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    57
apply (simp add: order_le_less)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    58
apply (blast intro: ack_less_mono2)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    59
done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
    60
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    61
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    62
text {* PROPERTY A 6 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    63
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    64
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
    65
proof (induct j)
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    66
  case 0 show ?case by simp
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    67
next
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    68
  case (Suc j) show ?case 
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    69
    by (auto intro!: ack_le_mono2)
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    70
      (metis Suc Suc_leI Suc_lessI less_ack2 linorder_not_less)
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25157
diff changeset
    71
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    72
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    73
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    74
text {* PROPERTY A 7-, the single-step lemma *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    75
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    76
lemma ack_less_ack_Suc1 [iff]: "ack i j < ack (Suc i) j"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    77
by (blast intro: ack_less_mono2 less_le_trans)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    78
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    79
19676
187234ec6050 renamed CONST to CONSTANT;
wenzelm
parents: 16731
diff changeset
    80
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
    81
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    82
lemma less_ack1 [iff]: "i < ack i j"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    83
apply (induct i)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    84
 apply simp_all
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    85
apply (blast intro: Suc_leI le_less_trans)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    86
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    87
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    88
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    89
text {* PROPERTY A 8 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    90
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    91
lemma ack_1 [simp]: "ack (Suc 0) j = j + 2"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    92
by (induct j) simp_all
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    93
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    94
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11464
diff changeset
    95
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
    96
  ack} is essential for the rewriting. *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
    97
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    98
lemma ack_2 [simp]: "ack (Suc (Suc 0)) j = 2 * j + 3"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
    99
by (induct j) simp_all
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   100
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   101
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   102
text {* PROPERTY A 7, monotonicity for @{text "<"} [not clear why
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   103
  @{thm [source] ack_1} is now needed first!] *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   104
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   105
lemma ack_less_mono1_aux: "ack i k < ack (Suc (i +i')) k"
34055
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   106
proof (induct i k rule: ack.induct)
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   107
  case (1 n) show ?case
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   108
    by (simp, metis ack_less_ack_Suc1 less_ack2 less_trans_Suc) 
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   109
next
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   110
  case (2 m) thus ?case by simp
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   111
next
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   112
  case (3 m n) thus ?case
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   113
    by (simp, blast intro: less_trans ack_less_mono2)
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   114
qed
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   115
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   116
lemma ack_less_mono1: "i < j ==> ack i k < ack j k"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   117
apply (drule less_imp_Suc_add)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   118
apply (blast intro!: ack_less_mono1_aux)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   119
done
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
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   122
text {* PROPERTY A 7', monotonicity for @{text "\<le>"} *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   123
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   124
lemma ack_le_mono1: "i \<le> j ==> ack i k \<le> ack j k"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   125
apply (simp add: order_le_less)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   126
apply (blast intro: ack_less_mono1)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   127
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   128
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   129
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   130
text {* PROPERTY A 10 *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   131
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   132
lemma ack_nest_bound: "ack i1 (ack i2 j) < ack (2 + (i1 + i2)) j"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   133
apply (simp add: numerals)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   134
apply (rule ack2_le_ack1 [THEN [2] less_le_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   135
apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   136
apply (rule le_add1 [THEN ack_le_mono1, THEN le_less_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   137
apply (rule ack_less_mono1 [THEN ack_less_mono2])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   138
apply (simp add: le_imp_less_Suc le_add2)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   139
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   140
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   141
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   142
text {* PROPERTY A 11 *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   143
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   144
lemma ack_add_bound: "ack i1 j + ack i2 j < ack (4 + (i1 + i2)) j"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   145
apply (rule less_trans [of _ "ack (Suc (Suc 0)) (ack (i1 + i2) j)"])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   146
 prefer 2
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   147
 apply (rule ack_nest_bound [THEN less_le_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   148
 apply (simp add: Suc3_eq_add_3)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   149
apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   150
apply (cut_tac i = i1 and m1 = i2 and k = j in le_add1 [THEN ack_le_mono1])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   151
apply (cut_tac i = "i2" and m1 = i1 and k = j in le_add2 [THEN ack_le_mono1])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   152
apply auto
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   153
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   154
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   155
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   156
text {* PROPERTY A 12.  Article uses existential quantifier but the ALF proof
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   157
  used @{text "k + 4"}.  Quantified version must be nested @{text
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   158
  "\<exists>k'. \<forall>i j. ..."} *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   159
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   160
lemma ack_add_bound2: "i < ack k j ==> i + j < ack (4 + k) j"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   161
apply (rule less_trans [of _ "ack k j + ack 0 j"])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   162
 apply (blast intro: add_less_mono less_ack2) 
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   163
apply (rule ack_add_bound [THEN less_le_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   164
apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   165
done
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   166
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   167
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   168
subsection{*Primitive Recursive Functions*}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   169
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   170
primrec hd0 :: "nat list => nat" where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   171
"hd0 [] = 0" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   172
"hd0 (m # ms) = m"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   173
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   174
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   175
text {* Inductive definition of the set of primitive recursive functions of type @{typ "nat list => nat"}. *}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   176
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   177
definition SC :: "nat list => nat" where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   178
"SC l = Suc (hd0 l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   179
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   180
definition CONSTANT :: "nat => nat list => nat" where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   181
"CONSTANT k l = k"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   182
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   183
definition PROJ :: "nat => nat list => nat" where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   184
"PROJ i l = hd0 (drop i l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   185
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   186
definition
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   187
COMP :: "(nat list => nat) => (nat list => nat) list => nat list => nat"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   188
where "COMP g fs l = g (map (\<lambda>f. f l) fs)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   189
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   190
definition PREC :: "(nat list => nat) => (nat list => nat) => nat list => nat"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   191
where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   192
  "PREC f g l =
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   193
    (case l of
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   194
      [] => 0
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   195
    | x # l' => nat_rec (f l') (\<lambda>y r. g (r # y # l')) x)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   196
  -- {* Note that @{term g} is applied first to @{term "PREC f g y"} and then to @{term y}! *}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   197
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   198
inductive PRIMREC :: "(nat list => nat) => bool" where
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   199
SC: "PRIMREC SC" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   200
CONSTANT: "PRIMREC (CONSTANT k)" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   201
PROJ: "PRIMREC (PROJ i)" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   202
COMP: "PRIMREC g ==> \<forall>f \<in> set fs. PRIMREC f ==> PRIMREC (COMP g fs)" |
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   203
PREC: "PRIMREC f ==> PRIMREC g ==> PRIMREC (PREC f g)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   204
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   205
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   206
text {* Useful special cases of evaluation *}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   207
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   208
lemma SC [simp]: "SC (x # l) = Suc x"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   209
by (simp add: SC_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   210
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   211
lemma CONSTANT [simp]: "CONSTANT k l = k"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   212
by (simp add: CONSTANT_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   213
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   214
lemma PROJ_0 [simp]: "PROJ 0 (x # l) = x"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   215
by (simp add: PROJ_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   216
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   217
lemma COMP_1 [simp]: "COMP g [f] l = g [f l]"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   218
by (simp add: COMP_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   219
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   220
lemma PREC_0 [simp]: "PREC f g (0 # l) = f l"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   221
by (simp add: PREC_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   222
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   223
lemma PREC_Suc [simp]: "PREC f g (Suc x # l) = g (PREC f g (x # l) # x # l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   224
by (simp add: PREC_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   225
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   226
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   227
text {* MAIN RESULT *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   228
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   229
lemma SC_case: "SC l < ack 1 (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   230
apply (unfold SC_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   231
apply (induct l)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   232
apply (simp_all add: le_add1 le_imp_less_Suc)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   233
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   234
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   235
lemma CONSTANT_case: "CONSTANT k l < ack k (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   236
by simp
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   237
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   238
lemma PROJ_case: "PROJ i l < ack 0 (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   239
apply (simp add: PROJ_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   240
apply (induct l arbitrary:i)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   241
 apply (auto simp add: drop_Cons split: nat.split)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   242
apply (blast intro: less_le_trans le_add2)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   243
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   244
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   245
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   246
text {* @{term COMP} case *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   247
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   248
lemma COMP_map_aux: "\<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (listsum l))
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   249
  ==> \<exists>k. \<forall>l. listsum (map (\<lambda>f. f l) fs) < ack k (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   250
apply (induct fs)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   251
 apply (rule_tac x = 0 in exI)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   252
 apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   253
apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   254
apply (blast intro: add_less_mono ack_add_bound less_trans)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   255
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   256
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   257
lemma COMP_case:
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   258
  "\<forall>l. g l < ack kg (listsum l) ==>
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   259
  \<forall>f \<in> set fs. PRIMREC f \<and> (\<exists>kf. \<forall>l. f l < ack kf (listsum l))
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   260
  ==> \<exists>k. \<forall>l. COMP g fs  l < ack k (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   261
apply (unfold COMP_def)
34055
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   262
apply (drule COMP_map_aux)
fdf294ee08b2 streamlined proofs
paulson
parents: 28480
diff changeset
   263
apply (meson ack_less_mono2 ack_nest_bound less_trans)
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   264
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   265
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   266
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   267
text {* @{term PREC} case *}
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   268
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   269
lemma PREC_case_aux:
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   270
  "\<forall>l. f l + listsum l < ack kf (listsum l) ==>
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   271
    \<forall>l. g l + listsum l < ack kg (listsum l) ==>
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   272
    PREC f g l + listsum l < ack (Suc (kf + kg)) (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   273
apply (unfold PREC_def)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   274
apply (case_tac l)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   275
 apply simp_all
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   276
 apply (blast intro: less_trans)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   277
apply (erule ssubst) -- {* get rid of the needless assumption *}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   278
apply (induct_tac a)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   279
 apply simp_all
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   280
 txt {* base case *}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   281
 apply (blast intro: le_add1 [THEN le_imp_less_Suc, THEN ack_less_mono1] less_trans)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   282
txt {* induction step *}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   283
apply (rule Suc_leI [THEN le_less_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   284
 apply (rule le_refl [THEN add_le_mono, THEN le_less_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   285
  prefer 2
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   286
  apply (erule spec)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   287
 apply (simp add: le_add2)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   288
txt {* final part of the simplification *}
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   289
apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   290
apply (rule le_add2 [THEN ack_le_mono1, THEN le_less_trans])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   291
apply (erule ack_less_mono2)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   292
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   293
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   294
lemma PREC_case:
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   295
  "\<forall>l. f l < ack kf (listsum l) ==>
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   296
    \<forall>l. g l < ack kg (listsum l) ==>
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   297
    \<exists>k. \<forall>l. PREC f g l < ack k (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
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
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   300
lemma ack_bounds_PRIMREC: "PRIMREC f ==> \<exists>k. \<forall>l. f l < ack k (listsum l)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   301
apply (erule PRIMREC.induct)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   302
    apply (blast intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)+
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   303
done
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8703
diff changeset
   304
27626
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   305
theorem ack_not_PRIMREC:
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   306
  "\<not> PRIMREC (\<lambda>l. case l of [] => 0 | x # l' => ack x x)"
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   307
apply (rule notI)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   308
apply (erule ack_bounds_PRIMREC [THEN exE])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   309
apply (rule less_irrefl [THEN notE])
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   310
apply (drule_tac x = "[x]" in spec)
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   311
apply simp
1a3507f86b39 beautified proofs
nipkow
parents: 26334
diff changeset
   312
done
3335
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   313
b0139b83a5ee New example ported from ZF
paulson
parents:
diff changeset
   314
end