| author | paulson | 
| Wed, 23 Jul 1997 11:52:22 +0200 | |
| changeset 3564 | f886dbd91ee5 | 
| parent 3366 | 2402c6ab1561 | 
| child 4360 | 40e5c97e988d | 
| permissions | -rw-r--r-- | 
| 923 | 1 | (* Title: HOL/Arith.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1993 University of Cambridge | |
| 5 | ||
| 3366 | 6 | Arithmetic operators + - and * (for div, mod and dvd, see Divides) | 
| 923 | 7 | *) | 
| 8 | ||
| 9 | Arith = Nat + | |
| 10 | ||
| 11 | instance | |
| 3366 | 12 |   nat :: {plus, minus, times, power}
 | 
| 923 | 13 | |
| 14 | consts | |
| 1370 
7361ac9b024d
removed quotes from types in consts and syntax sections
 clasohm parents: 
972diff
changeset | 15 | pred :: nat => nat | 
| 3308 
da002cef7090
Added overloaded function `size' for all datatypes.
 nipkow parents: 
3235diff
changeset | 16 | (* size of a datatype value; overloaded *) | 
| 
da002cef7090
Added overloaded function `size' for all datatypes.
 nipkow parents: 
3235diff
changeset | 17 | size :: 'a => nat | 
| 923 | 18 | |
| 19 | defs | |
| 2099 | 20 | pred_def "pred(m) == case m of 0 => 0 | Suc n => n" | 
| 1796 | 21 | |
| 2681 
93ed51a91622
definitions of +,-,* replaced by primrec definitions
 pusch parents: 
2099diff
changeset | 22 | primrec "op +" nat | 
| 3366 | 23 | add_0 "0 + n = n" | 
| 24 | add_Suc "Suc m + n = Suc(m + n)" | |
| 2681 
93ed51a91622
definitions of +,-,* replaced by primrec definitions
 pusch parents: 
2099diff
changeset | 25 | |
| 
93ed51a91622
definitions of +,-,* replaced by primrec definitions
 pusch parents: 
2099diff
changeset | 26 | primrec "op -" nat | 
| 3235 
351565b7321b
The diff laws must be named: we do "Delsimps [diff_Suc];"
 paulson parents: 
2887diff
changeset | 27 | diff_0 "m - 0 = m" | 
| 
351565b7321b
The diff laws must be named: we do "Delsimps [diff_Suc];"
 paulson parents: 
2887diff
changeset | 28 | diff_Suc "m - Suc n = pred(m - n)" | 
| 2681 
93ed51a91622
definitions of +,-,* replaced by primrec definitions
 pusch parents: 
2099diff
changeset | 29 | |
| 
93ed51a91622
definitions of +,-,* replaced by primrec definitions
 pusch parents: 
2099diff
changeset | 30 | primrec "op *" nat | 
| 3366 | 31 | mult_0 "0 * n = 0" | 
| 32 | mult_Suc "Suc m * n = n + (m * n)" | |
| 2681 
93ed51a91622
definitions of +,-,* replaced by primrec definitions
 pusch parents: 
2099diff
changeset | 33 | |
| 923 | 34 | end |