haftmann@29361
|
1 |
(* Title: Pure/Isar/locale.ML
|
haftmann@29361
|
2 |
Author: Clemens Ballarin, TU Muenchen
|
haftmann@29361
|
3 |
|
haftmann@29361
|
4 |
Locales -- Isar proof contexts as meta-level predicates, with local
|
haftmann@29361
|
5 |
syntax and implicit structures.
|
haftmann@29361
|
6 |
|
haftmann@29361
|
7 |
Draws basic ideas from Florian Kammueller's original version of
|
haftmann@29361
|
8 |
locales, but uses the richer infrastructure of Isar instead of the raw
|
haftmann@29361
|
9 |
meta-logic. Furthermore, structured import of contexts (with merge
|
haftmann@29361
|
10 |
and rename operations) are provided, as well as type-inference of the
|
haftmann@29361
|
11 |
signature parts, and predicate definitions of the specification text.
|
haftmann@29361
|
12 |
|
haftmann@29361
|
13 |
Interpretation enables the reuse of theorems of locales in other
|
haftmann@29361
|
14 |
contexts, namely those defined by theories, structured proofs and
|
haftmann@29361
|
15 |
locales themselves.
|
haftmann@29361
|
16 |
|
haftmann@29361
|
17 |
See also:
|
haftmann@29361
|
18 |
|
haftmann@29361
|
19 |
[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
|
haftmann@29361
|
20 |
In Stefano Berardi et al., Types for Proofs and Programs: International
|
haftmann@29361
|
21 |
Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
|
haftmann@29361
|
22 |
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
|
haftmann@29361
|
23 |
Dependencies between Locales. Technical Report TUM-I0607, Technische
|
haftmann@29361
|
24 |
Universitaet Muenchen, 2006.
|
haftmann@29361
|
25 |
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
|
haftmann@29361
|
26 |
Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
|
haftmann@29361
|
27 |
pages 31-43, 2006.
|
haftmann@29361
|
28 |
*)
|
haftmann@29361
|
29 |
|
haftmann@29361
|
30 |
signature LOCALE =
|
haftmann@29361
|
31 |
sig
|
haftmann@29361
|
32 |
type locale
|
haftmann@29361
|
33 |
|
haftmann@29361
|
34 |
val register_locale: bstring ->
|
haftmann@29361
|
35 |
(string * sort) list * (Binding.T * typ option * mixfix) list ->
|
haftmann@29361
|
36 |
term option * term list ->
|
haftmann@29441
|
37 |
thm option * thm option -> thm list ->
|
haftmann@29361
|
38 |
(declaration * stamp) list * (declaration * stamp) list ->
|
haftmann@29361
|
39 |
((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list ->
|
haftmann@29544
|
40 |
((string * morphism) * stamp) list -> theory -> theory
|
haftmann@29361
|
41 |
|
haftmann@29361
|
42 |
(* Locale name space *)
|
haftmann@29361
|
43 |
val intern: theory -> xstring -> string
|
haftmann@29361
|
44 |
val extern: theory -> string -> xstring
|
haftmann@29361
|
45 |
|
haftmann@29361
|
46 |
(* Specification *)
|
haftmann@29392
|
47 |
val defined: theory -> string -> bool
|
haftmann@29361
|
48 |
val params_of: theory -> string -> (Binding.T * typ option * mixfix) list
|
haftmann@29441
|
49 |
val intros_of: theory -> string -> thm option * thm option
|
haftmann@29441
|
50 |
val axioms_of: theory -> string -> thm list
|
haftmann@29544
|
51 |
val instance_of: theory -> string -> morphism -> term list
|
haftmann@29361
|
52 |
val specification_of: theory -> string -> term option * term list
|
haftmann@29361
|
53 |
val declarations_of: theory -> string -> declaration list * declaration list
|
haftmann@29544
|
54 |
val dependencies_of: theory -> string -> (string * morphism) list
|
haftmann@29361
|
55 |
|
haftmann@29361
|
56 |
(* Storing results *)
|
haftmann@29361
|
57 |
val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
|
haftmann@29361
|
58 |
Proof.context -> Proof.context
|
haftmann@29361
|
59 |
val add_type_syntax: string -> declaration -> Proof.context -> Proof.context
|
haftmann@29361
|
60 |
val add_term_syntax: string -> declaration -> Proof.context -> Proof.context
|
haftmann@29361
|
61 |
val add_declaration: string -> declaration -> Proof.context -> Proof.context
|
haftmann@29544
|
62 |
val add_dependency: string -> string * morphism -> theory -> theory
|
haftmann@29361
|
63 |
|
haftmann@29361
|
64 |
(* Activation *)
|
haftmann@29544
|
65 |
val activate_declarations: theory -> string * morphism ->
|
haftmann@29361
|
66 |
Proof.context -> Proof.context
|
haftmann@29544
|
67 |
val activate_global_facts: string * morphism -> theory -> theory
|
haftmann@29544
|
68 |
val activate_local_facts: string * morphism -> Proof.context -> Proof.context
|
haftmann@29361
|
69 |
val init: string -> theory -> Proof.context
|
haftmann@29361
|
70 |
|
haftmann@29361
|
71 |
(* Reasoning about locales *)
|
haftmann@29361
|
72 |
val witness_attrib: attribute
|
haftmann@29361
|
73 |
val intro_attrib: attribute
|
haftmann@29361
|
74 |
val unfold_attrib: attribute
|
haftmann@29361
|
75 |
val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
|
haftmann@29361
|
76 |
|
haftmann@29361
|
77 |
(* Registrations *)
|
haftmann@29544
|
78 |
val add_registration: string * (morphism * morphism) ->
|
haftmann@29361
|
79 |
theory -> theory
|
haftmann@29544
|
80 |
val amend_registration: morphism -> string * morphism ->
|
haftmann@29361
|
81 |
theory -> theory
|
haftmann@29544
|
82 |
val get_registrations: theory -> (string * morphism) list
|
haftmann@29361
|
83 |
|
haftmann@29361
|
84 |
(* Diagnostic *)
|
haftmann@29361
|
85 |
val print_locales: theory -> unit
|
haftmann@29361
|
86 |
val print_locale: theory -> bool -> bstring -> unit
|
haftmann@29361
|
87 |
end;
|
haftmann@29361
|
88 |
|
haftmann@29361
|
89 |
|
haftmann@29361
|
90 |
(*** Theorem list extensible via attribute --- to control intro_locales_tac ***)
|
haftmann@29361
|
91 |
|
haftmann@29361
|
92 |
(* FIXME: consider moving named_thms.ML up in the build hierarchy and using that *)
|
haftmann@29361
|
93 |
functor ThmsFun() =
|
haftmann@29361
|
94 |
struct
|
haftmann@29361
|
95 |
|
haftmann@29361
|
96 |
structure Data = GenericDataFun
|
haftmann@29361
|
97 |
(
|
haftmann@29361
|
98 |
type T = thm list;
|
haftmann@29361
|
99 |
val empty = [];
|
haftmann@29361
|
100 |
val extend = I;
|
haftmann@29361
|
101 |
fun merge _ = Thm.merge_thms;
|
haftmann@29361
|
102 |
);
|
haftmann@29361
|
103 |
|
haftmann@29361
|
104 |
val get = Data.get o Context.Proof;
|
haftmann@29361
|
105 |
val add = Thm.declaration_attribute (Data.map o Thm.add_thm);
|
haftmann@29361
|
106 |
|
haftmann@29361
|
107 |
end;
|
haftmann@29361
|
108 |
|
haftmann@29361
|
109 |
|
haftmann@29361
|
110 |
structure Locale: LOCALE =
|
haftmann@29361
|
111 |
struct
|
haftmann@29361
|
112 |
|
haftmann@29361
|
113 |
datatype ctxt = datatype Element.ctxt;
|
haftmann@29361
|
114 |
|
haftmann@29361
|
115 |
|
haftmann@29392
|
116 |
|
haftmann@29392
|
117 |
(*** Theory data ***)
|
haftmann@29361
|
118 |
|
haftmann@29361
|
119 |
datatype locale = Loc of {
|
haftmann@29392
|
120 |
(** static part **)
|
haftmann@29361
|
121 |
parameters: (string * sort) list * (Binding.T * typ option * mixfix) list,
|
haftmann@29361
|
122 |
(* type and term parameters *)
|
haftmann@29361
|
123 |
spec: term option * term list,
|
haftmann@29361
|
124 |
(* assumptions (as a single predicate expression) and defines *)
|
haftmann@29441
|
125 |
intros: thm option * thm option,
|
haftmann@29441
|
126 |
axioms: thm list,
|
haftmann@29392
|
127 |
(** dynamic part **)
|
haftmann@29361
|
128 |
decls: (declaration * stamp) list * (declaration * stamp) list,
|
haftmann@29361
|
129 |
(* type and term syntax declarations *)
|
haftmann@29361
|
130 |
notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * stamp) list,
|
haftmann@29361
|
131 |
(* theorem declarations *)
|
haftmann@29544
|
132 |
dependencies: ((string * morphism) * stamp) list
|
haftmann@29361
|
133 |
(* locale dependencies (sublocale relation) *)
|
haftmann@29392
|
134 |
};
|
haftmann@29361
|
135 |
|
haftmann@29441
|
136 |
fun mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies)) =
|
haftmann@29441
|
137 |
Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
|
haftmann@29441
|
138 |
decls = decls, notes = notes, dependencies = dependencies};
|
haftmann@29441
|
139 |
fun map_locale f (Loc {parameters, spec, intros, axioms, decls, notes, dependencies}) =
|
haftmann@29441
|
140 |
mk_locale (f ((parameters, spec, intros, axioms), ((decls, notes), dependencies)));
|
haftmann@29441
|
141 |
fun merge_locale (Loc {parameters, spec, intros, axioms, decls = (decls1, decls2),
|
haftmann@29441
|
142 |
notes, dependencies}, Loc {decls = (decls1', decls2'), notes = notes',
|
haftmann@29441
|
143 |
dependencies = dependencies', ...}) = mk_locale
|
haftmann@29441
|
144 |
((parameters, spec, intros, axioms),
|
haftmann@29441
|
145 |
(((merge (eq_snd op =) (decls1, decls1'), merge (eq_snd op =) (decls2, decls2')),
|
haftmann@29441
|
146 |
merge (eq_snd op =) (notes, notes')),
|
haftmann@29441
|
147 |
merge (eq_snd op =) (dependencies, dependencies')));
|
haftmann@29361
|
148 |
|
haftmann@29361
|
149 |
structure LocalesData = TheoryDataFun
|
haftmann@29361
|
150 |
(
|
haftmann@29392
|
151 |
type T = locale NameSpace.table;
|
haftmann@29361
|
152 |
val empty = NameSpace.empty_table;
|
haftmann@29361
|
153 |
val copy = I;
|
haftmann@29361
|
154 |
val extend = I;
|
haftmann@29392
|
155 |
fun merge _ = NameSpace.join_tables (K merge_locale);
|
haftmann@29361
|
156 |
);
|
haftmann@29361
|
157 |
|
haftmann@29361
|
158 |
val intern = NameSpace.intern o #1 o LocalesData.get;
|
haftmann@29361
|
159 |
val extern = NameSpace.extern o #1 o LocalesData.get;
|
haftmann@29361
|
160 |
|
haftmann@29392
|
161 |
fun get_locale thy = Symtab.lookup (#2 (LocalesData.get thy));
|
haftmann@29392
|
162 |
|
haftmann@29392
|
163 |
fun defined thy = is_some o get_locale thy;
|
haftmann@29361
|
164 |
|
haftmann@29361
|
165 |
fun the_locale thy name = case get_locale thy name
|
haftmann@29392
|
166 |
of SOME (Loc loc) => loc
|
haftmann@29361
|
167 |
| NONE => error ("Unknown locale " ^ quote name);
|
haftmann@29361
|
168 |
|
haftmann@29441
|
169 |
fun register_locale bname parameters spec intros axioms decls notes dependencies thy =
|
haftmann@29361
|
170 |
thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy) (Binding.name bname,
|
haftmann@29441
|
171 |
mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies))) #> snd);
|
haftmann@29361
|
172 |
|
haftmann@29392
|
173 |
fun change_locale name =
|
haftmann@29392
|
174 |
LocalesData.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
|
haftmann@29361
|
175 |
|
haftmann@29361
|
176 |
fun print_locales thy =
|
haftmann@29392
|
177 |
Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (LocalesData.get thy)))
|
haftmann@29392
|
178 |
|> Pretty.writeln;
|
haftmann@29361
|
179 |
|
haftmann@29361
|
180 |
|
haftmann@29361
|
181 |
(*** Primitive operations ***)
|
haftmann@29361
|
182 |
|
haftmann@29392
|
183 |
fun params_of thy = snd o #parameters o the_locale thy;
|
haftmann@29361
|
184 |
|
haftmann@29441
|
185 |
fun intros_of thy = #intros o the_locale thy;
|
haftmann@29441
|
186 |
|
haftmann@29441
|
187 |
fun axioms_of thy = #axioms o the_locale thy;
|
haftmann@29441
|
188 |
|
haftmann@29392
|
189 |
fun instance_of thy name morph = params_of thy name |>
|
haftmann@29392
|
190 |
map ((fn (b, T, _) => Free (Binding.base_name b, the T)) #> Morphism.term morph);
|
haftmann@29361
|
191 |
|
haftmann@29392
|
192 |
fun specification_of thy = #spec o the_locale thy;
|
haftmann@29361
|
193 |
|
haftmann@29392
|
194 |
fun declarations_of thy name = the_locale thy name |>
|
haftmann@29392
|
195 |
#decls |> apfst (map fst) |> apsnd (map fst);
|
haftmann@29361
|
196 |
|
haftmann@29544
|
197 |
fun dependencies_of thy name = the_locale thy name |>
|
haftmann@29544
|
198 |
#dependencies |> map fst;
|
haftmann@29544
|
199 |
|
haftmann@29361
|
200 |
|
haftmann@29361
|
201 |
(*** Activate context elements of locale ***)
|
haftmann@29361
|
202 |
|
haftmann@29361
|
203 |
(** Identifiers: activated locales in theory or proof context **)
|
haftmann@29361
|
204 |
|
haftmann@29361
|
205 |
type identifiers = (string * term list) list;
|
haftmann@29361
|
206 |
|
haftmann@29361
|
207 |
val empty = ([] : identifiers);
|
haftmann@29361
|
208 |
|
haftmann@29361
|
209 |
fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
|
haftmann@29361
|
210 |
|
haftmann@29361
|
211 |
local
|
haftmann@29361
|
212 |
|
haftmann@29361
|
213 |
datatype 'a delayed = Ready of 'a | ToDo of ('a delayed * 'a delayed);
|
haftmann@29361
|
214 |
|
haftmann@29361
|
215 |
structure IdentifiersData = GenericDataFun
|
haftmann@29361
|
216 |
(
|
haftmann@29361
|
217 |
type T = identifiers delayed;
|
haftmann@29361
|
218 |
val empty = Ready empty;
|
haftmann@29361
|
219 |
val extend = I;
|
haftmann@29361
|
220 |
fun merge _ = ToDo;
|
haftmann@29361
|
221 |
);
|
haftmann@29361
|
222 |
|
haftmann@29361
|
223 |
in
|
haftmann@29361
|
224 |
|
haftmann@29361
|
225 |
fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
|
haftmann@29361
|
226 |
| finish _ (Ready ids) = ids;
|
haftmann@29361
|
227 |
|
haftmann@29361
|
228 |
val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
|
haftmann@29361
|
229 |
(case IdentifiersData.get (Context.Theory thy) of
|
haftmann@29361
|
230 |
Ready _ => NONE |
|
haftmann@29361
|
231 |
ids => SOME (Context.theory_map (IdentifiersData.put (Ready (finish thy ids))) thy))
|
haftmann@29361
|
232 |
)));
|
haftmann@29361
|
233 |
|
haftmann@29361
|
234 |
fun get_global_idents thy =
|
haftmann@29361
|
235 |
let val (Ready ids) = (IdentifiersData.get o Context.Theory) thy in ids end;
|
haftmann@29361
|
236 |
val put_global_idents = Context.theory_map o IdentifiersData.put o Ready;
|
haftmann@29361
|
237 |
|
haftmann@29361
|
238 |
fun get_local_idents ctxt =
|
haftmann@29361
|
239 |
let val (Ready ids) = (IdentifiersData.get o Context.Proof) ctxt in ids end;
|
haftmann@29361
|
240 |
val put_local_idents = Context.proof_map o IdentifiersData.put o Ready;
|
haftmann@29361
|
241 |
|
haftmann@29361
|
242 |
end;
|
haftmann@29361
|
243 |
|
haftmann@29361
|
244 |
|
haftmann@29361
|
245 |
(** Resolve locale dependencies in a depth-first fashion **)
|
haftmann@29361
|
246 |
|
haftmann@29361
|
247 |
local
|
haftmann@29361
|
248 |
|
haftmann@29361
|
249 |
val roundup_bound = 120;
|
haftmann@29361
|
250 |
|
haftmann@29361
|
251 |
fun add thy depth (name, morph) (deps, marked) =
|
haftmann@29361
|
252 |
if depth > roundup_bound
|
haftmann@29361
|
253 |
then error "Roundup bound exceeded (sublocale relation probably not terminating)."
|
haftmann@29361
|
254 |
else
|
haftmann@29361
|
255 |
let
|
haftmann@29392
|
256 |
val {parameters = (_, params), dependencies, ...} = the_locale thy name;
|
haftmann@29361
|
257 |
val instance = instance_of thy name morph;
|
haftmann@29361
|
258 |
in
|
haftmann@29361
|
259 |
if member (ident_eq thy) marked (name, instance)
|
haftmann@29361
|
260 |
then (deps, marked)
|
haftmann@29361
|
261 |
else
|
haftmann@29361
|
262 |
let
|
haftmann@29361
|
263 |
val dependencies' =
|
haftmann@29361
|
264 |
map (fn ((name, morph'), _) => (name, morph' $> morph)) dependencies;
|
haftmann@29361
|
265 |
val marked' = (name, instance) :: marked;
|
haftmann@29361
|
266 |
val (deps', marked'') = fold_rev (add thy (depth + 1)) dependencies' ([], marked');
|
haftmann@29361
|
267 |
in
|
haftmann@29361
|
268 |
((name, morph) :: deps' @ deps, marked'')
|
haftmann@29361
|
269 |
end
|
haftmann@29361
|
270 |
end;
|
haftmann@29361
|
271 |
|
haftmann@29361
|
272 |
in
|
haftmann@29361
|
273 |
|
haftmann@29361
|
274 |
fun roundup thy activate_dep (name, morph) (marked, input) =
|
haftmann@29361
|
275 |
let
|
haftmann@29361
|
276 |
(* Find all dependencies incuding new ones (which are dependencies enriching
|
haftmann@29361
|
277 |
existing registrations). *)
|
haftmann@29361
|
278 |
val (dependencies, marked') = add thy 0 (name, morph) ([], empty);
|
haftmann@29361
|
279 |
(* Filter out exisiting fragments. *)
|
haftmann@29361
|
280 |
val dependencies' = filter_out (fn (name, morph) =>
|
haftmann@29361
|
281 |
member (ident_eq thy) marked (name, instance_of thy name morph)) dependencies;
|
haftmann@29361
|
282 |
in
|
haftmann@29361
|
283 |
(merge (ident_eq thy) (marked, marked'), input |> fold_rev (activate_dep thy) dependencies')
|
haftmann@29361
|
284 |
end;
|
haftmann@29361
|
285 |
|
haftmann@29361
|
286 |
end;
|
haftmann@29361
|
287 |
|
haftmann@29361
|
288 |
|
haftmann@29361
|
289 |
(* Declarations, facts and entire locale content *)
|
haftmann@29361
|
290 |
|
haftmann@29361
|
291 |
fun activate_decls thy (name, morph) ctxt =
|
haftmann@29361
|
292 |
let
|
haftmann@29392
|
293 |
val {decls = (typ_decls, term_decls), ...} = the_locale thy name;
|
haftmann@29361
|
294 |
in
|
haftmann@29361
|
295 |
ctxt |> fold_rev (fn (decl, _) => Context.proof_map (decl morph)) typ_decls |>
|
haftmann@29361
|
296 |
fold_rev (fn (decl, _) => Context.proof_map (decl morph)) term_decls
|
haftmann@29361
|
297 |
end;
|
haftmann@29361
|
298 |
|
haftmann@29361
|
299 |
fun activate_notes activ_elem transfer thy (name, morph) input =
|
haftmann@29361
|
300 |
let
|
haftmann@29392
|
301 |
val {notes, ...} = the_locale thy name;
|
haftmann@29361
|
302 |
fun activate ((kind, facts), _) input =
|
haftmann@29361
|
303 |
let
|
haftmann@29361
|
304 |
val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
|
haftmann@29361
|
305 |
in activ_elem (Notes (kind, facts')) input end;
|
haftmann@29361
|
306 |
in
|
haftmann@29361
|
307 |
fold_rev activate notes input
|
haftmann@29361
|
308 |
end;
|
haftmann@29361
|
309 |
|
haftmann@29361
|
310 |
fun activate_all name thy activ_elem transfer (marked, input) =
|
haftmann@29361
|
311 |
let
|
haftmann@29392
|
312 |
val {parameters = (_, params), spec = (asm, defs), ...} =
|
haftmann@29361
|
313 |
the_locale thy name;
|
haftmann@29361
|
314 |
in
|
haftmann@29361
|
315 |
input |>
|
haftmann@29361
|
316 |
(if not (null params) then activ_elem (Fixes params) else I) |>
|
haftmann@29361
|
317 |
(* FIXME type parameters *)
|
haftmann@29361
|
318 |
(if is_some asm then activ_elem (Assumes [(Attrib.empty_binding, [(the asm, [])])]) else I) |>
|
haftmann@29361
|
319 |
(if not (null defs)
|
haftmann@29361
|
320 |
then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
|
haftmann@29361
|
321 |
else I) |>
|
haftmann@29361
|
322 |
pair marked |> roundup thy (activate_notes activ_elem transfer) (name, Morphism.identity)
|
haftmann@29361
|
323 |
end;
|
haftmann@29361
|
324 |
|
haftmann@29361
|
325 |
|
haftmann@29361
|
326 |
(** Public activation functions **)
|
haftmann@29361
|
327 |
|
haftmann@29361
|
328 |
local
|
haftmann@29361
|
329 |
|
haftmann@29361
|
330 |
fun init_global_elem (Notes (kind, facts)) thy =
|
haftmann@29392
|
331 |
let
|
haftmann@29392
|
332 |
val facts' = Attrib.map_facts (Attrib.attribute_i thy) facts
|
haftmann@29392
|
333 |
in Old_Locale.global_note_qualified kind facts' thy |> snd end
|
haftmann@29361
|
334 |
|
haftmann@29361
|
335 |
fun init_local_elem (Fixes fixes) ctxt = ctxt |>
|
haftmann@29361
|
336 |
ProofContext.add_fixes_i fixes |> snd
|
haftmann@29361
|
337 |
| init_local_elem (Assumes assms) ctxt =
|
haftmann@29361
|
338 |
let
|
haftmann@29361
|
339 |
val assms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) assms
|
haftmann@29361
|
340 |
in
|
haftmann@29361
|
341 |
ctxt |> fold Variable.auto_fixes (maps (map fst o snd) assms') |>
|
haftmann@29361
|
342 |
ProofContext.add_assms_i Assumption.assume_export assms' |> snd
|
haftmann@29361
|
343 |
end
|
haftmann@29361
|
344 |
| init_local_elem (Defines defs) ctxt =
|
haftmann@29361
|
345 |
let
|
haftmann@29361
|
346 |
val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs
|
haftmann@29361
|
347 |
in
|
haftmann@29361
|
348 |
ctxt |> fold Variable.auto_fixes (map (fst o snd) defs') |>
|
haftmann@29361
|
349 |
ProofContext.add_assms_i LocalDefs.def_export (map (fn (attn, t) => (attn, [t])) defs') |>
|
haftmann@29361
|
350 |
snd
|
haftmann@29361
|
351 |
end
|
haftmann@29361
|
352 |
| init_local_elem (Notes (kind, facts)) ctxt =
|
haftmann@29361
|
353 |
let
|
haftmann@29361
|
354 |
val facts' = Attrib.map_facts (Attrib.attribute_i (ProofContext.theory_of ctxt)) facts
|
haftmann@29361
|
355 |
in Old_Locale.local_note_qualified kind facts' ctxt |> snd end
|
haftmann@29361
|
356 |
|
haftmann@29361
|
357 |
fun cons_elem false (Notes notes) elems = elems
|
haftmann@29361
|
358 |
| cons_elem _ elem elems = elem :: elems
|
haftmann@29361
|
359 |
|
haftmann@29361
|
360 |
in
|
haftmann@29361
|
361 |
|
haftmann@29361
|
362 |
fun activate_declarations thy dep ctxt =
|
haftmann@29502
|
363 |
roundup thy activate_decls dep (get_local_idents ctxt, ctxt) |-> put_local_idents;
|
haftmann@29361
|
364 |
|
haftmann@29361
|
365 |
fun activate_global_facts dep thy =
|
haftmann@29361
|
366 |
roundup thy (activate_notes init_global_elem Element.transfer_morphism)
|
haftmann@29502
|
367 |
dep (get_global_idents thy, thy) |-> put_global_idents;
|
haftmann@29361
|
368 |
|
haftmann@29361
|
369 |
fun activate_local_facts dep ctxt =
|
haftmann@29361
|
370 |
roundup (ProofContext.theory_of ctxt)
|
haftmann@29361
|
371 |
(activate_notes init_local_elem (Element.transfer_morphism o ProofContext.theory_of)) dep
|
haftmann@29502
|
372 |
(get_local_idents ctxt, ctxt) |-> put_local_idents;
|
haftmann@29361
|
373 |
|
haftmann@29361
|
374 |
fun init name thy =
|
haftmann@29361
|
375 |
activate_all name thy init_local_elem (Element.transfer_morphism o ProofContext.theory_of)
|
haftmann@29502
|
376 |
(empty, ProofContext.init thy) |-> put_local_idents;
|
haftmann@29361
|
377 |
|
haftmann@29361
|
378 |
fun print_locale thy show_facts name =
|
haftmann@29361
|
379 |
let
|
haftmann@29361
|
380 |
val name' = intern thy name;
|
haftmann@29361
|
381 |
val ctxt = init name' thy
|
haftmann@29361
|
382 |
in
|
haftmann@29361
|
383 |
Pretty.big_list "locale elements:"
|
haftmann@29361
|
384 |
(activate_all name' thy (cons_elem show_facts) (K (Element.transfer_morphism thy))
|
haftmann@29361
|
385 |
(empty, []) |> snd |> rev |>
|
haftmann@29361
|
386 |
map (Element.pretty_ctxt ctxt) |> map Pretty.chunks) |> Pretty.writeln
|
haftmann@29361
|
387 |
end
|
haftmann@29361
|
388 |
|
haftmann@29361
|
389 |
end;
|
haftmann@29361
|
390 |
|
haftmann@29361
|
391 |
|
haftmann@29361
|
392 |
(*** Registrations: interpretations in theories ***)
|
haftmann@29361
|
393 |
|
haftmann@29392
|
394 |
structure RegistrationsData = TheoryDataFun
|
haftmann@29361
|
395 |
(
|
haftmann@29544
|
396 |
type T = ((string * (morphism * morphism)) * stamp) list;
|
haftmann@29392
|
397 |
(* FIXME mixins need to be stamped *)
|
haftmann@29361
|
398 |
(* registrations, in reverse order of declaration *)
|
haftmann@29361
|
399 |
val empty = [];
|
haftmann@29361
|
400 |
val extend = I;
|
haftmann@29392
|
401 |
val copy = I;
|
haftmann@29361
|
402 |
fun merge _ data : T = Library.merge (eq_snd op =) data;
|
haftmann@29361
|
403 |
(* FIXME consolidate with dependencies, consider one data slot only *)
|
haftmann@29361
|
404 |
);
|
haftmann@29361
|
405 |
|
haftmann@29392
|
406 |
val get_registrations =
|
haftmann@29392
|
407 |
RegistrationsData.get #> map fst #> map (apsnd op $>);
|
haftmann@29361
|
408 |
|
haftmann@29392
|
409 |
fun add_registration (name, (base_morph, export)) thy =
|
haftmann@29392
|
410 |
roundup thy (fn _ => fn (name', morph') =>
|
haftmann@29392
|
411 |
(RegistrationsData.map o cons) ((name', (morph', export)), stamp ()))
|
haftmann@29502
|
412 |
(name, base_morph) (get_global_idents thy, thy) |> snd
|
haftmann@29502
|
413 |
(* FIXME |-> put_global_idents ?*);
|
haftmann@29361
|
414 |
|
haftmann@29392
|
415 |
fun amend_registration morph (name, base_morph) thy =
|
haftmann@29361
|
416 |
let
|
haftmann@29392
|
417 |
val regs = (RegistrationsData.get #> map fst) thy;
|
haftmann@29361
|
418 |
val base = instance_of thy name base_morph;
|
haftmann@29361
|
419 |
fun match (name', (morph', _)) =
|
haftmann@29361
|
420 |
name = name' andalso eq_list (op aconv) (base, instance_of thy name' morph');
|
haftmann@29361
|
421 |
val i = find_index match (rev regs);
|
haftmann@29392
|
422 |
val _ = if i = ~1 then error ("No registration of locale " ^
|
haftmann@29361
|
423 |
quote (extern thy name) ^ " and parameter instantiation " ^
|
haftmann@29361
|
424 |
space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available.")
|
haftmann@29361
|
425 |
else ();
|
haftmann@29361
|
426 |
in
|
haftmann@29392
|
427 |
RegistrationsData.map (nth_map (length regs - 1 - i)
|
haftmann@29361
|
428 |
(fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy
|
haftmann@29361
|
429 |
end;
|
haftmann@29361
|
430 |
|
haftmann@29361
|
431 |
|
haftmann@29502
|
432 |
|
haftmann@29361
|
433 |
(*** Storing results ***)
|
haftmann@29361
|
434 |
|
haftmann@29361
|
435 |
(* Theorems *)
|
haftmann@29361
|
436 |
|
haftmann@29361
|
437 |
fun add_thmss loc kind args ctxt =
|
haftmann@29361
|
438 |
let
|
haftmann@29361
|
439 |
val ([Notes args'], ctxt') = Element.activate_i [Notes (kind, args)] ctxt;
|
haftmann@29361
|
440 |
val ctxt'' = ctxt' |> ProofContext.theory (
|
haftmann@29392
|
441 |
(change_locale loc o apfst o apsnd) (cons (args', stamp ()))
|
haftmann@29392
|
442 |
#>
|
haftmann@29361
|
443 |
(* Registrations *)
|
haftmann@29361
|
444 |
(fn thy => fold_rev (fn (name, morph) =>
|
haftmann@29361
|
445 |
let
|
haftmann@29361
|
446 |
val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
|
haftmann@29361
|
447 |
Attrib.map_facts (Attrib.attribute_i thy)
|
haftmann@29361
|
448 |
in Old_Locale.global_note_qualified kind args'' #> snd end)
|
haftmann@29392
|
449 |
(get_registrations thy |> filter (fn (name, _) => name = loc)) thy))
|
haftmann@29361
|
450 |
in ctxt'' end;
|
haftmann@29361
|
451 |
|
haftmann@29361
|
452 |
|
haftmann@29361
|
453 |
(* Declarations *)
|
haftmann@29361
|
454 |
|
haftmann@29361
|
455 |
local
|
haftmann@29361
|
456 |
|
haftmann@29361
|
457 |
fun decl_attrib decl phi = Thm.declaration_attribute (K (decl phi));
|
haftmann@29361
|
458 |
|
haftmann@29361
|
459 |
fun add_decls add loc decl =
|
haftmann@29392
|
460 |
ProofContext.theory ((change_locale loc o apfst o apfst) (add (decl, stamp ())))
|
haftmann@29392
|
461 |
#>
|
haftmann@29361
|
462 |
add_thmss loc Thm.internalK
|
haftmann@29361
|
463 |
[((Binding.empty, [Attrib.internal (decl_attrib decl)]), [([Drule.dummy_thm], [])])];
|
haftmann@29361
|
464 |
|
haftmann@29361
|
465 |
in
|
haftmann@29361
|
466 |
|
haftmann@29361
|
467 |
val add_type_syntax = add_decls (apfst o cons);
|
haftmann@29361
|
468 |
val add_term_syntax = add_decls (apsnd o cons);
|
haftmann@29361
|
469 |
val add_declaration = add_decls (K I);
|
haftmann@29361
|
470 |
|
haftmann@29361
|
471 |
end;
|
haftmann@29361
|
472 |
|
haftmann@29361
|
473 |
(* Dependencies *)
|
haftmann@29361
|
474 |
|
haftmann@29392
|
475 |
fun add_dependency loc dep = (change_locale loc o apsnd) (cons (dep, stamp ()));
|
haftmann@29361
|
476 |
|
haftmann@29361
|
477 |
|
haftmann@29361
|
478 |
(*** Reasoning about locales ***)
|
haftmann@29361
|
479 |
|
haftmann@29361
|
480 |
(** Storage for witnesses, intro and unfold rules **)
|
haftmann@29361
|
481 |
|
haftmann@29361
|
482 |
structure Witnesses = ThmsFun();
|
haftmann@29361
|
483 |
structure Intros = ThmsFun();
|
haftmann@29361
|
484 |
structure Unfolds = ThmsFun();
|
haftmann@29361
|
485 |
|
haftmann@29361
|
486 |
val witness_attrib = Witnesses.add;
|
haftmann@29361
|
487 |
val intro_attrib = Intros.add;
|
haftmann@29361
|
488 |
val unfold_attrib = Unfolds.add;
|
haftmann@29361
|
489 |
|
haftmann@29361
|
490 |
(** Tactic **)
|
haftmann@29361
|
491 |
|
haftmann@29361
|
492 |
fun intro_locales_tac eager ctxt facts st =
|
haftmann@29361
|
493 |
Method.intros_tac
|
haftmann@29361
|
494 |
(Witnesses.get ctxt @ Intros.get ctxt @ (if eager then Unfolds.get ctxt else [])) facts st;
|
haftmann@29361
|
495 |
|
haftmann@29361
|
496 |
val _ = Context.>> (Context.map_theory
|
haftmann@29361
|
497 |
(Method.add_methods
|
haftmann@29361
|
498 |
[("intro_locales",
|
haftmann@29361
|
499 |
Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac false ctxt ORELSE'
|
haftmann@29361
|
500 |
Old_Locale.intro_locales_tac false ctxt)),
|
haftmann@29361
|
501 |
"back-chain introduction rules of locales without unfolding predicates"),
|
haftmann@29361
|
502 |
("unfold_locales",
|
haftmann@29361
|
503 |
Method.ctxt_args (fn ctxt => Method.METHOD (intro_locales_tac true ctxt ORELSE'
|
haftmann@29361
|
504 |
Old_Locale.intro_locales_tac true ctxt)),
|
haftmann@29361
|
505 |
"back-chain all introduction rules of locales")]));
|
haftmann@29361
|
506 |
|
haftmann@29361
|
507 |
end;
|
haftmann@29361
|
508 |
|