| author | wenzelm | 
| Sat, 18 Nov 2000 19:48:34 +0100 | |
| changeset 10495 | 284ee538991c | 
| parent 9647 | e9623f47275b | 
| child 12867 | 5c900a821a7c | 
| permissions | -rw-r--r-- | 
| 9647 | 1 | (* Title: HOL/ex/Natsum.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow & Lawrence C Paulson | |
| 4 | ||
| 5 | A summation operator. sum(f,n+1) is the sum of all f(i), i=0...n. | |
| 6 | ||
| 7 | Note: n is a natural number but the sum is an integer, | |
| 8 | and f maps integers to integers | |
| 9 | *) | |
| 10 | ||
| 11 | NatSum = Main + | |
| 12 | ||
| 13 | consts sum :: [i=>i, i] => i | |
| 14 | primrec | |
| 15 | "sum (f,0) = #0" | |
| 16 | "sum (f, succ(n)) = f($#n) $+ sum(f,n)" | |
| 17 | ||
| 18 | end |