src/HOL/Finite.thy
author nipkow
Sat, 27 Apr 1996 18:51:42 +0200
changeset 1699 0bcc8cab3461
parent 1556 2fd82cec17d4
child 3367 832c245d967c
permissions -rw-r--r--
Forgot to add Expr to IMP.

(*  Title:      HOL/Finite.thy
    ID:         $Id$
    Author:     Lawrence C Paulson & Tobias Nipkow
    Copyright   1995  University of Cambridge & TU Muenchen

Finite sets and their cardinality
*)

Finite = Arith +

consts Fin :: 'a set => 'a set set

inductive "Fin(A)"
  intrs
    emptyI  "{} : Fin(A)"
    insertI "[| a: A;  b: Fin(A) |] ==> insert a b : Fin(A)"

constdefs

  finite :: 'a set => bool
  "finite A == A : Fin(UNIV)"

  card :: 'a set => nat
  "card A == LEAST n. ? f. A = {f i |i. i<n}"

end