1526
|
1 |
(* Title: Pure/theory.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson and Markus Wenzel
|
|
4 |
Copyright 1996 University of Cambridge
|
|
5 |
|
|
6 |
Theories
|
|
7 |
*)
|
|
8 |
|
|
9 |
signature THEORY =
|
|
10 |
sig
|
|
11 |
type theory
|
|
12 |
exception THEORY of string * theory list
|
|
13 |
val rep_theory : theory ->
|
1539
|
14 |
{sign: Sign.sg, oraopt: (Sign.sg * exn -> term) option,
|
|
15 |
new_axioms: term Symtab.table, parents: theory list}
|
1526
|
16 |
val sign_of : theory -> Sign.sg
|
|
17 |
val syn_of : theory -> Syntax.syntax
|
|
18 |
val stamps_of_thy : theory -> string ref list
|
|
19 |
val parents_of : theory -> theory list
|
|
20 |
val subthy : theory * theory -> bool
|
|
21 |
val eq_thy : theory * theory -> bool
|
|
22 |
val proto_pure_thy : theory
|
|
23 |
val pure_thy : theory
|
|
24 |
val cpure_thy : theory
|
|
25 |
(*theory primitives*)
|
1539
|
26 |
val add_classes : (class * class list) list -> theory -> theory
|
|
27 |
val add_classrel : (class * class) list -> theory -> theory
|
|
28 |
val add_defsort : sort -> theory -> theory
|
|
29 |
val add_types : (string * int * mixfix) list -> theory -> theory
|
|
30 |
val add_tyabbrs : (string * string list * string * mixfix) list
|
1526
|
31 |
-> theory -> theory
|
1539
|
32 |
val add_tyabbrs_i : (string * string list * typ * mixfix) list
|
1526
|
33 |
-> theory -> theory
|
1539
|
34 |
val add_arities : (string * sort list * sort) list -> theory -> theory
|
|
35 |
val add_consts : (string * string * mixfix) list -> theory -> theory
|
|
36 |
val add_consts_i : (string * typ * mixfix) list -> theory -> theory
|
|
37 |
val add_syntax : (string * string * mixfix) list -> theory -> theory
|
|
38 |
val add_syntax_i : (string * typ * mixfix) list -> theory -> theory
|
|
39 |
val add_trfuns :
|
1526
|
40 |
(string * (Syntax.ast list -> Syntax.ast)) list *
|
|
41 |
(string * (term list -> term)) list *
|
|
42 |
(string * (term list -> term)) list *
|
|
43 |
(string * (Syntax.ast list -> Syntax.ast)) list -> theory -> theory
|
|
44 |
val add_trrules : (string * string)Syntax.trrule list -> theory -> theory
|
1539
|
45 |
val add_trrules_i : Syntax.ast Syntax.trrule list -> theory -> theory
|
1526
|
46 |
val cert_axm : Sign.sg -> string * term -> string * term
|
|
47 |
val read_axm : Sign.sg -> string * string -> string * term
|
|
48 |
val inferT_axm : Sign.sg -> string * term -> string * term
|
1539
|
49 |
val add_axioms : (string * string) list -> theory -> theory
|
|
50 |
val add_axioms_i : (string * term) list -> theory -> theory
|
|
51 |
val add_thyname : string -> theory -> theory
|
|
52 |
|
|
53 |
val set_oracle : (Sign.sg * exn -> term) -> theory -> theory
|
1526
|
54 |
|
|
55 |
val merge_theories : theory * theory -> theory
|
|
56 |
val merge_thy_list : bool -> theory list -> theory
|
|
57 |
end;
|
|
58 |
|
|
59 |
|
|
60 |
structure Theory : THEORY =
|
|
61 |
struct
|
|
62 |
(*** Theories ***)
|
|
63 |
|
|
64 |
datatype theory =
|
|
65 |
Theory of {
|
|
66 |
sign: Sign.sg,
|
1539
|
67 |
oraopt: (Sign.sg * exn -> term) option,
|
1526
|
68 |
new_axioms: term Symtab.table,
|
|
69 |
parents: theory list};
|
|
70 |
|
|
71 |
fun rep_theory (Theory args) = args;
|
|
72 |
|
|
73 |
(*errors involving theories*)
|
|
74 |
exception THEORY of string * theory list;
|
|
75 |
|
|
76 |
|
|
77 |
val sign_of = #sign o rep_theory;
|
|
78 |
val syn_of = #syn o Sign.rep_sg o sign_of;
|
|
79 |
|
|
80 |
(*stamps associated with a theory*)
|
|
81 |
val stamps_of_thy = #stamps o Sign.rep_sg o sign_of;
|
|
82 |
|
|
83 |
(*return the immediate ancestors*)
|
|
84 |
val parents_of = #parents o rep_theory;
|
|
85 |
|
|
86 |
|
|
87 |
(*compare theories*)
|
|
88 |
val subthy = Sign.subsig o pairself sign_of;
|
|
89 |
val eq_thy = Sign.eq_sg o pairself sign_of;
|
|
90 |
|
|
91 |
|
|
92 |
(* the Pure theories *)
|
|
93 |
|
|
94 |
val proto_pure_thy =
|
1539
|
95 |
Theory {sign = Sign.proto_pure, oraopt = None,
|
|
96 |
new_axioms = Symtab.null, parents = []};
|
1526
|
97 |
|
|
98 |
val pure_thy =
|
1539
|
99 |
Theory {sign = Sign.pure, oraopt = None,
|
|
100 |
new_axioms = Symtab.null, parents = []};
|
1526
|
101 |
|
|
102 |
val cpure_thy =
|
1539
|
103 |
Theory {sign = Sign.cpure, oraopt = None,
|
|
104 |
new_axioms = Symtab.null, parents = []};
|
1526
|
105 |
|
|
106 |
|
|
107 |
|
|
108 |
(** extend theory **)
|
|
109 |
|
|
110 |
fun err_dup_axms names =
|
|
111 |
error ("Duplicate axiom name(s) " ^ commas_quote names);
|
|
112 |
|
1539
|
113 |
fun ext_thy (thy as Theory {sign, oraopt, new_axioms, parents})
|
|
114 |
sign1 new_axms =
|
1526
|
115 |
let
|
|
116 |
val draft = Sign.is_draft sign;
|
|
117 |
val new_axioms1 =
|
|
118 |
Symtab.extend_new (if draft then new_axioms else Symtab.null, new_axms)
|
|
119 |
handle Symtab.DUPS names => err_dup_axms names;
|
|
120 |
val parents1 = if draft then parents else [thy];
|
|
121 |
in
|
1539
|
122 |
Theory {sign = sign1, oraopt = oraopt,
|
|
123 |
new_axioms = new_axioms1, parents = parents1}
|
1526
|
124 |
end;
|
|
125 |
|
|
126 |
|
|
127 |
(* extend signature of a theory *)
|
|
128 |
|
|
129 |
fun ext_sg extfun decls (thy as Theory {sign, ...}) =
|
|
130 |
ext_thy thy (extfun decls sign) [];
|
|
131 |
|
|
132 |
val add_classes = ext_sg Sign.add_classes;
|
|
133 |
val add_classrel = ext_sg Sign.add_classrel;
|
|
134 |
val add_defsort = ext_sg Sign.add_defsort;
|
|
135 |
val add_types = ext_sg Sign.add_types;
|
|
136 |
val add_tyabbrs = ext_sg Sign.add_tyabbrs;
|
|
137 |
val add_tyabbrs_i = ext_sg Sign.add_tyabbrs_i;
|
|
138 |
val add_arities = ext_sg Sign.add_arities;
|
|
139 |
val add_consts = ext_sg Sign.add_consts;
|
|
140 |
val add_consts_i = ext_sg Sign.add_consts_i;
|
|
141 |
val add_syntax = ext_sg Sign.add_syntax;
|
|
142 |
val add_syntax_i = ext_sg Sign.add_syntax_i;
|
|
143 |
val add_trfuns = ext_sg Sign.add_trfuns;
|
|
144 |
val add_trrules = ext_sg Sign.add_trrules;
|
|
145 |
val add_trrules_i = ext_sg Sign.add_trrules_i;
|
|
146 |
val add_thyname = ext_sg Sign.add_name;
|
|
147 |
|
|
148 |
|
|
149 |
(* prepare axioms *)
|
|
150 |
|
|
151 |
fun err_in_axm name =
|
|
152 |
error ("The error(s) above occurred in axiom " ^ quote name);
|
|
153 |
|
|
154 |
fun no_vars tm =
|
|
155 |
if null (term_vars tm) andalso null (term_tvars tm) then tm
|
|
156 |
else error "Illegal schematic variable(s) in term";
|
|
157 |
|
|
158 |
fun cert_axm sg (name, raw_tm) =
|
|
159 |
let
|
|
160 |
val (t, T, _) = Sign.certify_term sg raw_tm
|
|
161 |
handle TYPE arg => error (Sign.exn_type_msg sg arg)
|
|
162 |
| TERM (msg, _) => error msg;
|
|
163 |
in
|
|
164 |
assert (T = propT) "Term not of type prop";
|
|
165 |
(name, no_vars t)
|
|
166 |
end
|
|
167 |
handle ERROR => err_in_axm name;
|
|
168 |
|
|
169 |
(*Some duplication of code with read_def_cterm*)
|
|
170 |
fun read_axm sg (name, str) =
|
|
171 |
let val ts = Syntax.read (#syn (Sign.rep_sg sg)) propT str;
|
|
172 |
val (_, t, _) =
|
|
173 |
Sign.infer_types sg (K None) (K None) [] true (ts,propT);
|
|
174 |
in cert_axm sg (name,t) end;
|
|
175 |
|
|
176 |
fun inferT_axm sg (name, pre_tm) =
|
|
177 |
let val t = #2(Sign.infer_types sg (K None) (K None) [] true
|
|
178 |
([pre_tm], propT))
|
|
179 |
in (name, no_vars t) end
|
|
180 |
handle ERROR => err_in_axm name;
|
|
181 |
|
|
182 |
|
|
183 |
(* extend axioms of a theory *)
|
|
184 |
|
|
185 |
fun ext_axms prep_axm axms (thy as Theory {sign, ...}) =
|
|
186 |
let
|
|
187 |
val sign1 = Sign.make_draft sign;
|
|
188 |
val axioms = map (apsnd (Term.compress_term o Logic.varify) o
|
|
189 |
prep_axm sign)
|
|
190 |
axms;
|
|
191 |
in
|
|
192 |
ext_thy thy sign1 axioms
|
|
193 |
end;
|
|
194 |
|
|
195 |
val add_axioms = ext_axms read_axm;
|
|
196 |
val add_axioms_i = ext_axms cert_axm;
|
|
197 |
|
|
198 |
|
1539
|
199 |
(** Set oracle of theory **)
|
|
200 |
|
|
201 |
fun set_oracle oracle
|
|
202 |
(thy as Theory {sign, oraopt = None, new_axioms, parents}) =
|
|
203 |
if Sign.is_draft sign then
|
|
204 |
Theory {sign = sign,
|
|
205 |
oraopt = Some oracle,
|
|
206 |
new_axioms = new_axioms,
|
|
207 |
parents = parents}
|
|
208 |
else raise THEORY ("Can only set oracle of a draft", [thy])
|
|
209 |
| set_oracle _ thy = raise THEORY ("Oracle already set", [thy]);
|
|
210 |
|
1526
|
211 |
|
|
212 |
(** merge theories **)
|
|
213 |
|
|
214 |
fun merge_thy_list mk_draft thys =
|
|
215 |
let
|
|
216 |
fun is_union thy = forall (fn t => subthy (t, thy)) thys;
|
|
217 |
val is_draft = Sign.is_draft o sign_of;
|
|
218 |
|
|
219 |
fun add_sign (sg, Theory {sign, ...}) =
|
|
220 |
Sign.merge (sg, sign) handle TERM (msg, _) => error msg;
|
|
221 |
in
|
1539
|
222 |
case (find_first is_union thys, exists is_draft thys) of
|
1526
|
223 |
(Some thy, _) => thy
|
|
224 |
| (None, true) => raise THEORY ("Illegal merge of draft theories", thys)
|
|
225 |
| (None, false) => Theory {
|
|
226 |
sign =
|
|
227 |
(if mk_draft then Sign.make_draft else I)
|
|
228 |
(foldl add_sign (Sign.proto_pure, thys)),
|
1539
|
229 |
oraopt = None,
|
1526
|
230 |
new_axioms = Symtab.null,
|
1539
|
231 |
parents = thys}
|
1526
|
232 |
end;
|
|
233 |
|
|
234 |
fun merge_theories (thy1, thy2) = merge_thy_list false [thy1, thy2];
|
|
235 |
|
|
236 |
|
|
237 |
end;
|
|
238 |
|
|
239 |
open Theory;
|