src/ZF/ex/NatSum.thy
author wenzelm
Wed, 28 Nov 2001 00:39:33 +0100
changeset 12306 749a04f0cfb0
parent 9647 e9623f47275b
child 12867 5c900a821a7c
permissions -rw-r--r--
variant: preserve suffix of underscores (for skolem/internal names etc.);

(*  Title:      HOL/ex/Natsum.thy
    ID:         $Id$
    Author:     Tobias Nipkow & Lawrence C Paulson

A summation operator. sum(f,n+1) is the sum of all f(i), i=0...n.

Note: n is a natural number but the sum is an integer,
                            and f maps integers to integers
*)

NatSum = Main +

consts sum     :: [i=>i, i] => i
primrec 
  "sum (f,0) = #0"
  "sum (f, succ(n)) = f($#n) $+ sum(f,n)"

end