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