| 1477 |      1 | (*  Title:      FOLP/ex/nat.thy
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1477 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 0 |      4 |     Copyright   1992  University of Cambridge
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 17480 |      7 | header {* Theory of the natural numbers: Peano's axioms, primitive recursion *}
 | 
|  |      8 | 
 | 
|  |      9 | theory Nat
 | 
|  |     10 | imports FOLP
 | 
|  |     11 | begin
 | 
|  |     12 | 
 | 
|  |     13 | typedecl nat
 | 
|  |     14 | arities nat         :: "term"
 | 
| 0 |     15 | consts  "0"         :: "nat"    ("0")
 | 
|  |     16 |         Suc         :: "nat=>nat"
 | 
|  |     17 |         rec         :: "[nat, 'a, [nat,'a]=>'a] => 'a"
 | 
|  |     18 |         "+"         :: "[nat, nat] => nat"              (infixl 60)
 | 
|  |     19 | 
 | 
|  |     20 |   (*Proof terms*)
 | 
|  |     21 |        nrec         :: "[nat,p,[nat,p]=>p]=>p"
 | 
| 17480 |     22 |        ninj         :: "p=>p"
 | 
|  |     23 |        nneq         :: "p=>p"
 | 
|  |     24 |        rec0         :: "p"
 | 
|  |     25 |        recSuc       :: "p"
 | 
|  |     26 | 
 | 
|  |     27 | axioms
 | 
|  |     28 |   induct:     "[| b:P(0); !!x u. u:P(x) ==> c(x,u):P(Suc(x))
 | 
|  |     29 |               |] ==> nrec(n,b,c):P(n)"
 | 
| 0 |     30 | 
 | 
| 17480 |     31 |   Suc_inject: "p:Suc(m)=Suc(n) ==> ninj(p) : m=n"
 | 
|  |     32 |   Suc_neq_0:  "p:Suc(m)=0      ==> nneq(p) : R"
 | 
|  |     33 |   rec_0:      "rec0 : rec(0,a,f) = a"
 | 
|  |     34 |   rec_Suc:    "recSuc : rec(Suc(m), a, f) = f(m, rec(m,a,f))"
 | 
|  |     35 |   add_def:    "m+n == rec(m, n, %x y. Suc(y))"
 | 
| 0 |     36 | 
 | 
| 17480 |     37 |   nrecB0:     "b: A ==> nrec(0,b,c) = b : A"
 | 
|  |     38 |   nrecBSuc:   "c(n,nrec(n,b,c)) : A ==> nrec(Suc(n),b,c) = c(n,nrec(n,b,c)) : A"
 | 
|  |     39 | 
 | 
|  |     40 | ML {* use_legacy_bindings (the_context ()) *}
 | 
|  |     41 | 
 | 
| 0 |     42 | end
 |