author | wenzelm |
Fri, 28 Oct 2011 17:15:52 +0200 | |
changeset 45290 | f599ac41e7f5 |
parent 42375 | 774df7c59508 |
child 45390 | e29521ef9059 |
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 |
40 |
val extern: theory -> string -> xstring |
|
29392 | 41 |
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
|
42 |
val params_of: theory -> string -> ((string * typ) * mixfix) list |
29441 | 43 |
val intros_of: theory -> string -> thm option * thm option |
44 |
val axioms_of: theory -> string -> thm list |
|
29544 | 45 |
val instance_of: theory -> string -> morphism -> term list |
29361 | 46 |
val specification_of: theory -> string -> term option * term list |
47 |
||
48 |
(* Storing results *) |
|
49 |
val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list -> |
|
50 |
Proof.context -> Proof.context |
|
51 |
val add_declaration: string -> declaration -> Proof.context -> Proof.context |
|
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
52 |
val add_syntax_declaration: string -> 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 |
41270 | 74 |
val add_dependency: string -> string * morphism -> (morphism * bool) option -> |
75 |
morphism -> theory -> theory |
|
29361 | 76 |
|
77 |
(* Diagnostic *) |
|
37471 | 78 |
val all_locales: theory -> string list |
29361 | 79 |
val print_locales: theory -> unit |
30344
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
wenzelm
parents:
30223
diff
changeset
|
80 |
val print_locale: theory -> bool -> xstring -> unit |
38109 | 81 |
val print_registrations: Proof.context -> string -> unit |
41435 | 82 |
val print_dependencies: Proof.context -> bool -> morphism -> (string * morphism) list -> unit |
37897
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
83 |
val locale_deps: theory -> |
40782 | 84 |
{params: ((string * typ) * mixfix) list, axioms: term list, registrations: term list list} Graph.T |
37897
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
85 |
* term list list Symtab.table Symtab.table |
29361 | 86 |
end; |
87 |
||
88 |
structure Locale: LOCALE = |
|
89 |
struct |
|
90 |
||
91 |
datatype ctxt = datatype Element.ctxt; |
|
92 |
||
29392 | 93 |
|
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
94 |
(*** Locales ***) |
29361 | 95 |
|
41272 | 96 |
type mixins = (((morphism * bool) * serial) list) Inttab.table; |
97 |
(* table of mixin lists, per list mixins in reverse order of declaration; |
|
98 |
lists indexed by registration/dependency serial, |
|
99 |
entries for empty lists may be omitted *) |
|
100 |
||
101 |
fun lookup_mixins serial' mixins = the_default [] (Inttab.lookup mixins serial'); |
|
102 |
||
103 |
fun merge_mixins (mix1, mix2) = Inttab.join (K (Library.merge (eq_snd op =))) (mix1, mix2); |
|
104 |
||
105 |
fun insert_mixin serial' mixin = |
|
106 |
Inttab.map_default (serial', []) (cons (mixin, serial ())); |
|
107 |
||
108 |
fun rename_mixin (old, new) mix = |
|
109 |
case Inttab.lookup mix old of |
|
110 |
NONE => mix | |
|
111 |
SOME mxs => Inttab.delete old mix |> Inttab.update_new (new, mxs); |
|
112 |
||
113 |
fun compose_mixins mixins = |
|
114 |
fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity; |
|
115 |
||
29361 | 116 |
datatype locale = Loc of { |
29392 | 117 |
(** static part **) |
30755
7ef503d216c2
simplified internal locale parameters: maintain proper name and type, instead of binding and constraint;
wenzelm
parents:
30754
diff
changeset
|
118 |
parameters: (string * sort) list * ((string * typ) * mixfix) list, |
29361 | 119 |
(* type and term parameters *) |
120 |
spec: term option * term list, |
|
121 |
(* assumptions (as a single predicate expression) and defines *) |
|
29441 | 122 |
intros: thm option * thm option, |
123 |
axioms: thm list, |
|
29392 | 124 |
(** dynamic part **) |
36096 | 125 |
syntax_decls: (declaration * serial) list, |
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
126 |
(* syntax declarations *) |
36096 | 127 |
notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * serial) list, |
29361 | 128 |
(* theorem declarations *) |
36651 | 129 |
dependencies: ((string * (morphism * morphism)) * serial) list |
41272 | 130 |
(* locale dependencies (sublocale relation) in reverse order *), |
131 |
mixins: mixins |
|
132 |
(* mixin part of dependencies *) |
|
29392 | 133 |
}; |
29361 | 134 |
|
41272 | 135 |
fun mk_locale ((parameters, spec, intros, axioms), |
136 |
((syntax_decls, notes), (dependencies, mixins))) = |
|
29441 | 137 |
Loc {parameters = parameters, intros = intros, axioms = axioms, spec = spec, |
41272 | 138 |
syntax_decls = syntax_decls, notes = notes, dependencies = dependencies, mixins = mixins}; |
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
139 |
|
41272 | 140 |
fun map_locale f (Loc {parameters, spec, intros, axioms, |
141 |
syntax_decls, notes, dependencies, mixins}) = |
|
142 |
mk_locale (f ((parameters, spec, intros, axioms), |
|
143 |
((syntax_decls, notes), (dependencies, mixins)))); |
|
30754 | 144 |
|
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
145 |
fun merge_locale |
41272 | 146 |
(Loc {parameters, spec, intros, axioms, syntax_decls, notes, dependencies, mixins}, |
147 |
Loc {syntax_decls = syntax_decls', notes = notes', |
|
148 |
dependencies = dependencies', mixins = mixins', ...}) = |
|
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
149 |
mk_locale |
29441 | 150 |
((parameters, spec, intros, axioms), |
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
151 |
((merge (eq_snd op =) (syntax_decls, syntax_decls'), |
29441 | 152 |
merge (eq_snd op =) (notes, notes')), |
41272 | 153 |
(merge (eq_snd op =) (dependencies, dependencies'), |
154 |
(merge_mixins (mixins, mixins'))))); |
|
29361 | 155 |
|
33522 | 156 |
structure Locales = Theory_Data |
29361 | 157 |
( |
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
158 |
type T = locale Name_Space.table; |
33159 | 159 |
val empty : T = Name_Space.empty_table "locale"; |
29361 | 160 |
val extend = I; |
33522 | 161 |
val merge = Name_Space.join_tables (K merge_locale); |
29361 | 162 |
); |
163 |
||
33095
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
wenzelm
parents:
33092
diff
changeset
|
164 |
val intern = Name_Space.intern o #1 o Locales.get; |
42360 | 165 |
fun extern thy = Name_Space.extern (Proof_Context.init_global thy) (#1 (Locales.get thy)); |
29392 | 166 |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
167 |
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
|
168 |
val defined = Symtab.defined o #2 o Locales.get; |
29361 | 169 |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
170 |
fun the_locale thy name = |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
171 |
(case get_locale thy name of |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
172 |
SOME (Loc loc) => loc |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
173 |
| NONE => error ("Unknown locale " ^ quote name)); |
29361 | 174 |
|
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
175 |
fun register_locale binding parameters spec intros axioms syntax_decls notes dependencies thy = |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42360
diff
changeset
|
176 |
thy |> Locales.map (Name_Space.define (Proof_Context.init_global thy) true (Sign.naming_of thy) |
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
177 |
(binding, |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
178 |
mk_locale ((parameters, spec, intros, axioms), |
36096 | 179 |
((map (fn decl => (decl, serial ())) syntax_decls, map (fn n => (n, serial ())) notes), |
41272 | 180 |
(map (fn d => (d |> apsnd (rpair Morphism.identity), serial ())) dependencies, |
181 |
Inttab.empty)))) #> snd); |
|
182 |
(* FIXME Morphism.identity *) |
|
29361 | 183 |
|
29392 | 184 |
fun change_locale name = |
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
185 |
Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd; |
29361 | 186 |
|
187 |
||
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
188 |
(** Primitive operations **) |
29361 | 189 |
|
29392 | 190 |
fun params_of thy = snd o #parameters o the_locale thy; |
29361 | 191 |
|
29441 | 192 |
fun intros_of thy = #intros o the_locale thy; |
193 |
||
194 |
fun axioms_of thy = #axioms o the_locale thy; |
|
195 |
||
29392 | 196 |
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
|
197 |
map (Morphism.term morph o Free o #1); |
29361 | 198 |
|
29392 | 199 |
fun specification_of thy = #spec o the_locale thy; |
29361 | 200 |
|
29544 | 201 |
fun dependencies_of thy name = the_locale thy name |> |
41272 | 202 |
#dependencies; |
203 |
||
204 |
fun mixins_of thy name serial = the_locale thy name |> |
|
205 |
#mixins |> lookup_mixins serial; |
|
206 |
||
207 |
(* unused *) |
|
208 |
fun identity_on thy name morph = |
|
209 |
let val mk_instance = instance_of thy name |
|
210 |
in |
|
211 |
forall2 (curry Term.aconv_untyped) (mk_instance Morphism.identity) (mk_instance morph) |
|
212 |
end; |
|
29544 | 213 |
|
37133
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
214 |
(* Print instance and qualifiers *) |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
215 |
|
41435 | 216 |
fun pretty_reg ctxt (name, morph) = |
37133
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
217 |
let |
42360 | 218 |
val thy = Proof_Context.theory_of ctxt; |
37133
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
219 |
val name' = extern thy name; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
220 |
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
|
221 |
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
|
222 |
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
|
223 |
fun prt_term' t = |
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
wenzelm
parents:
38797
diff
changeset
|
224 |
if Config.get ctxt show_types |
37133
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
225 |
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
|
226 |
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
|
227 |
else prt_term t; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
228 |
fun prt_inst ts = |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
229 |
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
|
230 |
|
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
231 |
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
|
232 |
val ts = instance_of thy name morph; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
233 |
in |
40782 | 234 |
(case qs of |
235 |
[] => prt_inst ts |
|
236 |
| 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
|
237 |
end; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
238 |
|
29361 | 239 |
|
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
240 |
(*** Identifiers: activated locales in theory or proof context ***) |
29361 | 241 |
|
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
242 |
(* subsumption *) |
37103
6ea25bb157e1
Consistently use equality for registration lookup.
ballarin
parents:
37102
diff
changeset
|
243 |
fun ident_le thy ((n: string, ts), (m, ss)) = (m = n) andalso Pattern.matchess thy (ss, ts); |
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
244 |
(* smaller term is more general *) |
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
245 |
|
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
246 |
(* total order *) |
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
247 |
fun ident_ord ((n: string, ts), (m, ss)) = |
40782 | 248 |
(case fast_string_ord (m, n) of |
249 |
EQUAL => list_ord Term_Ord.fast_term_ord (ts, ss) |
|
250 |
| ord => ord); |
|
29361 | 251 |
|
252 |
local |
|
253 |
||
30754 | 254 |
datatype 'a delayed = Ready of 'a | ToDo of 'a delayed * 'a delayed; |
29361 | 255 |
|
33519 | 256 |
structure Identifiers = Generic_Data |
29361 | 257 |
( |
30754 | 258 |
type T = (string * term list) list delayed; |
259 |
val empty = Ready []; |
|
29361 | 260 |
val extend = I; |
33519 | 261 |
val merge = ToDo; |
29361 | 262 |
); |
263 |
||
37103
6ea25bb157e1
Consistently use equality for registration lookup.
ballarin
parents:
37102
diff
changeset
|
264 |
fun finish thy (ToDo (i1, i2)) = merge (ident_le thy) (finish thy i1, finish thy i2) |
29361 | 265 |
| finish _ (Ready ids) = ids; |
266 |
||
267 |
val _ = Context.>> (Context.map_theory (Theory.at_begin (fn thy => |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
268 |
(case Identifiers.get (Context.Theory thy) of |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
269 |
Ready _ => NONE |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
270 |
| ids => SOME (Context.theory_map (Identifiers.put (Ready (finish thy ids))) thy))))); |
29361 | 271 |
|
30764
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
wenzelm
parents:
30763
diff
changeset
|
272 |
in |
29361 | 273 |
|
30764
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
wenzelm
parents:
30763
diff
changeset
|
274 |
val get_idents = (fn Ready ids => ids) o Identifiers.get; |
3e3e7aa0cc7a
simplified Locale.activate operations, using generic context;
wenzelm
parents:
30763
diff
changeset
|
275 |
val put_idents = Identifiers.put o Ready; |
29361 | 276 |
|
277 |
end; |
|
278 |
||
279 |
||
280 |
(** Resolve locale dependencies in a depth-first fashion **) |
|
281 |
||
282 |
local |
|
283 |
||
284 |
val roundup_bound = 120; |
|
285 |
||
41272 | 286 |
fun add thy depth stem export (name, morph, mixins) (deps, marked) = |
29361 | 287 |
if depth > roundup_bound |
288 |
then error "Roundup bound exceeded (sublocale relation probably not terminating)." |
|
289 |
else |
|
290 |
let |
|
41272 | 291 |
val instance = instance_of thy name (morph $> stem $> export); |
29361 | 292 |
in |
37103
6ea25bb157e1
Consistently use equality for registration lookup.
ballarin
parents:
37102
diff
changeset
|
293 |
if member (ident_le thy) marked (name, instance) |
29361 | 294 |
then (deps, marked) |
295 |
else |
|
296 |
let |
|
41272 | 297 |
val full_morph = morph $> compose_mixins mixins $> stem; |
298 |
(* no inheritance of mixins, regardless of requests by clients *) |
|
299 |
val dependencies = dependencies_of thy name |> |
|
300 |
map (fn ((name', (morph', export')), serial') => |
|
301 |
(name', morph' $> export', mixins_of thy name serial')); |
|
29361 | 302 |
val marked' = (name, instance) :: marked; |
41272 | 303 |
val (deps', marked'') = |
304 |
fold_rev (add thy (depth + 1) full_morph export) dependencies |
|
305 |
([], marked'); |
|
29361 | 306 |
in |
41272 | 307 |
((name, full_morph) :: deps' @ deps, marked'') |
29361 | 308 |
end |
309 |
end; |
|
310 |
||
311 |
in |
|
312 |
||
33541 | 313 |
(* Note that while identifiers always have the external (exported) view, activate_dep |
32803 | 314 |
is presented with the internal view. *) |
315 |
||
316 |
fun roundup thy activate_dep export (name, morph) (marked, input) = |
|
29361 | 317 |
let |
41272 | 318 |
(* Find all dependencies including new ones (which are dependencies enriching |
29361 | 319 |
existing registrations). *) |
41272 | 320 |
val (dependencies, marked') = |
321 |
add thy 0 Morphism.identity export (name, morph, []) ([], []); |
|
32800 | 322 |
(* Filter out fragments from marked; these won't be activated. *) |
29361 | 323 |
val dependencies' = filter_out (fn (name, morph) => |
37103
6ea25bb157e1
Consistently use equality for registration lookup.
ballarin
parents:
37102
diff
changeset
|
324 |
member (ident_le thy) marked (name, instance_of thy name (morph $> export))) dependencies; |
29361 | 325 |
in |
37103
6ea25bb157e1
Consistently use equality for registration lookup.
ballarin
parents:
37102
diff
changeset
|
326 |
(merge (ident_le thy) (marked, marked'), input |> fold_rev activate_dep dependencies') |
29361 | 327 |
end; |
328 |
||
329 |
end; |
|
330 |
||
331 |
||
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
332 |
(*** Registrations: interpretations in theories or proof contexts ***) |
29361 | 333 |
|
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
334 |
structure Idtab = Table(type key = string * term list val ord = ident_ord); |
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
335 |
|
38107 | 336 |
structure Registrations = Generic_Data |
29361 | 337 |
( |
41272 | 338 |
type T = ((morphism * morphism) * serial) Idtab.table * mixins; |
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
339 |
(* registrations, indexed by locale name and instance; |
41272 | 340 |
unique registration serial points to mixin list *) |
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
341 |
val empty = (Idtab.empty, Inttab.empty); |
29361 | 342 |
val extend = I; |
37133
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
343 |
fun merge ((reg1, mix1), (reg2, mix2)) : T = |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
344 |
(Idtab.join (fn id => fn (r1 as (_, s1), r2 as (_, s2)) => |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
345 |
if s1 = s2 then raise Idtab.SAME else raise Idtab.DUP id) (reg1, reg2), |
41272 | 346 |
merge_mixins (mix1, mix2)) |
37133
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
347 |
handle Idtab.DUP id => |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
348 |
(* distinct interpretations with same base: merge their mixins *) |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
349 |
let |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
350 |
val (_, s1) = Idtab.lookup reg1 id |> the; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
351 |
val (morph2, s2) = Idtab.lookup reg2 id |> the; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
352 |
val reg2' = Idtab.update (id, (morph2, s1)) reg2; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
353 |
val _ = warning "Removed duplicate interpretation after retrieving its mixins."; |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
354 |
(* FIXME print interpretations, |
1d048c6940c8
Merge mixins of distinct interpretations with same base.
ballarin
parents:
37105
diff
changeset
|
355 |
which is not straightforward without theory context *) |
41272 | 356 |
in merge ((reg1, mix1), (reg2', rename_mixin (s2, s1) mix2)) end; |
29361 | 357 |
(* FIXME consolidate with dependencies, consider one data slot only *) |
358 |
); |
|
359 |
||
32801 | 360 |
|
361 |
(* Primitive operations *) |
|
362 |
||
37104
3877a6c45d57
Avoid recomputation of registration instance for lookup.
ballarin
parents:
37103
diff
changeset
|
363 |
fun add_reg thy export (name, morph) = |
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
364 |
Registrations.map (apfst (Idtab.insert (K false) |
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
365 |
((name, instance_of thy name (morph $> export)), ((morph, export), serial ())))); |
31988 | 366 |
|
36095
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
ballarin
parents:
36094
diff
changeset
|
367 |
fun add_mixin serial' mixin = |
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
ballarin
parents:
36094
diff
changeset
|
368 |
(* registration to be amended identified by its serial id *) |
41272 | 369 |
Registrations.map (apsnd (insert_mixin serial' mixin)); |
32801 | 370 |
|
38107 | 371 |
fun get_mixins context (name, morph) = |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
372 |
let |
38107 | 373 |
val thy = Context.theory_of context; |
374 |
val (regs, mixins) = Registrations.get context; |
|
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
375 |
in |
40782 | 376 |
(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
|
377 |
NONE => [] |
41272 | 378 |
| SOME (_, serial) => lookup_mixins serial mixins) |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
379 |
end; |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
380 |
|
38107 | 381 |
fun collect_mixins context (name, morph) = |
382 |
let |
|
383 |
val thy = Context.theory_of context; |
|
384 |
in |
|
385 |
roundup thy (fn dep => fn mixins => merge (eq_snd op =) (mixins, get_mixins context dep)) |
|
386 |
Morphism.identity (name, morph) ([(name, instance_of thy name morph)], []) |
|
387 |
|> snd |> filter (snd o fst) (* only inheritable mixins *) |
|
388 |
|> (fn x => merge (eq_snd op =) (x, get_mixins context (name, morph))) |
|
389 |
|> compose_mixins |
|
390 |
end; |
|
36091
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
ballarin
parents:
36090
diff
changeset
|
391 |
|
38107 | 392 |
fun get_registrations context select = Registrations.get context |
37491 | 393 |
|>> Idtab.dest |>> select |
36095
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
ballarin
parents:
36094
diff
changeset
|
394 |
(* with inherited mixins *) |
37105
e464f84f3680
Store registrations in efficient data structure.
ballarin
parents:
37104
diff
changeset
|
395 |
|-> (fn regs => fn _ => map (fn ((name, _), ((base, export) ,_)) => |
38107 | 396 |
(name, base $> (collect_mixins context (name, base $> export)) $> export)) regs); |
36091
055934ed89b0
A rough implementation of full mixin inheritance; additional unit tests.
ballarin
parents:
36090
diff
changeset
|
397 |
|
38111 | 398 |
fun registrations_of context name = |
38107 | 399 |
get_registrations context (filter (curry (op =) name o fst o fst)); |
37491 | 400 |
|
38107 | 401 |
fun all_registrations context = get_registrations context I; |
37973 | 402 |
|
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
403 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
404 |
(*** Activate context elements of locale ***) |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
405 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
406 |
(* Declarations, facts and entire locale content *) |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
407 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
408 |
fun activate_syntax_decls (name, morph) context = |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
409 |
let |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
410 |
val thy = Context.theory_of context; |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
411 |
val {syntax_decls, ...} = the_locale thy name; |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
412 |
in |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
413 |
context |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
414 |
|> fold_rev (fn (decl, _) => decl morph) syntax_decls |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
415 |
end; |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
416 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
417 |
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
|
418 |
let |
38107 | 419 |
val thy = Context.theory_of context; |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
420 |
val {notes, ...} = the_locale thy name; |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
421 |
fun activate ((kind, facts), _) input = |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
422 |
let |
40782 | 423 |
val mixin = |
424 |
(case export' of |
|
425 |
NONE => Morphism.identity |
|
426 |
| SOME export => collect_mixins context (name, morph $> export) $> export); |
|
427 |
val facts' = facts |
|
45290 | 428 |
|> Element.facts_map (Element.transform_ctxt (transfer input $> morph $> mixin)); |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
429 |
in activ_elem (Notes (kind, facts')) input end; |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
430 |
in |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
431 |
fold_rev activate notes input |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
432 |
end; |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
433 |
|
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
434 |
fun activate_all name thy activ_elem transfer (marked, input) = |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
435 |
let |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
436 |
val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name; |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
437 |
val input' = input |> |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
438 |
(not (null params) ? |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
439 |
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
|
440 |
(* FIXME type parameters *) |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
441 |
(case asm of SOME A => activ_elem (Assumes [(Attrib.empty_binding, [(A, [])])]) | _ => I) |> |
40782 | 442 |
(not (null defs) ? |
443 |
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
|
444 |
val activate = activate_notes activ_elem transfer (Context.Theory thy) NONE; |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
445 |
in |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
446 |
roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input') |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
447 |
end; |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
448 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
449 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
450 |
(** Public activation functions **) |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
451 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
452 |
fun activate_declarations dep = Context.proof_map (fn context => |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
453 |
let |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
454 |
val thy = Context.theory_of context; |
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
455 |
in |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
456 |
roundup thy activate_syntax_decls Morphism.identity dep (get_idents context, context) |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
457 |
|-> put_idents |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
458 |
end); |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
459 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
460 |
fun activate_facts export dep context = |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
461 |
let |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
462 |
val thy = Context.theory_of context; |
38316
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
ballarin
parents:
38211
diff
changeset
|
463 |
val activate = activate_notes Element.init (Element.transfer_morphism o Context.theory_of) context export; |
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
ballarin
parents:
38211
diff
changeset
|
464 |
in |
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
ballarin
parents:
38211
diff
changeset
|
465 |
roundup thy activate (case export of NONE => Morphism.identity | SOME export => export) |
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
ballarin
parents:
38211
diff
changeset
|
466 |
dep (get_idents context, context) |
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
ballarin
parents:
38211
diff
changeset
|
467 |
|-> put_idents |
88e774d09fbc
Revert performance improvement of 8ed3a5fb4d25 since it breaks notes element declarations.
ballarin
parents:
38211
diff
changeset
|
468 |
end; |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
469 |
|
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
470 |
fun init name thy = |
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
471 |
activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of) |
42360 | 472 |
([], Context.Proof (Proof_Context.init_global thy)) |-> put_idents |> Context.proof_of; |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
473 |
|
38211
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
474 |
|
8ed3a5fb4d25
Remove duplicate locale activation code; performance improvement.
ballarin
parents:
38111
diff
changeset
|
475 |
(*** Add and extend registrations ***) |
32801 | 476 |
|
38107 | 477 |
fun amend_registration (name, morph) mixin export context = |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
478 |
let |
38107 | 479 |
val thy = Context.theory_of context; |
480 |
val regs = Registrations.get context |> fst; |
|
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
481 |
val base = instance_of thy name (morph $> export); |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
482 |
in |
40782 | 483 |
(case Idtab.lookup regs (name, base) of |
484 |
NONE => |
|
485 |
error ("No interpretation of locale " ^ |
|
41272 | 486 |
quote (extern thy name) ^ " with\nparameter instantiation " ^ |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
487 |
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
|
488 |
" available") |
40782 | 489 |
| SOME (_, serial') => add_mixin serial' mixin context) |
32804
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
490 |
end; |
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
ballarin
diff
changeset
|
491 |
|
32801 | 492 |
(* Note that a registration that would be subsumed by an existing one will not be |
493 |
generated, and it will not be possible to amend it. *) |
|
494 |
||
38107 | 495 |
fun add_registration (name, base_morph) mixin export context = |
32801 | 496 |
let |
38107 | 497 |
val thy = Context.theory_of context; |
40782 | 498 |
val mix = (case mixin of NONE => Morphism.identity | SOME (mix, _) => mix); |
37102 | 499 |
val morph = base_morph $> mix; |
500 |
val inst = instance_of thy name morph; |
|
32801 | 501 |
in |
38107 | 502 |
if member (ident_le thy) (get_idents context) (name, inst) |
41272 | 503 |
then context (* FIXME amend mixins? *) |
32801 | 504 |
else |
38107 | 505 |
(get_idents context, context) |
36095
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
ballarin
parents:
36094
diff
changeset
|
506 |
(* add new registrations with inherited mixins *) |
37104
3877a6c45d57
Avoid recomputation of registration instance for lookup.
ballarin
parents:
37103
diff
changeset
|
507 |
|> 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
|
508 |
|> snd |
059c3568fdc8
Proper inheritance of mixins for activated facts and locale dependencies.
ballarin
parents:
36094
diff
changeset
|
509 |
(* add mixin *) |
40782 | 510 |
|> |
511 |
(case mixin of |
|
512 |
NONE => I |
|
513 |
| 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
|
514 |
(* 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
|
515 |
|> activate_facts (SOME export) (name, morph) |
32801 | 516 |
end; |
517 |
||
29361 | 518 |
|
31988 | 519 |
(*** Dependencies ***) |
520 |
||
41272 | 521 |
(* |
522 |
fun amend_dependency loc (name, morph) mixin export thy = |
|
523 |
let |
|
524 |
val deps = dependencies_of thy loc; |
|
525 |
in |
|
526 |
case AList.lookup (fn ((name, morph), ((name', (morph', _)), _)) => |
|
527 |
ident_ord ((name, instance_of thy name morph), (name', instance_of thy name' morph')) = EQUAL) deps (name, morph) of |
|
528 |
NONE => error ("Locale " ^ |
|
529 |
quote (extern thy name) ^ " with\parameter instantiation " ^ |
|
530 |
space_implode " " (map (quote o Syntax.string_of_term_global thy) morph) ^ |
|
531 |
" not a sublocale of " ^ quote (extern thy loc)) |
|
532 |
| SOME (_, serial') => change_locale ... |
|
533 |
end; |
|
534 |
*) |
|
535 |
||
41270 | 536 |
fun add_dependency loc (name, morph) mixin export thy = |
37102 | 537 |
let |
41272 | 538 |
val serial' = serial (); |
539 |
val thy' = thy |> |
|
540 |
(change_locale loc o apsnd) |
|
541 |
(apfst (cons ((name, (morph, export)), serial')) #> |
|
542 |
apsnd (case mixin of NONE => I | SOME mixin => insert_mixin serial' mixin)); |
|
38107 | 543 |
val context' = Context.Theory thy'; |
37102 | 544 |
val (_, regs) = fold_rev (roundup thy' cons export) |
38783
f171050ad3f8
For sublocale it is sufficient to reconsider ancestors of the target.
ballarin
parents:
38316
diff
changeset
|
545 |
(registrations_of context' loc) (get_idents (context'), []); |
37102 | 546 |
in |
547 |
thy' |
|
38107 | 548 |
|> fold_rev (fn dep => Context.theory_map (add_registration dep NONE export)) regs |
37102 | 549 |
end; |
31988 | 550 |
|
551 |
||
29361 | 552 |
(*** Storing results ***) |
553 |
||
554 |
(* Theorems *) |
|
555 |
||
556 |
fun add_thmss loc kind args ctxt = |
|
557 |
let |
|
30777
9960ff945c52
simplified Element.activate(_i): singleton version;
wenzelm
parents:
30775
diff
changeset
|
558 |
val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt; |
42360 | 559 |
val ctxt'' = ctxt' |> Proof_Context.background_theory |
38757
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents:
38756
diff
changeset
|
560 |
((change_locale loc o apfst o apsnd) (cons (args', serial ())) |
29392 | 561 |
#> |
29361 | 562 |
(* Registrations *) |
31988 | 563 |
(fn thy => fold_rev (fn (_, morph) => |
29361 | 564 |
let |
45290 | 565 |
val args'' = snd args' |> Element.facts_map (Element.transform_ctxt morph) |> |
29361 | 566 |
Attrib.map_facts (Attrib.attribute_i thy) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39134
diff
changeset
|
567 |
in Global_Theory.note_thmss kind args'' #> snd end) |
38111 | 568 |
(registrations_of (Context.Theory thy) loc) thy)) |
29361 | 569 |
in ctxt'' end; |
570 |
||
571 |
||
572 |
(* Declarations *) |
|
573 |
||
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
574 |
fun add_declaration loc decl = |
33643
b275f26a638b
eliminated obsolete "internal" kind -- collapsed to unspecific "";
wenzelm
parents:
33541
diff
changeset
|
575 |
add_thmss loc "" |
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
576 |
[((Binding.conceal Binding.empty, |
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
577 |
[Attrib.internal (fn phi => Thm.declaration_attribute (K (decl phi)))]), |
33278 | 578 |
[([Drule.dummy_thm], [])])]; |
29361 | 579 |
|
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
580 |
fun add_syntax_declaration loc decl = |
42360 | 581 |
Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ()))) |
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
33643
diff
changeset
|
582 |
#> add_declaration loc decl; |
29361 | 583 |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
584 |
|
29361 | 585 |
(*** Reasoning about locales ***) |
586 |
||
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
587 |
(* Storage for witnesses, intro and unfold rules *) |
29361 | 588 |
|
33519 | 589 |
structure Thms = Generic_Data |
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
590 |
( |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
591 |
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
|
592 |
val empty = ([], [], []); |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
593 |
val extend = I; |
33519 | 594 |
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
|
595 |
(Thm.merge_thms (witnesses1, witnesses2), |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
596 |
Thm.merge_thms (intros1, intros2), |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
597 |
Thm.merge_thms (unfolds1, unfolds2)); |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
598 |
); |
29361 | 599 |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
600 |
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
|
601 |
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
|
602 |
val get_unfolds = #3 o Thms.get o Context.Proof; |
29361 | 603 |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
604 |
val witness_add = |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
605 |
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
|
606 |
val intro_add = |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
607 |
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
|
608 |
val unfold_add = |
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
609 |
Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Thm.add_thm th z))); |
29361 | 610 |
|
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
611 |
|
40782 | 612 |
(* Tactics *) |
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
613 |
|
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
36091
diff
changeset
|
614 |
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
|
615 |
intros_tac |
30725
c23a5b3cd1b9
register_locale: produce stamps at the spot where elements are registered;
wenzelm
parents:
30585
diff
changeset
|
616 |
(get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else [])); |
29361 | 617 |
|
36093
0880493627ca
Graceful treatment of non-locale subgoals by methods unfold_locales and intro_locales.
ballarin
parents:
36091
diff
changeset
|
618 |
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
|
619 |
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
|
620 |
|
29361 | 621 |
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
|
622 |
(Method.setup (Binding.name "intro_locales") (Scan.succeed (METHOD o try_intro_locales_tac false)) |
30515 | 623 |
"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
|
624 |
Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o try_intro_locales_tac true)) |
30515 | 625 |
"back-chain all introduction rules of locales")); |
29361 | 626 |
|
37471 | 627 |
|
628 |
(*** diagnostic commands and interfaces ***) |
|
629 |
||
37897
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
630 |
val all_locales = Symtab.keys o snd o Locales.get; |
37471 | 631 |
|
632 |
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
|
633 |
Pretty.strs ("locales:" :: |
42360 | 634 |
map #1 (Name_Space.extern_table (Proof_Context.init_global thy) (Locales.get thy))) |
37471 | 635 |
|> Pretty.writeln; |
636 |
||
637 |
fun print_locale thy show_facts raw_name = |
|
638 |
let |
|
639 |
val name = intern thy raw_name; |
|
640 |
val ctxt = init name thy; |
|
641 |
fun cons_elem (elem as Notes _) = show_facts ? cons elem |
|
642 |
| cons_elem elem = cons elem; |
|
643 |
val elems = |
|
644 |
activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], []) |
|
645 |
|> snd |> rev; |
|
646 |
in |
|
647 |
Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems) |
|
648 |
|> Pretty.writeln |
|
649 |
end; |
|
650 |
||
38109 | 651 |
fun print_registrations ctxt raw_name = |
652 |
let |
|
42360 | 653 |
val thy = Proof_Context.theory_of ctxt; |
41435 | 654 |
val name = intern thy raw_name; |
655 |
val _ = the_locale thy name; (* error if locale unknown *) |
|
38109 | 656 |
in |
41435 | 657 |
(case registrations_of (Context.Proof ctxt) (* FIXME *) name of |
40782 | 658 |
[] => Pretty.str ("no interpretations") |
41435 | 659 |
| regs => Pretty.big_list "interpretations:" (map (pretty_reg ctxt) (rev regs))) |
660 |
end |> Pretty.writeln; |
|
661 |
||
662 |
fun print_dependencies ctxt clean export insts = |
|
663 |
let |
|
42360 | 664 |
val thy = Proof_Context.theory_of ctxt; |
41435 | 665 |
val idents = if clean then [] else get_idents (Context.Proof ctxt); |
666 |
in |
|
667 |
(case fold (roundup thy cons export) insts (idents, []) |> snd of |
|
668 |
[] => Pretty.str ("no dependencies") |
|
669 |
| deps => Pretty.big_list "dependencies:" (map (pretty_reg ctxt) (rev deps))) |
|
40782 | 670 |
end |> Pretty.writeln; |
37471 | 671 |
|
37897
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
672 |
fun locale_deps thy = |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
673 |
let |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
674 |
val names = all_locales thy |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
675 |
fun add_locale_node name = |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
676 |
let |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
677 |
val params = params_of thy name; |
40782 | 678 |
val axioms = |
679 |
these (Option.map (Logic.strip_imp_prems o Thm.prop_of) (fst (intros_of thy name))); |
|
680 |
val registrations = |
|
681 |
map (instance_of thy name o snd) (registrations_of (Context.Theory thy) name); |
|
682 |
in |
|
683 |
Graph.new_node (name, {params = params, axioms = axioms, registrations = registrations}) |
|
37897
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
684 |
end; |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
685 |
fun add_locale_deps name = |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
686 |
let |
40782 | 687 |
val dependencies = |
41272 | 688 |
(map o apsnd) (instance_of thy name o op $>) (dependencies_of thy name |> map fst); |
37897
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
689 |
in |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
690 |
fold (fn (super, ts) => fn (gr, deps) => (gr |> Graph.add_edge (super, name), |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
691 |
deps |> Symtab.map_default (super, Symtab.empty) (Symtab.cons_list (name, ts)))) |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
692 |
dependencies |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
693 |
end; |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
694 |
in |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
695 |
Graph.empty |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
696 |
|> fold add_locale_node names |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
697 |
|> rpair Symtab.empty |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
698 |
|> fold add_locale_deps names |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
699 |
end; |
c5ad6fec3470
abstract visualization of locale ingredients; all_locales yields proper locale identifiers
haftmann
parents:
37491
diff
changeset
|
700 |
|
29361 | 701 |
end; |