author | wenzelm |
Tue, 21 Sep 1999 17:31:20 +0200 | |
changeset 7566 | c5a3f980a7af |
parent 6053 | 8a1059aa01f0 |
child 9491 | 1a36151ee2fc |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/Finite.thy |
516 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
516 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Finite powerset operator |
|
7 |
*) |
|
8 |
||
806 | 9 |
Finite = Arith + Inductive + |
534 | 10 |
consts |
1478 | 11 |
Fin :: i=>i |
12 |
FiniteFun :: [i,i]=>i ("(_ -||>/ _)" [61, 60] 60) |
|
534 | 13 |
|
516 | 14 |
inductive |
15 |
domains "Fin(A)" <= "Pow(A)" |
|
16 |
intrs |
|
17 |
emptyI "0 : Fin(A)" |
|
18 |
consI "[| a: A; b: Fin(A) |] ==> cons(a,b) : Fin(A)" |
|
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
1478
diff
changeset
|
19 |
type_intrs empty_subsetI, cons_subsetI, PowI |
516 | 20 |
type_elims "[make_elim PowD]" |
534 | 21 |
|
22 |
inductive |
|
23 |
domains "FiniteFun(A,B)" <= "Fin(A*B)" |
|
24 |
intrs |
|
25 |
emptyI "0 : A -||> B" |
|
1155 | 26 |
consI "[| a: A; b: B; h: A -||> B; a ~: domain(h) |
1478 | 27 |
|] ==> cons(<a,b>,h) : A -||> B" |
534 | 28 |
type_intrs "Fin.intrs" |
516 | 29 |
end |