src/ZF/Arith.thy
author lcp
Mon, 20 Sep 1993 17:02:11 +0200
changeset 8 c3d2c6dcf3f0
parent 6 8ce8c4d13d4d
child 25 3ac1c0c0016e
permissions -rw-r--r--
Installation of new simplfier. Previously appeared to set up the old simplifier to rewrite with the partial ordering [=, something not possible with the new simplifier. But such rewriting appears not to have actually been used, and there were few complications. In terms.ML setloop was used to avoid infinite rewriting with the letrec rule. Congruence rules were deleted, and an occasional SIMP_TAC had to become asm_simp_tac.

(*  Title: 	ZF/arith.thy
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1992  University of Cambridge

Arithmetic operators and their definitions
*)

Arith = Epsilon +
consts
    rec  :: "[i, i, [i,i]=>i]=>i"
    "#*" :: "[i,i]=>i"      		(infixl 70)
    div  :: "[i,i]=>i"      		(infixl 70) 
    mod  :: "[i,i]=>i"      		(infixl 70)
    "#+" :: "[i,i]=>i"      		(infixl 65)
    "#-" :: "[i,i]=>i"      		(infixl 65)

rules
    rec_def  "rec(k,a,b) ==  transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))"

    add_def  "m#+n == rec(m, n, %u v.succ(v))"
    diff_def "m#-n == rec(n, m, %u v. rec(v, 0, %x y.x))"
    mult_def "m#*n == rec(m, 0, %u v. n #+ v)"
    mod_def  "m mod n == transrec(m, %j f. if(j:n, j, f`(j#-n)))"
    div_def  "m div n == transrec(m, %j f. if(j:n, 0, succ(f`(j#-n))))"
end