| author | wenzelm | 
| Mon, 08 May 2000 11:45:57 +0200 | |
| changeset 8831 | b824c0c55613 | 
| parent 8725 | 0e48ee5b52db | 
| child 8897 | fb1436ca3b2e | 
| 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  | 
||
| 5025 | 6  | 
The abstract type "theory" of theories.  | 
| 1526 | 7  | 
*)  | 
8  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
9  | 
signature BASIC_THEORY =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
10  | 
sig  | 
| 1526 | 11  | 
type theory  | 
12  | 
exception THEORY of string * theory list  | 
|
| 3996 | 13  | 
val rep_theory: theory ->  | 
| 3814 | 14  | 
    {sign: Sign.sg,
 | 
| 3996 | 15  | 
axioms: term Symtab.table,  | 
| 4996 | 16  | 
oracles: ((Sign.sg * Object.T -> term) * stamp) Symtab.table,  | 
| 4019 | 17  | 
parents: theory list,  | 
18  | 
ancestors: theory list}  | 
|
| 3996 | 19  | 
val sign_of: theory -> Sign.sg  | 
| 6188 | 20  | 
val is_draft: theory -> bool  | 
| 3996 | 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  | 
|
| 6311 | 27  | 
val read_def_axm: Sign.sg * (indexname -> typ option) * (indexname -> sort option) ->  | 
28  | 
string list -> string * string -> string * term  | 
|
| 3996 | 29  | 
val read_axm: Sign.sg -> string * string -> string * term  | 
30  | 
val inferT_axm: Sign.sg -> string * term -> string * term  | 
|
| 
3767
 
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  | 
| 3996 | 36  | 
val axiomK: string  | 
37  | 
val oracleK: string  | 
|
| 4912 | 38  | 
(*theory extension primitives*)  | 
| 3996 | 39  | 
val add_classes: (bclass * xclass list) list -> theory -> theory  | 
40  | 
val add_classes_i: (bclass * class list) list -> theory -> theory  | 
|
41  | 
val add_classrel: (xclass * xclass) list -> theory -> theory  | 
|
42  | 
val add_classrel_i: (class * class) list -> theory -> theory  | 
|
43  | 
val add_defsort: xsort -> theory -> theory  | 
|
44  | 
val add_defsort_i: sort -> theory -> theory  | 
|
45  | 
val add_types: (bstring * int * mixfix) list -> theory -> theory  | 
|
| 4846 | 46  | 
val add_nonterminals: bstring list -> theory -> theory  | 
| 3996 | 47  | 
val add_tyabbrs: (bstring * string list * string * mixfix) list  | 
| 1526 | 48  | 
-> theory -> theory  | 
| 3996 | 49  | 
val add_tyabbrs_i: (bstring * string list * typ * mixfix) list  | 
| 1526 | 50  | 
-> theory -> theory  | 
| 3996 | 51  | 
val add_arities: (xstring * xsort list * xsort) list -> theory -> theory  | 
52  | 
val add_arities_i: (string * sort list * sort) list -> theory -> theory  | 
|
53  | 
val add_consts: (bstring * string * mixfix) list -> theory -> theory  | 
|
54  | 
val add_consts_i: (bstring * typ * mixfix) list -> theory -> theory  | 
|
55  | 
val add_syntax: (bstring * string * mixfix) list -> theory -> theory  | 
|
56  | 
val add_syntax_i: (bstring * typ * mixfix) list -> theory -> theory  | 
|
57  | 
val add_modesyntax: string * bool -> (bstring * string * mixfix) list -> theory -> theory  | 
|
58  | 
val add_modesyntax_i: string * bool -> (bstring * typ * mixfix) list -> theory -> theory  | 
|
59  | 
val add_trfuns:  | 
|
| 4344 | 60  | 
(string * (Syntax.ast list -> Syntax.ast)) list *  | 
61  | 
(string * (term list -> term)) list *  | 
|
62  | 
(string * (term list -> term)) list *  | 
|
63  | 
(string * (Syntax.ast list -> Syntax.ast)) list -> theory -> theory  | 
|
| 3996 | 64  | 
val add_trfunsT:  | 
| 4344 | 65  | 
(string * (bool -> typ -> term list -> term)) list -> theory -> theory  | 
| 2693 | 66  | 
val add_tokentrfuns:  | 
| 6311 | 67  | 
(string * string * (string -> string * real)) list -> theory -> theory  | 
68  | 
val add_mode_tokentrfuns: string -> (string * (string -> string * real)) list  | 
|
69  | 
-> theory -> theory  | 
|
| 4617 | 70  | 
val add_trrules: (xstring * string) Syntax.trrule list -> theory -> theory  | 
| 3996 | 71  | 
val add_trrules_i: Syntax.ast Syntax.trrule list -> theory -> theory  | 
72  | 
val add_axioms: (bstring * string) list -> theory -> theory  | 
|
73  | 
val add_axioms_i: (bstring * term) list -> theory -> theory  | 
|
| 4996 | 74  | 
val add_oracle: bstring * (Sign.sg * Object.T -> term) -> theory -> theory  | 
| 3996 | 75  | 
val add_defs: (bstring * string) list -> theory -> theory  | 
76  | 
val add_defs_i: (bstring * term) list -> theory -> theory  | 
|
77  | 
val add_path: string -> theory -> theory  | 
|
| 4846 | 78  | 
val parent_path: theory -> theory  | 
79  | 
val root_path: theory -> theory  | 
|
| 8725 | 80  | 
val hide_space: string * xstring list -> theory -> theory  | 
81  | 
val hide_space_i: string * string list -> theory -> theory  | 
|
82  | 
val hide_classes: string list -> theory -> theory  | 
|
83  | 
val hide_types: string list -> theory -> theory  | 
|
84  | 
val hide_consts: string list -> theory -> theory  | 
|
| 3996 | 85  | 
val add_name: string -> theory -> theory  | 
| 5862 | 86  | 
val copy: theory -> theory  | 
| 6661 | 87  | 
val prep_ext: theory -> theory  | 
| 3996 | 88  | 
val prep_ext_merge: theory list -> theory  | 
| 4970 | 89  | 
val requires: theory -> string -> string -> unit  | 
| 6369 | 90  | 
val assert_super: theory -> theory -> theory  | 
| 3996 | 91  | 
val pre_pure: theory  | 
| 1526 | 92  | 
end;  | 
93  | 
||
| 6188 | 94  | 
signature PRIVATE_THEORY =  | 
| 5642 | 95  | 
sig  | 
96  | 
include THEORY  | 
|
| 6546 | 97  | 
val init_data: Object.kind -> (Object.T * (Object.T -> Object.T) * (Object.T -> Object.T) *  | 
| 5642 | 98  | 
(Object.T * Object.T -> Object.T) * (Sign.sg -> Object.T -> unit)) -> theory -> theory  | 
99  | 
val print_data: Object.kind -> theory -> unit  | 
|
100  | 
val get_data: Object.kind -> (Object.T -> 'a) -> theory -> 'a  | 
|
101  | 
  val put_data: Object.kind -> ('a -> Object.T) -> 'a -> theory -> theory
 | 
|
102  | 
end;  | 
|
| 1526 | 103  | 
|
| 5642 | 104  | 
|
| 6188 | 105  | 
structure Theory: PRIVATE_THEORY =  | 
| 1526 | 106  | 
struct  | 
| 2206 | 107  | 
|
| 3996 | 108  | 
|
| 2206 | 109  | 
(** datatype theory **)  | 
| 1526 | 110  | 
|
111  | 
datatype theory =  | 
|
112  | 
  Theory of {
 | 
|
113  | 
sign: Sign.sg,  | 
|
| 3996 | 114  | 
axioms: term Symtab.table,  | 
| 4996 | 115  | 
oracles: ((Sign.sg * Object.T -> term) * stamp) Symtab.table,  | 
| 4019 | 116  | 
parents: theory list,  | 
117  | 
ancestors: theory list};  | 
|
| 1526 | 118  | 
|
| 4996 | 119  | 
fun make_theory sign axms oras parents ancestors =  | 
| 4019 | 120  | 
  Theory {sign = sign, axioms = axms, oracles = oras,
 | 
121  | 
parents = parents, ancestors = ancestors};  | 
|
| 3996 | 122  | 
|
| 1526 | 123  | 
fun rep_theory (Theory args) = args;  | 
124  | 
||
| 3996 | 125  | 
val sign_of = #sign o rep_theory;  | 
| 6188 | 126  | 
val is_draft = Sign.is_draft o sign_of;  | 
| 3996 | 127  | 
val syn_of = #syn o Sign.rep_sg o sign_of;  | 
128  | 
val parents_of = #parents o rep_theory;  | 
|
| 4019 | 129  | 
val ancestors_of = #ancestors o rep_theory;  | 
| 3996 | 130  | 
|
| 1526 | 131  | 
(*errors involving theories*)  | 
132  | 
exception THEORY of string * theory list;  | 
|
133  | 
||
134  | 
(*compare theories*)  | 
|
135  | 
val subthy = Sign.subsig o pairself sign_of;  | 
|
136  | 
val eq_thy = Sign.eq_sg o pairself sign_of;  | 
|
137  | 
||
| 6369 | 138  | 
(*check for some theory*)  | 
| 4970 | 139  | 
fun requires thy name what =  | 
| 4846 | 140  | 
if exists (equal name) (Sign.stamp_names_of (sign_of thy)) then ()  | 
141  | 
  else error ("Require theory " ^ quote name ^ " as an ancestor for " ^ what);
 | 
|
| 1526 | 142  | 
|
| 6369 | 143  | 
fun assert_super thy1 thy2 =  | 
144  | 
if subthy (thy1, thy2) then thy2  | 
|
145  | 
  else raise THEORY ("Not a super theory", [thy1, thy2]);
 | 
|
146  | 
||
| 4846 | 147  | 
(*partial Pure theory*)  | 
| 4996 | 148  | 
val pre_pure = make_theory Sign.pre_pure Symtab.empty Symtab.empty [] [];  | 
| 1526 | 149  | 
|
150  | 
||
151  | 
||
152  | 
(** extend theory **)  | 
|
153  | 
||
| 8725 | 154  | 
(*name spaces*)  | 
| 3996 | 155  | 
val axiomK = "axiom";  | 
| 3814 | 156  | 
val oracleK = "oracle";  | 
157  | 
||
158  | 
||
159  | 
(* extend logical part of a theory *)  | 
|
160  | 
||
| 1526 | 161  | 
fun err_dup_axms names =  | 
162  | 
  error ("Duplicate axiom name(s) " ^ commas_quote names);
 | 
|
163  | 
||
| 3814 | 164  | 
fun err_dup_oras names =  | 
165  | 
  error ("Duplicate oracles " ^ commas_quote names);
 | 
|
166  | 
||
| 4996 | 167  | 
fun ext_theory thy ext_sg new_axms new_oras =  | 
| 1526 | 168  | 
let  | 
| 4019 | 169  | 
    val Theory {sign, axioms, oracles, parents, ancestors} = thy;
 | 
| 1526 | 170  | 
val draft = Sign.is_draft sign;  | 
| 3996 | 171  | 
val axioms' =  | 
| 4488 | 172  | 
Symtab.extend (if draft then axioms else Symtab.empty, new_axms)  | 
| 1526 | 173  | 
handle Symtab.DUPS names => err_dup_axms names;  | 
| 3814 | 174  | 
val oracles' =  | 
| 4488 | 175  | 
Symtab.extend (oracles, new_oras)  | 
| 3814 | 176  | 
handle Symtab.DUPS names => err_dup_oras names;  | 
| 4019 | 177  | 
val (parents', ancestors') =  | 
178  | 
if draft then (parents, ancestors) else ([thy], thy :: ancestors);  | 
|
| 1526 | 179  | 
in  | 
| 4996 | 180  | 
make_theory (ext_sg sign) axioms' oracles' parents' ancestors'  | 
| 1526 | 181  | 
end;  | 
182  | 
||
183  | 
||
184  | 
(* extend signature of a theory *)  | 
|
185  | 
||
| 4996 | 186  | 
fun ext_sign extfun decls thy = ext_theory thy (extfun decls) [] [];  | 
| 1526 | 187  | 
|
| 4996 | 188  | 
val add_classes = ext_sign Sign.add_classes;  | 
189  | 
val add_classes_i = ext_sign Sign.add_classes_i;  | 
|
190  | 
val add_classrel = ext_sign Sign.add_classrel;  | 
|
191  | 
val add_classrel_i = ext_sign Sign.add_classrel_i;  | 
|
192  | 
val add_defsort = ext_sign Sign.add_defsort;  | 
|
193  | 
val add_defsort_i = ext_sign Sign.add_defsort_i;  | 
|
194  | 
val add_types = ext_sign Sign.add_types;  | 
|
195  | 
val add_nonterminals = ext_sign Sign.add_nonterminals;  | 
|
196  | 
val add_tyabbrs = ext_sign Sign.add_tyabbrs;  | 
|
197  | 
val add_tyabbrs_i = ext_sign Sign.add_tyabbrs_i;  | 
|
198  | 
val add_arities = ext_sign Sign.add_arities;  | 
|
199  | 
val add_arities_i = ext_sign Sign.add_arities_i;  | 
|
200  | 
val add_consts = ext_sign Sign.add_consts;  | 
|
201  | 
val add_consts_i = ext_sign Sign.add_consts_i;  | 
|
202  | 
val add_syntax = ext_sign Sign.add_syntax;  | 
|
203  | 
val add_syntax_i = ext_sign Sign.add_syntax_i;  | 
|
204  | 
val add_modesyntax = curry (ext_sign Sign.add_modesyntax);  | 
|
205  | 
val add_modesyntax_i = curry (ext_sign Sign.add_modesyntax_i);  | 
|
206  | 
val add_trfuns = ext_sign Sign.add_trfuns;  | 
|
207  | 
val add_trfunsT = ext_sign Sign.add_trfunsT;  | 
|
208  | 
val add_tokentrfuns = ext_sign Sign.add_tokentrfuns;  | 
|
| 6311 | 209  | 
fun add_mode_tokentrfuns m = add_tokentrfuns o map (fn (s, f) => (m, s, f));  | 
| 4996 | 210  | 
val add_trrules = ext_sign Sign.add_trrules;  | 
211  | 
val add_trrules_i = ext_sign Sign.add_trrules_i;  | 
|
212  | 
val add_path = ext_sign Sign.add_path;  | 
|
| 4846 | 213  | 
val parent_path = add_path "..";  | 
214  | 
val root_path = add_path "/";  | 
|
| 4996 | 215  | 
val add_space = ext_sign Sign.add_space;  | 
| 8725 | 216  | 
val hide_space = ext_sign Sign.hide_space;  | 
217  | 
val hide_space_i = ext_sign Sign.hide_space_i;  | 
|
218  | 
val hide_classes = curry hide_space_i Sign.classK;  | 
|
219  | 
val hide_types = curry hide_space_i Sign.typeK;  | 
|
220  | 
val hide_consts = curry hide_space_i Sign.constK;  | 
|
| 4996 | 221  | 
val add_name = ext_sign Sign.add_name;  | 
| 6546 | 222  | 
val copy = ext_sign (K Sign.copy) ();  | 
| 6661 | 223  | 
val prep_ext = ext_sign (K Sign.prep_ext) ();  | 
| 3814 | 224  | 
|
| 3996 | 225  | 
|
| 6311 | 226  | 
|
| 3814 | 227  | 
(** add axioms **)  | 
228  | 
||
| 1526 | 229  | 
(* prepare axioms *)  | 
230  | 
||
231  | 
fun err_in_axm name =  | 
|
232  | 
  error ("The error(s) above occurred in axiom " ^ quote name);
 | 
|
233  | 
||
234  | 
fun no_vars tm =  | 
|
235  | 
if null (term_vars tm) andalso null (term_tvars tm) then tm  | 
|
236  | 
else error "Illegal schematic variable(s) in term";  | 
|
237  | 
||
238  | 
fun cert_axm sg (name, raw_tm) =  | 
|
239  | 
let  | 
|
240  | 
val (t, T, _) = Sign.certify_term sg raw_tm  | 
|
| 2979 | 241  | 
handle TYPE (msg, _, _) => error msg  | 
| 3996 | 242  | 
| TERM (msg, _) => error msg;  | 
| 1526 | 243  | 
in  | 
244  | 
assert (T = propT) "Term not of type prop";  | 
|
245  | 
(name, no_vars t)  | 
|
246  | 
end  | 
|
247  | 
handle ERROR => err_in_axm name;  | 
|
248  | 
||
| 5057 | 249  | 
(*some duplication of code with read_def_cterm*)  | 
| 6661 | 250  | 
fun read_def_axm (sg, types, sorts) used (name, str) =  | 
| 3814 | 251  | 
let  | 
252  | 
val ts = Syntax.read (#syn (Sign.rep_sg sg)) propT str;  | 
|
| 5057 | 253  | 
val (t, _) = Sign.infer_types sg types sorts used true (ts, propT);  | 
| 
1960
 
ae390b599213
Improved error handling: if there are syntax or type-checking
 
paulson 
parents: 
1539 
diff
changeset
 | 
254  | 
in cert_axm sg (name,t) end  | 
| 
 
ae390b599213
Improved error handling: if there are syntax or type-checking
 
paulson 
parents: 
1539 
diff
changeset
 | 
255  | 
handle ERROR => err_in_axm name;  | 
| 1526 | 256  | 
|
| 5057 | 257  | 
fun read_axm sg name_str = read_def_axm (sg, K None, K None) [] name_str;  | 
258  | 
||
| 1526 | 259  | 
fun inferT_axm sg (name, pre_tm) =  | 
| 3814 | 260  | 
let  | 
| 4251 | 261  | 
val (t, _) = Sign.infer_types sg (K None) (K None) [] true ([pre_tm], propT);  | 
| 3814 | 262  | 
in (name, no_vars t) end  | 
| 1526 | 263  | 
handle ERROR => err_in_axm name;  | 
264  | 
||
265  | 
||
266  | 
(* extend axioms of a theory *)  | 
|
267  | 
||
| 3814 | 268  | 
fun ext_axms prep_axm raw_axms (thy as Theory {sign, ...}) =
 | 
| 1526 | 269  | 
let  | 
| 3814 | 270  | 
val raw_axms' = map (apfst (Sign.full_name sign)) raw_axms;  | 
271  | 
val axioms =  | 
|
272  | 
map (apsnd (Term.compress_term o Logic.varify) o prep_axm sign) raw_axms';  | 
|
| 4996 | 273  | 
val ext_sg = Sign.add_space (axiomK, map fst axioms);  | 
| 1526 | 274  | 
in  | 
| 4996 | 275  | 
ext_theory thy ext_sg axioms []  | 
| 1526 | 276  | 
end;  | 
277  | 
||
278  | 
val add_axioms = ext_axms read_axm;  | 
|
279  | 
val add_axioms_i = ext_axms cert_axm;  | 
|
280  | 
||
281  | 
||
| 3814 | 282  | 
(* add oracle **)  | 
283  | 
||
284  | 
fun add_oracle (raw_name, oracle) (thy as Theory {sign, ...}) =
 | 
|
285  | 
let  | 
|
286  | 
val name = Sign.full_name sign raw_name;  | 
|
| 4996 | 287  | 
val ext_sg = Sign.add_space (oracleK, [name]);  | 
| 3814 | 288  | 
in  | 
| 4996 | 289  | 
ext_theory thy ext_sg [] [(name, (oracle, stamp ()))]  | 
| 3814 | 290  | 
end;  | 
291  | 
||
292  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
293  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
294  | 
(** add constant definitions **)  | 
| 
 
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  | 
(* all_axioms_of *)  | 
| 
 
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  | 
(*results may contain duplicates!*)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
299  | 
|
| 4019 | 300  | 
fun all_axioms_of thy =  | 
301  | 
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
 | 
302  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
303  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
304  | 
(* clash_types, clash_consts *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
305  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
306  | 
(*check if types have common instance (ignoring sorts)*)  | 
| 
 
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 clash_types ty1 ty2 =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
309  | 
let  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
310  | 
val ty1' = Type.varifyT ty1;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
311  | 
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
 | 
312  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
313  | 
Type.raw_unify (ty1', ty2')  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
314  | 
end;  | 
| 
 
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  | 
fun clash_consts (c1, ty1) (c2, ty2) =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
317  | 
c1 = c2 andalso clash_types ty1 ty2;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
318  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
319  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
320  | 
(* clash_defns *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
321  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
322  | 
fun clash_defn c_ty (name, tm) =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
323  | 
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
 | 
324  | 
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
 | 
325  | 
end handle TERM _ => None;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
326  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
327  | 
fun clash_defns c_ty axms =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
328  | 
distinct (mapfilter (clash_defn c_ty) axms);  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
329  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
330  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
331  | 
(* dest_defn *)  | 
| 
 
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  | 
fun dest_defn tm =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
334  | 
let  | 
| 
3787
 
67571f49ebe3
new internal forms: add_classes_i, add_classrel_i, add_defsort_i, add_arities_i
 
wenzelm 
parents: 
3767 
diff
changeset
 | 
335  | 
fun err msg = raise TERM (msg, [tm]);  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
336  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
337  | 
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
 | 
338  | 
handle TERM _ => err "Not a meta-equality (==)";  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
339  | 
val (head, args) = strip_comb lhs;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
340  | 
val (c, ty) = dest_Const head  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
341  | 
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
 | 
342  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
343  | 
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
 | 
344  | 
| occs_const (Abs (_, _, t)) = occs_const t  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
345  | 
| 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
 | 
346  | 
| occs_const _ = false;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
347  | 
|
| 4141 | 348  | 
fun dest_free (Free (x, _)) = x  | 
349  | 
      | dest_free (Const ("TYPE", Type ("itself", [TFree (x, _)]))) = x
 | 
|
350  | 
| dest_free _ = raise Match;  | 
|
351  | 
||
352  | 
val show_frees = commas_quote o map dest_free;  | 
|
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
353  | 
val show_tfrees = commas_quote o map fst;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
354  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
355  | 
val lhs_dups = duplicates args;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
356  | 
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
 | 
357  | 
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
 | 
358  | 
in  | 
| 4141 | 359  | 
if not (forall (can dest_free) args) then  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
360  | 
err "Arguments (on lhs) must be variables"  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
361  | 
else if not (null lhs_dups) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
362  | 
      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
 | 
363  | 
else if not (null rhs_extras) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
364  | 
      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
 | 
365  | 
else if not (null rhs_extrasT) then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
366  | 
      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
 | 
367  | 
else if occs_const rhs then  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
368  | 
      err ("Constant to be defined occurs on rhs")
 | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
369  | 
else (c, ty)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
370  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
371  | 
|
| 
 
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  | 
(* check_defn *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
374  | 
|
| 3814 | 375  | 
fun err_in_defn sg name msg =  | 
| 4974 | 376  | 
  (error_msg msg; error ("The error(s) above occurred in definition " ^
 | 
| 3814 | 377  | 
quote (Sign.full_name sg name)));  | 
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
378  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
379  | 
fun check_defn sign (axms, (name, tm)) =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
380  | 
let  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
381  | 
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
 | 
382  | 
[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
 | 
383  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
384  | 
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
 | 
385  | 
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
 | 
386  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
387  | 
val (c, ty) = dest_defn tm  | 
| 3814 | 388  | 
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
 | 
389  | 
val defns = clash_defns (c, ty) axms;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
390  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
391  | 
if not (null defns) then  | 
| 3814 | 392  | 
      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
 | 
393  | 
"\nclashes with " ^ show_defns c defns)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
394  | 
else (name, tm) :: axms  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
395  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
396  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
397  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
398  | 
(* add_defs *)  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
399  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
400  | 
fun ext_defns prep_axm raw_axms thy =  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
401  | 
let  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
402  | 
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
 | 
403  | 
val all_axms = all_axioms_of thy;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
404  | 
in  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
405  | 
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
 | 
406  | 
add_axioms_i axms thy  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
407  | 
end;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
408  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
409  | 
val add_defs_i = ext_defns cert_axm;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
410  | 
val add_defs = ext_defns read_axm;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
411  | 
|
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
412  | 
|
| 3878 | 413  | 
|
| 3865 | 414  | 
(** additional theory data **)  | 
415  | 
||
| 4996 | 416  | 
val init_data = curry (ext_sign Sign.init_data);  | 
417  | 
fun print_data kind = Sign.print_data kind o sign_of;  | 
|
418  | 
fun get_data kind f = Sign.get_data kind f o sign_of;  | 
|
419  | 
fun put_data kind f = ext_sign (Sign.put_data kind f);  | 
|
| 3865 | 420  | 
|
421  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
422  | 
|
| 6661 | 423  | 
(** merge theories **) (*exception ERROR*)  | 
| 4019 | 424  | 
|
425  | 
fun merge_sign (sg, thy) =  | 
|
| 4627 | 426  | 
Sign.nontriv_merge (sg, sign_of thy) handle TERM (msg, _) => error msg;  | 
| 1526 | 427  | 
|
| 3878 | 428  | 
(*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
 | 
429  | 
fun prep_ext_merge thys =  | 
| 3996 | 430  | 
if null thys then  | 
| 4019 | 431  | 
error "Merge: no parent theories"  | 
| 6188 | 432  | 
else if exists is_draft thys then  | 
| 4019 | 433  | 
error "Attempt to merge draft theories"  | 
| 3996 | 434  | 
else  | 
435  | 
let  | 
|
436  | 
val sign' =  | 
|
| 4019 | 437  | 
foldl merge_sign (sign_of (hd thys), tl thys)  | 
| 3996 | 438  | 
|> Sign.prep_ext  | 
439  | 
|> Sign.add_path "/";  | 
|
| 3814 | 440  | 
|
| 4488 | 441  | 
val axioms' = Symtab.empty;  | 
| 4019 | 442  | 
|
| 3996 | 443  | 
fun eq_ora ((_, (_, s1: stamp)), (_, (_, s2))) = s1 = s2;  | 
444  | 
val oracles' =  | 
|
445  | 
Symtab.make (gen_distinct eq_ora  | 
|
446  | 
(flat (map (Symtab.dest o #oracles o rep_theory) thys)))  | 
|
447  | 
handle Symtab.DUPS names => err_dup_oras names;  | 
|
| 4019 | 448  | 
|
449  | 
val parents' = gen_distinct eq_thy thys;  | 
|
450  | 
val ancestors' =  | 
|
451  | 
gen_distinct eq_thy (parents' @ flat (map ancestors_of thys));  | 
|
| 3996 | 452  | 
in  | 
| 4996 | 453  | 
make_theory sign' axioms' oracles' parents' ancestors'  | 
| 3996 | 454  | 
end;  | 
| 1526 | 455  | 
|
| 3885 | 456  | 
|
| 1526 | 457  | 
end;  | 
458  | 
||
| 
3767
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
459  | 
structure BasicTheory: BASIC_THEORY = Theory;  | 
| 
 
e2bb53d8dd26
moved theory stuff (add_defs etc.) here from drule.ML;
 
wenzelm 
parents: 
2979 
diff
changeset
 | 
460  | 
open BasicTheory;  |