src/Pure/Isar/bundle.ML
author wenzelm
Mon, 21 Jul 2025 16:21:37 +0200
changeset 82892 45107da819fc
parent 82588 62b4b9f336c5
permissions -rw-r--r--
eliminate odd Unicode characters (amending e9f3b94eb6a0, b69e4da2604b, 8f0b2daa7eaa, 8d1e295aab70);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/bundle.ML
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     3
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     4
Bundled declarations (notes etc.).
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     5
*)
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     6
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     7
signature BUNDLE =
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
     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
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
    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
7d3e818fe21f antiquotation for bundles
haftmann
parents: 72605
diff changeset
    15
  val extern: Proof.context -> string -> xstring
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    16
  val print_bundles: bool -> Proof.context -> unit
72605
a4cb880e873a type alias for mixin bundles
haftmann
parents: 72516
diff changeset
    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
a4cb880e873a type alias for mixin bundles
haftmann
parents: 72516
diff changeset
    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
a4cb880e873a type alias for mixin bundles
haftmann
parents: 72516
diff changeset
    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
a4cb880e873a type alias for mixin bundles
haftmann
parents: 72516
diff changeset
    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
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
    25
  val bundle: {open_bundle: bool} -> binding * Attrib.thms ->
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    26
    (binding * typ option * mixfix) list -> local_theory -> local_theory
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
    27
  val bundle_cmd: {open_bundle: bool} -> binding * (Facts.ref * Token.src list) list ->
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    28
    (binding * string option * mixfix) list -> local_theory -> local_theory
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
    29
  val init: {open_bundle: bool} -> binding -> theory -> local_theory
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
    30
end;
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
    31
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
    32
structure Bundle: BUNDLE =
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
    33
struct
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
    34
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    35
(** context data **)
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    36
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    37
structure Data = Generic_Data
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    38
(
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    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
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    41
  fun merge ((tab1, target1), (tab2, target2)) =
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    42
    (Name_Space.merge_tables (tab1, tab2), merge_options (target1, target2));
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    43
);
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    44
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    45
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    46
(* bundles *)
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    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
a4cb880e873a type alias for mixin bundles
haftmann
parents: 72516
diff changeset
    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
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    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
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    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
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
    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
7d3e818fe21f antiquotation for bundles
haftmann
parents: 72605
diff changeset
    69
fun extern ctxt = Name_Space.extern ctxt (Name_Space.space_of_table (get_all ctxt));
7d3e818fe21f antiquotation for bundles
haftmann
parents: 72605
diff changeset
    70
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    71
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    72
(* context and morphisms *)
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    73
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    74
val trim_context_bundle =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    75
  map (fn (fact, atts) => (map Thm.trim_context fact, (map o map) Token.trim_context atts));
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    76
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    77
fun transfer_bundle thy =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    78
  map (fn (fact, atts) => (map (Thm.transfer thy) fact, (map o map) (Token.transfer thy) atts));
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    79
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    80
fun transform_bundle phi =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    81
  map (fn (fact, atts) => (Morphism.fact phi fact, (map o map) (Token.transform phi) atts));
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    82
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    83
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
    84
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    85
(* target -- bundle under construction *)
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    86
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    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
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    90
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    91
val reset_target = (Context.theory_map o Data.map o apsnd o K) NONE;
67652
11716a084cae clarified signature;
wenzelm
parents: 67627
diff changeset
    92
val set_target = Context.theory_map o Data.map o apsnd o K o SOME o Attrib.trim_context_thms;
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    93
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    94
fun augment_target thms =
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    95
  Local_Theory.background_theory (fn thy => set_target (the_target thy @ thms) thy);
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    96
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    97
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    98
(* print bundles *)
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
    99
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   100
fun pretty_bundle ctxt (markup_name, bundle) =
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   101
  let
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   102
    val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt;
63275
wenzelm
parents: 63270
diff changeset
   103
    fun prt_thm_attribs atts th =
wenzelm
parents: 63270
diff changeset
   104
      Pretty.block (Pretty.breaks (prt_thm th :: Attrib.pretty_attribs ctxt atts));
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   105
    fun prt_thms (ths, []) = map prt_thm ths
63275
wenzelm
parents: 63270
diff changeset
   106
      | prt_thms (ths, atts) = map (prt_thm_attribs atts) ths;
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   107
  in
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   108
    Pretty.block ([Pretty.keyword1 "bundle", Pretty.str " ", Pretty.mark_str markup_name] @
63275
wenzelm
parents: 63270
diff changeset
   109
      (if null bundle then [] else Pretty.fbreaks (Pretty.str " =" :: maps prt_thms bundle)))
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   110
  end;
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   111
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   112
fun print_bundles verbose ctxt =
82588
wenzelm
parents: 82587
diff changeset
   113
  Name_Space.markup_table verbose ctxt (get_all ctxt)
wenzelm
parents: 82587
diff changeset
   114
  |> map (pretty_bundle ctxt)
wenzelm
parents: 82587
diff changeset
   115
  |> Pretty.chunks
wenzelm
parents: 82587
diff changeset
   116
  |> Pretty.writeln;
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   117
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   118
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   119
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   120
(** 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
   121
81116
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 81106
diff changeset
   122
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
   123
  [([Drule.dummy_thm],
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 81106
diff changeset
   124
    [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
   125
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 81106
diff changeset
   126
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
   127
  let
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 81106
diff changeset
   128
    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
   129
    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
   130
    val add1 = Syntax.effective_polarity ctxt add;
81514
98cb63b447c6 clarified 'unbundle' polarity, according to algebraic group laws;
wenzelm
parents: 81117
diff changeset
   131
    val bundle' = if add1 then bundle else rev (map (apsnd rev) bundle);
81116
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 81106
diff changeset
   132
  in
81104
wenzelm
parents: 78095
diff changeset
   133
    ctxt
wenzelm
parents: 78095
diff changeset
   134
    |> Context_Position.set_visible false
81514
98cb63b447c6 clarified 'unbundle' polarity, according to algebraic group laws;
wenzelm
parents: 81117
diff changeset
   135
    |> notes [(Binding.empty_atts, polarity_decls add1 @ bundle' @ polarity_decls add0)] |> snd
81104
wenzelm
parents: 78095
diff changeset
   136
    |> Context_Position.restore_visible ctxt
wenzelm
parents: 78095
diff changeset
   137
  end;
wenzelm
parents: 78095
diff changeset
   138
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   139
local
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   140
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   141
fun gen_unbundle loc prep args ctxt =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   142
  let
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   143
    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
   144
    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
   145
  in fold (apply_bundle loc) bundles ctxt end;
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   146
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   147
fun gen_includes prep = gen_unbundle false prep;
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   148
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   149
fun gen_include prep bs =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   150
  Proof.assert_forward #> Proof.map_context (gen_includes prep bs) #> Proof.reset_facts;
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   151
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   152
fun gen_including prep bs =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   153
  Proof.assert_backward #> Proof.map_context (gen_includes prep bs);
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   154
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   155
in
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   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 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
   158
val unbundle_cmd = gen_unbundle true check_name;
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   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 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
   161
val includes_cmd = gen_includes check_name;
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   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 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
   164
val include_cmd = gen_include check_name;
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   165
81116
0fb1e2dd4122 first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents: 81106
diff changeset
   166
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
   167
val including_cmd = gen_including check_name;
81105
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   168
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   169
end;
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   170
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   171
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   172
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   173
(** define bundle **)
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   174
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   175
fun define_bundle (b, bundle) context =
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   176
  let
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   177
    val (name, bundles') = get_all_generic context
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   178
      |> Name_Space.define context true (b, trim_context_bundle bundle);
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   179
    val context' = (Data.map o apfst o K) bundles' context;
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   180
  in (name, context') end;
f14fcf94e0e9 tuned module structure;
wenzelm
parents: 81104
diff changeset
   181
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   182
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   183
(* command *)
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   184
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   185
local
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   186
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   187
fun gen_bundle prep_fact prep_att add_fixes {open_bundle} (binding, raw_bundle) raw_fixes lthy =
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   188
  let
60469
d1ea37df7358 tuned signature;
wenzelm
parents: 60379
diff changeset
   189
    val (_, ctxt') = add_fixes raw_fixes lthy;
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   190
    val bundle0 = raw_bundle
53111
f7f1636ee2ba proper context;
wenzelm
parents: 50739
diff changeset
   191
      |> map (fn (fact, atts) => (prep_fact ctxt' fact, map (prep_att ctxt') atts));
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   192
    val bundle =
78066
wenzelm
parents: 78064
diff changeset
   193
      Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)]
wenzelm
parents: 78064
diff changeset
   194
      |> maps #2
78069
e5bd9b3c6f0f more standard treatment of morphism context;
wenzelm
parents: 78066
diff changeset
   195
      |> transform_bundle (Proof_Context.export_morphism ctxt' lthy)
e5bd9b3c6f0f more standard treatment of morphism context;
wenzelm
parents: 78066
diff changeset
   196
      |> trim_context_bundle;
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   197
  in
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 78069
diff changeset
   198
    lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding}
78069
e5bd9b3c6f0f more standard treatment of morphism context;
wenzelm
parents: 78066
diff changeset
   199
      (fn phi => fn context =>
e5bd9b3c6f0f more standard treatment of morphism context;
wenzelm
parents: 78066
diff changeset
   200
        let val psi = Morphism.set_trim_context'' context phi
e5bd9b3c6f0f more standard treatment of morphism context;
wenzelm
parents: 78066
diff changeset
   201
        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
   202
    |> open_bundle ? apply_bundle true (true, bundle)
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   203
  end;
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   204
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   205
in
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   206
60469
d1ea37df7358 tuned signature;
wenzelm
parents: 60379
diff changeset
   207
val bundle = gen_bundle (K I) (K I) Proof_Context.add_fixes;
d1ea37df7358 tuned signature;
wenzelm
parents: 60379
diff changeset
   208
val bundle_cmd = gen_bundle Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd;
47057
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   209
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   210
end;
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   211
12423b36fcc4 basic support for bundled declarations;
wenzelm
parents:
diff changeset
   212
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   213
(* target *)
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   214
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   215
local
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   216
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   217
fun bad_operation _ = error "Not possible in bundle target";
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   218
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   219
fun conclude {open_bundle, invisible} binding =
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   220
  Local_Theory.background_theory_result (fn thy =>
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   221
    let
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   222
      val (name, thy') = thy
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   223
        |> invisible ? Context_Position.set_visible_global false
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   224
        |> Context.Theory
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   225
        |> define_bundle (binding, the_target thy)
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   226
        ||> Context.the_theory;
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   227
      val bundle = get_global thy' name;
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   228
      val thy'' = thy'
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   229
        |> open_bundle ?
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   230
          (Context_Position.set_visible_global false #>
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   231
            Attrib.global_notes "" [(Binding.empty_atts, bundle)] #> snd)
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   232
        |> Context_Position.restore_visible_global thy
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   233
        |> reset_target;
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   234
    in (name, thy'') end);
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   235
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   236
fun pretty binding lthy =
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   237
  let
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   238
    val bundle = the_target (Proof_Context.theory_of lthy);
63275
wenzelm
parents: 63270
diff changeset
   239
    val (name, lthy') = lthy
wenzelm
parents: 63270
diff changeset
   240
      |> Local_Theory.raw_theory (Context_Position.set_visible_global false)
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   241
      |> conclude {open_bundle = false, invisible = true} binding;
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   242
    val thy_ctxt' = Proof_Context.init_global (Proof_Context.theory_of lthy');
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   243
    val markup_name =
72452
9017dfa56367 avoid _cmd suffix where no Isar command is involved
haftmann
parents: 72450
diff changeset
   244
      Name_Space.markup_extern thy_ctxt' (Name_Space.space_of_table (get_all thy_ctxt')) name;
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   245
  in [pretty_bundle lthy' (markup_name, bundle)] end;
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   246
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   247
fun bundle_notes kind facts lthy =
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   248
  let
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   249
    val bundle = facts
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   250
      |> maps (fn ((_, more_atts), thms) => map (fn (ths, atts) => (ths, atts @ more_atts)) thms);
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   251
  in
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   252
    lthy
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   253
    |> 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
   254
    |> Generic_Target.standard_notes (op <>) kind facts
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   255
    |> Attrib.local_notes kind facts
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   256
  end;
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   257
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 78069
diff changeset
   258
fun bundle_declaration pos decl lthy =
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   259
  lthy
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 78069
diff changeset
   260
  |> (augment_target o Attrib.internal_declaration pos)
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   261
    (Morphism.transform (Local_Theory.standard_morphism_theory lthy) decl)
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   262
  |> Generic_Target.standard_declaration (K true) decl;
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   263
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   264
in
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   265
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   266
fun init {open_bundle} binding thy =
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   267
  thy
72516
17dc99589a91 unified Local_Theory.init with Generic_Target.init
haftmann
parents: 72504
diff changeset
   268
  |> Local_Theory.init
66335
a849ce33923d treat exit separate from regular local theory operations
haftmann
parents: 63352
diff changeset
   269
     {background_naming = Sign.naming_of thy,
66337
5caea089dd17 more structural sharing between common target Generic_Target.init
haftmann
parents: 66336
diff changeset
   270
      setup = set_target [] #> Proof_Context.init_global,
81106
849efff7de15 provide 'open_bundle' command;
wenzelm
parents: 81105
diff changeset
   271
      conclude = conclude {open_bundle = open_bundle, invisible = false} binding #> #2}
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   272
     {define = bad_operation,
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   273
      notes = bundle_notes,
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   274
      abbrev = bad_operation,
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 78069
diff changeset
   275
      declaration = fn {pos, ...} => bundle_declaration pos,
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   276
      theory_registration = bad_operation,
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   277
      locale_dependency = bad_operation,
78095
bc42c074e58f tuned signature: more position information;
wenzelm
parents: 78069
diff changeset
   278
      pretty = pretty binding};
63270
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   279
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   280
end;
7dd3ee7ee422 support for bundle definition via target;
wenzelm
parents: 63267
diff changeset
   281
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47057
diff changeset
   282
end;