| author | nipkow |
| Fri, 25 Oct 1996 15:02:09 +0200 | |
| changeset 2129 | 2ffe6e24f38d |
| parent 1556 | 2fd82cec17d4 |
| child 3367 | 832c245d967c |
| permissions | -rw-r--r-- |
| 1475 | 1 |
(* Title: HOL/Finite.thy |
| 923 | 2 |
ID: $Id$ |
| 1531 | 3 |
Author: Lawrence C Paulson & Tobias Nipkow |
4 |
Copyright 1995 University of Cambridge & TU Muenchen |
|
| 923 | 5 |
|
| 1531 | 6 |
Finite sets and their cardinality |
| 923 | 7 |
*) |
8 |
||
| 1531 | 9 |
Finite = Arith + |
10 |
||
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
923
diff
changeset
|
11 |
consts Fin :: 'a set => 'a set set |
| 923 | 12 |
|
13 |
inductive "Fin(A)" |
|
14 |
intrs |
|
15 |
emptyI "{} : Fin(A)"
|
|
16 |
insertI "[| a: A; b: Fin(A) |] ==> insert a b : Fin(A)" |
|
17 |
||
| 1556 | 18 |
constdefs |
| 1531 | 19 |
|
| 1556 | 20 |
finite :: 'a set => bool |
21 |
"finite A == A : Fin(UNIV)" |
|
22 |
||
23 |
card :: 'a set => nat |
|
24 |
"card A == LEAST n. ? f. A = {f i |i. i<n}"
|
|
| 1531 | 25 |
|
| 923 | 26 |
end |