Provide internal function for printing a single interpretation.
1 (* Title: Pure/Isar/locale.ML
2 Author: Clemens Ballarin, TU Muenchen
4 Locales -- managed Isar proof contexts, based on Pure predicates.
6 Draws basic ideas from Florian Kammueller's original version of
7 locales, but uses the richer infrastructure of Isar instead of the raw
8 meta-logic. Furthermore, structured import of contexts (with merge
9 and rename operations) are provided, as well as type-inference of the
10 signature parts, and predicate definitions of the specification text.
12 Interpretation enables the reuse of theorems of locales in other
13 contexts, namely those defined by theories, structured proofs and
18 [1] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.
19 In Stefano Berardi et al., Types for Proofs and Programs: International
20 Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.
21 [2] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing
22 Dependencies between Locales. Technical Report TUM-I0607, Technische
23 Universitaet Muenchen, 2006.
24 [3] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and
25 Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,
31 (* Locale specification *)
32 val register_locale: binding ->
33 (string * sort) list * ((string * typ) * mixfix) list ->
34 term option * term list ->
35 thm option * thm option -> thm list ->
37 (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list ->
38 (string * morphism) list -> theory -> theory
39 val intern: theory -> xstring -> string
40 val extern: theory -> string -> xstring
41 val defined: theory -> string -> bool
42 val params_of: theory -> string -> ((string * typ) * mixfix) list
43 val intros_of: theory -> string -> thm option * thm option
44 val axioms_of: theory -> string -> thm list
45 val instance_of: theory -> string -> morphism -> term list
46 val specification_of: theory -> string -> term option * term list
49 val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
50 Proof.context -> Proof.context
51 val add_declaration: string -> declaration -> Proof.context -> Proof.context
52 val add_syntax_declaration: string -> declaration -> Proof.context -> Proof.context
55 val activate_declarations: string * morphism -> Proof.context -> Proof.context
56 val activate_facts: string * morphism -> Context.generic -> Context.generic
57 val init: string -> theory -> Proof.context
59 (* Reasoning about locales *)
60 val get_witnesses: Proof.context -> thm list
61 val get_intros: Proof.context -> thm list
62 val get_unfolds: Proof.context -> thm list
63 val witness_add: attribute
64 val intro_add: attribute
65 val unfold_add: attribute
66 val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
68 (* Registrations and dependencies *)
69 val add_registration: string * morphism -> (morphism * bool) option ->
70 morphism -> theory -> theory
71 val amend_registration: string * morphism -> morphism * bool ->
72 morphism -> theory -> theory
73 val add_dependency: string -> string * morphism -> morphism -> theory -> theory
76 val print_locales: theory -> unit
77 val print_locale: theory -> bool -> xstring -> unit
78 val print_registrations: theory -> string -> unit
81 structure Locale: LOCALE =
84 datatype ctxt = datatype Element.ctxt;
89 datatype locale = Loc of {
91 parameters: (string * sort) list * ((string * typ) * mixfix) list,
92 (* type and term parameters *)
93 spec: term option * term list,
94 (* assumptions (as a single predicate expression) and defines *)
95 intros: thm option * thm option,
98 syntax_decls: (declaration * serial) list,
99 (* syntax declarations *)
100 notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * serial) list,
101 (* theorem declarations *)
102 dependencies: ((string * (morphism * morphism)) * serial) list
103 (* locale dependencies (sublocale relation) *)
106 fun mk_locale ((parameters, spec, intros, axioms), ((syntax_decls, notes), dependencies)) =
107 Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
108 syntax_decls = syntax_decls, notes = notes, dependencies = dependencies};
110 fun map_locale f (Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies}) =
111 mk_locale (f ((parameters, spec, intros, axioms), ((syntax_decls, notes), dependencies)));
114 (Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies},
115 Loc {syntax_decls = syntax_decls', notes = notes', dependencies = dependencies', ...}) =
117 ((parameters, spec, intros, axioms),
118 ((merge (eq_snd op =) (syntax_decls, syntax_decls'),
119 merge (eq_snd op =) (notes, notes')),
120 merge (eq_snd op =) (dependencies, dependencies')));
122 structure Locales = Theory_Data
124 type T = locale Name_Space.table;
125 val empty : T = Name_Space.empty_table "locale";
127 val merge = Name_Space.join_tables (K merge_locale);
130 val intern = Name_Space.intern o #1 o Locales.get;
131 val extern = Name_Space.extern o #1 o Locales.get;
133 val get_locale = Symtab.lookup o #2 o Locales.get;
134 val defined = Symtab.defined o #2 o Locales.get;
136 fun the_locale thy name =
137 (case get_locale thy name of
138 SOME (Loc loc) => loc
139 | NONE => error ("Unknown locale " ^ quote name));
141 fun register_locale binding parameters spec intros axioms syntax_decls notes dependencies thy =
142 thy |> Locales.map (Name_Space.define true (Sign.naming_of thy)
144 mk_locale ((parameters, spec, intros, axioms),
145 ((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes),
146 map (fn d => (d |> apsnd (rpair Morphism.identity), serial ())) dependencies))) #> snd);
149 fun change_locale name =
150 Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
152 fun print_locales thy =
153 Pretty.strs ("locales:" :: map #1 (Name_Space.extern_table (Locales.get thy)))
157 (*** Primitive operations ***)
159 fun params_of thy = snd o #parameters o the_locale thy;
161 fun intros_of thy = #intros o the_locale thy;
163 fun axioms_of thy = #axioms o the_locale thy;
165 fun instance_of thy name morph = params_of thy name |>
166 map (Morphism.term morph o Free o #1);
168 fun specification_of thy = #spec o the_locale thy;
170 fun dependencies_of thy name = the_locale thy name |>
171 #dependencies |> map fst;
174 (*** Activate context elements of locale ***)
176 (** Identifiers: activated locales in theory or proof context **)
178 fun ident_eq thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts);
179 (* FIXME: this is ident_le, smaller term is more general *)
183 datatype 'a delayed = Ready of 'a | ToDo of 'a delayed * 'a delayed;
185 structure Identifiers = Generic_Data
187 type T = (string * term list) list delayed;
188 val empty = Ready [];
193 fun finish thy (ToDo (i1, i2)) = merge (ident_eq thy) (finish thy i1, finish thy i2)
194 | finish _ (Ready ids) = ids;
196 val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy =>
197 (case Identifiers.get (Context.Theory thy) of
199 | ids => SOME (Context.theory_map (Identifiers.put (Ready (finish thy ids))) thy)))));
203 val get_idents = (fn Ready ids => ids) o Identifiers.get;
204 val put_idents = Identifiers.put o Ready;
209 (** Resolve locale dependencies in a depth-first fashion **)
213 val roundup_bound = 120;
215 fun add thy depth export (name, morph) (deps, marked) =
216 if depth > roundup_bound
217 then error "Roundup bound exceeded (sublocale relation probably not terminating)."
220 val dependencies = dependencies_of thy name;
221 val instance = instance_of thy name (morph $> export);
223 if member (ident_eq thy) marked (name, instance)
227 val dependencies' = map (fn (name, (morph', export')) => (name, morph' $> export' $> morph)) dependencies;
228 val marked' = (name, instance) :: marked;
229 val (deps', marked'') = fold_rev (add thy (depth + 1) export) dependencies' ([], marked');
231 ((name, morph) :: deps' @ deps, marked'')
237 (* Note that while identifiers always have the external (exported) view, activate_dep
238 is presented with the internal view. *)
240 fun roundup thy activate_dep export (name, morph) (marked, input) =
242 (* Find all dependencies incuding new ones (which are dependencies enriching
243 existing registrations). *)
244 val (dependencies, marked') = add thy 0 export (name, morph) ([], []);
245 (* Filter out fragments from marked; these won't be activated. *)
246 val dependencies' = filter_out (fn (name, morph) =>
247 member (ident_eq thy) marked (name, instance_of thy name (morph $> export))) dependencies;
249 (merge (ident_eq thy) (marked, marked'), input |> fold_rev activate_dep dependencies')
255 (* Declarations, facts and entire locale content *)
257 fun activate_syntax_decls (name, morph) context =
259 val thy = Context.theory_of context;
260 val {syntax_decls, ...} = the_locale thy name;
263 |> fold_rev (fn (decl, _) => decl morph) syntax_decls
266 fun activate_notes activ_elem transfer thy (name, morph) input =
268 val {notes, ...} = the_locale thy name;
269 fun activate ((kind, facts), _) input =
271 val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph))
272 in activ_elem (Notes (kind, facts')) input end;
274 fold_rev activate notes input
277 fun activate_all name thy activ_elem transfer (marked, input) =
279 val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
280 val input' = input |>
282 activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |>
283 (* FIXME type parameters *)
284 (case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |>
286 then activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))
288 val activate = activate_notes activ_elem transfer thy;
290 roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input')
294 (** Public activation functions **)
296 fun activate_declarations dep = Context.proof_map (fn context =>
298 val thy = Context.theory_of context;
300 roundup thy activate_syntax_decls Morphism.identity dep (get_idents context, context)
304 fun activate_facts dep context =
306 val thy = Context.theory_of context;
307 val activate = activate_notes Element.init (Element.transfer_morphism o Context.theory_of) thy;
308 in roundup thy activate Morphism.identity dep (get_idents context, context) |-> put_idents end;
311 activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of)
312 ([], Context.Proof (ProofContext.init thy)) |-> put_idents |> Context.proof_of;
314 fun print_locale thy show_facts raw_name =
316 val name = intern thy raw_name;
317 val ctxt = init name thy;
318 fun cons_elem (elem as Notes _) = show_facts ? cons elem
319 | cons_elem elem = cons elem;
321 activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])
324 Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)
329 (*** Registrations: interpretations in theories ***)
331 structure Registrations = Theory_Data
333 type T = ((string * (morphism * morphism)) * serial) list *
334 (* registrations, in reverse order of declaration;
335 serial points to mixin list *)
336 (serial * ((morphism * bool) * serial) list) list;
337 (* alist of mixin lists, per list mixins in reverse order of declaration;
338 lists indexed by registration serial,
339 entries for empty lists may be omitted *)
340 val empty = ([], []);
342 fun merge ((r1, m1), (r2, m2)) : T =
343 (Library.merge (eq_snd op =) (r1, r2),
344 AList.join (op =) (K (Library.merge (eq_snd op =))) (m1, m2));
345 (* FIXME consolidate with dependencies, consider one data slot only *)
349 (* Primitive operations *)
351 fun add_reg export (dep, morph) =
352 Registrations.map (apfst (cons ((dep, (morph, export)), serial ())));
354 fun add_mixin serial' mixin =
355 (* registration to be amended identified by its serial id *)
356 Registrations.map (apsnd (AList.map_default (op =) (serial', []) (cons (mixin, serial ()))));
358 fun get_mixins thy (name, morph) =
360 val (regs, mixins) = Registrations.get thy;
362 case find_first (fn ((name', (morph', export')), _) => ident_eq thy
363 ((name', instance_of thy name' (morph' $> export')), (name, instance_of thy name morph))) (rev regs) of
365 | SOME (_, serial) => the_default [] (AList.lookup (op =) mixins serial)
368 fun collect_mixins thy (name, morph) =
369 roundup thy (fn dep => fn mixins =>
370 merge (eq_snd op =) (mixins, get_mixins thy dep)) Morphism.identity (name, morph) ([], [])
371 |> snd |> filter (snd o fst); (* only inheritable mixins *)
372 (* FIXME refactor usage *)
374 fun compose_mixins mixins =
375 fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity;
377 fun reg_morph mixins ((name, (base, export)), serial) =
378 let val mix = the_default [] (AList.lookup (op =) mixins serial) |> compose_mixins;
379 in (name, base $> mix $> export) end;
381 fun these_registrations thy name = Registrations.get thy
382 |>> filter (curry (op =) name o fst o fst)
383 (* with inherited mixins *)
384 |-> (fn regs => fn _ => map (fn ((name, (base, export)) ,_) =>
385 (name, base $> (collect_mixins thy (name, base $> export) |> compose_mixins) $> export)) regs);
387 fun all_registrations thy = Registrations.get thy
388 |-> (fn regs => fn mixins => map (reg_morph mixins) regs);
389 (* without inherited mixins *)
391 fun activate_notes' activ_elem transfer thy export (name, morph) input =
393 val {notes, ...} = the_locale thy name;
394 fun activate ((kind, facts), _) input =
396 val mixin = collect_mixins thy (name, morph $> export) |> compose_mixins;
397 val facts' = facts |> Element.facts_map (Element.morph_ctxt (transfer input $> morph $> mixin $> export))
398 in activ_elem (Notes (kind, facts')) input end;
400 fold_rev activate notes input
403 fun activate_facts' export dep context =
405 val thy = Context.theory_of context;
406 val activate = activate_notes' Element.init (Element.transfer_morphism o Context.theory_of) thy export;
407 in roundup thy activate export dep (get_idents context, context) |-> put_idents end;
412 fun pretty_reg thy (name, morph) =
414 val name' = extern thy name;
415 val ctxt = ProofContext.init thy;
416 fun prt_qual (qual, mand) = Pretty.str (qual ^ (if mand then "!" else "?"));
417 fun prt_quals qs = Pretty.separate "." (map prt_qual qs) |> Pretty.block;
418 val prt_term = Pretty.quote o Syntax.pretty_term ctxt;
419 fun prt_term' t = if !show_types
420 then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::",
421 Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)]
424 Pretty.block (Pretty.breaks (Pretty.str name' :: map prt_term' ts));
426 val qs = Binding.name "x" |> Morphism.binding morph |> Binding.prefix_of;
427 val ts = instance_of thy name morph;
431 | qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":",
432 Pretty.brk 1, prt_inst ts]
435 fun print_registrations thy raw_name =
436 (case these_registrations thy (intern thy raw_name) of
437 [] => Pretty.str ("no interpretations")
438 | regs => Pretty.big_list "interpretations:" (map (pretty_reg thy) (rev regs)))
442 (* Add and extend registrations *)
444 fun amend_registration (name, morph) mixin export thy =
446 val regs = Registrations.get thy |> fst;
447 val base = instance_of thy name (morph $> export);
448 fun match ((name', (morph', export')), _) =
449 name = name' andalso eq_list (op aconv) (base, instance_of thy name' (morph' $> export'));
451 case find_first match (rev regs) of
452 NONE => error ("No interpretation of locale " ^
453 quote (extern thy name) ^ " and\nparameter instantiation " ^
454 space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^
456 | SOME (_, serial') => add_mixin serial' mixin thy
459 (* Note that a registration that would be subsumed by an existing one will not be
460 generated, and it will not be possible to amend it. *)
462 fun add_registration (name, base_morph) mixin export thy =
464 val base = instance_of thy name base_morph;
466 if member (ident_eq thy) (get_idents (Context.Theory thy)) (name, base)
469 (get_idents (Context.Theory thy), thy)
470 (* add new registrations with inherited mixins *)
471 |> roundup thy (add_reg export) export (name, base_morph)
474 |> (case mixin of NONE => I
475 | SOME mixin => amend_registration (name, base_morph) mixin export)
476 (* activate import hierarchy as far as not already active *)
477 |> Context.theory_map (activate_facts' export (name, base_morph))
481 (*** Dependencies ***)
483 fun add_reg_activate_facts export (dep, morph) thy =
484 (get_idents (Context.Theory thy), thy)
485 |> roundup thy (add_reg export) export (dep, morph)
487 |> Context.theory_map (activate_facts' export (dep, morph));
489 fun add_dependency loc (dep, morph) export thy =
491 |> (change_locale loc o apsnd) (cons ((dep, (morph, export)), serial ()))
492 |> (fn thy => fold_rev (add_reg_activate_facts export)
493 (all_registrations thy) thy);
496 (*** Storing results ***)
500 fun add_thmss loc kind args ctxt =
502 val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;
503 val ctxt'' = ctxt' |> ProofContext.theory (
504 (change_locale loc o apfst o apsnd) (cons (args', serial ()))
507 (fn thy => fold_rev (fn (_, morph) =>
509 val args'' = snd args' |> Element.facts_map (Element.morph_ctxt morph) |>
510 Attrib.map_facts (Attrib.attribute_i thy)
511 in PureThy.note_thmss kind args'' #> snd end)
512 (these_registrations thy loc) thy))
518 fun add_declaration loc decl =
520 [((Binding.conceal Binding.empty,
521 [Attrib.internal (fn phi => Thm.declaration_attribute (K (decl phi)))]),
522 [([Drule.dummy_thm], [])])];
524 fun add_syntax_declaration loc decl =
525 ProofContext.theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
526 #> add_declaration loc decl;
529 (*** Reasoning about locales ***)
531 (* Storage for witnesses, intro and unfold rules *)
533 structure Thms = Generic_Data
535 type T = thm list * thm list * thm list;
536 val empty = ([], [], []);
538 fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) =
539 (Thm.merge_thms (witnesses1, witnesses2),
540 Thm.merge_thms (intros1, intros2),
541 Thm.merge_thms (unfolds1, unfolds2));
544 val get_witnesses = #1 o Thms.get o Context.Proof;
545 val get_intros = #2 o Thms.get o Context.Proof;
546 val get_unfolds = #3 o Thms.get o Context.Proof;
549 Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Thm.add_thm th x, y, z)));
551 Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Thm.add_thm th y, z)));
553 Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z)));
558 fun gen_intro_locales_tac intros_tac eager ctxt =
560 (get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));
562 val intro_locales_tac = gen_intro_locales_tac Method.intros_tac;
563 val try_intro_locales_tac= gen_intro_locales_tac Method.try_intros_tac;
565 val _ = Context.>> (Context.map_theory
566 (Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o try_intro_locales_tac false))
567 "back-chain introduction rules of locales without unfolding predicates" #>
568 Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o try_intro_locales_tac true))
569 "back-chain all introduction rules of locales"));