| 3115 |      1 | (*  Title:      FOL/ex/Nat.thy
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1473 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 0 |      4 |     Copyright   1992  University of Cambridge
 | 
|  |      5 | 
 | 
| 3115 |      6 | Examples for the manuals.
 | 
| 0 |      7 | 
 | 
|  |      8 | Theory of the natural numbers: Peano's axioms, primitive recursion
 | 
|  |      9 | *)
 | 
|  |     10 | 
 | 
|  |     11 | Nat = FOL +
 | 
| 352 |     12 | types   nat
 | 
|  |     13 | arities nat :: term
 | 
| 1322 |     14 | consts  "0" :: nat      ("0")
 | 
|  |     15 |         Suc :: nat=>nat  
 | 
|  |     16 |         rec :: [nat, 'a, [nat,'a]=>'a] => 'a  
 | 
|  |     17 |         "+" :: [nat, nat] => nat                (infixl 60)
 | 
| 0 |     18 | rules   induct      "[| P(0);  !!x. P(x) ==> P(Suc(x)) |]  ==> P(n)"
 | 
|  |     19 |         Suc_inject  "Suc(m)=Suc(n) ==> m=n"
 | 
|  |     20 |         Suc_neq_0   "Suc(m)=0      ==> R"
 | 
|  |     21 |         rec_0       "rec(0,a,f) = a"
 | 
|  |     22 |         rec_Suc     "rec(Suc(m), a, f) = f(m, rec(m,a,f))"
 | 
|  |     23 |         add_def     "m+n == rec(m, n, %x y. Suc(y))"
 | 
|  |     24 | end
 |