| author | wenzelm | 
| Tue, 22 Oct 2024 19:26:40 +0200 | |
| changeset 81232 | 9867b5cf3f7a | 
| parent 81117 | 0c898f7d9b15 | 
| child 81514 | 98cb63b447c6 | 
| permissions | -rw-r--r-- | 
| 47057 | 1  | 
(* Title: Pure/Isar/bundle.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Bundled declarations (notes etc.).  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature BUNDLE =  | 
|
8  | 
sig  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
9  | 
type name = bool * string  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
10  | 
type xname = (bool * Position.T) * (xstring * Position.T)  | 
| 
74261
 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 
wenzelm 
parents: 
74122 
diff
changeset
 | 
11  | 
val bundle_space: Context.generic -> Name_Space.T  | 
| 47057 | 12  | 
val check: Proof.context -> xstring * Position.T -> string  | 
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
13  | 
val check_name: Proof.context -> xname -> name  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
14  | 
val get: Proof.context -> string -> Attrib.thms  | 
| 74122 | 15  | 
val extern: Proof.context -> string -> xstring  | 
| 63270 | 16  | 
val print_bundles: bool -> Proof.context -> unit  | 
| 72605 | 17  | 
val unbundle: name list -> local_theory -> local_theory  | 
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
18  | 
val unbundle_cmd: xname list -> local_theory -> local_theory  | 
| 72605 | 19  | 
val includes: name list -> Proof.context -> Proof.context  | 
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
20  | 
val includes_cmd: xname list -> Proof.context -> Proof.context  | 
| 72605 | 21  | 
val include_: name list -> Proof.state -> Proof.state  | 
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
22  | 
val include_cmd: xname list -> Proof.state -> Proof.state  | 
| 72605 | 23  | 
val including: name list -> Proof.state -> Proof.state  | 
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
24  | 
val including_cmd: xname list -> Proof.state -> Proof.state  | 
| 81106 | 25  | 
  val bundle: {open_bundle: bool} -> binding * Attrib.thms ->
 | 
| 81105 | 26  | 
(binding * typ option * mixfix) list -> local_theory -> local_theory  | 
| 81106 | 27  | 
  val bundle_cmd: {open_bundle: bool} -> binding * (Facts.ref * Token.src list) list ->
 | 
| 81105 | 28  | 
(binding * string option * mixfix) list -> local_theory -> local_theory  | 
| 81106 | 29  | 
  val init: {open_bundle: bool} -> binding -> theory -> local_theory
 | 
| 47057 | 30  | 
end;  | 
31  | 
||
32  | 
structure Bundle: BUNDLE =  | 
|
33  | 
struct  | 
|
34  | 
||
| 63270 | 35  | 
(** context data **)  | 
36  | 
||
37  | 
structure Data = Generic_Data  | 
|
38  | 
(  | 
|
39  | 
type T = Attrib.thms Name_Space.table * Attrib.thms option;  | 
|
| 
74261
 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 
wenzelm 
parents: 
74122 
diff
changeset
 | 
40  | 
val empty : T = (Name_Space.empty_table Markup.bundleN, NONE);  | 
| 63270 | 41  | 
fun merge ((tab1, target1), (tab2, target2)) =  | 
42  | 
(Name_Space.merge_tables (tab1, tab2), merge_options (target1, target2));  | 
|
43  | 
);  | 
|
44  | 
||
45  | 
||
46  | 
(* bundles *)  | 
|
47  | 
||
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
48  | 
type name = bool * string;  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
49  | 
type xname = (bool * Position.T) * (xstring * Position.T);  | 
| 72605 | 50  | 
|
| 
72452
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
51  | 
val get_all_generic = #1 o Data.get;  | 
| 
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
52  | 
val get_all = get_all_generic o Context.Proof;  | 
| 63270 | 53  | 
|
| 
74261
 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 
wenzelm 
parents: 
74122 
diff
changeset
 | 
54  | 
val bundle_space = Name_Space.space_of_table o #1 o Data.get;  | 
| 
 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 
wenzelm 
parents: 
74122 
diff
changeset
 | 
55  | 
|
| 
72452
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
56  | 
fun check ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_all ctxt);  | 
| 63270 | 57  | 
|
| 
81117
 
0c898f7d9b15
ML antiquotation for formally-checked bundle names;
 
wenzelm 
parents: 
81116 
diff
changeset
 | 
58  | 
val _ = Theory.setup  | 
| 
 
0c898f7d9b15
ML antiquotation for formally-checked bundle names;
 
wenzelm 
parents: 
81116 
diff
changeset
 | 
59  | 
(ML_Antiquotation.inline_embedded \<^binding>\<open>bundle\<close>  | 
| 
 
0c898f7d9b15
ML antiquotation for formally-checked bundle names;
 
wenzelm 
parents: 
81116 
diff
changeset
 | 
60  | 
(Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, (name, pos)) =>  | 
| 
 
0c898f7d9b15
ML antiquotation for formally-checked bundle names;
 
wenzelm 
parents: 
81116 
diff
changeset
 | 
61  | 
ML_Syntax.print_string (check ctxt (name, pos)))));  | 
| 
 
0c898f7d9b15
ML antiquotation for formally-checked bundle names;
 
wenzelm 
parents: 
81116 
diff
changeset
 | 
62  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
63  | 
fun check_name ctxt ((b: bool, pos), arg) =  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
64  | 
(Context_Position.report ctxt pos Markup.quasi_keyword; (b, check ctxt arg));  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
65  | 
|
| 81106 | 66  | 
val get_global = Name_Space.get o get_all_generic o Context.Theory;  | 
| 
72452
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
67  | 
val get = Name_Space.get o get_all_generic o Context.Proof;  | 
| 
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
68  | 
|
| 74122 | 69  | 
fun extern ctxt = Name_Space.extern ctxt (Name_Space.space_of_table (get_all ctxt));  | 
70  | 
||
| 63270 | 71  | 
|
| 81105 | 72  | 
(* context and morphisms *)  | 
73  | 
||
74  | 
val trim_context_bundle =  | 
|
75  | 
map (fn (fact, atts) => (map Thm.trim_context fact, (map o map) Token.trim_context atts));  | 
|
76  | 
||
77  | 
fun transfer_bundle thy =  | 
|
78  | 
map (fn (fact, atts) => (map (Thm.transfer thy) fact, (map o map) (Token.transfer thy) atts));  | 
|
79  | 
||
80  | 
fun transform_bundle phi =  | 
|
81  | 
map (fn (fact, atts) => (Morphism.fact phi fact, (map o map) (Token.transform phi) atts));  | 
|
82  | 
||
83  | 
||
84  | 
||
| 63270 | 85  | 
(* target -- bundle under construction *)  | 
86  | 
||
87  | 
fun the_target thy =  | 
|
| 
77908
 
a6bd716a6124
tuned: concise combinators instead of bulky case-expressions;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
88  | 
#2 (Data.get (Context.Theory thy))  | 
| 
 
a6bd716a6124
tuned: concise combinators instead of bulky case-expressions;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
89  | 
|> \<^if_none>\<open>error "Missing bundle target"\<close>;  | 
| 63270 | 90  | 
|
91  | 
val reset_target = (Context.theory_map o Data.map o apsnd o K) NONE;  | 
|
| 67652 | 92  | 
val set_target = Context.theory_map o Data.map o apsnd o K o SOME o Attrib.trim_context_thms;  | 
| 63270 | 93  | 
|
94  | 
fun augment_target thms =  | 
|
95  | 
Local_Theory.background_theory (fn thy => set_target (the_target thy @ thms) thy);  | 
|
96  | 
||
97  | 
||
98  | 
(* print bundles *)  | 
|
99  | 
||
100  | 
fun pretty_bundle ctxt (markup_name, bundle) =  | 
|
101  | 
let  | 
|
102  | 
val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt;  | 
|
| 63275 | 103  | 
fun prt_thm_attribs atts th =  | 
104  | 
Pretty.block (Pretty.breaks (prt_thm th :: Attrib.pretty_attribs ctxt atts));  | 
|
| 63270 | 105  | 
fun prt_thms (ths, []) = map prt_thm ths  | 
| 63275 | 106  | 
| prt_thms (ths, atts) = map (prt_thm_attribs atts) ths;  | 
| 63270 | 107  | 
in  | 
108  | 
Pretty.block ([Pretty.keyword1 "bundle", Pretty.str " ", Pretty.mark_str markup_name] @  | 
|
| 63275 | 109  | 
(if null bundle then [] else Pretty.fbreaks (Pretty.str " =" :: maps prt_thms bundle)))  | 
| 63270 | 110  | 
end;  | 
111  | 
||
112  | 
fun print_bundles verbose ctxt =  | 
|
113  | 
Pretty.writeln_chunks  | 
|
| 
72452
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
114  | 
(map (pretty_bundle ctxt) (Name_Space.markup_table verbose ctxt (get_all ctxt)));  | 
| 63270 | 115  | 
|
116  | 
||
117  | 
||
| 81105 | 118  | 
(** open bundles **)  | 
| 
78064
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78061 
diff
changeset
 | 
119  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
120  | 
fun polarity_decls b =  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
121  | 
[([Drule.dummy_thm],  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
122  | 
[Attrib.internal \<^here> (K (Thm.declaration_attribute (K (Syntax.set_polarity_generic b))))])];  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
123  | 
|
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
124  | 
fun apply_bundle loc (add, bundle) ctxt =  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
125  | 
let  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
126  | 
val notes = if loc then Local_Theory.notes else Attrib.local_notes "";  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
127  | 
val add0 = Syntax.get_polarity ctxt;  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
128  | 
val add1 = Syntax.effective_polarity ctxt add;  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
129  | 
in  | 
| 81104 | 130  | 
ctxt  | 
131  | 
|> Context_Position.set_visible false  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
132  | 
|> notes [(Binding.empty_atts, polarity_decls add1 @ bundle @ polarity_decls add0)] |> snd  | 
| 81104 | 133  | 
|> Context_Position.restore_visible ctxt  | 
134  | 
end;  | 
|
135  | 
||
| 81105 | 136  | 
local  | 
137  | 
||
138  | 
fun gen_unbundle loc prep args ctxt =  | 
|
139  | 
let  | 
|
140  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
141  | 
val bundles = map (prep ctxt #> apsnd (get ctxt #> transfer_bundle thy)) args;  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
142  | 
in fold (apply_bundle loc) bundles ctxt end;  | 
| 81105 | 143  | 
|
144  | 
fun gen_includes prep = gen_unbundle false prep;  | 
|
145  | 
||
146  | 
fun gen_include prep bs =  | 
|
147  | 
Proof.assert_forward #> Proof.map_context (gen_includes prep bs) #> Proof.reset_facts;  | 
|
148  | 
||
149  | 
fun gen_including prep bs =  | 
|
150  | 
Proof.assert_backward #> Proof.map_context (gen_includes prep bs);  | 
|
151  | 
||
152  | 
in  | 
|
153  | 
||
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
154  | 
val unbundle = gen_unbundle true (K I);  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
155  | 
val unbundle_cmd = gen_unbundle true check_name;  | 
| 81105 | 156  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
157  | 
val includes = gen_includes (K I);  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
158  | 
val includes_cmd = gen_includes check_name;  | 
| 81105 | 159  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
160  | 
val include_ = gen_include (K I);  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
161  | 
val include_cmd = gen_include check_name;  | 
| 81105 | 162  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
163  | 
val including = gen_including (K I);  | 
| 
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
164  | 
val including_cmd = gen_including check_name;  | 
| 81105 | 165  | 
|
166  | 
end;  | 
|
167  | 
||
168  | 
||
169  | 
||
170  | 
(** define bundle **)  | 
|
171  | 
||
172  | 
fun define_bundle (b, bundle) context =  | 
|
173  | 
let  | 
|
174  | 
val (name, bundles') = get_all_generic context  | 
|
175  | 
|> Name_Space.define context true (b, trim_context_bundle bundle);  | 
|
176  | 
val context' = (Data.map o apfst o K) bundles' context;  | 
|
177  | 
in (name, context') end;  | 
|
178  | 
||
| 47057 | 179  | 
|
| 63270 | 180  | 
(* command *)  | 
| 47057 | 181  | 
|
182  | 
local  | 
|
183  | 
||
| 81106 | 184  | 
fun gen_bundle prep_fact prep_att add_fixes {open_bundle} (binding, raw_bundle) raw_fixes lthy =
 | 
| 47057 | 185  | 
let  | 
| 60469 | 186  | 
val (_, ctxt') = add_fixes raw_fixes lthy;  | 
| 47057 | 187  | 
val bundle0 = raw_bundle  | 
| 53111 | 188  | 
|> map (fn (fact, atts) => (prep_fact ctxt' fact, map (prep_att ctxt') atts));  | 
| 47057 | 189  | 
val bundle =  | 
| 78066 | 190  | 
Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)]  | 
191  | 
|> maps #2  | 
|
| 78069 | 192  | 
|> transform_bundle (Proof_Context.export_morphism ctxt' lthy)  | 
193  | 
|> trim_context_bundle;  | 
|
| 47057 | 194  | 
in  | 
| 78095 | 195  | 
    lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding}
 | 
| 78069 | 196  | 
(fn phi => fn context =>  | 
197  | 
let val psi = Morphism.set_trim_context'' context phi  | 
|
198  | 
in #2 (define_bundle (Morphism.binding psi binding, transform_bundle psi bundle) context) end)  | 
|
| 
81116
 
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
 
wenzelm 
parents: 
81106 
diff
changeset
 | 
199  | 
|> open_bundle ? apply_bundle true (true, bundle)  | 
| 47057 | 200  | 
end;  | 
201  | 
||
202  | 
in  | 
|
203  | 
||
| 60469 | 204  | 
val bundle = gen_bundle (K I) (K I) Proof_Context.add_fixes;  | 
205  | 
val bundle_cmd = gen_bundle Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd;  | 
|
| 47057 | 206  | 
|
207  | 
end;  | 
|
208  | 
||
209  | 
||
| 63270 | 210  | 
(* target *)  | 
211  | 
||
212  | 
local  | 
|
213  | 
||
214  | 
fun bad_operation _ = error "Not possible in bundle target";  | 
|
215  | 
||
| 81106 | 216  | 
fun conclude {open_bundle, invisible} binding =
 | 
| 63270 | 217  | 
Local_Theory.background_theory_result (fn thy =>  | 
| 81106 | 218  | 
let  | 
219  | 
val (name, thy') = thy  | 
|
220  | 
|> invisible ? Context_Position.set_visible_global false  | 
|
221  | 
|> Context.Theory  | 
|
222  | 
|> define_bundle (binding, the_target thy)  | 
|
223  | 
||> Context.the_theory;  | 
|
224  | 
val bundle = get_global thy' name;  | 
|
225  | 
val thy'' = thy'  | 
|
226  | 
|> open_bundle ?  | 
|
227  | 
(Context_Position.set_visible_global false #>  | 
|
228  | 
Attrib.global_notes "" [(Binding.empty_atts, bundle)] #> snd)  | 
|
229  | 
|> Context_Position.restore_visible_global thy  | 
|
230  | 
|> reset_target;  | 
|
231  | 
in (name, thy'') end);  | 
|
| 63270 | 232  | 
|
233  | 
fun pretty binding lthy =  | 
|
234  | 
let  | 
|
235  | 
val bundle = the_target (Proof_Context.theory_of lthy);  | 
|
| 63275 | 236  | 
val (name, lthy') = lthy  | 
237  | 
|> Local_Theory.raw_theory (Context_Position.set_visible_global false)  | 
|
| 81106 | 238  | 
      |> conclude {open_bundle = false, invisible = true} binding;
 | 
| 63270 | 239  | 
val thy_ctxt' = Proof_Context.init_global (Proof_Context.theory_of lthy');  | 
240  | 
val markup_name =  | 
|
| 
72452
 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 
haftmann 
parents: 
72450 
diff
changeset
 | 
241  | 
Name_Space.markup_extern thy_ctxt' (Name_Space.space_of_table (get_all thy_ctxt')) name;  | 
| 63270 | 242  | 
in [pretty_bundle lthy' (markup_name, bundle)] end;  | 
243  | 
||
244  | 
fun bundle_notes kind facts lthy =  | 
|
245  | 
let  | 
|
246  | 
val bundle = facts  | 
|
247  | 
|> maps (fn ((_, more_atts), thms) => map (fn (ths, atts) => (ths, atts @ more_atts)) thms);  | 
|
248  | 
in  | 
|
249  | 
lthy  | 
|
250  | 
|> augment_target (transform_bundle (Local_Theory.standard_morphism_theory lthy) bundle)  | 
|
| 
63278
 
6c963f1fc388
avoid duplicate Attrib.local_notes in aux. context;
 
wenzelm 
parents: 
63277 
diff
changeset
 | 
251  | 
|> Generic_Target.standard_notes (op <>) kind facts  | 
| 63270 | 252  | 
|> Attrib.local_notes kind facts  | 
253  | 
end;  | 
|
254  | 
||
| 78095 | 255  | 
fun bundle_declaration pos decl lthy =  | 
| 63270 | 256  | 
lthy  | 
| 78095 | 257  | 
|> (augment_target o Attrib.internal_declaration pos)  | 
| 63270 | 258  | 
(Morphism.transform (Local_Theory.standard_morphism_theory lthy) decl)  | 
259  | 
|> Generic_Target.standard_declaration (K true) decl;  | 
|
260  | 
||
261  | 
in  | 
|
262  | 
||
| 81106 | 263  | 
fun init {open_bundle} binding thy =
 | 
| 63270 | 264  | 
thy  | 
| 
72516
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72504 
diff
changeset
 | 
265  | 
|> Local_Theory.init  | 
| 
66335
 
a849ce33923d
treat exit separate from regular local theory operations
 
haftmann 
parents: 
63352 
diff
changeset
 | 
266  | 
     {background_naming = Sign.naming_of thy,
 | 
| 
66337
 
5caea089dd17
more structural sharing between common target Generic_Target.init
 
haftmann 
parents: 
66336 
diff
changeset
 | 
267  | 
setup = set_target [] #> Proof_Context.init_global,  | 
| 81106 | 268  | 
      conclude = conclude {open_bundle = open_bundle, invisible = false} binding #> #2}
 | 
| 63270 | 269  | 
     {define = bad_operation,
 | 
270  | 
notes = bundle_notes,  | 
|
271  | 
abbrev = bad_operation,  | 
|
| 78095 | 272  | 
      declaration = fn {pos, ...} => bundle_declaration pos,
 | 
| 63270 | 273  | 
theory_registration = bad_operation,  | 
274  | 
locale_dependency = bad_operation,  | 
|
| 78095 | 275  | 
pretty = pretty binding};  | 
| 63270 | 276  | 
|
277  | 
end;  | 
|
278  | 
||
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47057 
diff
changeset
 | 
279  | 
end;  |