src/Pure/Isar/local_theory.ML
changeset 72154 2b41b710f6ef
parent 72153 bdbd6ff5fd0b
child 72433 7e0e497dacbc
--- a/src/Pure/Isar/local_theory.ML	Fri Aug 14 14:25:08 2020 +0200
+++ b/src/Pure/Isar/local_theory.ML	Fri Aug 14 14:40:24 2020 +0200
@@ -75,7 +75,8 @@
   val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory
   val init_target: {background_naming: Name_Space.naming, after_close: local_theory -> local_theory} ->
     operations -> local_theory -> Binding.scope * local_theory
-  val open_target: local_theory -> Binding.scope * local_theory
+  val begin_target: local_theory -> Binding.scope * local_theory
+  val open_target: local_theory -> local_theory
   val close_target: local_theory -> local_theory
   val subtarget: (local_theory -> local_theory) -> local_theory -> local_theory
   val subtarget_result: (morphism -> 'a -> 'b) -> (local_theory -> 'a * local_theory) ->
@@ -393,21 +394,23 @@
     val lthy' = Data.map (cons entry) target;
   in (scope, lthy') end;
 
-fun open_target lthy =
+fun begin_target lthy =
   init_target {background_naming = background_naming_of lthy, after_close = I}
     (operations_of lthy) lthy;
 
+val open_target = begin_target #> #2;
+
 fun close_target lthy =
   let
     val _ = assert_not_bottom lthy;
     val ({after_close, ...} :: rest) = Data.get lthy;
   in lthy |> Data.put rest |> reset |> after_close end;
 
-fun subtarget body = open_target #> #2 #> body #> close_target;
+fun subtarget body = open_target #> body #> close_target;
 
 fun subtarget_result decl body lthy =
   let
-    val (x, inner_lthy) = lthy |> open_target |> #2 |> body;
+    val (x, inner_lthy) = lthy |> open_target |> body;
     val lthy' = inner_lthy |> close_target;
     val phi = Proof_Context.export_morphism inner_lthy lthy';
   in (decl phi x, lthy') end;