| author | paulson | 
| Fri, 18 Apr 1997 11:47:11 +0200 | |
| changeset 2981 | aa5aeb6467c6 | 
| parent 2855 | 36f75c4a0047 | 
| child 3768 | 67f4ac759100 | 
| permissions | -rw-r--r-- | 
| 1465 | 1  | 
(* Title: HOL/inductive.ML  | 
| 923 | 2  | 
ID: $Id$  | 
| 1465 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 923 | 4  | 
Copyright 1993 University of Cambridge  | 
5  | 
||
6  | 
(Co)Inductive Definitions for HOL  | 
|
7  | 
||
8  | 
Inductive definitions use least fixedpoints with standard products and sums  | 
|
9  | 
Coinductive definitions use greatest fixedpoints with Quine products and sums  | 
|
10  | 
||
11  | 
Sums are used only for mutual recursion;  | 
|
12  | 
Products are used only to derive "streamlined" induction rules for relations  | 
|
13  | 
*)  | 
|
14  | 
||
15  | 
fun gen_fp_oper a (X,T,t) =  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
16  | 
let val setT = Ind_Syntax.mk_setT T  | 
| 923 | 17  | 
in Const(a, (setT-->setT)-->setT) $ absfree(X, setT, t) end;  | 
18  | 
||
19  | 
structure Lfp_items =  | 
|
20  | 
struct  | 
|
| 
2855
 
36f75c4a0047
Now calls require_thy to ensure ancestors are present
 
paulson 
parents: 
1465 
diff
changeset
 | 
21  | 
val checkThy = (fn thy => require_thy thy "Lfp" "inductive definitions")  | 
| 1465 | 22  | 
val oper = gen_fp_oper "lfp"  | 
23  | 
val Tarski = def_lfp_Tarski  | 
|
24  | 
val induct = def_induct  | 
|
| 923 | 25  | 
end;  | 
26  | 
||
27  | 
structure Gfp_items =  | 
|
28  | 
struct  | 
|
| 
2855
 
36f75c4a0047
Now calls require_thy to ensure ancestors are present
 
paulson 
parents: 
1465 
diff
changeset
 | 
29  | 
val checkThy = (fn thy => require_thy thy "Gfp" "coinductive definitions")  | 
| 1465 | 30  | 
val oper = gen_fp_oper "gfp"  | 
31  | 
val Tarski = def_gfp_Tarski  | 
|
32  | 
val induct = def_Collect_coinduct  | 
|
| 923 | 33  | 
end;  | 
34  | 
||
35  | 
||
36  | 
functor Ind_section_Fun (Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end)  | 
|
37  | 
: sig include INTR_ELIM INDRULE end =  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
38  | 
let  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
39  | 
structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and  | 
| 1465 | 40  | 
Fp=Lfp_items);  | 
| 923 | 41  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
42  | 
structure Indrule = Indrule_Fun  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
43  | 
(structure Inductive=Inductive and Intr_elim=Intr_elim);  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
44  | 
in  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
45  | 
struct  | 
| 1465 | 46  | 
val thy = Intr_elim.thy  | 
47  | 
val defs = Intr_elim.defs  | 
|
48  | 
val mono = Intr_elim.mono  | 
|
49  | 
val intrs = Intr_elim.intrs  | 
|
50  | 
val elim = Intr_elim.elim  | 
|
51  | 
val mk_cases = Intr_elim.mk_cases  | 
|
| 
1430
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
52  | 
open Indrule  | 
| 
 
439e1476a7f8
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
923 
diff
changeset
 | 
53  | 
end  | 
| 923 | 54  | 
end;  | 
55  | 
||
56  | 
||
57  | 
structure Ind = Add_inductive_def_Fun (Lfp_items);  | 
|
58  | 
||
59  | 
||
60  | 
signature INDUCTIVE_STRING =  | 
|
61  | 
sig  | 
|
| 1465 | 62  | 
val thy_name : string (*name of the new theory*)  | 
63  | 
val srec_tms : string list (*recursion terms*)  | 
|
64  | 
val sintrs : string list (*desired introduction rules*)  | 
|
| 923 | 65  | 
end;  | 
66  | 
||
67  | 
||
| 
2855
 
36f75c4a0047
Now calls require_thy to ensure ancestors are present
 
paulson 
parents: 
1465 
diff
changeset
 | 
68  | 
(*Called by the theory sections or directly from ML*)  | 
| 923 | 69  | 
functor Inductive_Fun  | 
70  | 
(Inductive: sig include INDUCTIVE_STRING INDUCTIVE_ARG end)  | 
|
71  | 
: sig include INTR_ELIM INDRULE end =  | 
|
72  | 
Ind_section_Fun  | 
|
73  | 
(open Inductive Ind_Syntax  | 
|
74  | 
val sign = sign_of thy;  | 
|
75  | 
val rec_tms = map (readtm sign termTVar) srec_tms  | 
|
76  | 
and intr_tms = map (readtm sign propT) sintrs;  | 
|
77  | 
val thy = thy |> Ind.add_fp_def_i(rec_tms, intr_tms)  | 
|
78  | 
|> add_thyname thy_name);  | 
|
79  | 
||
80  | 
||
81  | 
||
82  | 
signature COINDRULE =  | 
|
83  | 
sig  | 
|
84  | 
val coinduct : thm  | 
|
85  | 
end;  | 
|
86  | 
||
87  | 
||
88  | 
functor CoInd_section_Fun  | 
|
89  | 
(Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end)  | 
|
90  | 
: sig include INTR_ELIM COINDRULE end =  | 
|
91  | 
struct  | 
|
92  | 
structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and Fp=Gfp_items);  | 
|
93  | 
||
94  | 
open Intr_elim  | 
|
95  | 
val coinduct = raw_induct  | 
|
96  | 
end;  | 
|
97  | 
||
98  | 
||
99  | 
structure CoInd = Add_inductive_def_Fun(Gfp_items);  |