author | berghofe |
Mon, 21 Jan 2002 14:43:38 +0100 | |
changeset 12822 | 073116d65bb9 |
parent 12372 | cd3a09c7dac9 |
child 13259 | 01fa0c8dbc92 |
permissions | -rw-r--r-- |
12175 | 1 |
(* Title: ZF/Inductive.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
(Co)Inductive Definitions for Zermelo-Fraenkel Set Theory. |
|
7 |
*) |
|
578
efc648d29dd0
ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
diff
changeset
|
8 |
|
12175 | 9 |
theory Inductive = Fixedpt + mono |
10 |
files |
|
12183 | 11 |
"ind_syntax.ML" |
12175 | 12 |
"Tools/cartprod.ML" |
13 |
"Tools/ind_cases.ML" |
|
14 |
"Tools/inductive_package.ML" |
|
15 |
"Tools/induct_tacs.ML" |
|
16 |
"Tools/primrec_package.ML": |
|
17 |
||
18 |
setup IndCases.setup |
|
12208 | 19 |
setup DatatypeTactics.setup |
2870
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
paulson
parents:
805
diff
changeset
|
20 |
|
12372 | 21 |
|
22 |
(*belongs to theory ZF*) |
|
23 |
declare bspec [dest?] |
|
24 |
||
25 |
(*belongs to theory upair*) |
|
26 |
declare UnI1 [elim?] UnI2 [elim?] |
|
27 |
||
2870
6d6fd10a9fdc
Now a non-trivial theory so that require_thy can find it
paulson
parents:
805
diff
changeset
|
28 |
end |