author | wenzelm |
Thu, 27 Sep 2001 18:46:32 +0200 | |
changeset 11599 | 12cc28aafb4d |
parent 9491 | 1a36151ee2fc |
child 12175 | 5cf58a1799a7 |
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 |
||
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
9 |
Finite = Inductive + Nat + |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
10 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
11 |
setup setup_datatypes |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
12 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
13 |
(*The natural numbers as a datatype*) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
14 |
rep_datatype |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
15 |
elim natE |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
16 |
induct nat_induct |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
17 |
case_eqns nat_case_0, nat_case_succ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
18 |
recursor_eqns recursor_0, recursor_succ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
19 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
6053
diff
changeset
|
20 |
|
534 | 21 |
consts |
1478 | 22 |
Fin :: i=>i |
23 |
FiniteFun :: [i,i]=>i ("(_ -||>/ _)" [61, 60] 60) |
|
534 | 24 |
|
516 | 25 |
inductive |
26 |
domains "Fin(A)" <= "Pow(A)" |
|
27 |
intrs |
|
28 |
emptyI "0 : Fin(A)" |
|
29 |
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
|
30 |
type_intrs empty_subsetI, cons_subsetI, PowI |
516 | 31 |
type_elims "[make_elim PowD]" |
534 | 32 |
|
33 |
inductive |
|
34 |
domains "FiniteFun(A,B)" <= "Fin(A*B)" |
|
35 |
intrs |
|
36 |
emptyI "0 : A -||> B" |
|
1155 | 37 |
consI "[| a: A; b: B; h: A -||> B; a ~: domain(h) |
1478 | 38 |
|] ==> cons(<a,b>,h) : A -||> B" |
534 | 39 |
type_intrs "Fin.intrs" |
516 | 40 |
end |