src/ZF/Arith.thy
author wenzelm
Sat, 01 Jul 2000 19:55:22 +0200
changeset 9230 17ae63f82ad8
parent 6070 032babd0120b
child 9491 1a36151ee2fc
permissions -rw-r--r--
GPLed;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/arith.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Arithmetic operators and their definitions
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1478
diff changeset
     9
Arith = Epsilon + 
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    10
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    11
setup setup_datatypes
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    12
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    13
(*The natural numbers as a datatype*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    14
rep_datatype 
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    15
  elim		natE
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    16
  induct	nat_induct
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    17
  case_eqns	nat_case_0, nat_case_succ
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    18
  recursor_eqns recursor_0, recursor_succ
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    19
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    20
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
consts
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    22
    "#*" :: [i,i]=>i                    (infixl 70)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    23
    div  :: [i,i]=>i                    (infixl 70) 
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    24
    mod  :: [i,i]=>i                    (infixl 70)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    25
    "#+" :: [i,i]=>i                    (infixl 65)
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    26
    "#-" :: [i,i]=>i                    (infixl 65)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    28
primrec
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    29
  add_0     "0 #+ n = n"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    30
  add_succ  "succ(m) #+ n = succ(m #+ n)"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    31
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    32
primrec
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    33
  diff_0     "m #- 0 = m"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    34
  diff_SUCC  "m #- succ(n) = nat_case(0, %x. x, m #- n)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    36
primrec
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    37
  mult_0    "0 #* n = 0"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    38
  mult_succ "succ(m) #* n = n #+ (m #* n)"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    39
 
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    40
defs
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    41
    mod_def  "m mod n == transrec(m, %j f. if j<n then j else f`(j#-n))"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 3840
diff changeset
    42
    div_def  "m div n == transrec(m, %j f. if j<n then 0 else succ(f`(j#-n)))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
end