| author | wenzelm | 
| Wed, 15 Oct 1997 15:18:41 +0200 | |
| changeset 3878 | 0258594baaa9 | 
| parent 3865 | 0035d1f97096 | 
| child 3885 | dccac762b0be | 
| permissions | -rw-r--r-- | 
| 1526 | 1  | 
(* Title: Pure/theory.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson and Markus Wenzel  | 
|
4  | 
Copyright 1996 University of Cambridge  | 
|
5  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
6  | 
Theories.  | 
| 1526 | 7  | 
*)  | 
8  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
9  | 
(*this part made pervasive*)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
10  | 
signature BASIC_THEORY =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
11  | 
sig  | 
| 1526 | 12  | 
type theory  | 
13  | 
exception THEORY of string * theory list  | 
|
14  | 
val rep_theory : theory ->  | 
|
| 3814 | 15  | 
    {sign: Sign.sg,
 | 
16  | 
new_axioms: term Symtab.table,  | 
|
17  | 
oracles: ((Sign.sg * exn -> term) * stamp) Symtab.table,  | 
|
18  | 
parents: theory list}  | 
|
| 1526 | 19  | 
val sign_of : theory -> Sign.sg  | 
20  | 
val syn_of : theory -> Syntax.syntax  | 
|
21  | 
val stamps_of_thy : theory -> string ref list  | 
|
22  | 
val parents_of : theory -> theory list  | 
|
23  | 
val subthy : theory * theory -> bool  | 
|
24  | 
val eq_thy : theory * theory -> bool  | 
|
25  | 
val proto_pure_thy : theory  | 
|
26  | 
val pure_thy : theory  | 
|
27  | 
val cpure_thy : theory  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
28  | 
val cert_axm : Sign.sg -> string * term -> string * term  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
29  | 
val read_axm : Sign.sg -> string * string -> string * term  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
30  | 
val inferT_axm : Sign.sg -> string * term -> string * term  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
31  | 
end  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
32  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
33  | 
signature THEORY =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
34  | 
sig  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
35  | 
include BASIC_THEORY  | 
| 3814 | 36  | 
val thmK : string  | 
37  | 
val oracleK : string  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
38  | 
(*theory extendsion primitives*)  | 
| 3806 | 39  | 
val add_classes : (xclass * xclass list) list -> theory -> theory  | 
40  | 
val add_classes_i : (xclass * class list) list -> theory -> theory  | 
|
41  | 
val add_classrel : (xclass * xclass) list -> theory -> theory  | 
|
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
42  | 
val add_classrel_i : (class * class) list -> theory -> theory  | 
| 3806 | 43  | 
val add_defsort : xsort -> theory -> theory  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
44  | 
val add_defsort_i : sort -> theory -> theory  | 
| 3806 | 45  | 
val add_types : (xstring * int * mixfix) list -> theory -> theory  | 
46  | 
val add_tyabbrs : (xstring * string list * string * mixfix) list  | 
|
| 1526 | 47  | 
-> theory -> theory  | 
| 3806 | 48  | 
val add_tyabbrs_i : (xstring * string list * typ * mixfix) list  | 
| 1526 | 49  | 
-> theory -> theory  | 
| 3806 | 50  | 
val add_arities : (xstring * xsort list * xsort) list -> theory -> theory  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
51  | 
val add_arities_i : (string * sort list * sort) list -> theory -> theory  | 
| 3806 | 52  | 
val add_consts : (xstring * string * mixfix) list -> theory -> theory  | 
53  | 
val add_consts_i : (xstring * typ * mixfix) list -> theory -> theory  | 
|
54  | 
val add_syntax : (xstring * string * mixfix) list -> theory -> theory  | 
|
55  | 
val add_syntax_i : (xstring * typ * mixfix) list -> theory -> theory  | 
|
56  | 
val add_modesyntax : string * bool -> (xstring * string * mixfix) list -> theory -> theory  | 
|
57  | 
val add_modesyntax_i : string * bool -> (xstring * typ * mixfix) list -> theory -> theory  | 
|
| 1539 | 58  | 
val add_trfuns :  | 
| 1526 | 59  | 
(string * (Syntax.ast list -> Syntax.ast)) list *  | 
60  | 
(string * (term list -> term)) list *  | 
|
61  | 
(string * (term list -> term)) list *  | 
|
62  | 
(string * (Syntax.ast list -> Syntax.ast)) list -> theory -> theory  | 
|
| 2385 | 63  | 
val add_trfunsT :  | 
64  | 
(string * (typ -> term list -> term)) list -> theory -> theory  | 
|
| 2693 | 65  | 
val add_tokentrfuns:  | 
66  | 
(string * string * (string -> string * int)) list -> theory -> theory  | 
|
| 1526 | 67  | 
val add_trrules : (string * string)Syntax.trrule list -> theory -> theory  | 
| 1539 | 68  | 
val add_trrules_i : Syntax.ast Syntax.trrule list -> theory -> theory  | 
| 3806 | 69  | 
val add_axioms : (xstring * string) list -> theory -> theory  | 
70  | 
val add_axioms_i : (xstring * term) list -> theory -> theory  | 
|
| 3814 | 71  | 
val add_oracle : string * (Sign.sg * exn -> term) -> theory -> theory  | 
| 3806 | 72  | 
val add_defs : (xstring * string) list -> theory -> theory  | 
73  | 
val add_defs_i : (xstring * term) list -> theory -> theory  | 
|
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
74  | 
val add_path : string -> theory -> theory  | 
| 3814 | 75  | 
val add_space : string * string list -> theory -> theory  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
76  | 
val add_name : string -> theory -> theory  | 
| 3878 | 77  | 
val init_data : string * exn * (exn -> exn) * (exn * exn -> exn) *  | 
78  | 
(string -> exn -> unit) -> theory -> theory  | 
|
| 3865 | 79  | 
val get_data : theory -> string -> exn  | 
80  | 
val put_data : string * exn -> theory -> theory  | 
|
| 3878 | 81  | 
val merge_list : theory list -> theory  | 
| 1526 | 82  | 
end;  | 
83  | 
||
84  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
85  | 
structure Theory: THEORY =  | 
| 1526 | 86  | 
struct  | 
| 2206 | 87  | 
|
88  | 
(** datatype theory **)  | 
|
| 1526 | 89  | 
|
90  | 
datatype theory =  | 
|
91  | 
  Theory of {
 | 
|
92  | 
sign: Sign.sg,  | 
|
93  | 
new_axioms: term Symtab.table,  | 
|
| 3814 | 94  | 
oracles: ((Sign.sg * exn -> term) * stamp) Symtab.table,  | 
| 1526 | 95  | 
parents: theory list};  | 
96  | 
||
97  | 
fun rep_theory (Theory args) = args;  | 
|
98  | 
||
99  | 
(*errors involving theories*)  | 
|
100  | 
exception THEORY of string * theory list;  | 
|
101  | 
||
102  | 
||
103  | 
val sign_of = #sign o rep_theory;  | 
|
104  | 
val syn_of = #syn o Sign.rep_sg o sign_of;  | 
|
105  | 
||
106  | 
(*stamps associated with a theory*)  | 
|
107  | 
val stamps_of_thy = #stamps o Sign.rep_sg o sign_of;  | 
|
108  | 
||
109  | 
(*return the immediate ancestors*)  | 
|
110  | 
val parents_of = #parents o rep_theory;  | 
|
111  | 
||
112  | 
||
113  | 
(*compare theories*)  | 
|
114  | 
val subthy = Sign.subsig o pairself sign_of;  | 
|
115  | 
val eq_thy = Sign.eq_sg o pairself sign_of;  | 
|
116  | 
||
117  | 
||
118  | 
(* the Pure theories *)  | 
|
119  | 
||
| 3878 | 120  | 
fun make_thy sign axms oras parents =  | 
121  | 
  Theory {sign = sign, new_axioms = axms, oracles = oras, parents = parents};
 | 
|
| 1526 | 122  | 
|
| 3878 | 123  | 
val proto_pure_thy = make_thy Sign.proto_pure Symtab.null Symtab.null [];  | 
124  | 
val pure_thy = make_thy Sign.pure Symtab.null Symtab.null [proto_pure_thy];  | 
|
125  | 
val cpure_thy = make_thy Sign.cpure Symtab.null Symtab.null [proto_pure_thy];  | 
|
| 1526 | 126  | 
|
127  | 
||
128  | 
||
129  | 
(** extend theory **)  | 
|
130  | 
||
| 3814 | 131  | 
(*name space kinds*)  | 
132  | 
val thmK = "thm";  | 
|
133  | 
val oracleK = "oracle";  | 
|
134  | 
||
135  | 
||
136  | 
(* extend logical part of a theory *)  | 
|
137  | 
||
| 1526 | 138  | 
fun err_dup_axms names =  | 
139  | 
  error ("Duplicate axiom name(s) " ^ commas_quote names);
 | 
|
140  | 
||
| 3814 | 141  | 
fun err_dup_oras names =  | 
142  | 
  error ("Duplicate oracles " ^ commas_quote names);
 | 
|
143  | 
||
144  | 
||
145  | 
fun ext_thy thy sign' new_axms new_oras =  | 
|
| 1526 | 146  | 
let  | 
| 3814 | 147  | 
    val Theory {sign, new_axioms, oracles, parents} = thy;
 | 
| 1526 | 148  | 
val draft = Sign.is_draft sign;  | 
| 3814 | 149  | 
val new_axioms' =  | 
| 1526 | 150  | 
Symtab.extend_new (if draft then new_axioms else Symtab.null, new_axms)  | 
151  | 
handle Symtab.DUPS names => err_dup_axms names;  | 
|
| 3814 | 152  | 
val oracles' =  | 
153  | 
Symtab.extend_new (oracles, new_oras)  | 
|
154  | 
handle Symtab.DUPS names => err_dup_oras names;  | 
|
155  | 
val parents' = if draft then parents else [thy];  | 
|
| 1526 | 156  | 
in  | 
| 3878 | 157  | 
make_thy sign' new_axioms' oracles' parents'  | 
| 1526 | 158  | 
end;  | 
159  | 
||
160  | 
||
161  | 
(* extend signature of a theory *)  | 
|
162  | 
||
163  | 
fun ext_sg extfun decls (thy as Theory {sign, ...}) =
 | 
|
| 3814 | 164  | 
ext_thy thy (extfun decls sign) [] [];  | 
| 1526 | 165  | 
|
| 2206 | 166  | 
val add_classes = ext_sg Sign.add_classes;  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
167  | 
val add_classes_i = ext_sg Sign.add_classes_i;  | 
| 2206 | 168  | 
val add_classrel = ext_sg Sign.add_classrel;  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
169  | 
val add_classrel_i = ext_sg Sign.add_classrel_i;  | 
| 2206 | 170  | 
val add_defsort = ext_sg Sign.add_defsort;  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
171  | 
val add_defsort_i = ext_sg Sign.add_defsort_i;  | 
| 2206 | 172  | 
val add_types = ext_sg Sign.add_types;  | 
173  | 
val add_tyabbrs = ext_sg Sign.add_tyabbrs;  | 
|
174  | 
val add_tyabbrs_i = ext_sg Sign.add_tyabbrs_i;  | 
|
175  | 
val add_arities = ext_sg Sign.add_arities;  | 
|
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
176  | 
val add_arities_i = ext_sg Sign.add_arities_i;  | 
| 2206 | 177  | 
val add_consts = ext_sg Sign.add_consts;  | 
178  | 
val add_consts_i = ext_sg Sign.add_consts_i;  | 
|
179  | 
val add_syntax = ext_sg Sign.add_syntax;  | 
|
180  | 
val add_syntax_i = ext_sg Sign.add_syntax_i;  | 
|
181  | 
val add_modesyntax = curry (ext_sg Sign.add_modesyntax);  | 
|
182  | 
val add_modesyntax_i = curry (ext_sg Sign.add_modesyntax_i);  | 
|
183  | 
val add_trfuns = ext_sg Sign.add_trfuns;  | 
|
| 2385 | 184  | 
val add_trfunsT = ext_sg Sign.add_trfunsT;  | 
| 2693 | 185  | 
val add_tokentrfuns = ext_sg Sign.add_tokentrfuns;  | 
| 2206 | 186  | 
val add_trrules = ext_sg Sign.add_trrules;  | 
187  | 
val add_trrules_i = ext_sg Sign.add_trrules_i;  | 
|
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
188  | 
val add_path = ext_sg Sign.add_path;  | 
| 
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
189  | 
val add_space = ext_sg Sign.add_space;  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
190  | 
val add_name = ext_sg Sign.add_name;  | 
| 1526 | 191  | 
|
192  | 
||
| 3814 | 193  | 
|
194  | 
(** add axioms **)  | 
|
195  | 
||
| 1526 | 196  | 
(* prepare axioms *)  | 
197  | 
||
198  | 
fun err_in_axm name =  | 
|
199  | 
  error ("The error(s) above occurred in axiom " ^ quote name);
 | 
|
200  | 
||
201  | 
fun no_vars tm =  | 
|
202  | 
if null (term_vars tm) andalso null (term_tvars tm) then tm  | 
|
203  | 
else error "Illegal schematic variable(s) in term";  | 
|
204  | 
||
205  | 
fun cert_axm sg (name, raw_tm) =  | 
|
206  | 
let  | 
|
207  | 
val (t, T, _) = Sign.certify_term sg raw_tm  | 
|
| 2979 | 208  | 
handle TYPE (msg, _, _) => error msg  | 
| 1526 | 209  | 
| TERM (msg, _) => error msg;  | 
210  | 
in  | 
|
211  | 
assert (T = propT) "Term not of type prop";  | 
|
212  | 
(name, no_vars t)  | 
|
213  | 
end  | 
|
214  | 
handle ERROR => err_in_axm name;  | 
|
215  | 
||
216  | 
(*Some duplication of code with read_def_cterm*)  | 
|
217  | 
fun read_axm sg (name, str) =  | 
|
| 3814 | 218  | 
let  | 
219  | 
val ts = Syntax.read (#syn (Sign.rep_sg sg)) propT str;  | 
|
220  | 
val (_, t, _) =  | 
|
221  | 
Sign.infer_types sg (K None) (K None) [] true (ts, propT);  | 
|
| 
1960
 
ae390b599213
Improved error handling: if there are syntax or type-checking
 
paulson 
parents: 
1539 
diff
changeset
 | 
222  | 
in cert_axm sg (name,t) end  | 
| 
 
ae390b599213
Improved error handling: if there are syntax or type-checking
 
paulson 
parents: 
1539 
diff
changeset
 | 
223  | 
handle ERROR => err_in_axm name;  | 
| 1526 | 224  | 
|
225  | 
fun inferT_axm sg (name, pre_tm) =  | 
|
| 3814 | 226  | 
let  | 
227  | 
val (_, t, _) =  | 
|
228  | 
Sign.infer_types sg (K None) (K None) [] true ([pre_tm], propT);  | 
|
229  | 
in (name, no_vars t) end  | 
|
| 1526 | 230  | 
handle ERROR => err_in_axm name;  | 
231  | 
||
232  | 
||
233  | 
(* extend axioms of a theory *)  | 
|
234  | 
||
| 3814 | 235  | 
fun ext_axms prep_axm raw_axms (thy as Theory {sign, ...}) =
 | 
| 1526 | 236  | 
let  | 
| 3814 | 237  | 
val raw_axms' = map (apfst (Sign.full_name sign)) raw_axms;  | 
238  | 
val axioms =  | 
|
239  | 
map (apsnd (Term.compress_term o Logic.varify) o prep_axm sign) raw_axms';  | 
|
240  | 
val sign' = Sign.add_space (thmK, map fst axioms) sign;  | 
|
| 1526 | 241  | 
in  | 
| 3814 | 242  | 
ext_thy thy sign' axioms []  | 
| 1526 | 243  | 
end;  | 
244  | 
||
245  | 
val add_axioms = ext_axms read_axm;  | 
|
246  | 
val add_axioms_i = ext_axms cert_axm;  | 
|
247  | 
||
248  | 
||
| 3814 | 249  | 
(* add oracle **)  | 
250  | 
||
251  | 
fun add_oracle (raw_name, oracle) (thy as Theory {sign, ...}) =
 | 
|
252  | 
let  | 
|
253  | 
val name = Sign.full_name sign raw_name;  | 
|
254  | 
val sign' = Sign.add_space (oracleK, [name]) sign;  | 
|
255  | 
in  | 
|
256  | 
ext_thy thy sign' [] [(name, (oracle, stamp ()))]  | 
|
257  | 
end;  | 
|
258  | 
||
259  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
260  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
261  | 
(** add constant definitions **)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
262  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
263  | 
(* all_axioms_of *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
264  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
265  | 
(*results may contain duplicates!*)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
266  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
267  | 
fun ancestry_of thy =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
268  | 
thy :: List.concat (map ancestry_of (parents_of thy));  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
269  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
270  | 
val all_axioms_of =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
271  | 
List.concat o map (Symtab.dest o #new_axioms o rep_theory) o ancestry_of;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
272  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
273  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
274  | 
(* clash_types, clash_consts *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
275  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
276  | 
(*check if types have common instance (ignoring sorts)*)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
277  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
278  | 
fun clash_types ty1 ty2 =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
279  | 
let  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
280  | 
val ty1' = Type.varifyT ty1;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
281  | 
val ty2' = incr_tvar (maxidx_of_typ ty1' + 1) (Type.varifyT ty2);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
282  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
283  | 
Type.raw_unify (ty1', ty2')  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
284  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
285  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
286  | 
fun clash_consts (c1, ty1) (c2, ty2) =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
287  | 
c1 = c2 andalso clash_types ty1 ty2;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
288  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
289  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
290  | 
(* clash_defns *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
291  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
292  | 
fun clash_defn c_ty (name, tm) =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
293  | 
let val (c, ty') = dest_Const (head_of (fst (Logic.dest_equals tm))) in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
294  | 
if clash_consts c_ty (c, ty') then Some (name, ty') else None  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
295  | 
end handle TERM _ => None;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
296  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
297  | 
fun clash_defns c_ty axms =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
298  | 
distinct (mapfilter (clash_defn c_ty) axms);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
299  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
300  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
301  | 
(* dest_defn *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
302  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
303  | 
fun dest_defn tm =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
304  | 
let  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
305  | 
fun err msg = raise TERM (msg, [tm]);  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
306  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
307  | 
val (lhs, rhs) = Logic.dest_equals (Logic.strip_imp_concl tm)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
308  | 
handle TERM _ => err "Not a meta-equality (==)";  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
309  | 
val (head, args) = strip_comb lhs;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
310  | 
val (c, ty) = dest_Const head  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
311  | 
handle TERM _ => err "Head of lhs not a constant";  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
312  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
313  | 
fun occs_const (Const c_ty') = (c_ty' = (c, ty))  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
314  | 
| occs_const (Abs (_, _, t)) = occs_const t  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
315  | 
| occs_const (t $ u) = occs_const t orelse occs_const u  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
316  | 
| occs_const _ = false;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
317  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
318  | 
val show_frees = commas_quote o map (fst o dest_Free);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
319  | 
val show_tfrees = commas_quote o map fst;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
320  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
321  | 
val lhs_dups = duplicates args;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
322  | 
val rhs_extras = gen_rems (op =) (term_frees rhs, args);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
323  | 
val rhs_extrasT = gen_rems (op =) (term_tfrees rhs, typ_tfrees ty);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
324  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
325  | 
if not (forall is_Free args) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
326  | 
err "Arguments (on lhs) must be variables"  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
327  | 
else if not (null lhs_dups) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
328  | 
      err ("Duplicate variables on lhs: " ^ show_frees lhs_dups)
 | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
329  | 
else if not (null rhs_extras) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
330  | 
      err ("Extra variables on rhs: " ^ show_frees rhs_extras)
 | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
331  | 
else if not (null rhs_extrasT) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
332  | 
      err ("Extra type variables on rhs: " ^ show_tfrees rhs_extrasT)
 | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
333  | 
else if occs_const rhs then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
334  | 
      err ("Constant to be defined occurs on rhs")
 | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
335  | 
else (c, ty)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
336  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
337  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
338  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
339  | 
(* check_defn *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
340  | 
|
| 3814 | 341  | 
fun err_in_defn sg name msg =  | 
342  | 
  (writeln msg; error ("The error(s) above occurred in definition " ^
 | 
|
343  | 
quote (Sign.full_name sg name)));  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
344  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
345  | 
fun check_defn sign (axms, (name, tm)) =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
346  | 
let  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
347  | 
fun show_const (c, ty) = quote (Pretty.string_of (Pretty.block  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
348  | 
[Pretty.str (c ^ " ::"), Pretty.brk 1, Sign.pretty_typ sign ty]));  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
349  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
350  | 
fun show_defn c (dfn, ty') = show_const (c, ty') ^ " in " ^ dfn;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
351  | 
fun show_defns c = cat_lines o map (show_defn c);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
352  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
353  | 
val (c, ty) = dest_defn tm  | 
| 3814 | 354  | 
handle TERM (msg, _) => err_in_defn sign name msg;  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
355  | 
val defns = clash_defns (c, ty) axms;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
356  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
357  | 
if not (null defns) then  | 
| 3814 | 358  | 
      err_in_defn sign name ("Definition of " ^ show_const (c, ty) ^
 | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
359  | 
"\nclashes with " ^ show_defns c defns)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
360  | 
else (name, tm) :: axms  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
361  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
362  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
363  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
364  | 
(* add_defs *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
365  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
366  | 
fun ext_defns prep_axm raw_axms thy =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
367  | 
let  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
368  | 
val axms = map (prep_axm (sign_of thy)) raw_axms;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
369  | 
val all_axms = all_axioms_of thy;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
370  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
371  | 
foldl (check_defn (sign_of thy)) (all_axms, axms);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
372  | 
add_axioms_i axms thy  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
373  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
374  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
375  | 
val add_defs_i = ext_defns cert_axm;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
376  | 
val add_defs = ext_defns read_axm;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
377  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
378  | 
|
| 3878 | 379  | 
|
| 3865 | 380  | 
(** additional theory data **)  | 
381  | 
||
382  | 
val init_data = ext_sg Sign.init_data;  | 
|
383  | 
val get_data = Sign.get_data o sign_of;  | 
|
384  | 
val put_data = ext_sg Sign.put_data;  | 
|
385  | 
||
386  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
387  | 
|
| 1526 | 388  | 
(** merge theories **)  | 
389  | 
||
| 3878 | 390  | 
(*merge list of theories from left to right, preparing extend*)  | 
391  | 
fun merge_list thys =  | 
|
| 1526 | 392  | 
let  | 
393  | 
fun is_union thy = forall (fn t => subthy (t, thy)) thys;  | 
|
394  | 
val is_draft = Sign.is_draft o sign_of;  | 
|
395  | 
||
396  | 
    fun add_sign (sg, Theory {sign, ...}) =
 | 
|
| 3878 | 397  | 
Sign.nontriv_merge (sg, sign) handle TERM (msg, _) => error msg;  | 
| 3814 | 398  | 
|
399  | 
    fun oracles_of (Theory {oracles, ...}) = Symtab.dest oracles;
 | 
|
400  | 
fun eq_ora ((_, (_, s1: stamp)), (_, (_, s2))) = s1 = s2;  | 
|
| 1526 | 401  | 
in  | 
| 3878 | 402  | 
if exists is_draft thys then  | 
403  | 
      raise THEORY ("Illegal merge of draft theories", thys)
 | 
|
404  | 
else  | 
|
405  | 
(case find_first is_union thys of  | 
|
406  | 
        Some (Theory {sign, oracles, new_axioms = _, parents = _}) =>
 | 
|
407  | 
make_thy (Sign.prep_ext sign) Symtab.null oracles thys  | 
|
408  | 
| None =>  | 
|
409  | 
make_thy  | 
|
410  | 
(Sign.prep_ext (foldl add_sign (Sign.proto_pure, thys)))  | 
|
411  | 
Symtab.null  | 
|
412  | 
(Symtab.make (gen_distinct eq_ora (flat (map oracles_of thys)))  | 
|
413  | 
handle Symtab.DUPS names => err_dup_oras names)  | 
|
414  | 
thys)  | 
|
| 1526 | 415  | 
end;  | 
416  | 
||
417  | 
||
418  | 
end;  | 
|
419  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
420  | 
structure BasicTheory: BASIC_THEORY = Theory;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
421  | 
open BasicTheory;  |