| author | wenzelm | 
| Sat, 01 Apr 2000 20:21:39 +0200 | |
| changeset 8657 | b9475dad85ed | 
| parent 7872 | 2e2d7e80fb07 | 
| child 8943 | a4f8be72f585 | 
| permissions | -rw-r--r-- | 
| 2608 | 1 | (* Title: HOL/NatDef.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | |
| 4 | Copyright 1991 University of Cambridge | |
| 5 | ||
| 6 | Definition of types ind and nat. | |
| 7 | ||
| 8 | Type nat is defined as a set Nat over type ind. | |
| 9 | *) | |
| 10 | ||
| 11 | NatDef = WF + | |
| 12 | ||
| 13 | (** type ind **) | |
| 14 | ||
| 3947 | 15 | global | 
| 16 | ||
| 2608 | 17 | types | 
| 18 | ind | |
| 19 | ||
| 20 | arities | |
| 21 | ind :: term | |
| 22 | ||
| 23 | consts | |
| 24 | Zero_Rep :: ind | |
| 25 | Suc_Rep :: ind => ind | |
| 26 | ||
| 27 | rules | |
| 28 | (*the axiom of infinity in 2 parts*) | |
| 29 | inj_Suc_Rep "inj(Suc_Rep)" | |
| 30 | Suc_Rep_not_Zero_Rep "Suc_Rep(x) ~= Zero_Rep" | |
| 31 | ||
| 32 | ||
| 33 | ||
| 34 | (** type nat **) | |
| 35 | ||
| 36 | (* type definition *) | |
| 37 | ||
| 38 | typedef (Nat) | |
| 39 |   nat = "lfp(%X. {Zero_Rep} Un (Suc_Rep``X))"   (lfp_def)
 | |
| 40 | ||
| 41 | instance | |
| 42 | nat :: ord | |
| 43 | ||
| 44 | ||
| 45 | (* abstract constants and syntax *) | |
| 46 | ||
| 47 | consts | |
| 48 |   "0"       :: nat                ("0")
 | |
| 49 | Suc :: nat => nat | |
| 50 | pred_nat :: "(nat * nat) set" | |
| 51 | ||
| 52 | syntax | |
| 53 |   "1"       :: nat                ("1")
 | |
| 54 |   "2"       :: nat                ("2")
 | |
| 55 | ||
| 56 | translations | |
| 5187 
55f07169cf5f
Removed nat_case, nat_rec, and natE (now provided by datatype
 berghofe parents: 
3947diff
changeset | 57 | "1" == "Suc 0" | 
| 
55f07169cf5f
Removed nat_case, nat_rec, and natE (now provided by datatype
 berghofe parents: 
3947diff
changeset | 58 | "2" == "Suc 1" | 
| 2608 | 59 | |
| 60 | ||
| 3947 | 61 | local | 
| 62 | ||
| 2608 | 63 | defs | 
| 64 | Zero_def "0 == Abs_Nat(Zero_Rep)" | |
| 65 | Suc_def "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))" | |
| 66 | ||
| 7872 | 67 | (*nat operations*) | 
| 3236 | 68 |   pred_nat_def  "pred_nat == {(m,n). n = Suc m}"
 | 
| 2608 | 69 | |
| 70 | less_def "m<n == (m,n):trancl(pred_nat)" | |
| 71 | ||
| 72 | le_def "m<=(n::nat) == ~(n<m)" | |
| 73 | ||
| 74 | end |