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