src/CCL/ex/Nat.thy
author wenzelm
Sat, 17 Sep 2005 17:35:26 +0200
changeset 17456 bcf7544875b2
parent 3837 d7f033c74b38
child 20140 98acc6d0fab6
permissions -rw-r--r--
converted to Isar theory format;

(*  Title:      CCL/ex/Nat.thy
    ID:         $Id$
    Author:     Martin Coen, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge
*)

header {* Programs defined over the natural numbers *}

theory Nat
imports Wfd
begin

consts

  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"

defs

  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)))
                        in sub(a,b)"
  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)"

  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.
                          ncase(m,ack(x,succ(zero)),%y. ack(x,ack(succ(x),y))))
                    in ack(a,b)"

ML {* use_legacy_bindings (the_context ()) *}

end