| author | paulson <lp15@cam.ac.uk> | 
| Fri, 18 Feb 2022 21:40:01 +0000 | |
| changeset 75101 | f0e2023f361a | 
| parent 74561 | 8e6c973003c8 | 
| child 77908 | a6bd716a6124 | 
| 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 | |
| 72605 | 9 | type name = string | 
| 74261 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 wenzelm parents: 
74122diff
changeset | 10 | val bundle_space: Context.generic -> Name_Space.T | 
| 47057 | 11 | val check: Proof.context -> xstring * Position.T -> string | 
| 72605 | 12 | val get: Proof.context -> name -> Attrib.thms | 
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 13 | val read: Proof.context -> xstring * Position.T -> Attrib.thms | 
| 74122 | 14 | val extern: Proof.context -> string -> xstring | 
| 63270 | 15 | val print_bundles: bool -> Proof.context -> unit | 
| 63267 | 16 | val bundle: binding * Attrib.thms -> | 
| 47057 | 17 | (binding * typ option * mixfix) list -> local_theory -> local_theory | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
56334diff
changeset | 18 | val bundle_cmd: binding * (Facts.ref * Token.src list) list -> | 
| 47057 | 19 | (binding * string option * mixfix) list -> local_theory -> local_theory | 
| 63270 | 20 | val init: binding -> theory -> local_theory | 
| 72605 | 21 | val unbundle: name list -> local_theory -> local_theory | 
| 63282 | 22 | val unbundle_cmd: (xstring * Position.T) list -> local_theory -> local_theory | 
| 72605 | 23 | val includes: name list -> Proof.context -> Proof.context | 
| 47067 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 wenzelm parents: 
47066diff
changeset | 24 | val includes_cmd: (xstring * Position.T) list -> Proof.context -> Proof.context | 
| 72605 | 25 | val include_: name list -> Proof.state -> Proof.state | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 26 | val include_cmd: (xstring * Position.T) list -> Proof.state -> Proof.state | 
| 72605 | 27 | val including: name list -> Proof.state -> Proof.state | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 28 | val including_cmd: (xstring * Position.T) list -> Proof.state -> Proof.state | 
| 47057 | 29 | end; | 
| 30 | ||
| 31 | structure Bundle: BUNDLE = | |
| 32 | struct | |
| 33 | ||
| 63270 | 34 | (** context data **) | 
| 35 | ||
| 36 | structure Data = Generic_Data | |
| 37 | ( | |
| 38 | type T = Attrib.thms Name_Space.table * Attrib.thms option; | |
| 74261 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 wenzelm parents: 
74122diff
changeset | 39 | val empty : T = (Name_Space.empty_table Markup.bundleN, NONE); | 
| 63270 | 40 | fun merge ((tab1, target1), (tab2, target2)) = | 
| 41 | (Name_Space.merge_tables (tab1, tab2), merge_options (target1, target2)); | |
| 42 | ); | |
| 43 | ||
| 44 | ||
| 45 | (* bundles *) | |
| 46 | ||
| 72605 | 47 | type name = string | 
| 48 | ||
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 49 | val get_all_generic = #1 o Data.get; | 
| 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 50 | val get_all = get_all_generic o Context.Proof; | 
| 63270 | 51 | |
| 74261 
d28a51dd9da6
export other entities, e.g. relevant for formal document output;
 wenzelm parents: 
74122diff
changeset | 52 | 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: 
74122diff
changeset | 53 | |
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 54 | fun check ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_all ctxt); | 
| 63270 | 55 | |
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 56 | val get = Name_Space.get o get_all_generic o Context.Proof; | 
| 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 57 | |
| 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 58 | fun read ctxt = get ctxt o check ctxt; | 
| 63270 | 59 | |
| 74122 | 60 | fun extern ctxt = Name_Space.extern ctxt (Name_Space.space_of_table (get_all ctxt)); | 
| 61 | ||
| 67627 | 62 | fun define_bundle (b, bundle) context = | 
| 63270 | 63 | let | 
| 67652 | 64 | val bundle' = Attrib.trim_context_thms bundle; | 
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 65 | val (name, bundles') = Name_Space.define context true (b, bundle') (get_all_generic context); | 
| 63270 | 66 | val context' = (Data.map o apfst o K) bundles' context; | 
| 67 | in (name, context') end; | |
| 68 | ||
| 69 | ||
| 70 | (* target -- bundle under construction *) | |
| 71 | ||
| 72 | fun the_target thy = | |
| 73 | (case #2 (Data.get (Context.Theory thy)) of | |
| 74 | SOME thms => thms | |
| 75 | | NONE => error "Missing bundle target"); | |
| 76 | ||
| 77 | val reset_target = (Context.theory_map o Data.map o apsnd o K) NONE; | |
| 67652 | 78 | val set_target = Context.theory_map o Data.map o apsnd o K o SOME o Attrib.trim_context_thms; | 
| 63270 | 79 | |
| 80 | fun augment_target thms = | |
| 81 | Local_Theory.background_theory (fn thy => set_target (the_target thy @ thms) thy); | |
| 82 | ||
| 83 | ||
| 84 | (* print bundles *) | |
| 85 | ||
| 86 | fun pretty_bundle ctxt (markup_name, bundle) = | |
| 87 | let | |
| 88 | val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt; | |
| 63275 | 89 | fun prt_thm_attribs atts th = | 
| 90 | Pretty.block (Pretty.breaks (prt_thm th :: Attrib.pretty_attribs ctxt atts)); | |
| 63270 | 91 | fun prt_thms (ths, []) = map prt_thm ths | 
| 63275 | 92 | | prt_thms (ths, atts) = map (prt_thm_attribs atts) ths; | 
| 63270 | 93 | in | 
| 94 | Pretty.block ([Pretty.keyword1 "bundle", Pretty.str " ", Pretty.mark_str markup_name] @ | |
| 63275 | 95 | (if null bundle then [] else Pretty.fbreaks (Pretty.str " =" :: maps prt_thms bundle))) | 
| 63270 | 96 | end; | 
| 97 | ||
| 98 | fun print_bundles verbose ctxt = | |
| 99 | Pretty.writeln_chunks | |
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 100 | (map (pretty_bundle ctxt) (Name_Space.markup_table verbose ctxt (get_all ctxt))); | 
| 63270 | 101 | |
| 102 | ||
| 103 | ||
| 104 | (** define bundle **) | |
| 47057 | 105 | |
| 63267 | 106 | fun transform_bundle phi = | 
| 61814 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 wenzelm parents: 
61268diff
changeset | 107 | map (fn (fact, atts) => (Morphism.fact phi fact, (map o map) (Token.transform phi) atts)); | 
| 47057 | 108 | |
| 109 | ||
| 63270 | 110 | (* command *) | 
| 47057 | 111 | |
| 112 | local | |
| 113 | ||
| 60469 | 114 | fun gen_bundle prep_fact prep_att add_fixes (binding, raw_bundle) raw_fixes lthy = | 
| 47057 | 115 | let | 
| 60469 | 116 | val (_, ctxt') = add_fixes raw_fixes lthy; | 
| 47057 | 117 | val bundle0 = raw_bundle | 
| 53111 | 118 | |> map (fn (fact, atts) => (prep_fact ctxt' fact, map (prep_att ctxt') atts)); | 
| 47057 | 119 | val bundle = | 
| 63352 | 120 | Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)] |> map snd |> flat | 
| 47057 | 121 | |> transform_bundle (Proof_Context.export_morphism ctxt' lthy); | 
| 122 | in | |
| 123 |     lthy |> Local_Theory.declaration {syntax = false, pervasive = true}
 | |
| 63270 | 124 | (fn phi => #2 o define_bundle (Morphism.binding phi binding, transform_bundle phi bundle)) | 
| 47057 | 125 | end; | 
| 126 | ||
| 127 | in | |
| 128 | ||
| 60469 | 129 | val bundle = gen_bundle (K I) (K I) Proof_Context.add_fixes; | 
| 130 | val bundle_cmd = gen_bundle Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; | |
| 47057 | 131 | |
| 132 | end; | |
| 133 | ||
| 134 | ||
| 63270 | 135 | (* target *) | 
| 136 | ||
| 137 | local | |
| 138 | ||
| 139 | fun bad_operation _ = error "Not possible in bundle target"; | |
| 140 | ||
| 63275 | 141 | fun conclude invisible binding = | 
| 63270 | 142 | Local_Theory.background_theory_result (fn thy => | 
| 63275 | 143 | thy | 
| 144 | |> invisible ? Context_Position.set_visible_global false | |
| 145 | |> Context.Theory | |
| 146 | |> define_bundle (binding, the_target thy) | |
| 63277 | 147 | ||> (Context.the_theory | 
| 148 | #> invisible ? Context_Position.restore_visible_global thy | |
| 149 | #> reset_target)); | |
| 63270 | 150 | |
| 151 | fun pretty binding lthy = | |
| 152 | let | |
| 153 | val bundle = the_target (Proof_Context.theory_of lthy); | |
| 63275 | 154 | val (name, lthy') = lthy | 
| 155 | |> Local_Theory.raw_theory (Context_Position.set_visible_global false) | |
| 156 | |> conclude true binding; | |
| 63270 | 157 | val thy_ctxt' = Proof_Context.init_global (Proof_Context.theory_of lthy'); | 
| 158 | val markup_name = | |
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 159 | Name_Space.markup_extern thy_ctxt' (Name_Space.space_of_table (get_all thy_ctxt')) name; | 
| 63270 | 160 | in [pretty_bundle lthy' (markup_name, bundle)] end; | 
| 161 | ||
| 162 | fun bundle_notes kind facts lthy = | |
| 163 | let | |
| 164 | val bundle = facts | |
| 165 | |> maps (fn ((_, more_atts), thms) => map (fn (ths, atts) => (ths, atts @ more_atts)) thms); | |
| 166 | in | |
| 167 | lthy | |
| 168 | |> augment_target (transform_bundle (Local_Theory.standard_morphism_theory lthy) bundle) | |
| 63278 
6c963f1fc388
avoid duplicate Attrib.local_notes in aux. context;
 wenzelm parents: 
63277diff
changeset | 169 | |> Generic_Target.standard_notes (op <>) kind facts | 
| 63270 | 170 | |> Attrib.local_notes kind facts | 
| 171 | end; | |
| 172 | ||
| 173 | fun bundle_declaration decl lthy = | |
| 174 | lthy | |
| 175 | |> (augment_target o Attrib.internal_declaration) | |
| 176 | (Morphism.transform (Local_Theory.standard_morphism_theory lthy) decl) | |
| 177 | |> Generic_Target.standard_declaration (K true) decl; | |
| 178 | ||
| 179 | in | |
| 180 | ||
| 181 | fun init binding thy = | |
| 182 | thy | |
| 72516 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 haftmann parents: 
72504diff
changeset | 183 | |> Local_Theory.init | 
| 66335 
a849ce33923d
treat exit separate from regular local theory operations
 haftmann parents: 
63352diff
changeset | 184 |      {background_naming = Sign.naming_of thy,
 | 
| 66337 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66336diff
changeset | 185 | setup = set_target [] #> Proof_Context.init_global, | 
| 
5caea089dd17
more structural sharing between common target Generic_Target.init
 haftmann parents: 
66336diff
changeset | 186 | conclude = conclude false binding #> #2} | 
| 63270 | 187 |      {define = bad_operation,
 | 
| 188 | notes = bundle_notes, | |
| 189 | abbrev = bad_operation, | |
| 190 | declaration = K bundle_declaration, | |
| 191 | theory_registration = bad_operation, | |
| 192 | locale_dependency = bad_operation, | |
| 66335 
a849ce33923d
treat exit separate from regular local theory operations
 haftmann parents: 
63352diff
changeset | 193 | pretty = pretty binding} | 
| 63270 | 194 | |
| 195 | end; | |
| 196 | ||
| 197 | ||
| 198 | ||
| 63282 | 199 | (** activate bundles **) | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 200 | |
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 201 | local | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 202 | |
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 203 | fun gen_activate notes prep_bundle args ctxt = | 
| 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 204 | let val decls = maps (prep_bundle ctxt) args in | 
| 63031 | 205 | ctxt | 
| 206 | |> Context_Position.set_visible false | |
| 63352 | 207 | |> notes [(Binding.empty_atts, decls)] |> #2 | 
| 63031 | 208 | |> Context_Position.restore_visible ctxt | 
| 209 | end; | |
| 47057 | 210 | |
| 72504 | 211 | fun gen_unbundle prep_bundle = gen_activate Local_Theory.notes prep_bundle; | 
| 212 | ||
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 213 | fun gen_includes prep_bundle = gen_activate (Attrib.local_notes "") prep_bundle; | 
| 63282 | 214 | |
| 72504 | 215 | fun gen_include prep_bundle bs = | 
| 216 | Proof.assert_forward #> Proof.map_context (gen_includes prep_bundle bs) #> Proof.reset_facts; | |
| 217 | ||
| 218 | fun gen_including prep_bundle bs = | |
| 219 | Proof.assert_backward #> Proof.map_context (gen_includes prep_bundle bs) | |
| 220 | ||
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 221 | in | 
| 47057 | 222 | |
| 72504 | 223 | val unbundle = gen_unbundle get; | 
| 224 | val unbundle_cmd = gen_unbundle read; | |
| 63282 | 225 | |
| 72452 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 226 | val includes = gen_includes get; | 
| 
9017dfa56367
avoid _cmd suffix where no Isar command is involved
 haftmann parents: 
72450diff
changeset | 227 | val includes_cmd = gen_includes read; | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 228 | |
| 72504 | 229 | val include_ = gen_include get; | 
| 230 | val include_cmd = gen_include read; | |
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 231 | |
| 72504 | 232 | val including = gen_including get; | 
| 233 | val including_cmd = gen_including read; | |
| 47067 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 wenzelm parents: 
47066diff
changeset | 234 | |
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47057diff
changeset | 235 | end; | 
| 47057 | 236 | |
| 237 | end; |