diff -r 151e76f0e3c7 -r bcf7544875b2 src/CCL/ex/Nat.thy --- a/src/CCL/ex/Nat.thy Sat Sep 17 14:02:31 2005 +0200 +++ b/src/CCL/ex/Nat.thy Sat Sep 17 17:35:26 2005 +0200 @@ -1,38 +1,46 @@ -(* Title: CCL/ex/nat.thy +(* Title: CCL/ex/Nat.thy ID: $Id$ Author: Martin Coen, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge - -Programs defined over the natural numbers *) -Nat = Wfd + +header {* Programs defined over the natural numbers *} + +theory Nat +imports Wfd +begin consts - not :: "i=>i" - "#+","#*","#-", - "##","#<","#<=" :: "[i,i]=>i" (infixr 60) - ackermann :: "[i,i]=>i" + not :: "i=>i" + "#+" :: "[i,i]=>i" (infixr 60) + "#*" :: "[i,i]=>i" (infixr 60) + "#-" :: "[i,i]=>i" (infixr 60) + "##" :: "[i,i]=>i" (infixr 60) + "#<" :: "[i,i]=>i" (infixr 60) + "#<=" :: "[i,i]=>i" (infixr 60) + ackermann :: "[i,i]=>i" -rules +defs - not_def "not(b) == if b then false else true" + not_def: "not(b) == if b then false else true" - add_def "a #+ b == nrec(a,b,%x g. succ(g))" - mult_def "a #* b == nrec(a,zero,%x g. b #+ g)" - sub_def "a #- b == letrec sub x y be ncase(y,x,%yy. ncase(x,zero,%xx. sub(xx,yy))) + add_def: "a #+ b == nrec(a,b,%x g. succ(g))" + mult_def: "a #* b == nrec(a,zero,%x g. b #+ g)" + sub_def: "a #- b == letrec sub x y be ncase(y,x,%yy. ncase(x,zero,%xx. sub(xx,yy))) in sub(a,b)" - le_def "a #<= b == letrec le x y be ncase(x,true,%xx. ncase(y,false,%yy. le(xx,yy))) + le_def: "a #<= b == letrec le x y be ncase(x,true,%xx. ncase(y,false,%yy. le(xx,yy))) in le(a,b)" - lt_def "a #< b == not(b #<= a)" + lt_def: "a #< b == not(b #<= a)" - div_def "a ## b == letrec div x y be if x #< y then zero else succ(div(x#-y,y)) + div_def: "a ## b == letrec div x y be if x #< y then zero else succ(div(x#-y,y)) in div(a,b)" - ack_def - "ackermann(a,b) == letrec ack n m be ncase(n,succ(m),%x. + ack_def: + "ackermann(a,b) == letrec ack n m be ncase(n,succ(m),%x. ncase(m,ack(x,succ(zero)),%y. ack(x,ack(succ(x),y)))) in ack(a,b)" +ML {* use_legacy_bindings (the_context ()) *} + end