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