12014
|
1 |
(* Title: Pure/Isar/locale.ML
|
11896
|
2 |
ID: $Id$
|
|
3 |
Author: Markus Wenzel, TU Muenchen
|
|
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE)
|
|
5 |
|
12058
|
6 |
Locales -- Isar proof contexts as meta-level predicates, with local
|
12289
|
7 |
syntax and implicit structures. Draws some basic ideas from Florian
|
|
8 |
Kammüller's original version of locales, but uses the richer
|
12063
|
9 |
infrastructure of Isar instead of the raw meta-logic.
|
11896
|
10 |
*)
|
|
11 |
|
|
12 |
signature LOCALE =
|
|
13 |
sig
|
12046
|
14 |
type context
|
|
15 |
datatype ('typ, 'term, 'fact, 'att) elem =
|
12058
|
16 |
Fixes of (string * 'typ option * mixfix option) list |
|
12046
|
17 |
Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
|
|
18 |
Defines of ((string * 'att list) * ('term * 'term list)) list |
|
12273
|
19 |
Notes of ((string * 'att list) * ('fact * 'att list) list) list
|
|
20 |
datatype expr =
|
|
21 |
Locale of string |
|
|
22 |
Rename of expr * string option list |
|
|
23 |
Merge of expr list
|
|
24 |
val empty: expr
|
|
25 |
datatype ('typ, 'term, 'fact, 'att) elem_expr =
|
|
26 |
Elem of ('typ, 'term, 'fact, 'att) elem | Expr of expr
|
12046
|
27 |
type 'att element
|
|
28 |
type 'att element_i
|
|
29 |
type locale
|
|
30 |
val intern: Sign.sg -> xstring -> string
|
12014
|
31 |
val cond_extern: Sign.sg -> string -> xstring
|
12502
|
32 |
val the_locale: theory -> string -> locale
|
12273
|
33 |
val attribute: ('att -> context attribute) -> ('typ, 'term, 'thm, 'att) elem_expr
|
|
34 |
-> ('typ, 'term, 'thm, context attribute) elem_expr
|
12502
|
35 |
val activate_context: context -> expr * context attribute element list ->
|
|
36 |
(context -> context) * (context -> context)
|
|
37 |
val activate_context_i: context -> expr * context attribute element_i list ->
|
|
38 |
(context -> context) * (context -> context)
|
12273
|
39 |
val add_locale: bstring -> expr -> context attribute element list -> theory -> theory
|
|
40 |
val add_locale_i: bstring -> expr -> context attribute element_i list -> theory -> theory
|
12289
|
41 |
val print_locales: theory -> unit
|
|
42 |
val print_locale: theory -> expr -> unit
|
12273
|
43 |
val add_thmss: string -> ((string * thm list) * context attribute list) list -> theory -> theory
|
11896
|
44 |
val setup: (theory -> theory) list
|
|
45 |
end;
|
|
46 |
|
12289
|
47 |
structure Locale: LOCALE =
|
11896
|
48 |
struct
|
|
49 |
|
|
50 |
|
12502
|
51 |
|
12273
|
52 |
(** locale elements and expressions **)
|
11896
|
53 |
|
12014
|
54 |
type context = ProofContext.context;
|
11896
|
55 |
|
12046
|
56 |
datatype ('typ, 'term, 'fact, 'att) elem =
|
12058
|
57 |
Fixes of (string * 'typ option * mixfix option) list |
|
12046
|
58 |
Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
|
|
59 |
Defines of ((string * 'att list) * ('term * 'term list)) list |
|
12273
|
60 |
Notes of ((string * 'att list) * ('fact * 'att list) list) list;
|
|
61 |
|
|
62 |
datatype expr =
|
|
63 |
Locale of string |
|
|
64 |
Rename of expr * string option list |
|
|
65 |
Merge of expr list;
|
11896
|
66 |
|
12273
|
67 |
val empty = Merge [];
|
|
68 |
|
|
69 |
datatype ('typ, 'term, 'fact, 'att) elem_expr =
|
|
70 |
Elem of ('typ, 'term, 'fact, 'att) elem | Expr of expr;
|
|
71 |
|
|
72 |
type 'att element = (string, string, string, 'att) elem_expr;
|
|
73 |
type 'att element_i = (typ, term, thm list, 'att) elem_expr;
|
12070
|
74 |
|
|
75 |
type locale =
|
12289
|
76 |
{import: expr, (*dynamic import*)
|
|
77 |
elems: ((typ, term, thm list, context attribute) elem * stamp) list, (*static content*)
|
12502
|
78 |
params: (string * typ option) list * string list, (*all vs. local params*)
|
12289
|
79 |
text: (string * typ) list * term list} (*logical representation*)
|
12063
|
80 |
|
12289
|
81 |
fun make_locale import elems params text =
|
|
82 |
{import = import, elems = elems, params = params, text = text}: locale;
|
12063
|
83 |
|
11896
|
84 |
|
|
85 |
|
|
86 |
(** theory data **)
|
|
87 |
|
|
88 |
structure LocalesArgs =
|
|
89 |
struct
|
12014
|
90 |
val name = "Isar/locales";
|
12063
|
91 |
type T = NameSpace.T * locale Symtab.table;
|
11896
|
92 |
|
12063
|
93 |
val empty = (NameSpace.empty, Symtab.empty);
|
|
94 |
val copy = I;
|
12118
|
95 |
val prep_ext = I;
|
12289
|
96 |
|
|
97 |
(*joining of locale elements: only facts may be added later!*)
|
|
98 |
fun join ({import, elems, params, text}: locale, {elems = elems', ...}: locale) =
|
|
99 |
Some (make_locale import (gen_merge_lists eq_snd elems elems') params text);
|
12273
|
100 |
fun merge ((space1, locs1), (space2, locs2)) =
|
12289
|
101 |
(NameSpace.merge (space1, space2), Symtab.join join (locs1, locs2));
|
|
102 |
|
12273
|
103 |
fun print _ (space, locs) =
|
|
104 |
Pretty.strs ("locales:" :: map (NameSpace.cond_extern space o #1) (Symtab.dest locs))
|
12014
|
105 |
|> Pretty.writeln;
|
11896
|
106 |
end;
|
|
107 |
|
|
108 |
structure LocalesData = TheoryDataFun(LocalesArgs);
|
|
109 |
val print_locales = LocalesData.print;
|
|
110 |
|
12289
|
111 |
val intern = NameSpace.intern o #1 o LocalesData.get_sg;
|
|
112 |
val cond_extern = NameSpace.cond_extern o #1 o LocalesData.get_sg;
|
|
113 |
|
12277
|
114 |
|
|
115 |
(* access locales *)
|
|
116 |
|
12063
|
117 |
fun declare_locale name =
|
|
118 |
LocalesData.map (apfst (fn space => (NameSpace.extend (space, [name]))));
|
11896
|
119 |
|
12273
|
120 |
fun put_locale name loc = LocalesData.map (apsnd (fn locs => Symtab.update ((name, loc), locs)));
|
12063
|
121 |
fun get_locale thy name = Symtab.lookup (#2 (LocalesData.get thy), name);
|
11896
|
122 |
|
12014
|
123 |
fun the_locale thy name =
|
|
124 |
(case get_locale thy name of
|
|
125 |
Some loc => loc
|
|
126 |
| None => error ("Unknown locale " ^ quote name));
|
11896
|
127 |
|
12046
|
128 |
|
12277
|
129 |
(* diagnostics *)
|
12273
|
130 |
|
12277
|
131 |
fun err_in_locale ctxt msg ids =
|
|
132 |
let
|
12289
|
133 |
fun prt_id (name, parms) = [Pretty.block (Pretty.breaks (map Pretty.str (name :: parms)))];
|
|
134 |
val prt_ids = flat (separate [Pretty.str " +", Pretty.brk 1] (map prt_id ids));
|
12502
|
135 |
val err_msg =
|
|
136 |
if null ids then msg
|
|
137 |
else msg ^ "\n" ^ Pretty.string_of (Pretty.block
|
|
138 |
(Pretty.str "The error(s) above occurred in locale:" :: Pretty.brk 1 :: prt_ids));
|
|
139 |
in raise ProofContext.CONTEXT (err_msg, ctxt) end;
|
12063
|
140 |
|
12277
|
141 |
|
|
142 |
|
|
143 |
(** operations on locale elements **)
|
|
144 |
|
12289
|
145 |
(* prepare elements *)
|
12263
|
146 |
|
12307
|
147 |
local
|
|
148 |
|
|
149 |
fun prep_name ctxt (name, atts) =
|
|
150 |
if NameSpace.is_qualified name then
|
|
151 |
raise ProofContext.CONTEXT ("Illegal qualified name: " ^ quote name, ctxt)
|
|
152 |
else (name, atts);
|
|
153 |
|
|
154 |
fun prep_elem prep_vars prep_propp prep_thms ctxt =
|
12063
|
155 |
fn Fixes fixes =>
|
12307
|
156 |
let val vars = snd (foldl_map prep_vars (ctxt, map (fn (x, T, _) => ([x], T)) fixes))
|
|
157 |
in Fixes (map (fn (([x'], T'), (_, _, mx)) => (x', T', mx)) (vars ~~ fixes)) end
|
12063
|
158 |
| Assumes asms =>
|
12307
|
159 |
Assumes (map (prep_name ctxt o #1) asms ~~ snd (prep_propp (ctxt, map #2 asms)))
|
12063
|
160 |
| Defines defs =>
|
12307
|
161 |
let val propps = snd (prep_propp (ctxt, map (fn (_, (t, ps)) => [(t, (ps, []))]) defs)) in
|
|
162 |
Defines (map (prep_name ctxt o #1) defs ~~ map (fn [(t', (ps', []))] => (t', ps')) propps)
|
|
163 |
end
|
12063
|
164 |
| Notes facts =>
|
12307
|
165 |
Notes (map (fn (a, bs) => (prep_name ctxt a, map (apfst (prep_thms ctxt)) bs)) facts);
|
|
166 |
|
12502
|
167 |
in
|
|
168 |
|
12307
|
169 |
fun read_elem x = prep_elem ProofContext.read_vars ProofContext.read_propp ProofContext.get_thms x;
|
|
170 |
fun cert_elem x = prep_elem ProofContext.cert_vars ProofContext.cert_propp (K I) x;
|
|
171 |
|
12273
|
172 |
fun read_expr ctxt (Locale xname) = Locale (intern (ProofContext.sign_of ctxt) xname)
|
|
173 |
| read_expr ctxt (Merge exprs) = Merge (map (read_expr ctxt) exprs)
|
12289
|
174 |
| read_expr ctxt (Rename (expr, xs)) = Rename (read_expr ctxt expr, xs);
|
12273
|
175 |
|
|
176 |
fun read_element ctxt (Elem e) = Elem (read_elem ctxt e)
|
|
177 |
| read_element ctxt (Expr e) = Expr (read_expr ctxt e);
|
12063
|
178 |
|
12307
|
179 |
fun cert_element ctxt (Elem e) = Elem (cert_elem ctxt e)
|
|
180 |
| cert_element ctxt (Expr e) = Expr e;
|
12063
|
181 |
|
12502
|
182 |
end;
|
|
183 |
|
|
184 |
|
|
185 |
(* internalize attributes *)
|
|
186 |
|
|
187 |
local fun read_att attrib (x, srcs) = (x, map attrib srcs) in
|
|
188 |
|
12273
|
189 |
fun attribute _ (Elem (Fixes fixes)) = Elem (Fixes fixes)
|
12277
|
190 |
| attribute attrib (Elem (Assumes asms)) = Elem (Assumes (map (apfst (read_att attrib)) asms))
|
|
191 |
| attribute attrib (Elem (Defines defs)) = Elem (Defines (map (apfst (read_att attrib)) defs))
|
12273
|
192 |
| attribute attrib (Elem (Notes facts)) =
|
12277
|
193 |
Elem (Notes (map (apfst (read_att attrib) o apsnd (map (read_att attrib))) facts))
|
12273
|
194 |
| attribute _ (Expr expr) = Expr expr;
|
12063
|
195 |
|
|
196 |
end;
|
12046
|
197 |
|
|
198 |
|
12277
|
199 |
(* renaming *)
|
12263
|
200 |
|
|
201 |
fun rename ren x = if_none (assoc_string (ren, x)) x;
|
|
202 |
|
|
203 |
fun rename_term ren (Free (x, T)) = Free (rename ren x, T)
|
|
204 |
| rename_term ren (t $ u) = rename_term ren t $ rename_term ren u
|
|
205 |
| rename_term ren (Abs (x, T, t)) = Abs (x, T, rename_term ren t)
|
|
206 |
| rename_term _ a = a;
|
|
207 |
|
|
208 |
fun rename_thm ren th =
|
|
209 |
let
|
|
210 |
val {sign, hyps, prop, maxidx, ...} = Thm.rep_thm th;
|
|
211 |
val cert = Thm.cterm_of sign;
|
12502
|
212 |
val (xs, Ts) = Library.split_list (foldl Term.add_frees ([], prop :: hyps));
|
12263
|
213 |
val xs' = map (rename ren) xs;
|
|
214 |
fun cert_frees names = map (cert o Free) (names ~~ Ts);
|
|
215 |
fun cert_vars names = map (cert o Var o apfst (rpair (maxidx + 1))) (names ~~ Ts);
|
|
216 |
in
|
|
217 |
if xs = xs' then th
|
|
218 |
else
|
|
219 |
th
|
|
220 |
|> Drule.implies_intr_list (map cert hyps)
|
|
221 |
|> Drule.forall_intr_list (cert_frees xs)
|
|
222 |
|> Drule.forall_elim_list (cert_vars xs)
|
|
223 |
|> Thm.instantiate ([], cert_vars xs ~~ cert_frees xs')
|
|
224 |
|> (fn th' => Drule.implies_elim_list th' (map (Thm.assume o cert o rename_term ren) hyps))
|
|
225 |
end;
|
|
226 |
|
|
227 |
fun rename_elem ren (Fixes fixes) = Fixes (map (fn (x, T, mx) =>
|
12273
|
228 |
(rename ren x, T, if mx = None then mx else Some Syntax.NoSyn)) fixes) (*drops syntax!*)
|
12263
|
229 |
| rename_elem ren (Assumes asms) = Assumes (map (apsnd (map (fn (t, (ps, qs)) =>
|
|
230 |
(rename_term ren t, (map (rename_term ren) ps, map (rename_term ren) qs))))) asms)
|
|
231 |
| rename_elem ren (Defines defs) = Defines (map (apsnd (fn (t, ps) =>
|
|
232 |
(rename_term ren t, map (rename_term ren) ps))) defs)
|
12273
|
233 |
| rename_elem ren (Notes facts) = Notes (map (apsnd (map (apfst (map (rename_thm ren))))) facts);
|
12263
|
234 |
|
12307
|
235 |
fun qualify_elem prfx elem =
|
|
236 |
let
|
12323
|
237 |
fun qualify (arg as ((name, atts), x)) =
|
|
238 |
if name = "" then arg
|
|
239 |
else ((NameSpace.pack (filter_out (equal "") (prfx @ [name])), atts), x);
|
12307
|
240 |
in
|
|
241 |
(case elem of
|
|
242 |
Fixes fixes => Fixes fixes
|
|
243 |
| Assumes asms => Assumes (map qualify asms)
|
|
244 |
| Defines defs => Defines (map qualify defs)
|
|
245 |
| Notes facts => Notes (map qualify facts))
|
|
246 |
end;
|
|
247 |
|
12263
|
248 |
|
12502
|
249 |
(* type instantiation *)
|
|
250 |
|
|
251 |
fun inst_type [] T = T
|
|
252 |
| inst_type env T = Term.map_type_tfree (fn v => if_none (assoc (env, v)) (TFree v)) T;
|
|
253 |
|
|
254 |
fun inst_term [] t = t
|
|
255 |
| inst_term env t = Term.map_term_types (inst_type env) t;
|
|
256 |
|
|
257 |
fun inst_thm [] th = th
|
|
258 |
| inst_thm env th =
|
|
259 |
let
|
|
260 |
val {sign, hyps, prop, maxidx, ...} = Thm.rep_thm th;
|
|
261 |
val cert = Thm.cterm_of sign and certT = Thm.ctyp_of sign;
|
|
262 |
val names = foldr Term.add_term_tfree_names (prop :: hyps, []);
|
|
263 |
val env' = filter (fn ((a, _), _) => a mem_string names) env;
|
|
264 |
in
|
|
265 |
if null env' then th
|
|
266 |
else
|
|
267 |
th
|
|
268 |
|> Drule.implies_intr_list (map cert hyps)
|
|
269 |
|> Drule.tvars_intr_list names
|
|
270 |
|> (fn (th', al) => th' |>
|
|
271 |
Thm.instantiate ((map (fn ((a, _), T) => (the (assoc (al, a)), certT T)) env'), []))
|
|
272 |
|> (fn th'' => Drule.implies_elim_list th''
|
|
273 |
(map (Thm.assume o cert o inst_term env') hyps))
|
|
274 |
end;
|
|
275 |
|
|
276 |
fun inst_elem env (Fixes fixes) =
|
|
277 |
Fixes (map (fn (x, T, mx) => (x, apsome (inst_type env) T, mx)) fixes)
|
|
278 |
| inst_elem env (Assumes asms) = Assumes (map (apsnd (map (fn (t, (ps, qs)) =>
|
|
279 |
(inst_term env t, (map (inst_term env) ps, map (inst_term env) qs))))) asms)
|
|
280 |
| inst_elem env (Defines defs) = Defines (map (apsnd (fn (t, ps) =>
|
|
281 |
(inst_term env t, map (inst_term env) ps))) defs)
|
|
282 |
| inst_elem env (Notes facts) = Notes (map (apsnd (map (apfst (map (inst_thm env))))) facts);
|
|
283 |
|
|
284 |
|
12277
|
285 |
(* evaluation *)
|
11896
|
286 |
|
12502
|
287 |
fun frozen_tvars ctxt Ts =
|
|
288 |
let
|
|
289 |
val tvars = rev (foldl Term.add_tvarsT ([], Ts));
|
|
290 |
val tfrees = map TFree
|
|
291 |
(Term.invent_type_names (ProofContext.used_types ctxt) (length tvars) ~~ map #2 tvars);
|
|
292 |
in map #1 tvars ~~ tfrees end;
|
|
293 |
|
|
294 |
fun unify_parms ctxt raw_parmss =
|
|
295 |
let
|
|
296 |
val tsig = Sign.tsig_of (ProofContext.sign_of ctxt);
|
|
297 |
val maxidx = length raw_parmss;
|
|
298 |
val idx_parmss = (0 upto maxidx - 1) ~~ raw_parmss;
|
|
299 |
|
|
300 |
fun varify i = Term.map_type_tfree (fn (a, S) => TVar ((a, i), S));
|
|
301 |
fun varify_parms (i, ps) =
|
|
302 |
mapfilter (fn (_, None) => None | (x, Some T) => Some (x, varify i T)) ps;
|
|
303 |
val parms = flat (map varify_parms idx_parmss);
|
|
304 |
|
|
305 |
fun unify T ((env, maxidx), U) = Type.unify tsig maxidx env (U, T); (*should never fail*)
|
|
306 |
fun unify_list (envir, T :: Us) = foldl (unify T) (envir, Us)
|
|
307 |
| unify_list (envir, []) = envir;
|
|
308 |
val (unifier, _) = foldl unify_list
|
|
309 |
((Vartab.empty, maxidx), map #2 (Symtab.dest (Symtab.make_multi parms)));
|
|
310 |
|
|
311 |
val parms' = map (apsnd (Envir.norm_type unifier)) (gen_distinct eq_fst parms);
|
|
312 |
val unifier' = Vartab.extend (unifier, frozen_tvars ctxt (map #2 parms'));
|
|
313 |
|
|
314 |
fun inst_parms (i, ps) =
|
|
315 |
foldr Term.add_typ_tfrees (mapfilter snd ps, [])
|
|
316 |
|> mapfilter (fn (a, S) =>
|
|
317 |
let val T = Envir.norm_type unifier' (TVar ((a, i), S))
|
|
318 |
in if T = TFree (a, S) then None else Some ((a, S), T) end);
|
|
319 |
in map inst_parms idx_parmss end;
|
|
320 |
|
|
321 |
fun unique_parms ctxt elemss =
|
|
322 |
let
|
|
323 |
val param_decls =
|
|
324 |
flat (map (fn ((name, (ps, qs)), _) => map (rpair (name, ps)) qs) elemss)
|
|
325 |
|> Symtab.make_multi |> Symtab.dest;
|
|
326 |
in
|
|
327 |
(case find_first (fn (_, ids) => length ids > 1) param_decls of
|
|
328 |
Some (q, ids) => err_in_locale ctxt ("Multiple declaration of parameter " ^ quote q)
|
|
329 |
(map (apsnd (map fst)) ids)
|
|
330 |
| None => map (apfst (apsnd #1)) elemss)
|
|
331 |
end;
|
|
332 |
|
|
333 |
fun inst_types _ [elems] = [elems]
|
|
334 |
| inst_types ctxt elemss =
|
|
335 |
let
|
|
336 |
val envs = unify_parms ctxt (map (#2 o #1) elemss);
|
|
337 |
fun inst (((name, ps), elems), env) =
|
|
338 |
((name, map (apsnd (apsome (inst_type env))) ps), map (inst_elem env) elems);
|
|
339 |
in map inst (elemss ~~ envs) end;
|
|
340 |
|
|
341 |
|
12273
|
342 |
fun eval_expr ctxt expr =
|
12014
|
343 |
let
|
|
344 |
val thy = ProofContext.theory_of ctxt;
|
12263
|
345 |
|
12289
|
346 |
fun renaming (Some x :: xs) (y :: ys) = (y, x) :: renaming xs ys
|
|
347 |
| renaming (None :: xs) (y :: ys) = renaming xs ys
|
12273
|
348 |
| renaming [] _ = []
|
12289
|
349 |
| renaming xs [] = raise ERROR_MESSAGE ("Too many arguments in renaming: " ^
|
|
350 |
commas (map (fn None => "_" | Some x => quote x) xs));
|
|
351 |
|
|
352 |
fun rename_parms ren (name, ps) =
|
|
353 |
let val ps' = map (rename ren) ps in
|
|
354 |
(case duplicates ps' of [] => (name, ps')
|
|
355 |
| dups => err_in_locale ctxt ("Duplicate parameters: " ^ commas_quote dups) [(name, ps')])
|
|
356 |
end;
|
12263
|
357 |
|
12273
|
358 |
fun identify ((ids, parms), Locale name) =
|
12289
|
359 |
let
|
|
360 |
val {import, params, ...} = the_locale thy name;
|
|
361 |
val ps = map #1 (#1 params);
|
|
362 |
in
|
12273
|
363 |
if (name, ps) mem ids then (ids, parms)
|
12277
|
364 |
else
|
12289
|
365 |
let val (ids', parms') = identify ((ids, parms), import); (*acyclic dependencies!*)
|
|
366 |
in (ids' @ [(name, ps)], merge_lists parms' ps) end
|
12273
|
367 |
end
|
|
368 |
| identify ((ids, parms), Rename (e, xs)) =
|
|
369 |
let
|
|
370 |
val (ids', parms') = identify (([], []), e);
|
12289
|
371 |
val ren = renaming xs parms' handle ERROR_MESSAGE msg => err_in_locale ctxt msg ids';
|
|
372 |
val ids'' = distinct (map (rename_parms ren) ids');
|
|
373 |
val parms'' = distinct (flat (map #2 ids''));
|
|
374 |
in (merge_lists ids ids'', merge_lists parms parms'') end
|
12273
|
375 |
| identify (arg, Merge es) = foldl identify (arg, es);
|
12014
|
376 |
|
12307
|
377 |
fun eval (name, xs) =
|
12273
|
378 |
let
|
12502
|
379 |
val {params = (ps, qs), elems, ...} = the_locale thy name;
|
12307
|
380 |
val ren = filter_out (op =) (map #1 ps ~~ xs);
|
12502
|
381 |
val (params', elems') =
|
|
382 |
if null ren then ((ps, qs), map #1 elems)
|
|
383 |
else ((map (apfst (rename ren)) ps, map (rename ren) qs),
|
|
384 |
map (rename_elem ren o #1) elems);
|
|
385 |
val elems'' = map (qualify_elem [NameSpace.base name, space_implode "_" xs]) elems';
|
|
386 |
in ((name, params'), elems'') end;
|
12307
|
387 |
|
12502
|
388 |
val raw_elemss = unique_parms ctxt (map eval (#1 (identify (([], []), expr))));
|
|
389 |
val elemss = inst_types ctxt raw_elemss;
|
|
390 |
in elemss end;
|
12046
|
391 |
|
12070
|
392 |
|
12273
|
393 |
|
|
394 |
(** activation **)
|
12070
|
395 |
|
12502
|
396 |
(* internalize elems *)
|
|
397 |
|
|
398 |
fun declare_elem gen =
|
|
399 |
let
|
|
400 |
val gen_typ = if gen then Term.map_type_tfree (Type.param []) else I;
|
|
401 |
val gen_term = if gen then Term.map_term_types gen_typ else I;
|
|
402 |
|
|
403 |
fun declare (Fixes fixes) = ProofContext.add_syntax fixes o
|
|
404 |
ProofContext.fix_direct (map (fn (x, T, _) => ([x], apsome gen_typ T)) fixes)
|
|
405 |
| declare (Assumes asms) = (fn ctxt => #1 (ProofContext.bind_propp_i
|
|
406 |
(ctxt, map (map (fn (t, (ps, ps')) =>
|
|
407 |
(gen_term t, (map gen_term ps, map gen_term ps'))) o #2) asms)))
|
|
408 |
| declare (Defines defs) = (fn ctxt => #1 (ProofContext.bind_propp_i
|
|
409 |
(ctxt, map (fn (_, (t, ps)) => [(gen_term t, (map gen_term ps, []))]) defs)))
|
|
410 |
| declare (Notes _) = I;
|
|
411 |
in declare end;
|
12307
|
412 |
|
12273
|
413 |
fun activate_elem (Fixes fixes) = ProofContext.add_syntax fixes o
|
12502
|
414 |
ProofContext.fix_direct (map (fn (x, T, _) => ([x], T)) fixes)
|
12273
|
415 |
| activate_elem (Assumes asms) =
|
|
416 |
#1 o ProofContext.assume_i ProofContext.export_assume asms o
|
|
417 |
ProofContext.fix_frees (flat (map (map #1 o #2) asms))
|
|
418 |
| activate_elem (Defines defs) = (fn ctxt => #1 (ProofContext.assume_i ProofContext.export_def
|
|
419 |
(map (fn ((name, atts), (t, ps)) =>
|
|
420 |
let val (c, t') = ProofContext.cert_def ctxt t
|
|
421 |
in ((if name = "" then Thm.def_name c else name, atts), [(t', (ps, []))]) end) defs) ctxt))
|
|
422 |
| activate_elem (Notes facts) = #1 o ProofContext.have_thmss facts;
|
12263
|
423 |
|
12273
|
424 |
|
12502
|
425 |
fun perform_elems f named_elems = ProofContext.qualified (fn context =>
|
|
426 |
foldl (fn (ctxt, ((name, ps), es)) =>
|
|
427 |
foldl (fn (c, e) => f e c) (ctxt, es) handle ProofContext.CONTEXT (msg, ctxt) =>
|
12307
|
428 |
err_in_locale ctxt msg [(name, map fst ps)]) (context, named_elems));
|
|
429 |
|
12502
|
430 |
fun declare_elemss gen = perform_elems (declare_elem gen);
|
|
431 |
fun activate_elemss x = perform_elems activate_elem x;
|
|
432 |
|
12307
|
433 |
|
12502
|
434 |
(* context specifications: import expression + external elements *)
|
|
435 |
|
|
436 |
local
|
|
437 |
|
|
438 |
fun close_frees ctxt t =
|
|
439 |
let val frees = rev (filter_out (ProofContext.is_fixed ctxt o #1) (Term.add_frees ([], t)))
|
|
440 |
in Term.list_all_free (frees, t) end;
|
|
441 |
|
|
442 |
(*quantify dangling frees, strip term bindings*)
|
|
443 |
fun closeup ctxt (Assumes asms) = Assumes (asms |> map (fn (a, propps) =>
|
|
444 |
(a, map (fn (t, _) => (close_frees ctxt t, ([], []))) propps)))
|
|
445 |
| closeup ctxt (Defines defs) = Defines (defs |> map (fn (a, (t, _)) =>
|
|
446 |
(a, (close_frees ctxt (#2 (ProofContext.cert_def ctxt t)), []))))
|
|
447 |
| closeup ctxt elem = elem;
|
|
448 |
|
|
449 |
fun prepare_context prep_elem prep_expr close context (import, elements) =
|
|
450 |
let
|
|
451 |
fun prep_element (ctxt, Elem raw_elem) =
|
|
452 |
let val elem = (if close then closeup ctxt else I) (prep_elem ctxt raw_elem)
|
|
453 |
in (ctxt |> declare_elem false elem, [(("", []), [elem])]) end
|
|
454 |
| prep_element (ctxt, Expr raw_expr) =
|
|
455 |
let
|
|
456 |
val expr = prep_expr ctxt raw_expr;
|
|
457 |
val named_elemss = eval_expr ctxt expr;
|
|
458 |
in (ctxt |> declare_elemss true named_elemss, named_elemss) end;
|
12277
|
459 |
|
12502
|
460 |
val (import_ctxt, import_elemss) = prep_element (context, Expr import);
|
|
461 |
val (elements_ctxt, elements_elemss) =
|
|
462 |
apsnd flat (foldl_map prep_element (import_ctxt, elements));
|
|
463 |
|
|
464 |
val xs = flat (map (map #1 o (#2 o #1)) (import_elemss @ elements_elemss));
|
|
465 |
val env = frozen_tvars elements_ctxt (mapfilter (ProofContext.default_type elements_ctxt) xs);
|
|
466 |
|
|
467 |
fun inst_elems ((name, ps), elems) = ((name, ps), elems); (* FIXME *)
|
|
468 |
|
|
469 |
in (map inst_elems import_elemss, map inst_elems elements_elemss) end;
|
12273
|
470 |
|
12502
|
471 |
fun gen_activate_context prep_elem prep_expr ctxt args =
|
|
472 |
pairself activate_elemss (prepare_context prep_elem prep_expr false ctxt args);
|
|
473 |
|
|
474 |
in
|
|
475 |
|
|
476 |
val read_context = prepare_context read_elem read_expr true;
|
|
477 |
val cert_context = prepare_context cert_elem (K I) true;
|
|
478 |
val activate_context = gen_activate_context read_elem read_expr;
|
|
479 |
val activate_context_i = gen_activate_context cert_elem (K I);
|
|
480 |
|
|
481 |
fun activate_locale name ctxt =
|
|
482 |
#1 (activate_context_i ctxt (Locale name, [])) ctxt;
|
|
483 |
|
|
484 |
end;
|
11896
|
485 |
|
|
486 |
|
|
487 |
|
12070
|
488 |
(** print locale **)
|
|
489 |
|
12307
|
490 |
fun print_locale thy raw_expr =
|
12070
|
491 |
let
|
|
492 |
val sg = Theory.sign_of thy;
|
12289
|
493 |
val thy_ctxt = ProofContext.init thy;
|
12277
|
494 |
|
12502
|
495 |
val elemss = #1 (read_context thy_ctxt (raw_expr, []));
|
|
496 |
val ctxt = activate_elemss elemss thy_ctxt;
|
12070
|
497 |
|
12307
|
498 |
val prt_typ = Pretty.quote o ProofContext.pretty_typ ctxt;
|
|
499 |
val prt_term = Pretty.quote o ProofContext.pretty_term ctxt;
|
|
500 |
val prt_thm = Pretty.quote o ProofContext.pretty_thm ctxt;
|
12070
|
501 |
|
|
502 |
fun prt_syn syn =
|
|
503 |
let val s = (case syn of None => "(structure)" | Some mx => Syntax.string_of_mixfix mx)
|
|
504 |
in if s = "" then [] else [Pretty.brk 4, Pretty.str s] end;
|
|
505 |
fun prt_fix (x, Some T, syn) = Pretty.block (Pretty.str (x ^ " ::") :: Pretty.brk 1 ::
|
|
506 |
prt_typ T :: Pretty.brk 1 :: prt_syn syn)
|
|
507 |
| prt_fix (x, None, syn) = Pretty.block (Pretty.str x :: Pretty.brk 1 :: prt_syn syn);
|
|
508 |
|
12307
|
509 |
fun prt_name "" = [Pretty.brk 1]
|
|
510 |
| prt_name name = [Pretty.str (ProofContext.cond_extern ctxt name ^ ":"), Pretty.brk 1];
|
|
511 |
fun prt_asm ((a, _), ts) = Pretty.block (prt_name a @ Pretty.breaks (map (prt_term o fst) ts));
|
|
512 |
fun prt_def ((a, _), (t, _)) = Pretty.block (prt_name a @ [prt_term t]);
|
|
513 |
fun prt_fact ((a, _), ths) = Pretty.block
|
|
514 |
(prt_name a @ Pretty.breaks (map prt_thm (flat (map fst ths))));
|
12070
|
515 |
|
12289
|
516 |
fun items _ [] = []
|
|
517 |
| items prfx (x :: xs) = Pretty.block [Pretty.str prfx, Pretty.brk 1, x] :: items " and" xs;
|
|
518 |
fun prt_elem (Fixes fixes) = items "fixes" (map prt_fix fixes)
|
|
519 |
| prt_elem (Assumes asms) = items "assumes" (map prt_asm asms)
|
|
520 |
| prt_elem (Defines defs) = items "defines" (map prt_def defs)
|
|
521 |
| prt_elem (Notes facts) = items "notes" (map prt_fact facts);
|
12277
|
522 |
in
|
12502
|
523 |
Pretty.big_list "locale elements:" (map (Pretty.chunks o prt_elem) (flat (map #2 elemss)))
|
12307
|
524 |
|> Pretty.writeln
|
12277
|
525 |
end;
|
12070
|
526 |
|
|
527 |
|
|
528 |
|
11896
|
529 |
(** define locales **)
|
|
530 |
|
12063
|
531 |
(* add_locale(_i) *)
|
11896
|
532 |
|
12502
|
533 |
local
|
|
534 |
|
|
535 |
fun gen_add_locale prep_context prep_expr bname raw_import raw_body thy =
|
12063
|
536 |
let
|
|
537 |
val sign = Theory.sign_of thy;
|
11896
|
538 |
val name = Sign.full_name sign bname;
|
12502
|
539 |
val _ = conditional (is_some (get_locale thy name)) (fn () =>
|
|
540 |
error ("Duplicate definition of locale " ^ quote name));
|
11896
|
541 |
|
12273
|
542 |
val thy_ctxt = ProofContext.init thy;
|
|
543 |
|
12502
|
544 |
val (import_elemss, body_elemss) = prep_context thy_ctxt (raw_import, raw_body);
|
12273
|
545 |
val import = prep_expr thy_ctxt raw_import;
|
12502
|
546 |
val import_elemss = eval_expr thy_ctxt import;
|
12273
|
547 |
|
12502
|
548 |
val import_ctxt = thy_ctxt |> activate_elemss import_elemss;
|
|
549 |
val body_ctxt = import_ctxt |> activate_elemss body_elemss;
|
12273
|
550 |
|
12502
|
551 |
val elems = flat (map #2 body_elemss);
|
|
552 |
val (import_parms, body_parms) = pairself (flat o map (#2 o #1)) (import_elemss, body_elemss);
|
12273
|
553 |
val text = ([], []); (* FIXME *)
|
12063
|
554 |
in
|
|
555 |
thy
|
|
556 |
|> declare_locale name
|
12289
|
557 |
|> put_locale name (make_locale import (map (fn e => (e, stamp ())) elems)
|
12502
|
558 |
(import_parms @ body_parms, map #1 body_parms) text)
|
12063
|
559 |
end;
|
|
560 |
|
12502
|
561 |
in
|
|
562 |
|
|
563 |
val add_locale = gen_add_locale read_context read_expr;
|
|
564 |
val add_locale_i = gen_add_locale cert_context (K I);
|
|
565 |
|
|
566 |
end;
|
12063
|
567 |
|
11896
|
568 |
|
12063
|
569 |
|
12070
|
570 |
(** store results **)
|
11896
|
571 |
|
12143
|
572 |
fun add_thmss name args thy =
|
12063
|
573 |
let
|
12289
|
574 |
val {import, params, elems, text} = the_locale thy name;
|
12143
|
575 |
val note = Notes (map (fn ((a, ths), atts) =>
|
|
576 |
((a, atts), [(map (curry Thm.name_thm a) ths, [])])) args);
|
12063
|
577 |
in
|
12502
|
578 |
thy |> ProofContext.init |> activate_locale name |> activate_elem note; (*test attributes!*)
|
12289
|
579 |
thy |> put_locale name (make_locale import (elems @ [(note, stamp ())]) params text)
|
12063
|
580 |
end;
|
11896
|
581 |
|
12063
|
582 |
|
11896
|
583 |
|
|
584 |
(** locale theory setup **)
|
12063
|
585 |
|
11896
|
586 |
val setup =
|
|
587 |
[LocalesData.init];
|
|
588 |
|
|
589 |
end;
|