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