src/ZF/ex/Primrec.ML
author lcp
Tue, 25 Apr 1995 11:14:03 +0200
changeset 1072 0140ff702b23
parent 782 200a16083201
child 1461 6bcb44e4d6e5
permissions -rw-r--r--
updated version
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/ex/Primrec
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     6
Primitive Recursive Functions
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
Proof adopted from
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
Nora Szasz, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
A Machine Checked Proof that Ackermann's Function is not Primitive Recursive,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
In: Huet & Plotkin, eds., Logical Environments (CUP, 1993), 317-338.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
See also E. Mendelson, Introduction to Mathematical Logic.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
(Van Nostrand, 1964), page 250, exercise 11.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
open Primrec;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    19
val pr_typechecks = 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    20
    nat_typechecks @ list.intrs @ 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
    [lam_type, list_case_type, drop_type, map_type, apply_type, rec_type];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    22
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    23
(** Useful special cases of evaluation ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    25
val pr_ss = arith_ss 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
    addsimps list.case_eqns
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
    addsimps [list_rec_Nil, list_rec_Cons, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
	      drop_0, drop_Nil, drop_succ_Cons,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
	      map_Nil, map_Cons]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
    setsolver (type_auto_tac pr_typechecks);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
goalw Primrec.thy [SC_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
    "!!x l. [| x:nat;  l: list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    34
by (asm_simp_tac pr_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    35
qed "SC";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    36
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    37
goalw Primrec.thy [CONST_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
    "!!l. [| l: list(nat) |] ==> CONST(k) ` l = k";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
by (asm_simp_tac pr_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    40
qed "CONST";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    42
goalw Primrec.thy [PROJ_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    43
    "!!l. [| x: nat;  l: list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    44
by (asm_simp_tac pr_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    45
qed "PROJ_0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    46
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    47
goalw Primrec.thy [COMP_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    48
    "!!l. [| l: list(nat) |] ==> COMP(g,[f]) ` l = g` [f`l]";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    49
by (asm_simp_tac pr_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    50
qed "COMP_1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    51
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    52
goalw Primrec.thy [PREC_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    53
    "!!l. l: list(nat) ==> PREC(f,g) ` (Cons(0,l)) = f`l";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    54
by (asm_simp_tac pr_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    55
qed "PREC_0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    56
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    57
goalw Primrec.thy [PREC_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    58
    "!!l. [| x:nat;  l: list(nat) |] ==>  \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    59
\         PREC(f,g) ` (Cons(succ(x),l)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    60
\         g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    61
by (asm_simp_tac pr_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    62
qed "PREC_succ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    63
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    64
(*** Inductive definition of the PR functions ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    65
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    66
(* c: primrec ==> c: list(nat) -> nat *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    67
val primrec_into_fun = primrec.dom_subset RS subsetD;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    68
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    69
val pr_ss = pr_ss 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    70
    setsolver (type_auto_tac ([primrec_into_fun] @ 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    71
			      pr_typechecks @ primrec.intrs));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    72
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    73
goalw Primrec.thy [ACK_def] "!!i. i:nat ==> ACK(i): primrec";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    74
by (etac nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    75
by (ALLGOALS (asm_simp_tac pr_ss));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    76
qed "ACK_in_primrec";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    77
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    78
val ack_typechecks =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    79
    [ACK_in_primrec, primrec_into_fun RS apply_type,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    80
     add_type, list_add_type, nat_into_Ord] @ 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    81
    nat_typechecks @ list.intrs @ primrec.intrs;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    82
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    83
(*strict typechecking for the Ackermann proof; instantiates no vars*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    84
fun tc_tac rls =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    85
    REPEAT
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    86
      (SOMEGOAL (test_assume_tac ORELSE' match_tac (rls @ ack_typechecks)));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    87
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    88
goal Primrec.thy "!!i j. [| i:nat;  j:nat |] ==>  ack(i,j): nat";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    89
by (tc_tac []);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    90
qed "ack_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    91
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    92
(** Ackermann's function cases **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    93
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    94
(*PROPERTY A 1*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    95
goalw Primrec.thy [ACK_def] "!!j. j:nat ==> ack(0,j) = succ(j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    96
by (asm_simp_tac (pr_ss addsimps [SC]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    97
qed "ack_0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    98
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    99
(*PROPERTY A 2*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   100
goalw Primrec.thy [ACK_def] "ack(succ(i), 0) = ack(i,1)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   101
by (asm_simp_tac (pr_ss addsimps [CONST,PREC_0]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   102
qed "ack_succ_0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   103
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   104
(*PROPERTY A 3*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   105
(*Could be proved in Primrec0, like the previous two cases, but using
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   106
  primrec_into_fun makes type-checking easier!*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   107
goalw Primrec.thy [ACK_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   108
    "!!i j. [| i:nat;  j:nat |] ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   109
\           ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   110
by (asm_simp_tac (pr_ss addsimps [CONST,PREC_succ,COMP_1,PROJ_0]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   111
qed "ack_succ_succ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   112
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   113
val ack_ss = 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   114
    pr_ss addsimps [ack_0, ack_succ_0, ack_succ_succ, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   115
		    ack_type, nat_into_Ord];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   116
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   117
(*PROPERTY A 4*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   118
goal Primrec.thy "!!i. i:nat ==> ALL j:nat. j < ack(i,j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   119
by (etac nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   120
by (asm_simp_tac ack_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   121
by (rtac ballI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   122
by (eres_inst_tac [("n","j")] nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   123
by (DO_GOAL [rtac (nat_0I RS nat_0_le RS lt_trans),
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   124
	     asm_simp_tac ack_ss] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   125
by (DO_GOAL [etac (succ_leI RS lt_trans1),
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   126
	     asm_simp_tac ack_ss] 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   127
qed "lt_ack2_lemma";
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   128
bind_thm ("lt_ack2", (lt_ack2_lemma RS bspec));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   129
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   130
(*PROPERTY A 5-, the single-step lemma*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   131
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   132
by (etac nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   133
by (ALLGOALS (asm_simp_tac (ack_ss addsimps [lt_ack2])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   134
qed "ack_lt_ack_succ2";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   135
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   136
(*PROPERTY A 5, monotonicity for < *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   137
goal Primrec.thy "!!i j k. [| j<k; i:nat; k:nat |] ==> ack(i,j) < ack(i,k)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   138
by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   139
by (etac succ_lt_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   140
by (assume_tac 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   141
by (rtac lt_trans 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   142
by (REPEAT (ares_tac ([ack_lt_ack_succ2, ack_type] @ pr_typechecks) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   143
qed "ack_lt_mono2";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   144
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   145
(*PROPERTY A 5', monotonicity for le *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   146
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   147
    "!!i j k. [| j le k;  i: nat;  k:nat |] ==> ack(i,j) le ack(i,k)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   148
by (res_inst_tac [("f", "%j.ack(i,j)")] Ord_lt_mono_imp_le_mono 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   149
by (REPEAT (ares_tac [ack_lt_mono2, ack_type RS nat_into_Ord] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   150
qed "ack_le_mono2";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   151
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   152
(*PROPERTY A 6*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   153
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   154
    "!!i j. [| i:nat;  j:nat |] ==> ack(i, succ(j)) le ack(succ(i), j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   155
by (nat_ind_tac "j" [] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   156
by (ALLGOALS (asm_simp_tac ack_ss));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   157
by (rtac ack_le_mono2 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   158
by (rtac (lt_ack2 RS succ_leI RS le_trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   159
by (REPEAT (ares_tac (ack_typechecks) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   160
qed "ack2_le_ack1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   161
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   162
(*PROPERTY A 7-, the single-step lemma*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   163
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   164
by (rtac (ack_lt_mono2 RS lt_trans2) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   165
by (rtac ack2_le_ack1 4);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   166
by (REPEAT (ares_tac ([nat_le_refl, ack_type] @ pr_typechecks) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   167
qed "ack_lt_ack_succ1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   168
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   169
(*PROPERTY A 7, monotonicity for < *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   170
goal Primrec.thy "!!i j k. [| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   171
by (forward_tac [lt_nat_in_nat] 1 THEN assume_tac 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   172
by (etac succ_lt_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   173
by (assume_tac 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   174
by (rtac lt_trans 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   175
by (REPEAT (ares_tac ([ack_lt_ack_succ1, ack_type] @ pr_typechecks) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   176
qed "ack_lt_mono1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   177
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   178
(*PROPERTY A 7', monotonicity for le *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   179
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   180
    "!!i j k. [| i le j; j:nat; k:nat |] ==> ack(i,k) le ack(j,k)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   181
by (res_inst_tac [("f", "%j.ack(j,k)")] Ord_lt_mono_imp_le_mono 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   182
by (REPEAT (ares_tac [ack_lt_mono1, ack_type RS nat_into_Ord] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   183
qed "ack_le_mono1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   184
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   185
(*PROPERTY A 8*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   186
goal Primrec.thy "!!j. j:nat ==> ack(1,j) = succ(succ(j))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   187
by (etac nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   188
by (ALLGOALS (asm_simp_tac ack_ss));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   189
qed "ack_1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   190
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   191
(*PROPERTY A 9*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   192
goal Primrec.thy "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   193
by (etac nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   194
by (ALLGOALS (asm_simp_tac (ack_ss addsimps [ack_1, add_succ_right])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   195
qed "ack_2";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   196
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   197
(*PROPERTY A 10*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   198
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   199
    "!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   200
\               ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   201
by (rtac (ack2_le_ack1 RSN (2,lt_trans2)) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   202
by (asm_simp_tac ack_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   203
by (rtac (add_le_self RS ack_le_mono1 RS lt_trans1) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   204
by (rtac (add_le_self2 RS ack_lt_mono1 RS ack_lt_mono2) 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   205
by (tc_tac []);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   206
qed "ack_nest_bound";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   207
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   208
(*PROPERTY A 11*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   209
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   210
    "!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   211
\          ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   212
by (res_inst_tac [("j", "ack(succ(1), ack(i1 #+ i2, j))")] lt_trans 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   213
by (asm_simp_tac (ack_ss addsimps [ack_2]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   214
by (rtac (ack_nest_bound RS lt_trans2) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   215
by (asm_simp_tac ack_ss 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   216
by (rtac (add_le_mono RS leI RS leI) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   217
by (REPEAT (ares_tac ([add_le_self, add_le_self2, ack_le_mono1] @
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   218
                      ack_typechecks) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   219
qed "ack_add_bound";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   220
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   221
(*PROPERTY A 12.  Article uses existential quantifier but the ALF proof
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   222
  used k#+4.  Quantified version must be nested EX k'. ALL i,j... *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   223
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   224
    "!!i j k. [| i < ack(k,j);  j:nat;  k:nat |] ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   225
\             i#+j < ack(succ(succ(succ(succ(k)))), j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   226
by (res_inst_tac [("j", "ack(k,j) #+ ack(0,j)")] lt_trans 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   227
by (rtac (ack_add_bound RS lt_trans2) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   228
by (asm_simp_tac (ack_ss addsimps [add_0_right]) 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   229
by (REPEAT (ares_tac ([add_lt_mono, lt_ack2] @ ack_typechecks) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   230
qed "ack_add_bound2";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   231
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   232
(*** MAIN RESULT ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   233
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   234
val ack2_ss =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   235
    ack_ss addsimps [list_add_Nil, list_add_Cons, list_add_type, nat_into_Ord];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   236
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   237
goalw Primrec.thy [SC_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   238
    "!!l. l: list(nat) ==> SC ` l < ack(1, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   239
by (etac list.elim 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   240
by (asm_simp_tac (ack2_ss addsimps [succ_iff]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   241
by (asm_simp_tac (ack2_ss addsimps [ack_1, add_le_self]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   242
qed "SC_case";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   243
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   244
(*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   245
goal Primrec.thy "!!j. [| i:nat; j:nat |] ==> i < ack(i,j)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   246
by (etac nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   247
by (asm_simp_tac (ack_ss addsimps [nat_0_le]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   248
by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   249
by (tc_tac []);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   250
qed "lt_ack1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   251
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   252
goalw Primrec.thy [CONST_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   253
    "!!l. [| l: list(nat);  k: nat |] ==> CONST(k) ` l < ack(k, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   254
by (asm_simp_tac (ack2_ss addsimps [lt_ack1]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   255
qed "CONST_case";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   256
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   257
goalw Primrec.thy [PROJ_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   258
    "!!l. l: list(nat) ==> ALL i:nat. PROJ(i) ` l < ack(0, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   259
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   260
by (etac list.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   261
by (asm_simp_tac (ack2_ss addsimps [nat_0_le]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   262
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   263
by (rtac ballI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   264
by (eres_inst_tac [("n","x")] natE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   265
by (asm_simp_tac (ack2_ss addsimps [add_le_self]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   266
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   267
by (etac (bspec RS lt_trans2) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   268
by (rtac (add_le_self2 RS succ_leI) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   269
by (tc_tac []);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   270
qed "PROJ_case_lemma";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   271
val PROJ_case = PROJ_case_lemma RS bspec;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   272
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   273
(** COMP case **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   274
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   275
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   276
 "!!fs. fs : list({f: primrec .					\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   277
\              	   EX kf:nat. ALL l:list(nat). 			\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   278
\		    	      f`l < ack(kf, list_add(l))})	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   279
\      ==> EX k:nat. ALL l: list(nat). 				\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   280
\                list_add(map(%f. f ` l, fs)) < ack(k, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   281
by (etac list.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   282
by (DO_GOAL [res_inst_tac [("x","0")] bexI,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   283
	     asm_simp_tac (ack2_ss addsimps [lt_ack1, nat_0_le]),
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   284
	     resolve_tac nat_typechecks] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   285
by (safe_tac ZF_cs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   286
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   287
by (rtac (ballI RS bexI) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   288
by (rtac (add_lt_mono RS lt_trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   289
by (REPEAT (FIRSTGOAL (etac bspec)));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   290
by (rtac ack_add_bound 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   291
by (tc_tac []);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   292
qed "COMP_map_lemma";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   293
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   294
goalw Primrec.thy [COMP_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   295
 "!!g. [| g: primrec;  kg: nat;					\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   296
\         ALL l:list(nat). g`l < ack(kg, list_add(l));		\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   297
\         fs : list({f: primrec .				\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   298
\                    EX kf:nat. ALL l:list(nat). 		\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   299
\		    	f`l < ack(kf, list_add(l))}) 		\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   300
\      |] ==> EX k:nat. ALL l: list(nat). COMP(g,fs)`l < ack(k, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   301
by (asm_simp_tac ZF_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   302
by (forward_tac [list_CollectD] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   303
by (etac (COMP_map_lemma RS bexE) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   304
by (rtac (ballI RS bexI) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   305
by (etac (bspec RS lt_trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   306
by (rtac lt_trans 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   307
by (rtac ack_nest_bound 3);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   308
by (etac (bspec RS ack_lt_mono2) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   309
by (tc_tac [map_type]);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   310
qed "COMP_case";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   311
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   312
(** PREC case **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   313
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   314
goalw Primrec.thy [PREC_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   315
 "!!f g. [| ALL l:list(nat). f`l #+ list_add(l) < ack(kf, list_add(l));	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   316
\           ALL l:list(nat). g`l #+ list_add(l) < ack(kg, list_add(l));	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   317
\           f: primrec;  kf: nat;					\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   318
\           g: primrec;  kg: nat;					\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   319
\           l: list(nat)						\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   320
\        |] ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   321
by (etac list.elim 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   322
by (asm_simp_tac (ack2_ss addsimps [[nat_le_refl, lt_ack2] MRS lt_trans]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   323
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   324
by (etac ssubst 1);  (*get rid of the needless assumption*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   325
by (eres_inst_tac [("n","a")] nat_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   326
(*base case*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   327
by (DO_GOAL [asm_simp_tac ack2_ss, rtac lt_trans, etac bspec,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   328
	     assume_tac, rtac (add_le_self RS ack_lt_mono1),
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   329
	     REPEAT o ares_tac (ack_typechecks)] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   330
(*ind step*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   331
by (asm_simp_tac (ack2_ss addsimps [add_succ_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   332
by (rtac (succ_leI RS lt_trans1) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   333
by (res_inst_tac [("j", "g ` ?ll #+ ?mm")] lt_trans1 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   334
by (etac bspec 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   335
by (rtac (nat_le_refl RS add_le_mono) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   336
by (tc_tac []);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   337
by (asm_simp_tac (ack2_ss addsimps [add_le_self2]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   338
(*final part of the simplification*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   339
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   340
by (rtac (add_le_self2 RS ack_le_mono1 RS lt_trans1) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   341
by (etac ack_lt_mono2 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   342
by (tc_tac []);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   343
qed "PREC_case_lemma";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   344
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   345
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   346
 "!!f g. [| f: primrec;  kf: nat;				\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   347
\           g: primrec;  kg: nat;				\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   348
\           ALL l:list(nat). f`l < ack(kf, list_add(l));	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   349
\           ALL l:list(nat). g`l < ack(kg, list_add(l)) 	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   350
\        |] ==> EX k:nat. ALL l: list(nat). 			\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   351
\		    PREC(f,g)`l< ack(k, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   352
by (rtac (ballI RS bexI) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   353
by (rtac ([add_le_self, PREC_case_lemma] MRS lt_trans1) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   354
by (REPEAT
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   355
    (SOMEGOAL
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   356
     (FIRST' [test_assume_tac,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   357
	      match_tac (ack_typechecks),
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   358
	      rtac (ack_add_bound2 RS ballI) THEN' etac bspec])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   359
qed "PREC_case";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   360
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   361
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   362
    "!!f. f:primrec ==> EX k:nat. ALL l:list(nat). f`l < ack(k, list_add(l))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   363
by (etac primrec.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   364
by (safe_tac ZF_cs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   365
by (DEPTH_SOLVE
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   366
    (ares_tac ([SC_case, CONST_case, PROJ_case, COMP_case, PREC_case,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   367
		       bexI, ballI] @ nat_typechecks) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   368
qed "ack_bounds_primrec";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   369
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   370
goal Primrec.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   371
    "~ (lam l:list(nat). list_case(0, %x xs. ack(x,x), l)) : primrec";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   372
by (rtac notI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   373
by (etac (ack_bounds_primrec RS bexE) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   374
by (rtac lt_irrefl 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   375
by (dres_inst_tac [("x", "[x]")] bspec 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   376
by (asm_simp_tac ack2_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   377
by (asm_full_simp_tac (ack2_ss addsimps [add_0_right]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   378
qed "ack_not_primrec";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   379