| author | blanchet | 
| Wed, 18 Apr 2012 22:39:35 +0200 | |
| changeset 47558 | 55b42f9af99d | 
| parent 47249 | c0481c3c2a6c | 
| child 49569 | 7b6aaf446496 | 
| permissions | -rw-r--r-- | 
| 29361 | 1 | (* Title: Pure/Isar/locale.ML | 
| 2 | Author: Clemens Ballarin, TU Muenchen | |
| 3 | ||
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 4 | Locales -- managed Isar proof contexts, based on Pure predicates. | 
| 29361 | 5 | |
| 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. | |
| 11 | ||
| 12 | Interpretation enables the reuse of theorems of locales in other | |
| 13 | contexts, namely those defined by theories, structured proofs and | |
| 14 | locales themselves. | |
| 15 | ||
| 16 | See also: | |
| 17 | ||
| 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, | |
| 26 | pages 31-43, 2006. | |
| 27 | *) | |
| 28 | ||
| 29 | signature LOCALE = | |
| 30 | sig | |
| 29576 | 31 | (* Locale specification *) | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30223diff
changeset | 32 | val register_locale: binding -> | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30754diff
changeset | 33 | (string * sort) list * ((string * typ) * mixfix) list -> | 
| 29361 | 34 | term option * term list -> | 
| 29441 | 35 | thm option * thm option -> thm list -> | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 36 | declaration list -> | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 37 | (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list -> | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 38 | (string * morphism) list -> theory -> theory | 
| 29361 | 39 | val intern: theory -> xstring -> string | 
| 45488 
6d71d9e52369
pass positions for named targets, for formal links in the document model;
 wenzelm parents: 
45390diff
changeset | 40 | val check: theory -> xstring * Position.T -> string | 
| 29361 | 41 | val extern: theory -> string -> xstring | 
| 29392 | 42 | val defined: theory -> string -> bool | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30754diff
changeset | 43 | val params_of: theory -> string -> ((string * typ) * mixfix) list | 
| 29441 | 44 | val intros_of: theory -> string -> thm option * thm option | 
| 45 | val axioms_of: theory -> string -> thm list | |
| 29544 | 46 | val instance_of: theory -> string -> morphism -> term list | 
| 29361 | 47 | val specification_of: theory -> string -> term option * term list | 
| 48 | ||
| 49 | (* Storing results *) | |
| 50 | val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list -> | |
| 51 | Proof.context -> Proof.context | |
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 52 | val add_declaration: string -> bool -> declaration -> Proof.context -> Proof.context | 
| 29361 | 53 | |
| 54 | (* Activation *) | |
| 30764 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 wenzelm parents: 
30763diff
changeset | 55 | val activate_declarations: string * morphism -> Proof.context -> Proof.context | 
| 38316 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 ballarin parents: 
38211diff
changeset | 56 | val activate_facts: morphism option -> string * morphism -> Context.generic -> Context.generic | 
| 29361 | 57 | val init: string -> theory -> Proof.context | 
| 58 | ||
| 59 | (* Reasoning about locales *) | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 60 | val get_witnesses: Proof.context -> thm list | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 61 | val get_intros: Proof.context -> thm list | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 62 | val get_unfolds: Proof.context -> thm list | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 63 | val witness_add: attribute | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 64 | val intro_add: attribute | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 65 | val unfold_add: attribute | 
| 29361 | 66 | val intro_locales_tac: bool -> Proof.context -> thm list -> tactic | 
| 67 | ||
| 31988 | 68 | (* Registrations and dependencies *) | 
| 32845 
d2d0b9b1a69d
Avoid administrative overhead for identity mixins.
 ballarin parents: 
32804diff
changeset | 69 | val add_registration: string * morphism -> (morphism * bool) option -> | 
| 38107 | 70 | morphism -> Context.generic -> Context.generic | 
| 32845 
d2d0b9b1a69d
Avoid administrative overhead for identity mixins.
 ballarin parents: 
32804diff
changeset | 71 | val amend_registration: string * morphism -> morphism * bool -> | 
| 38107 | 72 | morphism -> Context.generic -> Context.generic | 
| 38111 | 73 | val registrations_of: Context.generic -> string -> (string * morphism) list | 
| 46880 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 74 | val all_registrations_of: Context.generic -> (string * morphism) list | 
| 41270 | 75 | val add_dependency: string -> string * morphism -> (morphism * bool) option -> | 
| 76 | morphism -> theory -> theory | |
| 29361 | 77 | |
| 78 | (* Diagnostic *) | |
| 37471 | 79 | val all_locales: theory -> string list | 
| 29361 | 80 | val print_locales: theory -> unit | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46906diff
changeset | 81 | val print_locale: theory -> bool -> xstring * Position.T -> unit | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46906diff
changeset | 82 | val print_registrations: Proof.context -> xstring * Position.T -> unit | 
| 41435 | 83 | val print_dependencies: Proof.context -> bool -> morphism -> (string * morphism) list -> unit | 
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 84 | val locale_deps: theory -> | 
| 40782 | 85 |     {params: ((string * typ) * mixfix) list, axioms: term list, registrations: term list list} Graph.T
 | 
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 86 | * term list list Symtab.table Symtab.table | 
| 29361 | 87 | end; | 
| 88 | ||
| 89 | structure Locale: LOCALE = | |
| 90 | struct | |
| 91 | ||
| 92 | datatype ctxt = datatype Element.ctxt; | |
| 93 | ||
| 29392 | 94 | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 95 | (*** Locales ***) | 
| 29361 | 96 | |
| 41272 | 97 | type mixins = (((morphism * bool) * serial) list) Inttab.table; | 
| 98 | (* table of mixin lists, per list mixins in reverse order of declaration; | |
| 99 | lists indexed by registration/dependency serial, | |
| 100 | entries for empty lists may be omitted *) | |
| 101 | ||
| 46859 | 102 | fun lookup_mixins serial' mixins = Inttab.lookup_list mixins serial'; | 
| 41272 | 103 | |
| 46859 | 104 | fun merge_mixins mixs : mixins = Inttab.merge_list (eq_snd op =) mixs; | 
| 41272 | 105 | |
| 46859 | 106 | fun insert_mixin serial' mixin = Inttab.cons_list (serial', (mixin, serial ())); | 
| 41272 | 107 | |
| 108 | fun rename_mixin (old, new) mix = | |
| 46858 | 109 | (case Inttab.lookup mix old of | 
| 110 | NONE => mix | |
| 111 | | SOME mxs => Inttab.delete old mix |> Inttab.update_new (new, mxs)); | |
| 41272 | 112 | |
| 113 | fun compose_mixins mixins = | |
| 114 | fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity; | |
| 115 | ||
| 29361 | 116 | datatype locale = Loc of {
 | 
| 29392 | 117 | (** static part **) | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30754diff
changeset | 118 | parameters: (string * sort) list * ((string * typ) * mixfix) list, | 
| 29361 | 119 | (* type and term parameters *) | 
| 120 | spec: term option * term list, | |
| 121 | (* assumptions (as a single predicate expression) and defines *) | |
| 29441 | 122 | intros: thm option * thm option, | 
| 123 | axioms: thm list, | |
| 29392 | 124 | (** dynamic part **) | 
| 36096 | 125 | syntax_decls: (declaration * serial) list, | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 126 | (* syntax declarations *) | 
| 36096 | 127 | notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * serial) list, | 
| 29361 | 128 | (* theorem declarations *) | 
| 36651 | 129 | dependencies: ((string * (morphism * morphism)) * serial) list | 
| 41272 | 130 | (* locale dependencies (sublocale relation) in reverse order *), | 
| 131 | mixins: mixins | |
| 132 | (* mixin part of dependencies *) | |
| 29392 | 133 | }; | 
| 29361 | 134 | |
| 41272 | 135 | fun mk_locale ((parameters, spec, intros, axioms), | 
| 136 | ((syntax_decls, notes), (dependencies, mixins))) = | |
| 29441 | 137 |   Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec,
 | 
| 41272 | 138 | syntax_decls = syntax_decls, notes = notes, dependencies = dependencies, mixins = mixins}; | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 139 | |
| 41272 | 140 | fun map_locale f (Loc {parameters, spec, intros, axioms,
 | 
| 141 | syntax_decls, notes, dependencies, mixins}) = | |
| 142 | mk_locale (f ((parameters, spec, intros, axioms), | |
| 143 | ((syntax_decls, notes), (dependencies, mixins)))); | |
| 30754 | 144 | |
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 145 | fun merge_locale | 
| 41272 | 146 |  (Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies, mixins},
 | 
| 147 |   Loc {syntax_decls = syntax_decls', notes = notes',
 | |
| 148 | dependencies = dependencies', mixins = mixins', ...}) = | |
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 149 | mk_locale | 
| 29441 | 150 | ((parameters, spec, intros, axioms), | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 151 | ((merge (eq_snd op =) (syntax_decls, syntax_decls'), | 
| 29441 | 152 | merge (eq_snd op =) (notes, notes')), | 
| 41272 | 153 | (merge (eq_snd op =) (dependencies, dependencies'), | 
| 154 | (merge_mixins (mixins, mixins'))))); | |
| 29361 | 155 | |
| 33522 | 156 | structure Locales = Theory_Data | 
| 29361 | 157 | ( | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 158 | type T = locale Name_Space.table; | 
| 33159 | 159 | val empty : T = Name_Space.empty_table "locale"; | 
| 29361 | 160 | val extend = I; | 
| 33522 | 161 | val merge = Name_Space.join_tables (K merge_locale); | 
| 29361 | 162 | ); | 
| 163 | ||
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 164 | val intern = Name_Space.intern o #1 o Locales.get; | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46978diff
changeset | 165 | fun check thy = #1 o Name_Space.check (Context.Theory thy) (Locales.get thy); | 
| 42360 | 166 | fun extern thy = Name_Space.extern (Proof_Context.init_global thy) (#1 (Locales.get thy)); | 
| 29392 | 167 | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 168 | val get_locale = Symtab.lookup o #2 o Locales.get; | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 169 | val defined = Symtab.defined o #2 o Locales.get; | 
| 29361 | 170 | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 171 | fun the_locale thy name = | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 172 | (case get_locale thy name of | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 173 | SOME (Loc loc) => loc | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 174 |   | NONE => error ("Unknown locale " ^ quote name));
 | 
| 29361 | 175 | |
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 176 | fun register_locale binding parameters spec intros axioms syntax_decls notes dependencies thy = | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46978diff
changeset | 177 | thy |> Locales.map (Name_Space.define (Context.Theory thy) true | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 178 | (binding, | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 179 | mk_locale ((parameters, spec, intros, axioms), | 
| 36096 | 180 | ((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes), | 
| 41272 | 181 | (map (fn d => (d |> apsnd (rpair Morphism.identity), serial ())) dependencies, | 
| 182 | Inttab.empty)))) #> snd); | |
| 183 | (* FIXME Morphism.identity *) | |
| 29361 | 184 | |
| 29392 | 185 | fun change_locale name = | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 186 | Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd; | 
| 29361 | 187 | |
| 188 | ||
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 189 | (** Primitive operations **) | 
| 29361 | 190 | |
| 29392 | 191 | fun params_of thy = snd o #parameters o the_locale thy; | 
| 29361 | 192 | |
| 29441 | 193 | fun intros_of thy = #intros o the_locale thy; | 
| 194 | ||
| 195 | fun axioms_of thy = #axioms o the_locale thy; | |
| 196 | ||
| 29392 | 197 | fun instance_of thy name morph = params_of thy name |> | 
| 30755 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 wenzelm parents: 
30754diff
changeset | 198 | map (Morphism.term morph o Free o #1); | 
| 29361 | 199 | |
| 29392 | 200 | fun specification_of thy = #spec o the_locale thy; | 
| 29361 | 201 | |
| 29544 | 202 | fun dependencies_of thy name = the_locale thy name |> | 
| 41272 | 203 | #dependencies; | 
| 204 | ||
| 205 | fun mixins_of thy name serial = the_locale thy name |> | |
| 206 | #mixins |> lookup_mixins serial; | |
| 207 | ||
| 46880 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 208 | (* FIXME unused *) | 
| 41272 | 209 | fun identity_on thy name morph = | 
| 210 | let val mk_instance = instance_of thy name | |
| 46880 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 211 | in ListPair.all Term.aconv_untyped (mk_instance Morphism.identity, mk_instance morph) end; | 
| 29544 | 212 | |
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 213 | (* Print instance and qualifiers *) | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 214 | |
| 41435 | 215 | fun pretty_reg ctxt (name, morph) = | 
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 216 | let | 
| 42360 | 217 | val thy = Proof_Context.theory_of ctxt; | 
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 218 | val name' = extern thy name; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 219 | fun prt_qual (qual, mand) = Pretty.str (qual ^ (if mand then "!" else "?")); | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 220 | fun prt_quals qs = Pretty.separate "." (map prt_qual qs) |> Pretty.block; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 221 | val prt_term = Pretty.quote o Syntax.pretty_term ctxt; | 
| 39134 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
38797diff
changeset | 222 | fun prt_term' t = | 
| 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 wenzelm parents: 
38797diff
changeset | 223 | if Config.get ctxt show_types | 
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 224 | then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::", | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 225 | Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)] | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 226 | else prt_term t; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 227 | fun prt_inst ts = | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 228 | Pretty.block (Pretty.breaks (Pretty.str name' :: map prt_term' ts)); | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 229 | |
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 230 | val qs = Binding.name "x" |> Morphism.binding morph |> Binding.prefix_of; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 231 | val ts = instance_of thy name morph; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 232 | in | 
| 40782 | 233 | (case qs of | 
| 234 | [] => prt_inst ts | |
| 235 | | qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":", Pretty.brk 1, prt_inst ts]) | |
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 236 | end; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 237 | |
| 29361 | 238 | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 239 | (*** Identifiers: activated locales in theory or proof context ***) | 
| 29361 | 240 | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 241 | type idents = term list list Symtab.table; (* name ~> instance (grouped by name) *) | 
| 37105 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 242 | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 243 | val empty_idents : idents = Symtab.empty; | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 244 | val insert_idents = Symtab.insert_list (eq_list (op aconv)); | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 245 | val merge_idents = Symtab.merge_list (eq_list (op aconv)); | 
| 29361 | 246 | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 247 | fun redundant_ident thy idents (name, instance) = | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 248 | exists (fn pat => Pattern.matchess thy (pat, instance)) (Symtab.lookup_list idents name); | 
| 29361 | 249 | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 250 | structure Idents = Generic_Data | 
| 29361 | 251 | ( | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 252 | type T = idents; | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 253 | val empty = empty_idents; | 
| 29361 | 254 | val extend = I; | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 255 | val merge = merge_idents; | 
| 29361 | 256 | ); | 
| 257 | ||
| 258 | ||
| 259 | (** Resolve locale dependencies in a depth-first fashion **) | |
| 260 | ||
| 261 | local | |
| 262 | ||
| 263 | val roundup_bound = 120; | |
| 264 | ||
| 41272 | 265 | fun add thy depth stem export (name, morph, mixins) (deps, marked) = | 
| 29361 | 266 | if depth > roundup_bound | 
| 267 | then error "Roundup bound exceeded (sublocale relation probably not terminating)." | |
| 268 | else | |
| 269 | let | |
| 41272 | 270 | val instance = instance_of thy name (morph $> stem $> export); | 
| 29361 | 271 | in | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 272 | if redundant_ident thy marked (name, instance) then (deps, marked) | 
| 29361 | 273 | else | 
| 274 | let | |
| 41272 | 275 | val full_morph = morph $> compose_mixins mixins $> stem; | 
| 276 | (* no inheritance of mixins, regardless of requests by clients *) | |
| 277 | val dependencies = dependencies_of thy name |> | |
| 278 | map (fn ((name', (morph', export')), serial') => | |
| 279 | (name', morph' $> export', mixins_of thy name serial')); | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 280 | val marked' = insert_idents (name, instance) marked; | 
| 41272 | 281 | val (deps', marked'') = | 
| 282 | fold_rev (add thy (depth + 1) full_morph export) dependencies | |
| 283 | ([], marked'); | |
| 29361 | 284 | in | 
| 41272 | 285 | ((name, full_morph) :: deps' @ deps, marked'') | 
| 29361 | 286 | end | 
| 287 | end; | |
| 288 | ||
| 289 | in | |
| 290 | ||
| 33541 | 291 | (* Note that while identifiers always have the external (exported) view, activate_dep | 
| 46870 | 292 | is presented with the internal view. *) | 
| 32803 | 293 | |
| 294 | fun roundup thy activate_dep export (name, morph) (marked, input) = | |
| 29361 | 295 | let | 
| 41272 | 296 | (* Find all dependencies including new ones (which are dependencies enriching | 
| 29361 | 297 | existing registrations). *) | 
| 41272 | 298 | val (dependencies, marked') = | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 299 | add thy 0 Morphism.identity export (name, morph, []) ([], empty_idents); | 
| 32800 | 300 | (* Filter out fragments from marked; these won't be activated. *) | 
| 29361 | 301 | val dependencies' = filter_out (fn (name, morph) => | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 302 | redundant_ident thy marked (name, instance_of thy name (morph $> export))) dependencies; | 
| 29361 | 303 | in | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 304 | (merge_idents (marked, marked'), input |> fold_rev activate_dep dependencies') | 
| 29361 | 305 | end; | 
| 306 | ||
| 307 | end; | |
| 308 | ||
| 309 | ||
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 310 | (*** Registrations: interpretations in theories or proof contexts ***) | 
| 29361 | 311 | |
| 46860 
fe8d6532e1c1
clarified total_ident_ord, swapping first argument back to normal (unlike e464f84f3680) -- NB: "fast" ord is erratic anyway;
 wenzelm parents: 
46859diff
changeset | 312 | val total_ident_ord = prod_ord fast_string_ord (list_ord Term_Ord.fast_term_ord); | 
| 
fe8d6532e1c1
clarified total_ident_ord, swapping first argument back to normal (unlike e464f84f3680) -- NB: "fast" ord is erratic anyway;
 wenzelm parents: 
46859diff
changeset | 313 | |
| 
fe8d6532e1c1
clarified total_ident_ord, swapping first argument back to normal (unlike e464f84f3680) -- NB: "fast" ord is erratic anyway;
 wenzelm parents: 
46859diff
changeset | 314 | structure Idtab = Table(type key = string * term list val ord = total_ident_ord); | 
| 37105 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 315 | |
| 38107 | 316 | structure Registrations = Generic_Data | 
| 29361 | 317 | ( | 
| 41272 | 318 | type T = ((morphism * morphism) * serial) Idtab.table * mixins; | 
| 37105 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 319 | (* registrations, indexed by locale name and instance; | 
| 41272 | 320 | unique registration serial points to mixin list *) | 
| 37105 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 321 | val empty = (Idtab.empty, Inttab.empty); | 
| 29361 | 322 | val extend = I; | 
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 323 | fun merge ((reg1, mix1), (reg2, mix2)) : T = | 
| 45589 
bb944d58ac19
simplified Locale.add_thmss, after partial evaluation of attributes;
 wenzelm parents: 
45587diff
changeset | 324 | (Idtab.join (fn id => fn ((_, s1), (_, s2)) => | 
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 325 | if s1 = s2 then raise Idtab.SAME else raise Idtab.DUP id) (reg1, reg2), | 
| 41272 | 326 | merge_mixins (mix1, mix2)) | 
| 37133 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 327 | handle Idtab.DUP id => | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 328 | (* distinct interpretations with same base: merge their mixins *) | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 329 | let | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 330 | val (_, s1) = Idtab.lookup reg1 id |> the; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 331 | val (morph2, s2) = Idtab.lookup reg2 id |> the; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 332 | val reg2' = Idtab.update (id, (morph2, s1)) reg2; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 333 | val _ = warning "Removed duplicate interpretation after retrieving its mixins."; | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 334 | (* FIXME print interpretations, | 
| 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 ballarin parents: 
37105diff
changeset | 335 | which is not straightforward without theory context *) | 
| 41272 | 336 | in merge ((reg1, mix1), (reg2', rename_mixin (s2, s1) mix2)) end; | 
| 29361 | 337 | (* FIXME consolidate with dependencies, consider one data slot only *) | 
| 338 | ); | |
| 339 | ||
| 32801 | 340 | |
| 341 | (* Primitive operations *) | |
| 342 | ||
| 37104 
3877a6c45d57
Avoid recomputation of registration instance for lookup.
 ballarin parents: 
37103diff
changeset | 343 | fun add_reg thy export (name, morph) = | 
| 37105 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 344 | Registrations.map (apfst (Idtab.insert (K false) | 
| 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 345 | ((name, instance_of thy name (morph $> export)), ((morph, export), serial ())))); | 
| 31988 | 346 | |
| 36095 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 ballarin parents: 
36094diff
changeset | 347 | fun add_mixin serial' mixin = | 
| 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 ballarin parents: 
36094diff
changeset | 348 | (* registration to be amended identified by its serial id *) | 
| 41272 | 349 | Registrations.map (apsnd (insert_mixin serial' mixin)); | 
| 32801 | 350 | |
| 38107 | 351 | fun get_mixins context (name, morph) = | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 352 | let | 
| 38107 | 353 | val thy = Context.theory_of context; | 
| 354 | val (regs, mixins) = Registrations.get context; | |
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 355 | in | 
| 40782 | 356 | (case Idtab.lookup regs (name, instance_of thy name morph) of | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 357 | NONE => [] | 
| 41272 | 358 | | SOME (_, serial) => lookup_mixins serial mixins) | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 359 | end; | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 360 | |
| 38107 | 361 | fun collect_mixins context (name, morph) = | 
| 362 | let | |
| 363 | val thy = Context.theory_of context; | |
| 364 | in | |
| 365 | roundup thy (fn dep => fn mixins => merge (eq_snd op =) (mixins, get_mixins context dep)) | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 366 | Morphism.identity (name, morph) | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 367 | (insert_idents (name, instance_of thy name morph) empty_idents, []) | 
| 38107 | 368 | |> snd |> filter (snd o fst) (* only inheritable mixins *) | 
| 369 | |> (fn x => merge (eq_snd op =) (x, get_mixins context (name, morph))) | |
| 370 | |> compose_mixins | |
| 371 | end; | |
| 36091 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 ballarin parents: 
36090diff
changeset | 372 | |
| 46858 | 373 | fun get_registrations context select = | 
| 374 | Registrations.get context | |
| 37491 | 375 | |>> Idtab.dest |>> select | 
| 36095 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 ballarin parents: 
36094diff
changeset | 376 | (* with inherited mixins *) | 
| 37105 
e464f84f3680
Store registrations in efficient data structure.
 ballarin parents: 
37104diff
changeset | 377 | |-> (fn regs => fn _ => map (fn ((name, _), ((base, export) ,_)) => | 
| 38107 | 378 | (name, base $> (collect_mixins context (name, base $> export)) $> export)) regs); | 
| 36091 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 ballarin parents: 
36090diff
changeset | 379 | |
| 38111 | 380 | fun registrations_of context name = | 
| 38107 | 381 | get_registrations context (filter (curry (op =) name o fst o fst)); | 
| 37491 | 382 | |
| 46880 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 383 | fun all_registrations_of context = get_registrations context I; | 
| 37973 | 384 | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 385 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 386 | (*** Activate context elements of locale ***) | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 387 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 388 | (* Declarations, facts and entire locale content *) | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 389 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 390 | fun activate_syntax_decls (name, morph) context = | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 391 | let | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 392 | val thy = Context.theory_of context; | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 393 |     val {syntax_decls, ...} = the_locale thy name;
 | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 394 | in | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 395 | context | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 396 | |> fold_rev (fn (decl, _) => decl morph) syntax_decls | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 397 | end; | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 398 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 399 | fun activate_notes activ_elem transfer context export' (name, morph) input = | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 400 | let | 
| 38107 | 401 | val thy = Context.theory_of context; | 
| 46880 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 402 | val mixin = | 
| 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 403 | (case export' of | 
| 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 404 | NONE => Morphism.identity | 
| 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 405 | | SOME export => collect_mixins context (name, morph $> export) $> export); | 
| 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 406 | val morph' = transfer input $> morph $> mixin; | 
| 46881 
b81f1de9f57e
activate_notes in parallel -- to speedup main operation of locale interpretation;
 wenzelm parents: 
46880diff
changeset | 407 | val notes' = | 
| 46978 
23a59a495934
tuned grouping -- merely indicate order of magnitude;
 wenzelm parents: 
46922diff
changeset | 408 | grouped 100 Par_List.map | 
| 46893 
d5bb4c212df1
some grouping of Par_List operations, to adjust granularity;
 wenzelm parents: 
46881diff
changeset | 409 | (Element.transform_ctxt morph' o Notes o #1) (#notes (the_locale thy name)); | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 410 | in | 
| 46880 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 411 | fold_rev (fn elem => fn res => activ_elem (Element.transform_ctxt (transfer res) elem) res) | 
| 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 wenzelm parents: 
46870diff
changeset | 412 | notes' input | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 413 | end; | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 414 | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 415 | fun activate_all name thy activ_elem transfer (marked, input) = | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 416 | let | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 417 |     val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
 | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 418 | val input' = input |> | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 419 | (not (null params) ? | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 420 | activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |> | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 421 | (* FIXME type parameters *) | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 422 | (case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |> | 
| 40782 | 423 | (not (null defs) ? | 
| 424 | activ_elem (Defines (map (fn def => (Attrib.empty_binding, (def, []))) defs))); | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 425 | val activate = activate_notes activ_elem transfer (Context.Theory thy) NONE; | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 426 | in | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 427 | roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input') | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 428 | end; | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 429 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 430 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 431 | (** Public activation functions **) | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 432 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 433 | fun activate_declarations dep = Context.proof_map (fn context => | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 434 | let | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 435 | val thy = Context.theory_of context; | 
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 436 | in | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 437 | roundup thy activate_syntax_decls Morphism.identity dep (Idents.get context, context) | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 438 | |-> Idents.put | 
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 439 | end); | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 440 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 441 | fun activate_facts export dep context = | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 442 | let | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 443 | val thy = Context.theory_of context; | 
| 45587 | 444 | val activate = | 
| 445 | activate_notes Element.init (Element.transfer_morphism o Context.theory_of) context export; | |
| 38316 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 ballarin parents: 
38211diff
changeset | 446 | in | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 447 | roundup thy activate (the_default Morphism.identity export) dep (Idents.get context, context) | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 448 | |-> Idents.put | 
| 38316 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 ballarin parents: 
38211diff
changeset | 449 | end; | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 450 | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 451 | fun init name thy = | 
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 452 | activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of) | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 453 | (empty_idents, Context.Proof (Proof_Context.init_global thy)) | 
| 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 454 | |-> Idents.put |> Context.proof_of; | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 455 | |
| 38211 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 456 | |
| 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 ballarin parents: 
38111diff
changeset | 457 | (*** Add and extend registrations ***) | 
| 32801 | 458 | |
| 38107 | 459 | fun amend_registration (name, morph) mixin export context = | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 460 | let | 
| 38107 | 461 | val thy = Context.theory_of context; | 
| 462 | val regs = Registrations.get context |> fst; | |
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 463 | val base = instance_of thy name (morph $> export); | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 464 | in | 
| 40782 | 465 | (case Idtab.lookup regs (name, base) of | 
| 466 | NONE => | |
| 467 |         error ("No interpretation of locale " ^
 | |
| 41272 | 468 | quote (extern thy name) ^ " with\nparameter instantiation " ^ | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 469 | space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 470 | " available") | 
| 40782 | 471 | | SOME (_, serial') => add_mixin serial' mixin context) | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 472 | end; | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarindiff
changeset | 473 | |
| 32801 | 474 | (* Note that a registration that would be subsumed by an existing one will not be | 
| 475 | generated, and it will not be possible to amend it. *) | |
| 476 | ||
| 38107 | 477 | fun add_registration (name, base_morph) mixin export context = | 
| 32801 | 478 | let | 
| 38107 | 479 | val thy = Context.theory_of context; | 
| 40782 | 480 | val mix = (case mixin of NONE => Morphism.identity | SOME (mix, _) => mix); | 
| 37102 | 481 | val morph = base_morph $> mix; | 
| 482 | val inst = instance_of thy name morph; | |
| 32801 | 483 | in | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 484 | if redundant_ident thy (Idents.get context) (name, inst) | 
| 41272 | 485 | then context (* FIXME amend mixins? *) | 
| 32801 | 486 | else | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 487 | (Idents.get context, context) | 
| 36095 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 ballarin parents: 
36094diff
changeset | 488 | (* add new registrations with inherited mixins *) | 
| 46870 | 489 | |> (snd o roundup thy (add_reg thy export) export (name, morph)) | 
| 36095 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 ballarin parents: 
36094diff
changeset | 490 | (* add mixin *) | 
| 46870 | 491 | |> (case mixin of NONE => I | SOME mixin => amend_registration (name, morph) mixin export) | 
| 36095 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 ballarin parents: 
36094diff
changeset | 492 | (* activate import hierarchy as far as not already active *) | 
| 38316 
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
 ballarin parents: 
38211diff
changeset | 493 | |> activate_facts (SOME export) (name, morph) | 
| 32801 | 494 | end; | 
| 495 | ||
| 29361 | 496 | |
| 31988 | 497 | (*** Dependencies ***) | 
| 498 | ||
| 41272 | 499 | (* | 
| 500 | fun amend_dependency loc (name, morph) mixin export thy = | |
| 501 | let | |
| 502 | val deps = dependencies_of thy loc; | |
| 503 | in | |
| 504 | case AList.lookup (fn ((name, morph), ((name', (morph', _)), _)) => | |
| 46860 
fe8d6532e1c1
clarified total_ident_ord, swapping first argument back to normal (unlike e464f84f3680) -- NB: "fast" ord is erratic anyway;
 wenzelm parents: 
46859diff
changeset | 505 | total_ident_ord ((name, instance_of thy name morph), (name', instance_of thy name' morph')) = EQUAL) deps (name, morph) of | 
| 41272 | 506 |         NONE => error ("Locale " ^
 | 
| 507 | quote (extern thy name) ^ " with\parameter instantiation " ^ | |
| 508 | space_implode " " (map (quote o Syntax.string_of_term_global thy) morph) ^ | |
| 509 | " not a sublocale of " ^ quote (extern thy loc)) | |
| 510 | | SOME (_, serial') => change_locale ... | |
| 511 | end; | |
| 512 | *) | |
| 513 | ||
| 41270 | 514 | fun add_dependency loc (name, morph) mixin export thy = | 
| 37102 | 515 | let | 
| 41272 | 516 | val serial' = serial (); | 
| 517 | val thy' = thy |> | |
| 518 | (change_locale loc o apsnd) | |
| 519 | (apfst (cons ((name, (morph, export)), serial')) #> | |
| 520 | apsnd (case mixin of NONE => I | SOME mixin => insert_mixin serial' mixin)); | |
| 38107 | 521 | val context' = Context.Theory thy'; | 
| 46858 | 522 | val (_, regs) = | 
| 523 | fold_rev (roundup thy' cons export) | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 524 | (registrations_of context' loc) (Idents.get context', []); | 
| 37102 | 525 | in | 
| 526 | thy' | |
| 38107 | 527 | |> fold_rev (fn dep => Context.theory_map (add_registration dep NONE export)) regs | 
| 37102 | 528 | end; | 
| 31988 | 529 | |
| 530 | ||
| 29361 | 531 | (*** Storing results ***) | 
| 532 | ||
| 533 | (* Theorems *) | |
| 534 | ||
| 46906 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 535 | fun add_thmss _ _ [] ctxt = ctxt | 
| 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 536 | | add_thmss loc kind facts ctxt = | 
| 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 537 | ctxt | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47246diff
changeset | 538 | |> Attrib.local_notes kind facts |> snd | 
| 46906 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 539 | |> Proof_Context.background_theory | 
| 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 540 | ((change_locale loc o apfst o apsnd) (cons ((kind, facts), serial ())) #> | 
| 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 541 | (* Registrations *) | 
| 47249 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47246diff
changeset | 542 | (fn thy => | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47246diff
changeset | 543 | fold_rev (fn (_, morph) => | 
| 
c0481c3c2a6c
added Attrib.global_notes/local_notes/generic_notes convenience;
 wenzelm parents: 
47246diff
changeset | 544 | snd o Attrib.global_notes kind (Element.transform_facts morph facts)) | 
| 46906 
3c1787d46935
suppress vacous notes elements, with subtle change of semantics: 'interpret' no longer pulls-in unnamed facts "by fact";
 wenzelm parents: 
46893diff
changeset | 545 | (registrations_of (Context.Theory thy) loc) thy)); | 
| 29361 | 546 | |
| 547 | ||
| 548 | (* Declarations *) | |
| 549 | ||
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 550 | local | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 551 | |
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 552 | fun add_decl loc decl = | 
| 33643 
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
 wenzelm parents: 
33541diff
changeset | 553 | add_thmss loc "" | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 554 | [((Binding.conceal Binding.empty, | 
| 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
33643diff
changeset | 555 | [Attrib.internal (fn phi => Thm.declaration_attribute (K (decl phi)))]), | 
| 33278 | 556 | [([Drule.dummy_thm], [])])]; | 
| 29361 | 557 | |
| 47246 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 558 | in | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 559 | |
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 560 | fun add_declaration loc syntax decl = | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 561 | syntax ? | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 562 | Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ()))) | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 563 | #> add_decl loc decl; | 
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 564 | |
| 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 wenzelm parents: 
47005diff
changeset | 565 | end; | 
| 29361 | 566 | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 567 | |
| 29361 | 568 | (*** Reasoning about locales ***) | 
| 569 | ||
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 570 | (* Storage for witnesses, intro and unfold rules *) | 
| 29361 | 571 | |
| 33519 | 572 | structure Thms = Generic_Data | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 573 | ( | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 574 | type T = thm list * thm list * thm list; | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 575 | val empty = ([], [], []); | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 576 | val extend = I; | 
| 33519 | 577 | fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) = | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 578 | (Thm.merge_thms (witnesses1, witnesses2), | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 579 | Thm.merge_thms (intros1, intros2), | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 580 | Thm.merge_thms (unfolds1, unfolds2)); | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 581 | ); | 
| 29361 | 582 | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 583 | val get_witnesses = #1 o Thms.get o Context.Proof; | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 584 | val get_intros = #2 o Thms.get o Context.Proof; | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 585 | val get_unfolds = #3 o Thms.get o Context.Proof; | 
| 29361 | 586 | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 587 | val witness_add = | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 588 | Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Thm.add_thm th x, y, z))); | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 589 | val intro_add = | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 590 | Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Thm.add_thm th y, z))); | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 591 | val unfold_add = | 
| 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 592 | Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z))); | 
| 29361 | 593 | |
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 594 | |
| 40782 | 595 | (* Tactics *) | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 596 | |
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 597 | fun gen_intro_locales_tac intros_tac eager ctxt = | 
| 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 598 | intros_tac | 
| 30725 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 wenzelm parents: 
30585diff
changeset | 599 | (get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else [])); | 
| 29361 | 600 | |
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 601 | val intro_locales_tac = gen_intro_locales_tac Method.intros_tac; | 
| 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 602 | val try_intro_locales_tac= gen_intro_locales_tac Method.try_intros_tac; | 
| 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 603 | |
| 29361 | 604 | val _ = Context.>> (Context.map_theory | 
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 605 | (Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o try_intro_locales_tac false)) | 
| 30515 | 606 | "back-chain introduction rules of locales without unfolding predicates" #> | 
| 36093 
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
 ballarin parents: 
36091diff
changeset | 607 | Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o try_intro_locales_tac true)) | 
| 30515 | 608 | "back-chain all introduction rules of locales")); | 
| 29361 | 609 | |
| 37471 | 610 | |
| 611 | (*** diagnostic commands and interfaces ***) | |
| 612 | ||
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 613 | val all_locales = Symtab.keys o snd o Locales.get; | 
| 37471 | 614 | |
| 615 | fun print_locales thy = | |
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
41435diff
changeset | 616 |   Pretty.strs ("locales:" ::
 | 
| 42360 | 617 | map #1 (Name_Space.extern_table (Proof_Context.init_global thy) (Locales.get thy))) | 
| 37471 | 618 | |> Pretty.writeln; | 
| 619 | ||
| 620 | fun print_locale thy show_facts raw_name = | |
| 621 | let | |
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46906diff
changeset | 622 | val name = check thy raw_name; | 
| 37471 | 623 | val ctxt = init name thy; | 
| 624 | fun cons_elem (elem as Notes _) = show_facts ? cons elem | |
| 625 | | cons_elem elem = cons elem; | |
| 626 | val elems = | |
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 627 | activate_all name thy cons_elem (K (Element.transfer_morphism thy)) (empty_idents, []) | 
| 37471 | 628 | |> snd |> rev; | 
| 629 | in | |
| 630 | Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems) | |
| 631 | |> Pretty.writeln | |
| 632 | end; | |
| 633 | ||
| 38109 | 634 | fun print_registrations ctxt raw_name = | 
| 635 | let | |
| 42360 | 636 | val thy = Proof_Context.theory_of ctxt; | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
46906diff
changeset | 637 | val name = check thy raw_name; | 
| 38109 | 638 | in | 
| 41435 | 639 | (case registrations_of (Context.Proof ctxt) (* FIXME *) name of | 
| 46859 | 640 | [] => Pretty.str "no interpretations" | 
| 41435 | 641 | | regs => Pretty.big_list "interpretations:" (map (pretty_reg ctxt) (rev regs))) | 
| 642 | end |> Pretty.writeln; | |
| 643 | ||
| 644 | fun print_dependencies ctxt clean export insts = | |
| 645 | let | |
| 42360 | 646 | val thy = Proof_Context.theory_of ctxt; | 
| 46862 
ef45df55127d
clarified idents for activated locales, with subtle change of semantics: insert/merge wrt. term equality, but avoid redundant_ident on retrieval;
 wenzelm parents: 
46860diff
changeset | 647 | val idents = if clean then empty_idents else Idents.get (Context.Proof ctxt); | 
| 41435 | 648 | in | 
| 649 | (case fold (roundup thy cons export) insts (idents, []) |> snd of | |
| 46859 | 650 | [] => Pretty.str "no dependencies" | 
| 41435 | 651 | | deps => Pretty.big_list "dependencies:" (map (pretty_reg ctxt) (rev deps))) | 
| 40782 | 652 | end |> Pretty.writeln; | 
| 37471 | 653 | |
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 654 | fun locale_deps thy = | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 655 | let | 
| 46858 | 656 | val names = all_locales thy; | 
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 657 | fun add_locale_node name = | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 658 | let | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 659 | val params = params_of thy name; | 
| 40782 | 660 | val axioms = | 
| 661 | these (Option.map (Logic.strip_imp_prems o Thm.prop_of) (fst (intros_of thy name))); | |
| 662 | val registrations = | |
| 663 | map (instance_of thy name o snd) (registrations_of (Context.Theory thy) name); | |
| 664 | in | |
| 665 |         Graph.new_node (name, {params = params, axioms = axioms, registrations = registrations})
 | |
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 666 | end; | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 667 | fun add_locale_deps name = | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 668 | let | 
| 40782 | 669 | val dependencies = | 
| 45587 | 670 | map (apsnd (instance_of thy name o op $>) o fst) (dependencies_of thy name); | 
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 671 | in | 
| 46858 | 672 | fold (fn (super, ts) => fn (gr, deps) => | 
| 673 | (gr |> Graph.add_edge (super, name), | |
| 37897 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 674 | deps |> Symtab.map_default (super, Symtab.empty) (Symtab.cons_list (name, ts)))) | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 675 | dependencies | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 676 | end; | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 677 | in | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 678 | Graph.empty | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 679 | |> fold add_locale_node names | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 680 | |> rpair Symtab.empty | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 681 | |> fold add_locale_deps names | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 682 | end; | 
| 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 haftmann parents: 
37491diff
changeset | 683 | |
| 29361 | 684 | end; |