| author | nipkow | 
| Thu, 03 Feb 2022 10:33:55 +0100 | |
| changeset 75061 | 57df04e4f018 | 
| parent 74563 | 042041c0ebeb | 
| child 77738 | e64428b6b170 | 
| 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  | 
|
| 66581 | 8  | 
meta-logic. Furthermore, structured composition of contexts (with merge  | 
9  | 
and instantiation) is provided, as well as type-inference of the  | 
|
10  | 
signature parts and predicate definitions of the specification text.  | 
|
| 29361 | 11  | 
|
| 66581 | 12  | 
Interpretation enables the transfer of declarations and theorems to other  | 
13  | 
contexts, namely those defined by theories, structured proofs and locales  | 
|
14  | 
themselves.  | 
|
15  | 
||
16  | 
A comprehensive account of locales is available:  | 
|
17  | 
||
18  | 
[1] Clemens Ballarin. Locales: a module system for mathematical theories.  | 
|
| 66588 | 19  | 
Journal of Automated Reasoning, 52(2):123-153, 2014.  | 
| 29361 | 20  | 
|
21  | 
See also:  | 
|
22  | 
||
| 66581 | 23  | 
[2] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar.  | 
| 29361 | 24  | 
In Stefano Berardi et al., Types for Proofs and Programs: International  | 
25  | 
Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004.  | 
|
| 66581 | 26  | 
[3] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing  | 
| 29361 | 27  | 
Dependencies between Locales. Technical Report TUM-I0607, Technische  | 
28  | 
Universitaet Muenchen, 2006.  | 
|
| 66581 | 29  | 
[4] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and  | 
| 29361 | 30  | 
Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108,  | 
31  | 
pages 31-43, 2006.  | 
|
32  | 
*)  | 
|
33  | 
||
34  | 
signature LOCALE =  | 
|
35  | 
sig  | 
|
| 29576 | 36  | 
(* Locale specification *)  | 
| 
30344
 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 
wenzelm 
parents: 
30223 
diff
changeset
 | 
37  | 
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
 | 
38  | 
(string * sort) list * ((string * typ) * mixfix) list ->  | 
| 29361 | 39  | 
term option * term list ->  | 
| 29441 | 40  | 
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
 | 
41  | 
Element.context_i list ->  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
42  | 
declaration list ->  | 
| 67652 | 43  | 
(string * Attrib.fact list) list ->  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
44  | 
(string * morphism) list -> theory -> theory  | 
| 68862 | 45  | 
val locale_space: theory -> Name_Space.T  | 
| 
55728
 
aaf024d63b35
reverted c05d3e22adaf: Locale.intern is still required by AFP/Simpl;
 
wenzelm 
parents: 
55695 
diff
changeset
 | 
46  | 
val intern: theory -> xstring -> string  | 
| 
45488
 
6d71d9e52369
pass positions for named targets, for formal links in the document model;
 
wenzelm 
parents: 
45390 
diff
changeset
 | 
47  | 
val check: theory -> xstring * Position.T -> string  | 
| 29361 | 48  | 
val extern: theory -> string -> xstring  | 
| 53041 | 49  | 
val markup_name: Proof.context -> string -> string  | 
| 52103 | 50  | 
val pretty_name: Proof.context -> string -> Pretty.T  | 
| 29392 | 51  | 
val defined: theory -> string -> bool  | 
| 69068 | 52  | 
val parameters_of: theory -> string -> (string * sort) list * ((string * typ) * mixfix) list  | 
| 
30755
 
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
 
wenzelm 
parents: 
30754 
diff
changeset
 | 
53  | 
val params_of: theory -> string -> ((string * typ) * mixfix) list  | 
| 29441 | 54  | 
val intros_of: theory -> string -> thm option * thm option  | 
55  | 
val axioms_of: theory -> string -> thm list  | 
|
| 29544 | 56  | 
val instance_of: theory -> string -> morphism -> term list  | 
| 29361 | 57  | 
val specification_of: theory -> string -> term option * term list  | 
| 68861 | 58  | 
val hyp_spec_of: theory -> string -> Element.context_i list  | 
| 29361 | 59  | 
|
60  | 
(* Storing results *)  | 
|
| 67654 | 61  | 
val add_facts: string -> string -> Attrib.fact list -> Proof.context -> Proof.context  | 
| 
47246
 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 
wenzelm 
parents: 
47005 
diff
changeset
 | 
62  | 
val add_declaration: string -> bool -> declaration -> Proof.context -> Proof.context  | 
| 29361 | 63  | 
|
64  | 
(* Activation *)  | 
|
| 67666 | 65  | 
val activate_facts: morphism option -> string * morphism -> Context.generic -> Context.generic  | 
| 
30764
 
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
66  | 
val activate_declarations: string * morphism -> Proof.context -> Proof.context  | 
| 29361 | 67  | 
val init: string -> theory -> Proof.context  | 
68  | 
||
69  | 
(* Reasoning about locales *)  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
70  | 
val get_witnesses: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
71  | 
val get_intros: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
72  | 
val get_unfolds: Proof.context -> thm list  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
73  | 
val witness_add: attribute  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
74  | 
val intro_add: attribute  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
75  | 
val unfold_add: attribute  | 
| 69017 | 76  | 
  val intro_locales_tac: {strict: bool, eager: bool} -> Proof.context -> thm list -> tactic
 | 
| 29361 | 77  | 
|
| 31988 | 78  | 
(* Registrations and dependencies *)  | 
| 69058 | 79  | 
  type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism}
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
80  | 
val amend_registration: registration -> Context.generic -> Context.generic  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
81  | 
val add_registration: registration -> Context.generic -> Context.generic  | 
| 38111 | 82  | 
val registrations_of: Context.generic -> string -> (string * morphism) list  | 
| 73845 | 83  | 
val add_dependency: string -> registration -> theory -> theory  | 
| 29361 | 84  | 
|
85  | 
(* Diagnostic *)  | 
|
| 
69029
 
aec64b88e708
clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
 
wenzelm 
parents: 
69017 
diff
changeset
 | 
86  | 
val get_locales: theory -> string list  | 
| 
69059
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
87  | 
val pretty_locales: theory -> bool -> Pretty.T  | 
| 
69057
 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69056 
diff
changeset
 | 
88  | 
val pretty_locale: theory -> bool -> string -> Pretty.T  | 
| 
69059
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
89  | 
val pretty_registrations: Proof.context -> string -> Pretty.T  | 
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
90  | 
  val pretty_locale_deps: theory -> {name: string, parents: string list, body: Pretty.T} list
 | 
| 69063 | 91  | 
type locale_dependency =  | 
92  | 
    {source: string, target: string, prefix: (string * bool) list, morphism: morphism,
 | 
|
93  | 
pos: Position.T, serial: serial}  | 
|
94  | 
val dest_dependencies: theory list -> theory -> locale_dependency list  | 
|
| 
70622
 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 
ballarin 
parents: 
70608 
diff
changeset
 | 
95  | 
val tracing : Proof.context -> string -> unit  | 
| 29361 | 96  | 
end;  | 
97  | 
||
98  | 
structure Locale: LOCALE =  | 
|
99  | 
struct  | 
|
100  | 
||
101  | 
datatype ctxt = datatype Element.ctxt;  | 
|
102  | 
||
| 29392 | 103  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
104  | 
(*** Locales ***)  | 
| 29361 | 105  | 
|
| 69063 | 106  | 
type dep = {name: string, morphisms: morphism * morphism, pos: Position.T, serial: serial};
 | 
107  | 
fun eq_dep (dep1: dep, dep2: dep) = #serial dep1 = #serial dep2;  | 
|
| 69061 | 108  | 
|
109  | 
fun make_dep (name, morphisms) : dep =  | 
|
| 69063 | 110  | 
  {name = name, morphisms = morphisms, pos = Position.thread_data (), serial = serial ()};
 | 
| 69061 | 111  | 
|
| 69046 | 112  | 
(*table of mixin lists, per list mixins in reverse order of declaration;  | 
113  | 
lists indexed by registration/dependency serial,  | 
|
114  | 
entries for empty lists may be omitted*)  | 
|
115  | 
type mixin = (morphism * bool) * serial;  | 
|
116  | 
type mixins = mixin list Inttab.table;  | 
|
| 41272 | 117  | 
|
| 69047 | 118  | 
fun lookup_mixins (mixins: mixins) serial' = Inttab.lookup_list mixins serial';  | 
| 41272 | 119  | 
|
| 69046 | 120  | 
val merge_mixins: mixins * mixins -> mixins = Inttab.merge_list (eq_snd op =);  | 
121  | 
||
122  | 
fun insert_mixin serial' mixin : mixins -> mixins = Inttab.cons_list (serial', (mixin, serial ()));  | 
|
| 41272 | 123  | 
|
| 69046 | 124  | 
fun rename_mixin (old, new) (mixins: mixins) =  | 
125  | 
(case Inttab.lookup mixins old of  | 
|
126  | 
NONE => mixins  | 
|
127  | 
| SOME mixin => Inttab.delete old mixins |> Inttab.update_new (new, mixin));  | 
|
| 41272 | 128  | 
|
| 69046 | 129  | 
fun compose_mixins (mixins: mixin list) =  | 
| 41272 | 130  | 
fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity;  | 
131  | 
||
| 29361 | 132  | 
datatype locale = Loc of {
 | 
| 63029 | 133  | 
(* static part *)  | 
134  | 
||
135  | 
(*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
 | 
136  | 
parameters: (string * sort) list * ((string * typ) * mixfix) list,  | 
| 63029 | 137  | 
(*assumptions (as a single predicate expression) and defines*)  | 
| 29361 | 138  | 
spec: term option * term list,  | 
| 29441 | 139  | 
intros: thm option * thm option,  | 
140  | 
axioms: thm list,  | 
|
| 63029 | 141  | 
(*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
 | 
142  | 
hyp_spec: Element.context_i list,  | 
| 63029 | 143  | 
|
144  | 
(* dynamic part *)  | 
|
145  | 
||
146  | 
(*syntax declarations*)  | 
|
| 36096 | 147  | 
syntax_decls: (declaration * serial) list,  | 
| 63029 | 148  | 
(*theorem declarations*)  | 
| 67652 | 149  | 
notes: ((string * Attrib.fact list) * serial) list,  | 
| 63029 | 150  | 
(*locale dependencies (sublocale relation) in reverse order*)  | 
| 69061 | 151  | 
dependencies: dep list,  | 
| 63029 | 152  | 
(*mixin part of dependencies*)  | 
| 41272 | 153  | 
mixins: mixins  | 
| 29392 | 154  | 
};  | 
| 29361 | 155  | 
|
| 
59296
 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 
haftmann 
parents: 
58028 
diff
changeset
 | 
156  | 
fun mk_locale ((parameters, spec, intros, axioms, hyp_spec),  | 
| 41272 | 157  | 
((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
 | 
158  | 
  Loc {parameters = parameters, spec = spec, intros = intros, axioms = axioms, hyp_spec = hyp_spec,
 | 
| 41272 | 159  | 
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
 | 
160  | 
|
| 
59296
 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 
haftmann 
parents: 
58028 
diff
changeset
 | 
161  | 
fun map_locale f (Loc {parameters, spec, intros, axioms, hyp_spec,
 | 
| 41272 | 162  | 
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
 | 
163  | 
mk_locale (f ((parameters, spec, intros, axioms, hyp_spec),  | 
| 41272 | 164  | 
((syntax_decls, notes), (dependencies, mixins))));  | 
| 30754 | 165  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
166  | 
fun merge_locale  | 
| 
59296
 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 
haftmann 
parents: 
58028 
diff
changeset
 | 
167  | 
 (Loc {parameters, spec, intros, axioms, hyp_spec, syntax_decls, notes, dependencies, mixins},
 | 
| 41272 | 168  | 
  Loc {syntax_decls = syntax_decls', notes = notes',
 | 
169  | 
dependencies = dependencies', mixins = mixins', ...}) =  | 
|
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
170  | 
mk_locale  | 
| 
59296
 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 
haftmann 
parents: 
58028 
diff
changeset
 | 
171  | 
((parameters, spec, intros, axioms, hyp_spec),  | 
| 
35798
 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 
wenzelm 
parents: 
33643 
diff
changeset
 | 
172  | 
((merge (eq_snd op =) (syntax_decls, syntax_decls'),  | 
| 29441 | 173  | 
merge (eq_snd op =) (notes, notes')),  | 
| 69063 | 174  | 
(merge eq_dep (dependencies, dependencies'),  | 
| 41272 | 175  | 
(merge_mixins (mixins, mixins')))));  | 
| 29361 | 176  | 
|
| 33522 | 177  | 
structure Locales = Theory_Data  | 
| 29361 | 178  | 
(  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
179  | 
type T = locale Name_Space.table;  | 
| 74112 | 180  | 
val empty : T = Name_Space.empty_table Markup.localeN;  | 
| 33522 | 181  | 
val merge = Name_Space.join_tables (K merge_locale);  | 
| 29361 | 182  | 
);  | 
183  | 
||
| 56025 | 184  | 
val locale_space = Name_Space.space_of_table o Locales.get;  | 
185  | 
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
 | 
186  | 
fun check thy = #1 o Name_Space.check (Context.Theory thy) (Locales.get thy);  | 
| 53041 | 187  | 
|
| 74559 | 188  | 
val _ = Theory.setup  | 
189  | 
(ML_Antiquotation.inline_embedded \<^binding>\<open>locale\<close>  | 
|
| 74563 | 190  | 
(Args.theory -- Scan.lift Parse.embedded_position >>  | 
| 74559 | 191  | 
(ML_Syntax.print_string o uncurry check)));  | 
192  | 
||
| 56025 | 193  | 
fun extern thy = Name_Space.extern (Proof_Context.init_global thy) (locale_space thy);  | 
| 29392 | 194  | 
|
| 53041 | 195  | 
fun markup_extern ctxt =  | 
| 56025 | 196  | 
Name_Space.markup_extern ctxt (locale_space (Proof_Context.theory_of ctxt));  | 
| 53041 | 197  | 
|
198  | 
fun markup_name ctxt name = markup_extern ctxt name |-> Markup.markup;  | 
|
199  | 
fun pretty_name ctxt name = markup_extern ctxt name |> Pretty.mark_str;  | 
|
| 52103 | 200  | 
|
| 59884 | 201  | 
val get_locale = Name_Space.lookup o Locales.get;  | 
| 56025 | 202  | 
val defined = is_some oo get_locale;  | 
| 29361 | 203  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
204  | 
fun the_locale thy name =  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
205  | 
(case get_locale thy name of  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
206  | 
SOME (Loc loc) => loc  | 
| 
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
207  | 
  | NONE => error ("Unknown locale " ^ quote name));
 | 
| 29361 | 208  | 
|
| 69061 | 209  | 
fun register_locale  | 
210  | 
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
 | 
211  | 
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
 | 
212  | 
(binding,  | 
| 61093 | 213  | 
mk_locale ((parameters, spec, (apply2 o Option.map) Thm.trim_context intros,  | 
214  | 
map Thm.trim_context axioms, hyp_spec),  | 
|
| 36096 | 215  | 
((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes),  | 
| 69061 | 216  | 
(map (fn (name, morph) => make_dep (name, (morph, Morphism.identity))) dependencies,  | 
| 41272 | 217  | 
Inttab.empty)))) #> snd);  | 
218  | 
(* FIXME Morphism.identity *)  | 
|
| 29361 | 219  | 
|
| 29392 | 220  | 
fun change_locale name =  | 
| 56025 | 221  | 
Locales.map o Name_Space.map_table_entry name o map_locale o apsnd;  | 
| 29361 | 222  | 
|
223  | 
||
| 68862 | 224  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
225  | 
(** Primitive operations **)  | 
| 29361 | 226  | 
|
| 69068 | 227  | 
fun parameters_of thy = #parameters o the_locale thy;  | 
228  | 
val params_of = #2 oo parameters_of;  | 
|
| 29361 | 229  | 
|
| 61093 | 230  | 
fun intros_of thy = (apply2 o Option.map) (Thm.transfer thy) o #intros o the_locale thy;  | 
| 29441 | 231  | 
|
| 61093 | 232  | 
fun axioms_of thy = map (Thm.transfer thy) o #axioms o the_locale thy;  | 
| 29441 | 233  | 
|
| 29392 | 234  | 
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
 | 
235  | 
map (Morphism.term morph o Free o #1);  | 
| 29361 | 236  | 
|
| 29392 | 237  | 
fun specification_of thy = #spec o the_locale thy;  | 
| 29361 | 238  | 
|
| 68861 | 239  | 
fun hyp_spec_of thy = #hyp_spec o the_locale thy;  | 
240  | 
||
241  | 
fun dependencies_of thy = #dependencies o the_locale thy;  | 
|
| 41272 | 242  | 
|
| 69047 | 243  | 
fun mixins_of thy name serial = lookup_mixins (#mixins (the_locale thy name)) serial;  | 
| 41272 | 244  | 
|
| 29544 | 245  | 
|
| 
37133
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
246  | 
(* Print instance and qualifiers *)  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
247  | 
|
| 69052 | 248  | 
fun pretty_reg_inst ctxt qs (name, ts) =  | 
| 
37133
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
249  | 
let  | 
| 61731 | 250  | 
fun print_qual (qual, mandatory) = qual ^ (if mandatory then "" else "?");  | 
| 52431 | 251  | 
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
 | 
252  | 
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
 | 
253  | 
fun prt_term' t =  | 
| 
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
38797 
diff
changeset
 | 
254  | 
if Config.get ctxt show_types  | 
| 
37133
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
255  | 
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
 | 
256  | 
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
 | 
257  | 
else prt_term t;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
258  | 
fun prt_inst ts =  | 
| 52103 | 259  | 
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
 | 
260  | 
in  | 
| 40782 | 261  | 
(case qs of  | 
262  | 
[] => prt_inst ts  | 
|
263  | 
| 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
 | 
264  | 
end;  | 
| 
 
1d048c6940c8
Merge mixins of distinct interpretations with same base.
 
ballarin 
parents: 
37105 
diff
changeset
 | 
265  | 
|
| 69052 | 266  | 
fun pretty_reg ctxt export (name, morph) =  | 
267  | 
let  | 
|
268  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
269  | 
val morph' = morph $> export;  | 
|
| 69062 | 270  | 
val qs = Morphism.binding_prefix morph';  | 
| 69052 | 271  | 
val ts = instance_of thy name morph';  | 
272  | 
in pretty_reg_inst ctxt qs (name, ts) end;  | 
|
273  | 
||
| 29361 | 274  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
275  | 
(*** Identifiers: activated locales in theory or proof context ***)  | 
| 29361 | 276  | 
|
| 
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
 | 
277  | 
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
 | 
278  | 
|
| 
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
 | 
279  | 
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
 | 
280  | 
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
 | 
281  | 
val merge_idents = Symtab.merge_list (eq_list (op aconv));  | 
| 29361 | 282  | 
|
| 
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
 | 
283  | 
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
 | 
284  | 
exists (fn pat => Pattern.matchess thy (pat, instance)) (Symtab.lookup_list idents name);  | 
| 29361 | 285  | 
|
| 
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
 | 
286  | 
structure Idents = Generic_Data  | 
| 29361 | 287  | 
(  | 
| 
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
 | 
288  | 
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
 | 
289  | 
val empty = empty_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
 | 
290  | 
val merge = merge_idents;  | 
| 29361 | 291  | 
);  | 
292  | 
||
293  | 
||
294  | 
(** Resolve locale dependencies in a depth-first fashion **)  | 
|
295  | 
||
296  | 
local  | 
|
297  | 
||
298  | 
val roundup_bound = 120;  | 
|
299  | 
||
| 51515 | 300  | 
fun add thy depth stem export (name, morph) (deps, marked) =  | 
| 29361 | 301  | 
if depth > roundup_bound  | 
302  | 
then error "Roundup bound exceeded (sublocale relation probably not terminating)."  | 
|
303  | 
else  | 
|
304  | 
let  | 
|
| 41272 | 305  | 
val instance = instance_of thy name (morph $> stem $> export);  | 
| 29361 | 306  | 
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
 | 
307  | 
if redundant_ident thy marked (name, instance) then (deps, marked)  | 
| 29361 | 308  | 
else  | 
309  | 
let  | 
|
| 69061 | 310  | 
(*no inheritance of mixins, regardless of requests by clients*)  | 
311  | 
val dependencies =  | 
|
312  | 
            dependencies_of thy name |> map (fn dep as {morphisms = (morph', export'), ...} =>
 | 
|
313  | 
(#name dep, morph' $> export' $> compose_mixins (mixins_of thy name (#serial 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
 | 
314  | 
val marked' = insert_idents (name, instance) marked;  | 
| 41272 | 315  | 
val (deps', marked'') =  | 
| 51515 | 316  | 
fold_rev (add thy (depth + 1) (morph $> stem) export) dependencies  | 
| 41272 | 317  | 
([], marked');  | 
| 69061 | 318  | 
in ((name, morph $> stem) :: deps' @ deps, marked'') end  | 
| 29361 | 319  | 
end;  | 
320  | 
||
321  | 
in  | 
|
322  | 
||
| 33541 | 323  | 
(* Note that while identifiers always have the external (exported) view, activate_dep  | 
| 46870 | 324  | 
is presented with the internal view. *)  | 
| 32803 | 325  | 
|
326  | 
fun roundup thy activate_dep export (name, morph) (marked, input) =  | 
|
| 29361 | 327  | 
let  | 
| 41272 | 328  | 
(* Find all dependencies including new ones (which are dependencies enriching  | 
| 29361 | 329  | 
existing registrations). *)  | 
| 41272 | 330  | 
val (dependencies, marked') =  | 
| 51515 | 331  | 
add thy 0 Morphism.identity export (name, morph) ([], empty_idents);  | 
| 32800 | 332  | 
(* Filter out fragments from marked; these won't be activated. *)  | 
| 29361 | 333  | 
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
 | 
334  | 
redundant_ident thy marked (name, instance_of thy name (morph $> export))) dependencies;  | 
| 29361 | 335  | 
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
 | 
336  | 
(merge_idents (marked, marked'), input |> fold_rev activate_dep dependencies')  | 
| 29361 | 337  | 
end;  | 
338  | 
||
339  | 
end;  | 
|
340  | 
||
341  | 
||
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
342  | 
(*** Registrations: interpretations in theories or proof contexts ***)  | 
| 29361 | 343  | 
|
| 
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
 | 
344  | 
val total_ident_ord = prod_ord fast_string_ord (list_ord Term_Ord.fast_term_ord);  | 
| 69051 | 345  | 
structure Idtab = Table(type key = string * term list val ord = total_ident_ord);  | 
| 
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
 | 
346  | 
|
| 69053 | 347  | 
type reg = {morphisms: morphism * morphism, pos: Position.T, serial: serial};
 | 
| 69051 | 348  | 
type regs = reg Idtab.table;  | 
349  | 
||
350  | 
val join_regs : regs * regs -> regs =  | 
|
351  | 
Idtab.join (fn id => fn (reg1, reg2) =>  | 
|
352  | 
if #serial reg1 = #serial reg2 then raise Idtab.SAME else raise Idtab.DUP id);  | 
|
| 
37105
 
e464f84f3680
Store registrations in efficient data structure.
 
ballarin 
parents: 
37104 
diff
changeset
 | 
353  | 
|
| 69052 | 354  | 
(* FIXME consolidate with locale dependencies, consider one data slot only *)  | 
355  | 
structure Global_Registrations = Theory_Data'  | 
|
| 29361 | 356  | 
(  | 
| 69046 | 357  | 
(*registrations, indexed by locale name and instance;  | 
358  | 
unique registration serial points to mixin list*)  | 
|
| 69051 | 359  | 
type T = regs * mixins;  | 
| 69052 | 360  | 
val empty: T = (Idtab.empty, Inttab.empty);  | 
361  | 
fun merge old_thys =  | 
|
362  | 
let  | 
|
363  | 
fun recursive_merge ((regs1, mixins1), (regs2, mixins2)) : T =  | 
|
364  | 
(join_regs (regs1, regs2), merge_mixins (mixins1, mixins2))  | 
|
365  | 
handle Idtab.DUP id =>  | 
|
366  | 
(*distinct interpretations with same base: merge their mixins*)  | 
|
367  | 
let  | 
|
| 69053 | 368  | 
val reg1 = Idtab.lookup regs1 id |> the;  | 
| 69052 | 369  | 
val reg2 = Idtab.lookup regs2 id |> the;  | 
| 69053 | 370  | 
val reg2' =  | 
371  | 
             {morphisms = #morphisms reg2,
 | 
|
372  | 
pos = Position.thread_data (),  | 
|
373  | 
serial = #serial reg1};  | 
|
374  | 
val regs2' = Idtab.update (id, reg2') regs2;  | 
|
375  | 
val mixins2' = rename_mixin (#serial reg2, #serial reg1) mixins2;  | 
|
| 69052 | 376  | 
val _ =  | 
| 69053 | 377  | 
              warning ("Removed duplicate interpretation after retrieving its mixins" ^
 | 
378  | 
Position.here_list [#pos reg1, #pos reg2] ^ ":\n " ^  | 
|
379  | 
Pretty.string_of (pretty_reg_inst (Syntax.init_pretty_global (#1 old_thys)) [] id));  | 
|
| 69052 | 380  | 
in recursive_merge ((regs1, mixins1), (regs2', mixins2')) end;  | 
381  | 
in recursive_merge end;  | 
|
| 29361 | 382  | 
);  | 
383  | 
||
| 69052 | 384  | 
structure Local_Registrations = Proof_Data  | 
385  | 
(  | 
|
386  | 
type T = Global_Registrations.T;  | 
|
387  | 
val init = Global_Registrations.get;  | 
|
388  | 
);  | 
|
389  | 
||
390  | 
val get_registrations = Context.cases Global_Registrations.get Local_Registrations.get;  | 
|
391  | 
||
392  | 
fun map_registrations f (Context.Theory thy) = Context.Theory (Global_Registrations.map f thy)  | 
|
393  | 
| map_registrations f (Context.Proof ctxt) = Context.Proof (Local_Registrations.map f ctxt);  | 
|
394  | 
||
| 32801 | 395  | 
|
396  | 
(* Primitive operations *)  | 
|
397  | 
||
| 
37104
 
3877a6c45d57
Avoid recomputation of registration instance for lookup.
 
ballarin 
parents: 
37103 
diff
changeset
 | 
398  | 
fun add_reg thy export (name, morph) =  | 
| 69051 | 399  | 
let  | 
| 69053 | 400  | 
    val reg = {morphisms = (morph, export), pos = Position.thread_data (), serial = serial ()};
 | 
| 69051 | 401  | 
val id = (name, instance_of thy name (morph $> export));  | 
| 69052 | 402  | 
in (map_registrations o apfst) (Idtab.insert (K false) (id, reg)) end;  | 
| 31988 | 403  | 
|
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
404  | 
fun add_mixin serial' mixin =  | 
| 
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
405  | 
(* registration to be amended identified by its serial id *)  | 
| 69052 | 406  | 
(map_registrations o apsnd) (insert_mixin serial' mixin);  | 
| 69046 | 407  | 
|
| 69052 | 408  | 
val get_regs = #1 o get_registrations;  | 
| 32801 | 409  | 
|
| 38107 | 410  | 
fun get_mixins context (name, morph) =  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
411  | 
let  | 
| 38107 | 412  | 
val thy = Context.theory_of context;  | 
| 69052 | 413  | 
val (regs, mixins) = get_registrations context;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
414  | 
in  | 
| 40782 | 415  | 
(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
 | 
416  | 
NONE => []  | 
| 69051 | 417  | 
    | SOME {serial, ...} => lookup_mixins mixins serial)
 | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
418  | 
end;  | 
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
419  | 
|
| 38107 | 420  | 
fun collect_mixins context (name, morph) =  | 
421  | 
let  | 
|
422  | 
val thy = Context.theory_of context;  | 
|
423  | 
in  | 
|
424  | 
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
 | 
425  | 
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
 | 
426  | 
(insert_idents (name, instance_of thy name morph) empty_idents, [])  | 
| 38107 | 427  | 
|> snd |> filter (snd o fst) (* only inheritable mixins *)  | 
428  | 
|> (fn x => merge (eq_snd op =) (x, get_mixins context (name, morph)))  | 
|
429  | 
|> compose_mixins  | 
|
430  | 
end;  | 
|
| 
36091
 
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
 
ballarin 
parents: 
36090 
diff
changeset
 | 
431  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
432  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
433  | 
(*** Activate context elements of locale ***)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
434  | 
|
| 70608 | 435  | 
fun activate_err msg kind (name, morph) context =  | 
436  | 
  cat_error msg ("The above error(s) occurred while activating " ^ kind ^ " of locale instance\n" ^
 | 
|
| 
63723
 
dacc380ab327
Improved error reporting when activating a locale instance.
 
ballarin 
parents: 
63352 
diff
changeset
 | 
437  | 
(pretty_reg (Context.proof_of context) Morphism.identity (name, morph) |>  | 
| 
 
dacc380ab327
Improved error reporting when activating a locale instance.
 
ballarin 
parents: 
63352 
diff
changeset
 | 
438  | 
Pretty.string_of));  | 
| 
 
dacc380ab327
Improved error reporting when activating a locale instance.
 
ballarin 
parents: 
63352 
diff
changeset
 | 
439  | 
|
| 67667 | 440  | 
fun init_element elem context =  | 
441  | 
context  | 
|
442  | 
|> Context.mapping I (Thm.unchecked_hyps #> Context_Position.not_really)  | 
|
443  | 
|> Element.init elem  | 
|
444  | 
|> Context.mapping I (fn ctxt =>  | 
|
445  | 
let val ctxt0 = Context.proof_of context  | 
|
446  | 
in ctxt |> Context_Position.restore_visible ctxt0 |> Thm.restore_hyps ctxt0 end);  | 
|
447  | 
||
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
448  | 
|
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
449  | 
(* Potentially lazy notes *)  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
450  | 
|
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
451  | 
fun make_notes kind = map (fn ((b, atts), facts) =>  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
452  | 
if null atts andalso forall (null o #2) facts  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
453  | 
then Lazy_Notes (kind, (b, Lazy.value (maps #1 facts)))  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
454  | 
else Notes (kind, [((b, atts), facts)]));  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
455  | 
|
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
456  | 
fun lazy_notes thy loc =  | 
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
457  | 
rev (#notes (the_locale thy loc))  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
458  | 
|> maps (fn ((kind, notes), _) => make_notes kind notes);  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
459  | 
|
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
460  | 
fun consolidate_notes elems =  | 
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
461  | 
elems  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
462  | 
|> map_filter (fn Lazy_Notes (_, (_, ths)) => SOME ths | _ => NONE)  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
463  | 
|> Lazy.consolidate  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
464  | 
|> ignore;  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
465  | 
|
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
466  | 
fun force_notes (Lazy_Notes (kind, (b, ths))) = Notes (kind, [((b, []), [(Lazy.force ths, [])])])  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
467  | 
| force_notes elem = elem;  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
468  | 
|
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
469  | 
|
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
470  | 
(* Declarations, facts and entire locale content *)  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
471  | 
|
| 70608 | 472  | 
val trace_locales =  | 
473  | 
  Attrib.setup_config_bool (Binding.make ("trace_locales", \<^here>)) (K false);
 | 
|
474  | 
||
| 
70622
 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 
ballarin 
parents: 
70608 
diff
changeset
 | 
475  | 
fun tracing context msg =  | 
| 
 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 
ballarin 
parents: 
70608 
diff
changeset
 | 
476  | 
if Config.get context trace_locales then Output.tracing msg else ();  | 
| 
 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 
ballarin 
parents: 
70608 
diff
changeset
 | 
477  | 
|
| 70608 | 478  | 
fun trace kind (name, morph) context =  | 
| 
70622
 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 
ballarin 
parents: 
70608 
diff
changeset
 | 
479  | 
  tracing (Context.proof_of context) ("Activating " ^ kind ^ " of " ^
 | 
| 
 
2fb2e7661e16
Integrate locale activation fallback diagnostics with 'trace_locales'.
 
ballarin 
parents: 
70608 
diff
changeset
 | 
480  | 
(pretty_reg (Context.proof_of context) Morphism.identity (name, morph) |> Pretty.string_of));  | 
| 70608 | 481  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
482  | 
fun activate_syntax_decls (name, morph) context =  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
483  | 
let  | 
| 70608 | 484  | 
val _ = trace "syntax" (name, morph) context;  | 
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
485  | 
val thy = Context.theory_of context;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
486  | 
    val {syntax_decls, ...} = the_locale thy name;
 | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
487  | 
in  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
488  | 
context  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
489  | 
|> fold_rev (fn (decl, _) => decl morph) syntax_decls  | 
| 70608 | 490  | 
handle ERROR msg => activate_err msg "syntax" (name, morph) context  | 
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
491  | 
end;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
492  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
493  | 
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
 | 
494  | 
let  | 
| 38107 | 495  | 
val thy = Context.theory_of context;  | 
| 
46880
 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 
wenzelm 
parents: 
46870 
diff
changeset
 | 
496  | 
val mixin =  | 
| 
 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 
wenzelm 
parents: 
46870 
diff
changeset
 | 
497  | 
(case export' of  | 
| 
 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 
wenzelm 
parents: 
46870 
diff
changeset
 | 
498  | 
NONE => Morphism.identity  | 
| 
 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 
wenzelm 
parents: 
46870 
diff
changeset
 | 
499  | 
| SOME export => collect_mixins context (name, morph $> export) $> export);  | 
| 
 
af8e516953ce
refined activate_notes: simultaneous transformation before activation;
 
wenzelm 
parents: 
46870 
diff
changeset
 | 
500  | 
val morph' = transfer input $> morph $> mixin;  | 
| 
67680
 
175a070e9dd8
eliminated questionable Par_List.map -- locale interpretation is mostly lazy (see also b81f1de9f57e);
 
wenzelm 
parents: 
67677 
diff
changeset
 | 
501  | 
val notes' = map (Element.transform_ctxt morph') (lazy_notes thy name);  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
502  | 
in  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
503  | 
(notes', input) |-> fold (fn elem => fn res =>  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
504  | 
activ_elem (Element.transform_ctxt (transfer res) elem) res)  | 
| 70608 | 505  | 
end handle ERROR msg => activate_err msg "facts" (name, morph) context;  | 
506  | 
||
507  | 
fun activate_notes_trace activ_elem transfer context export' (name, morph) context' =  | 
|
508  | 
let  | 
|
509  | 
val _ = trace "facts" (name, morph) context';  | 
|
510  | 
in  | 
|
511  | 
activate_notes activ_elem transfer context export' (name, morph) context'  | 
|
512  | 
end;  | 
|
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
513  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
514  | 
fun activate_all name thy activ_elem transfer (marked, input) =  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
515  | 
let  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
516  | 
    val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name;
 | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
517  | 
val input' = input |>  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
518  | 
(not (null params) ?  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
519  | 
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
 | 
520  | 
(* FIXME type parameters *)  | 
| 63352 | 521  | 
(case asm of SOME A => activ_elem (Assumes [(Binding.empty_atts, [(A, [])])]) | _ => I) |>  | 
| 40782 | 522  | 
(not (null defs) ?  | 
| 63352 | 523  | 
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
 | 
524  | 
val activate = activate_notes activ_elem transfer (Context.Theory thy) NONE;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
525  | 
in  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
526  | 
roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input')  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
527  | 
end;  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
528  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
529  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
530  | 
(** Public activation functions **)  | 
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
531  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
532  | 
fun activate_facts export dep context =  | 
| 
63030
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
533  | 
context  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
534  | 
|> Context_Position.set_visible_generic false  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
535  | 
|> pair (Idents.get context)  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
536  | 
|> roundup (Context.theory_of context)  | 
| 70608 | 537  | 
(activate_notes_trace init_element Morphism.transfer_morphism'' context export)  | 
| 
63030
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
538  | 
(the_default Morphism.identity export) dep  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
539  | 
|-> Idents.put  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
540  | 
|> Context_Position.restore_visible_generic context;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
541  | 
|
| 67666 | 542  | 
fun activate_declarations dep = Context.proof_map (fn context =>  | 
543  | 
context  | 
|
544  | 
|> Context_Position.set_visible_generic false  | 
|
545  | 
|> pair (Idents.get context)  | 
|
546  | 
|> roundup (Context.theory_of context) activate_syntax_decls Morphism.identity dep  | 
|
547  | 
|-> Idents.put  | 
|
548  | 
|> Context_Position.restore_visible_generic context);  | 
|
549  | 
||
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
550  | 
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
 | 
551  | 
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
 | 
552  | 
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
 | 
553  | 
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
 | 
554  | 
in  | 
| 
63030
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
555  | 
context  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
556  | 
|> Context_Position.set_visible_generic false  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
557  | 
|> pair empty_idents  | 
| 67667 | 558  | 
|> activate_all name thy init_element Morphism.transfer_morphism''  | 
| 
63030
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
559  | 
|-> (fn marked' => Idents.put (merge_idents (marked, marked')))  | 
| 
 
4e7eff53bee7
avoid massive multiplication of reports due to interpretation;
 
wenzelm 
parents: 
63029 
diff
changeset
 | 
560  | 
|> 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
 | 
561  | 
|> 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
 | 
562  | 
end;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
563  | 
|
| 
38211
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
564  | 
|
| 
 
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
 
ballarin 
parents: 
38111 
diff
changeset
 | 
565  | 
(*** Add and extend registrations ***)  | 
| 32801 | 566  | 
|
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
567  | 
type registration = Locale.registration;  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
568  | 
|
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
569  | 
fun amend_registration {mixin = NONE, ...} context = context
 | 
| 69058 | 570  | 
  | amend_registration {inst = (name, morph), mixin = SOME mixin, export} context =
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
571  | 
let  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
572  | 
val thy = Context.theory_of context;  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
573  | 
val ctxt = Context.proof_of context;  | 
| 53041 | 574  | 
|
| 69046 | 575  | 
val regs = get_regs context;  | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
576  | 
val base = instance_of thy name (morph $> export);  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
577  | 
val serial' =  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
578  | 
(case Idtab.lookup regs (name, base) of  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
579  | 
NONE =>  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
580  | 
              error ("No interpretation of locale " ^ quote (markup_name ctxt name) ^
 | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
581  | 
" with\nparameter instantiation " ^  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
582  | 
space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
583  | 
" available")  | 
| 69051 | 584  | 
          | SOME {serial = serial', ...} => serial');
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
585  | 
in  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
586  | 
add_mixin serial' mixin context  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
68029 
diff
changeset
 | 
587  | 
end;  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
diff
changeset
 | 
588  | 
|
| 32801 | 589  | 
(* Note that a registration that would be subsumed by an existing one will not be  | 
590  | 
generated, and it will not be possible to amend it. *)  | 
|
591  | 
||
| 69058 | 592  | 
fun add_registration {inst = (name, base_morph), mixin, export} context =
 | 
| 32801 | 593  | 
let  | 
| 38107 | 594  | 
val thy = Context.theory_of context;  | 
| 
68855
 
59ce31e15c33
more careful treatment position: existing facts refer to interpretation command, future facts refer to themselves (see also 4270da306442);
 
wenzelm 
parents: 
68853 
diff
changeset
 | 
595  | 
val pos_morph = Morphism.binding_morphism "position" (Binding.set_pos (Position.thread_data ()));  | 
| 68856 | 596  | 
val mix_morph = (case mixin of NONE => base_morph | SOME (mix, _) => base_morph $> mix);  | 
597  | 
val inst = instance_of thy name mix_morph;  | 
|
| 51727 | 598  | 
val idents = Idents.get context;  | 
| 32801 | 599  | 
in  | 
| 68856 | 600  | 
if redundant_ident thy idents (name, inst) then context (* FIXME amend mixins? *)  | 
| 32801 | 601  | 
else  | 
| 51727 | 602  | 
(idents, context)  | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
603  | 
(* add new registrations with inherited mixins *)  | 
| 68856 | 604  | 
|> roundup thy (add_reg thy export) export (name, mix_morph) |> #2  | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
605  | 
(* add mixin *)  | 
| 69058 | 606  | 
      |> amend_registration {inst = (name, mix_morph), mixin = mixin, export = export}
 | 
| 
36095
 
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
 
ballarin 
parents: 
36094 
diff
changeset
 | 
607  | 
(* activate import hierarchy as far as not already active *)  | 
| 68856 | 608  | 
|> activate_facts (SOME export) (name, mix_morph $> pos_morph)  | 
| 32801 | 609  | 
end;  | 
610  | 
||
| 29361 | 611  | 
|
| 
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
 | 
612  | 
|
| 31988 | 613  | 
(*** Dependencies ***)  | 
614  | 
||
| 69056 | 615  | 
fun registrations_of context loc =  | 
616  | 
  Idtab.fold_rev (fn ((name, _), {morphisms, ...}) =>
 | 
|
617  | 
name = loc ? cons (name, morphisms)) (get_regs context) []  | 
|
618  | 
(*with inherited mixins*)  | 
|
619  | 
|> map (fn (name, (base, export)) =>  | 
|
620  | 
(name, base $> (collect_mixins context (name, base $> export)) $> export));  | 
|
621  | 
||
| 73845 | 622  | 
fun add_dependency loc {inst = (name, morph), mixin, export} thy =
 | 
| 37102 | 623  | 
let  | 
| 69061 | 624  | 
val dep = make_dep (name, (morph, export));  | 
| 69060 | 625  | 
val add_dep =  | 
| 69061 | 626  | 
apfst (cons dep) #>  | 
627  | 
apsnd (case mixin of NONE => I | SOME mixin => insert_mixin (#serial dep) mixin);  | 
|
| 69060 | 628  | 
val thy' = change_locale loc (apsnd add_dep) thy;  | 
| 38107 | 629  | 
val context' = Context.Theory thy';  | 
| 46858 | 630  | 
val (_, regs) =  | 
631  | 
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
 | 
632  | 
(registrations_of context' loc) (Idents.get context', []);  | 
| 68853 | 633  | 
in  | 
| 73845 | 634  | 
    fold_rev (fn inst => Context.theory_map (add_registration {inst = inst, mixin = NONE, export = export}))
 | 
| 69060 | 635  | 
regs thy'  | 
| 68853 | 636  | 
end;  | 
| 31988 | 637  | 
|
| 72505 | 638  | 
|
| 31988 | 639  | 
|
| 29361 | 640  | 
(*** Storing results ***)  | 
641  | 
||
| 67654 | 642  | 
fun add_facts loc kind facts ctxt =  | 
| 61079 | 643  | 
if null facts then ctxt  | 
644  | 
else  | 
|
645  | 
let  | 
|
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
646  | 
val stored_notes = ((kind, map Attrib.trim_context_fact facts), serial ());  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
647  | 
val applied_notes = make_notes kind facts;  | 
| 61088 | 648  | 
|
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
649  | 
fun apply_notes morph = applied_notes |> fold (fn elem => fn context =>  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
650  | 
let val elem' = Element.transform_ctxt (Morphism.transfer_morphism'' context $> morph) elem  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
651  | 
in Element.init elem' context end);  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
652  | 
val apply_registrations = Context.theory_map (fn context =>  | 
| 
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
653  | 
fold_rev (apply_notes o #2) (registrations_of context loc) context);  | 
| 61079 | 654  | 
in  | 
655  | 
ctxt  | 
|
| 67655 | 656  | 
|> Attrib.local_notes kind facts |> #2  | 
| 61079 | 657  | 
|> Proof_Context.background_theory  | 
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
658  | 
((change_locale loc o apfst o apsnd) (cons stored_notes) #> apply_registrations)  | 
| 61079 | 659  | 
end;  | 
660  | 
||
| 
47246
 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 
wenzelm 
parents: 
47005 
diff
changeset
 | 
661  | 
fun add_declaration loc syntax decl =  | 
| 
 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 
wenzelm 
parents: 
47005 
diff
changeset
 | 
662  | 
syntax ?  | 
| 
 
2bbab021c0e6
clarified Named_Target.target_declaration: propagate through other levels as well;
 
wenzelm 
parents: 
47005 
diff
changeset
 | 
663  | 
Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))  | 
| 67654 | 664  | 
#> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration decl)];  | 
| 29361 | 665  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
666  | 
|
| 29361 | 667  | 
(*** Reasoning about locales ***)  | 
668  | 
||
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
669  | 
(* Storage for witnesses, intro and unfold rules *)  | 
| 29361 | 670  | 
|
| 33519 | 671  | 
structure Thms = Generic_Data  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
672  | 
(  | 
| 74151 | 673  | 
type T = thm Item_Net.T * thm Item_Net.T * thm Item_Net.T;  | 
| 74152 | 674  | 
val empty = (Thm.item_net, Thm.item_net, Thm.item_net);  | 
| 33519 | 675  | 
fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) =  | 
| 74151 | 676  | 
(Item_Net.merge (witnesses1, witnesses2),  | 
677  | 
Item_Net.merge (intros1, intros2),  | 
|
678  | 
Item_Net.merge (unfolds1, unfolds2));  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
679  | 
);  | 
| 29361 | 680  | 
|
| 61093 | 681  | 
fun get_thms which ctxt =  | 
| 67649 | 682  | 
map (Thm.transfer' ctxt) (which (Thms.get (Context.Proof ctxt)));  | 
| 61093 | 683  | 
|
| 74151 | 684  | 
val get_witnesses = get_thms (Item_Net.content o #1);  | 
685  | 
val get_intros = get_thms (Item_Net.content o #2);  | 
|
686  | 
val get_unfolds = get_thms (Item_Net.content o #3);  | 
|
| 29361 | 687  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
688  | 
val witness_add =  | 
| 61093 | 689  | 
Thm.declaration_attribute (fn th =>  | 
| 74151 | 690  | 
Thms.map (fn (x, y, z) => (Item_Net.update (Thm.trim_context th) x, y, z)));  | 
| 61093 | 691  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
692  | 
val intro_add =  | 
| 61093 | 693  | 
Thm.declaration_attribute (fn th =>  | 
| 74151 | 694  | 
Thms.map (fn (x, y, z) => (x, Item_Net.update (Thm.trim_context th) y, z)));  | 
| 61093 | 695  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
696  | 
val unfold_add =  | 
| 61093 | 697  | 
Thm.declaration_attribute (fn th =>  | 
| 74151 | 698  | 
Thms.map (fn (x, y, z) => (x, y, Item_Net.update (Thm.trim_context th) z)));  | 
| 29361 | 699  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
700  | 
|
| 40782 | 701  | 
(* Tactics *)  | 
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
702  | 
|
| 69017 | 703  | 
fun intro_locales_tac {strict, eager} ctxt =
 | 
704  | 
(if strict then Method.intros_tac else Method.try_intros_tac) ctxt  | 
|
| 
30725
 
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
 
wenzelm 
parents: 
30585 
diff
changeset
 | 
705  | 
(get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else []));  | 
| 29361 | 706  | 
|
| 53171 | 707  | 
val _ = Theory.setup  | 
| 69017 | 708  | 
(Method.setup \<^binding>\<open>intro_locales\<close>  | 
709  | 
    (Scan.succeed (METHOD o intro_locales_tac {strict = false, eager = false}))
 | 
|
| 30515 | 710  | 
"back-chain introduction rules of locales without unfolding predicates" #>  | 
| 69017 | 711  | 
Method.setup \<^binding>\<open>unfold_locales\<close>  | 
712  | 
    (Scan.succeed (METHOD o intro_locales_tac {strict = false, eager = true}))
 | 
|
| 53171 | 713  | 
"back-chain all introduction rules of locales");  | 
| 29361 | 714  | 
|
| 37471 | 715  | 
|
716  | 
(*** diagnostic commands and interfaces ***)  | 
|
717  | 
||
| 
69029
 
aec64b88e708
clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
 
wenzelm 
parents: 
69017 
diff
changeset
 | 
718  | 
fun get_locales thy = map #1 (Name_Space.dest_table (Locales.get thy));  | 
| 
 
aec64b88e708
clarified locale content: proper args with types for interpretation/axioms and typargs derived from the result;
 
wenzelm 
parents: 
69017 
diff
changeset
 | 
719  | 
|
| 
69059
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
720  | 
fun pretty_locales thy verbose =  | 
| 50301 | 721  | 
Pretty.block  | 
722  | 
(Pretty.breaks  | 
|
723  | 
(Pretty.str "locales:" ::  | 
|
724  | 
map (Pretty.mark_str o #1)  | 
|
| 
69059
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
725  | 
(Name_Space.markup_table verbose (Proof_Context.init_global thy) (Locales.get thy))));  | 
| 37471 | 726  | 
|
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
727  | 
fun pretty_locale thy show_facts name =  | 
| 37471 | 728  | 
let  | 
| 52103 | 729  | 
val locale_ctxt = init name thy;  | 
| 37471 | 730  | 
fun cons_elem (elem as Notes _) = show_facts ? cons elem  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
731  | 
| cons_elem (elem as Lazy_Notes _) = show_facts ? cons elem  | 
| 37471 | 732  | 
| cons_elem elem = cons elem;  | 
733  | 
val elems =  | 
|
| 53087 | 734  | 
activate_all name thy cons_elem (K (Morphism.transfer_morphism thy)) (empty_idents, [])  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
735  | 
|> snd |> rev  | 
| 
67677
 
ac4b475fc8c3
use lazy notes for locale context init and later additions of facts;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
736  | 
|> tap consolidate_notes  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
737  | 
|> map force_notes;  | 
| 37471 | 738  | 
in  | 
| 
69057
 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69056 
diff
changeset
 | 
739  | 
Pretty.block (Pretty.keyword1 "locale" :: Pretty.brk 1 :: pretty_name locale_ctxt name ::  | 
| 
 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69056 
diff
changeset
 | 
740  | 
maps (fn elem => [Pretty.fbrk, Pretty.chunks (Element.pretty_ctxt locale_ctxt elem)]) elems)  | 
| 37471 | 741  | 
end;  | 
742  | 
||
| 
69059
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
743  | 
fun pretty_registrations ctxt name =  | 
| 
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
744  | 
(case registrations_of (Context.Proof ctxt) name of  | 
| 
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
745  | 
[] => Pretty.str "no interpretations"  | 
| 
 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69058 
diff
changeset
 | 
746  | 
| regs => Pretty.big_list "interpretations:" (map (pretty_reg ctxt Morphism.identity) (rev regs)));  | 
| 41435 | 747  | 
|
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
748  | 
fun pretty_locale_deps thy =  | 
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
749  | 
let  | 
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
750  | 
fun make_node name =  | 
| 
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
47249 
diff
changeset
 | 
751  | 
     {name = name,
 | 
| 69061 | 752  | 
parents = map #name (dependencies_of thy name),  | 
| 
69057
 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 
wenzelm 
parents: 
69056 
diff
changeset
 | 
753  | 
body = pretty_locale thy false name};  | 
| 56025 | 754  | 
val names = sort_strings (Name_Space.fold_table (cons o #1) (Locales.get thy) []);  | 
755  | 
in map make_node names end;  | 
|
| 
37897
 
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
 
haftmann 
parents: 
37491 
diff
changeset
 | 
756  | 
|
| 69063 | 757  | 
type locale_dependency =  | 
758  | 
  {source: string, target: string, prefix: (string * bool) list, morphism: morphism,
 | 
|
759  | 
pos: Position.T, serial: serial};  | 
|
760  | 
||
761  | 
fun dest_dependencies prev_thys thy =  | 
|
762  | 
let  | 
|
763  | 
fun remove_prev loc prev_thy deps =  | 
|
764  | 
(case get_locale prev_thy loc of  | 
|
765  | 
NONE => deps  | 
|
766  | 
      | SOME (Loc {dependencies = prev_deps, ...}) =>
 | 
|
767  | 
if eq_list eq_dep (prev_deps, deps) then []  | 
|
768  | 
else subtract eq_dep prev_deps deps);  | 
|
769  | 
fun result loc (dep: dep) =  | 
|
770  | 
let val morphism = op $> (#morphisms dep) in  | 
|
771  | 
       {source = #name dep,
 | 
|
772  | 
target = loc,  | 
|
773  | 
prefix = Morphism.binding_prefix morphism,  | 
|
774  | 
morphism = morphism,  | 
|
775  | 
pos = #pos dep,  | 
|
776  | 
serial = #serial dep}  | 
|
777  | 
end;  | 
|
778  | 
    fun add (loc, Loc {dependencies = deps, ...}) =
 | 
|
779  | 
fold (cons o result loc) (fold (remove_prev loc) prev_thys deps);  | 
|
780  | 
in  | 
|
781  | 
Name_Space.fold_table add (Locales.get thy) []  | 
|
782  | 
|> sort (int_ord o apply2 #serial)  | 
|
783  | 
end;  | 
|
784  | 
||
| 29361 | 785  | 
end;  |