src/Pure/Pure.thy
changeset 63270 7dd3ee7ee422
parent 63180 ddfd021884b4
child 63271 ecaa677d20bc
--- a/src/Pure/Pure.thy	Thu Jun 09 12:21:15 2016 +0200
+++ b/src/Pure/Pure.thy	Thu Jun 09 15:41:49 2016 +0200
@@ -32,6 +32,7 @@
     "parse_ast_translation" "parse_translation" "print_translation"
     "typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML"
   and "bundle" :: thy_decl
+  and "bundle_target" :: thy_decl_block
   and "include" "including" :: prf_decl
   and "print_bundles" :: diag
   and "context" "locale" "experiment" :: thy_decl_block
@@ -496,6 +497,10 @@
       >> (uncurry Bundle.bundle_cmd));
 
 val _ =
+  Outer_Syntax.command @{command_keyword bundle_target} "define bundle of declarations"
+    (Parse.binding --| Parse.begin >> (Toplevel.begin_local_theory true o Bundle.init));
+
+val _ =
   Outer_Syntax.command @{command_keyword include}
     "include declarations from bundle in proof body"
     (Scan.repeat1 (Parse.position Parse.name) >> (Toplevel.proof o Bundle.include_cmd));