haftmann@29361
|
1 |
(* Title: Pure/Isar/locale.ML
|
haftmann@29361
|
2 |
Author: Clemens Ballarin, TU Muenchen
|
haftmann@29361
|
3 |
|
wenzelm@30725
|
4 |
Locales -- managed Isar proof contexts, based on Pure predicates.
|
haftmann@29361
|
5 |
|
haftmann@29361
|
6 |
Draws basic ideas from Florian Kammueller's original version of
|
haftmann@29361
|
7 |
locales, but uses the richer infrastructure of Isar instead of the raw
|
haftmann@29361
|
8 |
meta-logic. Furthermore, structured import of contexts (with merge
|
haftmann@29361
|
9 |
and rename operations) are provided, as well as type-inference of the
|
haftmann@29361
|
10 |
signature parts, and predicate definitions of the specification text.
|
haftmann@29361
|
11 |
|
haftmann@29361
|
12 |
Interpretation enables the reuse of theorems of locales in other
|
haftmann@29361
|
13 |
contexts, namely those defined by theories, structured proofs and
|
haftmann@29361
|
14 |
locales themselves.
|
haftmann@29361
|
15 |
|
haftmann@29361
|
16 |
See also:
|
haftmann@29361
|
17 |
|
haftmann@29361
|
18 |
[1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
|
haftmann@29361
|
19 |
In Stefano Berardi et al., Types for Proofs and Programs: International
|
haftmann@29361
|
20 |
Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
|
haftmann@29361
|
21 |
[2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
|
haftmann@29361
|
22 |
Dependencies between Locales. Technical Report TUM-I0607, Technische
|
haftmann@29361
|
23 |
Universitaet Muenchen, 2006.
|
haftmann@29361
|
24 |
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
|
haftmann@29361
|
25 |
Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
|
haftmann@29361
|
26 |
pages 31-43, 2006.
|
haftmann@29361
|
27 |
*)
|
haftmann@29361
|
28 |
|
haftmann@29361
|
29 |
signature LOCALE =
|
haftmann@29361
|
30 |
sig
|
haftmann@29576
|
31 |
(* Locale specification *)
|
wenzelm@30344
|
32 |
val register_locale: binding ->
|
wenzelm@30755
|
33 |
(string * sort) list * ((string * typ) * mixfix) list ->
|
haftmann@29361
|
34 |
term option * term list ->
|
haftmann@29441
|
35 |
thm option * thm option -> thm list ->
|
wenzelm@35798
|
36 |
declaration list ->
|
wenzelm@30725
|
37 |
(string * (Attrib.binding * (thm list * Attrib.src list) list) list) list ->
|
wenzelm@30725
|
38 |
(string * morphism) list -> theory -> theory
|
haftmann@29361
|
39 |
val intern: theory -> xstring -> string
|
haftmann@29361
|
40 |
val extern: theory -> string -> xstring
|
haftmann@29392
|
41 |
val defined: theory -> string -> bool
|
wenzelm@30755
|
42 |
val params_of: theory -> string -> ((string * typ) * mixfix) list
|
haftmann@29441
|
43 |
val intros_of: theory -> string -> thm option * thm option
|
haftmann@29441
|
44 |
val axioms_of: theory -> string -> thm list
|
haftmann@29544
|
45 |
val instance_of: theory -> string -> morphism -> term list
|
haftmann@29361
|
46 |
val specification_of: theory -> string -> term option * term list
|
haftmann@29361
|
47 |
|
haftmann@29361
|
48 |
(* Storing results *)
|
haftmann@29361
|
49 |
val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
|
haftmann@29361
|
50 |
Proof.context -> Proof.context
|
haftmann@29361
|
51 |
val add_declaration: string -> declaration -> Proof.context -> Proof.context
|
wenzelm@35798
|
52 |
val add_syntax_declaration: string -> declaration -> Proof.context -> Proof.context
|
haftmann@29361
|
53 |
|
haftmann@29361
|
54 |
(* Activation *)
|
wenzelm@30764
|
55 |
val activate_declarations: string * morphism -> Proof.context -> Proof.context
|
wenzelm@30764
|
56 |
val activate_facts: string * morphism -> Context.generic -> Context.generic
|
haftmann@29361
|
57 |
val init: string -> theory -> Proof.context
|
haftmann@29361
|
58 |
|
haftmann@29361
|
59 |
(* Reasoning about locales *)
|
wenzelm@30725
|
60 |
val get_witnesses: Proof.context -> thm list
|
wenzelm@30725
|
61 |
val get_intros: Proof.context -> thm list
|
wenzelm@30725
|
62 |
val get_unfolds: Proof.context -> thm list
|
wenzelm@30725
|
63 |
val witness_add: attribute
|
wenzelm@30725
|
64 |
val intro_add: attribute
|
wenzelm@30725
|
65 |
val unfold_add: attribute
|
haftmann@29361
|
66 |
val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
|
haftmann@29361
|
67 |
|
haftmann@31988
|
68 |
(* Registrations and dependencies *)
|
ballarin@32845
|
69 |
val add_registration: string * morphism -> (morphism * bool) option ->
|
ballarin@32845
|
70 |
morphism -> theory -> theory
|
ballarin@32845
|
71 |
val amend_registration: string * morphism -> morphism * bool ->
|
ballarin@32845
|
72 |
morphism -> theory -> theory
|
haftmann@32074
|
73 |
val add_dependency: string -> string * morphism -> morphism -> theory -> theory
|
haftmann@29361
|
74 |
|
haftmann@29361
|
75 |
(* Diagnostic *)
|
haftmann@29361
|
76 |
val print_locales: theory -> unit
|
wenzelm@30344
|
77 |
val print_locale: theory -> bool -> xstring -> unit
|
ballarin@32804
|
78 |
val print_registrations: theory -> string -> unit
|
haftmann@29361
|
79 |
end;
|
haftmann@29361
|
80 |
|
haftmann@29361
|
81 |
structure Locale: LOCALE =
|
haftmann@29361
|
82 |
struct
|
haftmann@29361
|
83 |
|
haftmann@29361
|
84 |
datatype ctxt = datatype Element.ctxt;
|
haftmann@29361
|
85 |
|
haftmann@29392
|
86 |
|
haftmann@29392
|
87 |
(*** Theory data ***)
|
haftmann@29361
|
88 |
|
haftmann@29361
|
89 |
datatype locale = Loc of {
|
haftmann@29392
|
90 |
(** static part **)
|
wenzelm@30755
|
91 |
parameters: (string * sort) list * ((string * typ) * mixfix) list,
|
haftmann@29361
|
92 |
(* type and term parameters *)
|
haftmann@29361
|
93 |
spec: term option * term list,
|
haftmann@29361
|
94 |
(* assumptions (as a single predicate expression) and defines *)
|
haftmann@29441
|
95 |
intros: thm option * thm option,
|
haftmann@29441
|
96 |
axioms: thm list,
|
haftmann@29392
|
97 |
(** dynamic part **)
|
ballarin@36096
|
98 |
syntax_decls: (declaration * serial) list,
|
wenzelm@35798
|
99 |
(* syntax declarations *)
|
ballarin@36096
|
100 |
notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * serial) list,
|
haftmann@29361
|
101 |
(* theorem declarations *)
|
ballarin@36651
|
102 |
dependencies: ((string * (morphism * morphism)) * serial) list
|
haftmann@29361
|
103 |
(* locale dependencies (sublocale relation) *)
|
haftmann@29392
|
104 |
};
|
haftmann@29361
|
105 |
|
wenzelm@35798
|
106 |
fun mk_locale ((parameters, spec, intros, axioms), ((syntax_decls, notes), dependencies)) =
|
haftmann@29441
|
107 |
Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
|
wenzelm@35798
|
108 |
syntax_decls = syntax_decls, notes = notes, dependencies = dependencies};
|
wenzelm@35798
|
109 |
|
wenzelm@35798
|
110 |
fun map_locale f (Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies}) =
|
wenzelm@35798
|
111 |
mk_locale (f ((parameters, spec, intros, axioms), ((syntax_decls, notes), dependencies)));
|
wenzelm@30754
|
112 |
|
wenzelm@35798
|
113 |
fun merge_locale
|
wenzelm@35798
|
114 |
(Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies},
|
wenzelm@35798
|
115 |
Loc {syntax_decls = syntax_decls', notes = notes', dependencies = dependencies', ...}) =
|
wenzelm@35798
|
116 |
mk_locale
|
haftmann@29441
|
117 |
((parameters, spec, intros, axioms),
|
wenzelm@35798
|
118 |
((merge (eq_snd op =) (syntax_decls, syntax_decls'),
|
haftmann@29441
|
119 |
merge (eq_snd op =) (notes, notes')),
|
haftmann@29441
|
120 |
merge (eq_snd op =) (dependencies, dependencies')));
|
haftmann@29361
|
121 |
|
wenzelm@33522
|
122 |
structure Locales = Theory_Data
|
haftmann@29361
|
123 |
(
|
wenzelm@33095
|
124 |
type T = locale Name_Space.table;
|
wenzelm@33159
|
125 |
val empty : T = Name_Space.empty_table "locale";
|
haftmann@29361
|
126 |
val extend = I;
|
wenzelm@33522
|
127 |
val merge = Name_Space.join_tables (K merge_locale);
|
haftmann@29361
|
128 |
);
|
haftmann@29361
|
129 |
|
wenzelm@33095
|
130 |
val intern = Name_Space.intern o #1 o Locales.get;
|
wenzelm@33095
|
131 |
val extern = Name_Space.extern o #1 o Locales.get;
|
haftmann@29392
|
132 |
|
wenzelm@30725
|
133 |
val get_locale = Symtab.lookup o #2 o Locales.get;
|
wenzelm@30725
|
134 |
val defined = Symtab.defined o #2 o Locales.get;
|
haftmann@29361
|
135 |
|
wenzelm@30725
|
136 |
fun the_locale thy name =
|
wenzelm@30725
|
137 |
(case get_locale thy name of
|
wenzelm@30725
|
138 |
SOME (Loc loc) => loc
|
wenzelm@30725
|
139 |
| NONE => error ("Unknown locale " ^ quote name));
|
haftmann@29361
|
140 |
|
wenzelm@35798
|
141 |
fun register_locale binding parameters spec intros axioms syntax_decls notes dependencies thy =
|
wenzelm@33095
|
142 |
thy |> Locales.map (Name_Space.define true (Sign.naming_of thy)
|
wenzelm@30725
|
143 |
(binding,
|
wenzelm@30725
|
144 |
mk_locale ((parameters, spec, intros, axioms),
|
ballarin@36096
|
145 |
((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes),
|
ballarin@36651
|
146 |
map (fn d => (d |> apsnd (rpair Morphism.identity), serial ())) dependencies))) #> snd);
|
ballarin@36651
|
147 |
(* FIXME *)
|
haftmann@29361
|
148 |
|
haftmann@29392
|
149 |
fun change_locale name =
|
wenzelm@30725
|
150 |
Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
|
haftmann@29361
|
151 |
|
haftmann@29361
|
152 |
fun print_locales thy =
|
wenzelm@33095
|
153 |
Pretty.strs ("locales:" :: map #1 (Name_Space.extern_table (Locales.get thy)))
|
haftmann@29392
|
154 |
|> Pretty.writeln;
|
haftmann@29361
|
155 |
|
haftmann@29361
|
156 |
|
haftmann@29361
|
157 |
(*** Primitive operations ***)
|
haftmann@29361
|
158 |
|
haftmann@29392
|
159 |
fun params_of thy = snd o #parameters o the_locale thy;
|
haftmann@29361
|
160 |
|
haftmann@29441
|
161 |
fun intros_of thy = #intros o the_locale thy;
|
haftmann@29441
|
162 |
|
haftmann@29441
|
163 |
fun axioms_of thy = #axioms o the_locale thy;
|
haftmann@29441
|
164 |
|
haftmann@29392
|
165 |
fun instance_of thy name morph = params_of thy name |>
|
wenzelm@30755
|
166 |
map (Morphism.term morph o Free o #1);
|
haftmann@29361
|
167 |
|
haftmann@29392
|
168 |
fun specification_of thy = #spec o the_locale thy;
|
haftmann@29361
|
169 |
|
haftmann@29544
|
170 |
fun dependencies_of thy name = the_locale thy name |>
|
haftmann@29544
|
171 |
#dependencies |> map fst;
|
haftmann@29544
|
172 |
|
haftmann@29361
|
173 |
|
haftmann@29361
|
174 |
(*** Activate context elements of locale ***)
|
haftmann@29361
|
175 |
|
haftmann@29361
|
176 |
(** Identifiers: activated locales in theory or proof context **)
|
haftmann@29361
|
177 |
|
haftmann@29361
|
178 |
fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
|
ballarin@32804
|
179 |
(* FIXME: this is ident_le, smaller term is more general *)
|
haftmann@29361
|
180 |
|
haftmann@29361
|
181 |
local
|
haftmann@29361
|
182 |
|
wenzelm@30754
|
183 |
datatype 'a delayed = Ready of 'a | ToDo of 'a delayed * 'a delayed;
|
haftmann@29361
|
184 |
|
wenzelm@33519
|
185 |
structure Identifiers = Generic_Data
|
haftmann@29361
|
186 |
(
|
wenzelm@30754
|
187 |
type T = (string * term list) list delayed;
|
wenzelm@30754
|
188 |
val empty = Ready [];
|
haftmann@29361
|
189 |
val extend = I;
|
wenzelm@33519
|
190 |
val merge = ToDo;
|
haftmann@29361
|
191 |
);
|
haftmann@29361
|
192 |
|
haftmann@29361
|
193 |
fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
|
haftmann@29361
|
194 |
| finish _ (Ready ids) = ids;
|
haftmann@29361
|
195 |
|
haftmann@29361
|
196 |
val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
|
wenzelm@30725
|
197 |
(case Identifiers.get (Context.Theory thy) of
|
wenzelm@30725
|
198 |
Ready _ => NONE
|
wenzelm@30725
|
199 |
| ids => SOME (Context.theory_map (Identifiers.put (Ready (finish thy ids))) thy)))));
|
haftmann@29361
|
200 |
|
wenzelm@30764
|
201 |
in
|
haftmann@29361
|
202 |
|
wenzelm@30764
|
203 |
val get_idents = (fn Ready ids => ids) o Identifiers.get;
|
wenzelm@30764
|
204 |
val put_idents = Identifiers.put o Ready;
|
haftmann@29361
|
205 |
|
haftmann@29361
|
206 |
end;
|
haftmann@29361
|
207 |
|
haftmann@29361
|
208 |
|
haftmann@29361
|
209 |
(** Resolve locale dependencies in a depth-first fashion **)
|
haftmann@29361
|
210 |
|
haftmann@29361
|
211 |
local
|
haftmann@29361
|
212 |
|
haftmann@29361
|
213 |
val roundup_bound = 120;
|
haftmann@29361
|
214 |
|
ballarin@32803
|
215 |
fun add thy depth export (name, morph) (deps, marked) =
|
haftmann@29361
|
216 |
if depth > roundup_bound
|
haftmann@29361
|
217 |
then error "Roundup bound exceeded (sublocale relation probably not terminating)."
|
haftmann@29361
|
218 |
else
|
haftmann@29361
|
219 |
let
|
wenzelm@30754
|
220 |
val dependencies = dependencies_of thy name;
|
ballarin@32803
|
221 |
val instance = instance_of thy name (morph $> export);
|
haftmann@29361
|
222 |
in
|
haftmann@29361
|
223 |
if member (ident_eq thy) marked (name, instance)
|
haftmann@29361
|
224 |
then (deps, marked)
|
haftmann@29361
|
225 |
else
|
haftmann@29361
|
226 |
let
|
ballarin@36651
|
227 |
val dependencies' = map (fn (name, (morph', export')) => (name, morph' $> export' $> morph)) dependencies;
|
haftmann@29361
|
228 |
val marked' = (name, instance) :: marked;
|
ballarin@32803
|
229 |
val (deps', marked'') = fold_rev (add thy (depth + 1) export) dependencies' ([], marked');
|
haftmann@29361
|
230 |
in
|
haftmann@29361
|
231 |
((name, morph) :: deps' @ deps, marked'')
|
haftmann@29361
|
232 |
end
|
haftmann@29361
|
233 |
end;
|
haftmann@29361
|
234 |
|
haftmann@29361
|
235 |
in
|
haftmann@29361
|
236 |
|
ballarin@33541
|
237 |
(* Note that while identifiers always have the external (exported) view, activate_dep
|
ballarin@32803
|
238 |
is presented with the internal view. *)
|
ballarin@32803
|
239 |
|
ballarin@32803
|
240 |
fun roundup thy activate_dep export (name, morph) (marked, input) =
|
haftmann@29361
|
241 |
let
|
haftmann@29361
|
242 |
(* Find all dependencies incuding new ones (which are dependencies enriching
|
haftmann@29361
|
243 |
existing registrations). *)
|
ballarin@32803
|
244 |
val (dependencies, marked') = add thy 0 export (name, morph) ([], []);
|
ballarin@32800
|
245 |
(* Filter out fragments from marked; these won't be activated. *)
|
haftmann@29361
|
246 |
val dependencies' = filter_out (fn (name, morph) =>
|
ballarin@32803
|
247 |
member (ident_eq thy) marked (name, instance_of thy name (morph $> export))) dependencies;
|
haftmann@29361
|
248 |
in
|
wenzelm@30773
|
249 |
(merge (ident_eq thy) (marked, marked'), input |> fold_rev activate_dep dependencies')
|
haftmann@29361
|
250 |
end;
|
haftmann@29361
|
251 |
|
haftmann@29361
|
252 |
end;
|
haftmann@29361
|
253 |
|
haftmann@29361
|
254 |
|
haftmann@29361
|
255 |
(* Declarations, facts and entire locale content *)
|
haftmann@29361
|
256 |
|
wenzelm@35798
|
257 |
fun activate_syntax_decls (name, morph) context =
|
haftmann@29361
|
258 |
let
|
wenzelm@30764
|
259 |
val thy = Context.theory_of context;
|
wenzelm@35798
|
260 |
val {syntax_decls, ...} = the_locale thy name;
|
haftmann@29361
|
261 |
in
|
wenzelm@30764
|
262 |
context
|
wenzelm@35798
|
263 |
|> fold_rev (fn (decl, _) => decl morph) syntax_decls
|
haftmann@29361
|
264 |
end;
|
haftmann@29361
|
265 |
|
haftmann@29361
|
266 |
fun activate_notes activ_elem transfer thy (name, morph) input =
|
haftmann@29361
|
267 |
let
|
haftmann@29392
|
268 |
val {notes, ...} = the_locale thy name;
|
haftmann@29361
|
269 |
fun activate ((kind, facts), _) input =
|
haftmann@29361
|
270 |
let
|
haftmann@29361
|
271 |
val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
|
haftmann@29361
|
272 |
in activ_elem (Notes (kind, facts')) input end;
|
haftmann@29361
|
273 |
in
|
haftmann@29361
|
274 |
fold_rev activate notes input
|
haftmann@29361
|
275 |
end;
|
haftmann@29361
|
276 |
|
haftmann@29361
|
277 |
fun activate_all name thy activ_elem transfer (marked, input) =
|
haftmann@29361
|
278 |
let
|
wenzelm@30764
|
279 |
val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
|
wenzelm@30764
|
280 |
val input' = input |>
|
wenzelm@30755
|
281 |
(not (null params) ?
|
wenzelm@30755
|
282 |
activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |>
|
haftmann@29361
|
283 |
(* FIXME type parameters *)
|
wenzelm@30764
|
284 |
(case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |>
|
haftmann@29361
|
285 |
(if not (null defs)
|
haftmann@29361
|
286 |
then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
|
wenzelm@30764
|
287 |
else I);
|
wenzelm@30773
|
288 |
val activate = activate_notes activ_elem transfer thy;
|
wenzelm@30764
|
289 |
in
|
ballarin@32803
|
290 |
roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input')
|
haftmann@29361
|
291 |
end;
|
haftmann@29361
|
292 |
|
haftmann@29361
|
293 |
|
haftmann@29361
|
294 |
(** Public activation functions **)
|
haftmann@29361
|
295 |
|
wenzelm@30775
|
296 |
fun activate_declarations dep = Context.proof_map (fn context =>
|
wenzelm@30764
|
297 |
let
|
wenzelm@30764
|
298 |
val thy = Context.theory_of context;
|
wenzelm@35798
|
299 |
in
|
wenzelm@35798
|
300 |
roundup thy activate_syntax_decls Morphism.identity dep (get_idents context, context)
|
wenzelm@35798
|
301 |
|-> put_idents
|
wenzelm@35798
|
302 |
end);
|
haftmann@29361
|
303 |
|
wenzelm@30764
|
304 |
fun activate_facts dep context =
|
wenzelm@30764
|
305 |
let
|
wenzelm@30764
|
306 |
val thy = Context.theory_of context;
|
wenzelm@30775
|
307 |
val activate = activate_notes Element.init (Element.transfer_morphism o Context.theory_of) thy;
|
ballarin@32803
|
308 |
in roundup thy activate Morphism.identity dep (get_idents context, context) |-> put_idents end;
|
haftmann@29361
|
309 |
|
haftmann@29361
|
310 |
fun init name thy =
|
wenzelm@30775
|
311 |
activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of)
|
wenzelm@36610
|
312 |
([], Context.Proof (ProofContext.init_global thy)) |-> put_idents |> Context.proof_of;
|
haftmann@29361
|
313 |
|
wenzelm@30764
|
314 |
fun print_locale thy show_facts raw_name =
|
haftmann@29361
|
315 |
let
|
wenzelm@30764
|
316 |
val name = intern thy raw_name;
|
wenzelm@30764
|
317 |
val ctxt = init name thy;
|
wenzelm@30764
|
318 |
fun cons_elem (elem as Notes _) = show_facts ? cons elem
|
wenzelm@30764
|
319 |
| cons_elem elem = cons elem;
|
wenzelm@30764
|
320 |
val elems =
|
wenzelm@30764
|
321 |
activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])
|
wenzelm@30764
|
322 |
|> snd |> rev;
|
haftmann@29361
|
323 |
in
|
wenzelm@30764
|
324 |
Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)
|
wenzelm@30764
|
325 |
|> Pretty.writeln
|
wenzelm@30764
|
326 |
end;
|
haftmann@29361
|
327 |
|
haftmann@29361
|
328 |
|
haftmann@29361
|
329 |
(*** Registrations: interpretations in theories ***)
|
haftmann@29361
|
330 |
|
wenzelm@33522
|
331 |
structure Registrations = Theory_Data
|
haftmann@29361
|
332 |
(
|
ballarin@36088
|
333 |
type T = ((string * (morphism * morphism)) * serial) list *
|
ballarin@36090
|
334 |
(* registrations, in reverse order of declaration;
|
ballarin@36090
|
335 |
serial points to mixin list *)
|
ballarin@36088
|
336 |
(serial * ((morphism * bool) * serial) list) list;
|
ballarin@36090
|
337 |
(* alist of mixin lists, per list mixins in reverse order of declaration;
|
ballarin@36090
|
338 |
lists indexed by registration serial,
|
ballarin@36090
|
339 |
entries for empty lists may be omitted *)
|
ballarin@32801
|
340 |
val empty = ([], []);
|
haftmann@29361
|
341 |
val extend = I;
|
wenzelm@33522
|
342 |
fun merge ((r1, m1), (r2, m2)) : T =
|
ballarin@32801
|
343 |
(Library.merge (eq_snd op =) (r1, r2),
|
ballarin@32801
|
344 |
AList.join (op =) (K (Library.merge (eq_snd op =))) (m1, m2));
|
haftmann@29361
|
345 |
(* FIXME consolidate with dependencies, consider one data slot only *)
|
haftmann@29361
|
346 |
);
|
haftmann@29361
|
347 |
|
ballarin@32801
|
348 |
|
ballarin@32801
|
349 |
(* Primitive operations *)
|
ballarin@32801
|
350 |
|
ballarin@36095
|
351 |
fun add_reg export (dep, morph) =
|
ballarin@36095
|
352 |
Registrations.map (apfst (cons ((dep, (morph, export)), serial ())));
|
haftmann@31988
|
353 |
|
ballarin@36095
|
354 |
fun add_mixin serial' mixin =
|
ballarin@36095
|
355 |
(* registration to be amended identified by its serial id *)
|
ballarin@36095
|
356 |
Registrations.map (apsnd (AList.map_default (op =) (serial', []) (cons (mixin, serial ()))));
|
ballarin@32801
|
357 |
|
ballarin@32804
|
358 |
fun get_mixins thy (name, morph) =
|
ballarin@32804
|
359 |
let
|
ballarin@32804
|
360 |
val (regs, mixins) = Registrations.get thy;
|
ballarin@32804
|
361 |
in
|
ballarin@32804
|
362 |
case find_first (fn ((name', (morph', export')), _) => ident_eq thy
|
ballarin@32804
|
363 |
((name', instance_of thy name' (morph' $> export')), (name, instance_of thy name morph))) (rev regs) of
|
ballarin@32804
|
364 |
NONE => []
|
ballarin@36088
|
365 |
| SOME (_, serial) => the_default [] (AList.lookup (op =) mixins serial)
|
ballarin@32804
|
366 |
end;
|
ballarin@32804
|
367 |
|
ballarin@32804
|
368 |
fun collect_mixins thy (name, morph) =
|
ballarin@32804
|
369 |
roundup thy (fn dep => fn mixins =>
|
ballarin@32804
|
370 |
merge (eq_snd op =) (mixins, get_mixins thy dep)) Morphism.identity (name, morph) ([], [])
|
ballarin@32804
|
371 |
|> snd |> filter (snd o fst); (* only inheritable mixins *)
|
ballarin@36095
|
372 |
(* FIXME refactor usage *)
|
ballarin@32804
|
373 |
|
ballarin@36091
|
374 |
fun compose_mixins mixins =
|
ballarin@36091
|
375 |
fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity;
|
ballarin@36091
|
376 |
|
ballarin@36091
|
377 |
fun reg_morph mixins ((name, (base, export)), serial) =
|
ballarin@36091
|
378 |
let val mix = the_default [] (AList.lookup (op =) mixins serial) |> compose_mixins;
|
ballarin@36091
|
379 |
in (name, base $> mix $> export) end;
|
ballarin@36091
|
380 |
|
ballarin@36091
|
381 |
fun these_registrations thy name = Registrations.get thy
|
ballarin@36091
|
382 |
|>> filter (curry (op =) name o fst o fst)
|
ballarin@36095
|
383 |
(* with inherited mixins *)
|
ballarin@36091
|
384 |
|-> (fn regs => fn _ => map (fn ((name, (base, export)) ,_) =>
|
ballarin@36091
|
385 |
(name, base $> (collect_mixins thy (name, base $> export) |> compose_mixins) $> export)) regs);
|
ballarin@36091
|
386 |
|
ballarin@36919
|
387 |
fun all_registrations thy = Registrations.get thy
|
ballarin@36919
|
388 |
|-> (fn regs => fn mixins => map (reg_morph mixins) regs);
|
ballarin@36919
|
389 |
(* without inherited mixins *)
|
ballarin@32804
|
390 |
|
ballarin@32804
|
391 |
fun activate_notes' activ_elem transfer thy export (name, morph) input =
|
ballarin@32804
|
392 |
let
|
ballarin@32804
|
393 |
val {notes, ...} = the_locale thy name;
|
ballarin@32804
|
394 |
fun activate ((kind, facts), _) input =
|
ballarin@32804
|
395 |
let
|
ballarin@32804
|
396 |
val mixin = collect_mixins thy (name, morph $> export) |> compose_mixins;
|
ballarin@32804
|
397 |
val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph $> mixin $> export))
|
ballarin@32804
|
398 |
in activ_elem (Notes (kind, facts')) input end;
|
ballarin@32804
|
399 |
in
|
ballarin@32804
|
400 |
fold_rev activate notes input
|
ballarin@32804
|
401 |
end;
|
ballarin@32804
|
402 |
|
ballarin@32804
|
403 |
fun activate_facts' export dep context =
|
ballarin@32804
|
404 |
let
|
ballarin@32804
|
405 |
val thy = Context.theory_of context;
|
ballarin@32804
|
406 |
val activate = activate_notes' Element.init (Element.transfer_morphism o Context.theory_of) thy export;
|
ballarin@32804
|
407 |
in roundup thy activate export dep (get_idents context, context) |-> put_idents end;
|
ballarin@32804
|
408 |
|
ballarin@32804
|
409 |
|
ballarin@32804
|
410 |
(* Diagnostic *)
|
ballarin@32804
|
411 |
|
ballarin@36652
|
412 |
fun pretty_reg thy (name, morph) =
|
ballarin@32804
|
413 |
let
|
ballarin@32804
|
414 |
val name' = extern thy name;
|
wenzelm@36610
|
415 |
val ctxt = ProofContext.init_global thy;
|
ballarin@32804
|
416 |
fun prt_qual (qual, mand) = Pretty.str (qual ^ (if mand then "!" else "?"));
|
ballarin@32804
|
417 |
fun prt_quals qs = Pretty.separate "." (map prt_qual qs) |> Pretty.block;
|
ballarin@32804
|
418 |
val prt_term = Pretty.quote o Syntax.pretty_term ctxt;
|
ballarin@32804
|
419 |
fun prt_term' t = if !show_types
|
ballarin@32804
|
420 |
then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::",
|
ballarin@32804
|
421 |
Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)]
|
ballarin@32804
|
422 |
else prt_term t;
|
ballarin@32804
|
423 |
fun prt_inst ts =
|
ballarin@32804
|
424 |
Pretty.block (Pretty.breaks (Pretty.str name' :: map prt_term' ts));
|
ballarin@36652
|
425 |
|
ballarin@36652
|
426 |
val qs = Binding.name "x" |> Morphism.binding morph |> Binding.prefix_of;
|
ballarin@36652
|
427 |
val ts = instance_of thy name morph;
|
ballarin@32804
|
428 |
in
|
ballarin@36652
|
429 |
case qs of
|
ballarin@36652
|
430 |
[] => prt_inst ts
|
ballarin@36652
|
431 |
| qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":",
|
ballarin@36652
|
432 |
Pretty.brk 1, prt_inst ts]
|
ballarin@32804
|
433 |
end;
|
ballarin@32801
|
434 |
|
ballarin@36652
|
435 |
fun print_registrations thy raw_name =
|
ballarin@36652
|
436 |
(case these_registrations thy (intern thy raw_name) of
|
ballarin@36652
|
437 |
[] => Pretty.str ("no interpretations")
|
ballarin@36652
|
438 |
| regs => Pretty.big_list "interpretations:" (map (pretty_reg thy) (rev regs)))
|
ballarin@36652
|
439 |
|> Pretty.writeln;
|
ballarin@36652
|
440 |
|
ballarin@32801
|
441 |
|
ballarin@32801
|
442 |
(* Add and extend registrations *)
|
ballarin@32801
|
443 |
|
ballarin@32804
|
444 |
fun amend_registration (name, morph) mixin export thy =
|
ballarin@32804
|
445 |
let
|
ballarin@32804
|
446 |
val regs = Registrations.get thy |> fst;
|
ballarin@32804
|
447 |
val base = instance_of thy name (morph $> export);
|
ballarin@32804
|
448 |
fun match ((name', (morph', export')), _) =
|
ballarin@32804
|
449 |
name = name' andalso eq_list (op aconv) (base, instance_of thy name' (morph' $> export'));
|
ballarin@32804
|
450 |
in
|
ballarin@32804
|
451 |
case find_first match (rev regs) of
|
ballarin@32804
|
452 |
NONE => error ("No interpretation of locale " ^
|
ballarin@32804
|
453 |
quote (extern thy name) ^ " and\nparameter instantiation " ^
|
ballarin@32804
|
454 |
space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^
|
ballarin@32804
|
455 |
" available")
|
ballarin@36095
|
456 |
| SOME (_, serial') => add_mixin serial' mixin thy
|
ballarin@32804
|
457 |
end;
|
ballarin@32804
|
458 |
|
ballarin@32801
|
459 |
(* Note that a registration that would be subsumed by an existing one will not be
|
ballarin@32801
|
460 |
generated, and it will not be possible to amend it. *)
|
ballarin@32801
|
461 |
|
ballarin@32804
|
462 |
fun add_registration (name, base_morph) mixin export thy =
|
ballarin@32801
|
463 |
let
|
ballarin@32801
|
464 |
val base = instance_of thy name base_morph;
|
ballarin@32801
|
465 |
in
|
ballarin@32801
|
466 |
if member (ident_eq thy) (get_idents (Context.Theory thy)) (name, base)
|
ballarin@32801
|
467 |
then thy
|
ballarin@32801
|
468 |
else
|
ballarin@36095
|
469 |
(get_idents (Context.Theory thy), thy)
|
ballarin@36095
|
470 |
(* add new registrations with inherited mixins *)
|
ballarin@36919
|
471 |
|> roundup thy (add_reg export) export (name, base_morph)
|
ballarin@36095
|
472 |
|> snd
|
ballarin@36095
|
473 |
(* add mixin *)
|
ballarin@36095
|
474 |
|> (case mixin of NONE => I
|
ballarin@36919
|
475 |
| SOME mixin => amend_registration (name, base_morph) mixin export)
|
ballarin@36095
|
476 |
(* activate import hierarchy as far as not already active *)
|
ballarin@36919
|
477 |
|> Context.theory_map (activate_facts' export (name, base_morph))
|
ballarin@32801
|
478 |
end;
|
ballarin@32801
|
479 |
|
haftmann@29361
|
480 |
|
haftmann@31988
|
481 |
(*** Dependencies ***)
|
haftmann@31988
|
482 |
|
ballarin@36095
|
483 |
fun add_reg_activate_facts export (dep, morph) thy =
|
ballarin@36095
|
484 |
(get_idents (Context.Theory thy), thy)
|
ballarin@36095
|
485 |
|> roundup thy (add_reg export) export (dep, morph)
|
ballarin@36095
|
486 |
|> snd
|
ballarin@36095
|
487 |
|> Context.theory_map (activate_facts' export (dep, morph));
|
ballarin@36095
|
488 |
|
haftmann@32074
|
489 |
fun add_dependency loc (dep, morph) export thy =
|
haftmann@31988
|
490 |
thy
|
ballarin@36651
|
491 |
|> (change_locale loc o apsnd) (cons ((dep, (morph, export)), serial ()))
|
ballarin@36095
|
492 |
|> (fn thy => fold_rev (add_reg_activate_facts export)
|
haftmann@31988
|
493 |
(all_registrations thy) thy);
|
haftmann@31988
|
494 |
|
haftmann@31988
|
495 |
|
haftmann@29361
|
496 |
(*** Storing results ***)
|
haftmann@29361
|
497 |
|
haftmann@29361
|
498 |
(* Theorems *)
|
haftmann@29361
|
499 |
|
haftmann@29361
|
500 |
fun add_thmss loc kind args ctxt =
|
haftmann@29361
|
501 |
let
|
wenzelm@30777
|
502 |
val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;
|
haftmann@29361
|
503 |
val ctxt'' = ctxt' |> ProofContext.theory (
|
ballarin@36088
|
504 |
(change_locale loc o apfst o apsnd) (cons (args', serial ()))
|
haftmann@29392
|
505 |
#>
|
haftmann@29361
|
506 |
(* Registrations *)
|
haftmann@31988
|
507 |
(fn thy => fold_rev (fn (_, morph) =>
|
haftmann@29361
|
508 |
let
|
haftmann@29361
|
509 |
val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
|
haftmann@29361
|
510 |
Attrib.map_facts (Attrib.attribute_i thy)
|
wenzelm@30438
|
511 |
in PureThy.note_thmss kind args'' #> snd end)
|
haftmann@31988
|
512 |
(these_registrations thy loc) thy))
|
haftmann@29361
|
513 |
in ctxt'' end;
|
haftmann@29361
|
514 |
|
haftmann@29361
|
515 |
|
haftmann@29361
|
516 |
(* Declarations *)
|
haftmann@29361
|
517 |
|
wenzelm@35798
|
518 |
fun add_declaration loc decl =
|
wenzelm@33643
|
519 |
add_thmss loc ""
|
wenzelm@35798
|
520 |
[((Binding.conceal Binding.empty,
|
wenzelm@35798
|
521 |
[Attrib.internal (fn phi => Thm.declaration_attribute (K (decl phi)))]),
|
wenzelm@33278
|
522 |
[([Drule.dummy_thm], [])])];
|
haftmann@29361
|
523 |
|
wenzelm@35798
|
524 |
fun add_syntax_declaration loc decl =
|
ballarin@36096
|
525 |
ProofContext.theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
|
wenzelm@35798
|
526 |
#> add_declaration loc decl;
|
haftmann@29361
|
527 |
|
wenzelm@30725
|
528 |
|
haftmann@29361
|
529 |
(*** Reasoning about locales ***)
|
haftmann@29361
|
530 |
|
wenzelm@30725
|
531 |
(* Storage for witnesses, intro and unfold rules *)
|
haftmann@29361
|
532 |
|
wenzelm@33519
|
533 |
structure Thms = Generic_Data
|
wenzelm@30725
|
534 |
(
|
wenzelm@30725
|
535 |
type T = thm list * thm list * thm list;
|
wenzelm@30725
|
536 |
val empty = ([], [], []);
|
wenzelm@30725
|
537 |
val extend = I;
|
wenzelm@33519
|
538 |
fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) =
|
wenzelm@30725
|
539 |
(Thm.merge_thms (witnesses1, witnesses2),
|
wenzelm@30725
|
540 |
Thm.merge_thms (intros1, intros2),
|
wenzelm@30725
|
541 |
Thm.merge_thms (unfolds1, unfolds2));
|
wenzelm@30725
|
542 |
);
|
haftmann@29361
|
543 |
|
wenzelm@30725
|
544 |
val get_witnesses = #1 o Thms.get o Context.Proof;
|
wenzelm@30725
|
545 |
val get_intros = #2 o Thms.get o Context.Proof;
|
wenzelm@30725
|
546 |
val get_unfolds = #3 o Thms.get o Context.Proof;
|
haftmann@29361
|
547 |
|
wenzelm@30725
|
548 |
val witness_add =
|
wenzelm@30725
|
549 |
Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Thm.add_thm th x, y, z)));
|
wenzelm@30725
|
550 |
val intro_add =
|
wenzelm@30725
|
551 |
Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Thm.add_thm th y, z)));
|
wenzelm@30725
|
552 |
val unfold_add =
|
wenzelm@30725
|
553 |
Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z)));
|
haftmann@29361
|
554 |
|
wenzelm@30725
|
555 |
|
wenzelm@30725
|
556 |
(* Tactic *)
|
wenzelm@30725
|
557 |
|
ballarin@36093
|
558 |
fun gen_intro_locales_tac intros_tac eager ctxt =
|
ballarin@36093
|
559 |
intros_tac
|
wenzelm@30725
|
560 |
(get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));
|
haftmann@29361
|
561 |
|
ballarin@36093
|
562 |
val intro_locales_tac = gen_intro_locales_tac Method.intros_tac;
|
ballarin@36093
|
563 |
val try_intro_locales_tac= gen_intro_locales_tac Method.try_intros_tac;
|
ballarin@36093
|
564 |
|
haftmann@29361
|
565 |
val _ = Context.>> (Context.map_theory
|
ballarin@36093
|
566 |
(Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o try_intro_locales_tac false))
|
wenzelm@30515
|
567 |
"back-chain introduction rules of locales without unfolding predicates" #>
|
ballarin@36093
|
568 |
Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o try_intro_locales_tac true))
|
wenzelm@30515
|
569 |
"back-chain all introduction rules of locales"));
|
haftmann@29361
|
570 |
|
haftmann@29361
|
571 |
end;
|