| 1478 |      1 | (*  Title:      ZF/arith.thy
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 1478 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 0 |      4 |     Copyright   1992  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Arithmetic operators and their definitions
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 2469 |      9 | Arith = Epsilon + 
 | 
| 6070 |     10 | 
 | 
|  |     11 | setup setup_datatypes
 | 
|  |     12 | 
 | 
|  |     13 | (*The natural numbers as a datatype*)
 | 
|  |     14 | rep_datatype 
 | 
|  |     15 |   elim		natE
 | 
|  |     16 |   induct	nat_induct
 | 
|  |     17 |   case_eqns	nat_case_0, nat_case_succ
 | 
|  |     18 |   recursor_eqns recursor_0, recursor_succ
 | 
|  |     19 | 
 | 
|  |     20 | 
 | 
| 0 |     21 | consts
 | 
| 1478 |     22 |     "#*" :: [i,i]=>i                    (infixl 70)
 | 
|  |     23 |     div  :: [i,i]=>i                    (infixl 70) 
 | 
|  |     24 |     mod  :: [i,i]=>i                    (infixl 70)
 | 
|  |     25 |     "#+" :: [i,i]=>i                    (infixl 65)
 | 
|  |     26 |     "#-" :: [i,i]=>i                    (infixl 65)
 | 
| 0 |     27 | 
 | 
| 6070 |     28 | primrec
 | 
|  |     29 |   add_0     "0 #+ n = n"
 | 
|  |     30 |   add_succ  "succ(m) #+ n = succ(m #+ n)"
 | 
|  |     31 | 
 | 
|  |     32 | primrec
 | 
|  |     33 |   diff_0     "m #- 0 = m"
 | 
|  |     34 |   diff_SUCC  "m #- succ(n) = nat_case(0, %x. x, m #- n)"
 | 
| 0 |     35 | 
 | 
| 6070 |     36 | primrec
 | 
|  |     37 |   mult_0    "0 #* n = 0"
 | 
|  |     38 |   mult_succ "succ(m) #* n = n #+ (m #* n)"
 | 
|  |     39 |  
 | 
|  |     40 | defs
 | 
|  |     41 |     mod_def  "m mod n == transrec(m, %j f. if j<n then j else f`(j#-n))"
 | 
|  |     42 |     div_def  "m div n == transrec(m, %j f. if j<n then 0 else succ(f`(j#-n)))"
 | 
| 0 |     43 | end
 |