| author | lcp | 
| Mon, 21 Mar 1994 11:02:57 +0100 | |
| changeset 286 | e7efbf03562b | 
| parent 120 | 09287f26bfb8 | 
| permissions | -rw-r--r-- | 
| 0 | 1  | 
(* Title: ZF/datatype.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1993 University of Cambridge  | 
|
5  | 
||
| 120 | 6  | 
(Co)Datatype Definitions for Zermelo-Fraenkel Set Theory  | 
| 0 | 7  | 
*)  | 
8  | 
||
9  | 
||
10  | 
(*Datatype definitions use least fixedpoints, standard products and sums*)  | 
|
11  | 
functor Datatype_Fun (Const: CONSTRUCTOR)  | 
|
12  | 
: sig include CONSTRUCTOR_RESULT INTR_ELIM INDRULE end =  | 
|
13  | 
struct  | 
|
14  | 
structure Constructor = Constructor_Fun (structure Const=Const and  | 
|
15  | 
Pr=Standard_Prod and Su=Standard_Sum);  | 
|
16  | 
open Const Constructor;  | 
|
17  | 
||
18  | 
structure Inductive = Inductive_Fun  | 
|
19  | 
(val thy = con_thy;  | 
|
20  | 
val rec_doms = (map #1 rec_specs) ~~ (map #2 rec_specs);  | 
|
21  | 
val sintrs = sintrs;  | 
|
22  | 
val monos = monos;  | 
|
23  | 
val con_defs = con_defs;  | 
|
24  | 
val type_intrs = type_intrs;  | 
|
25  | 
val type_elims = type_elims);  | 
|
26  | 
||
27  | 
open Inductive  | 
|
28  | 
end;  | 
|
29  | 
||
30  | 
||
| 120 | 31  | 
(*Codatatype definitions use greatest fixedpoints, Quine products and sums*)  | 
32  | 
functor CoDatatype_Fun (Const: CONSTRUCTOR)  | 
|
33  | 
: sig include CONSTRUCTOR_RESULT INTR_ELIM COINDRULE end =  | 
|
| 0 | 34  | 
struct  | 
35  | 
structure Constructor = Constructor_Fun (structure Const=Const and  | 
|
36  | 
Pr=Quine_Prod and Su=Quine_Sum);  | 
|
37  | 
open Const Constructor;  | 
|
38  | 
||
| 120 | 39  | 
structure CoInductive = CoInductive_Fun  | 
| 0 | 40  | 
(val thy = con_thy;  | 
41  | 
val rec_doms = (map #1 rec_specs) ~~ (map #2 rec_specs);  | 
|
42  | 
val sintrs = sintrs;  | 
|
43  | 
val monos = monos;  | 
|
44  | 
val con_defs = con_defs;  | 
|
45  | 
val type_intrs = type_intrs;  | 
|
46  | 
val type_elims = type_elims);  | 
|
47  | 
||
| 120 | 48  | 
open CoInductive  | 
| 0 | 49  | 
end;  | 
50  | 
||
51  | 
||
52  | 
||
53  | 
(*For most datatypes involving univ*)  | 
|
| 
70
 
8a29f8b4aca1
ZF/ind-syntax/fold_con_tac: deleted, since fold_tac now works
 
lcp 
parents: 
55 
diff
changeset
 | 
54  | 
val datatype_intrs =  | 
| 0 | 55  | 
[SigmaI, InlI, InrI,  | 
56  | 
Pair_in_univ, Inl_in_univ, Inr_in_univ,  | 
|
| 55 | 57  | 
zero_in_univ, A_into_univ, nat_into_univ, UnCI];  | 
| 0 | 58  | 
|
| 55 | 59  | 
(*Needed for mutual recursion*)  | 
| 
70
 
8a29f8b4aca1
ZF/ind-syntax/fold_con_tac: deleted, since fold_tac now works
 
lcp 
parents: 
55 
diff
changeset
 | 
60  | 
val datatype_elims = [make_elim InlD, make_elim InrD];  | 
| 0 | 61  | 
|
| 120 | 62  | 
(*For most codatatypes involving quniv*)  | 
63  | 
val codatatype_intrs =  | 
|
| 0 | 64  | 
[QSigmaI, QInlI, QInrI,  | 
65  | 
QPair_in_quniv, QInl_in_quniv, QInr_in_quniv,  | 
|
| 55 | 66  | 
zero_in_quniv, A_into_quniv, nat_into_quniv, UnCI];  | 
| 0 | 67  | 
|
| 120 | 68  | 
val codatatype_elims = [make_elim QInlD, make_elim QInrD];  | 
| 55 | 69  |