src/Pure/Isar/local_theory.ML
author wenzelm
Fri, 31 Aug 2012 16:35:30 +0200
changeset 49042 01041f7bf9b4
parent 47281 d6c76b1823fb
child 49058 2924a83a4a0b
permissions -rw-r--r--
more precise register_proofs for local goals; tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/local_theory.ML
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     3
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
     4
Local theory operations, with abstract target context.
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     5
*)
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
     6
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
     7
type local_theory = Proof.context;
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 36610
diff changeset
     8
type generic_theory = Context.generic;
18951
4f5f6c632127 type local_theory = Proof.context;
wenzelm
parents: 18876
diff changeset
     9
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    10
signature LOCAL_THEORY =
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    11
sig
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    12
  type operations
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
    13
  val assert: local_theory -> local_theory
47273
ea089b484157 better restore to first target, not last target;
wenzelm
parents: 47271
diff changeset
    14
  val restore: local_theory -> local_theory
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    15
  val level: Proof.context -> int
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
    16
  val assert_bottom: bool -> local_theory -> local_theory
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    17
  val open_target: Name_Space.naming -> operations -> local_theory -> local_theory
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    18
  val close_target: local_theory -> local_theory
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
    19
  val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
    20
  val naming_of: local_theory -> Name_Space.naming
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
    21
  val full_name: local_theory -> binding -> string
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
    22
  val map_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
    23
  val conceal: local_theory -> local_theory
33724
5ee13e0428d2 uniform new_group/reset_group;
wenzelm
parents: 33671
diff changeset
    24
  val new_group: local_theory -> local_theory
5ee13e0428d2 uniform new_group/reset_group;
wenzelm
parents: 33671
diff changeset
    25
  val reset_group: local_theory -> local_theory
33276
f2bc8bc6e73d added restore_naming;
wenzelm
parents: 33166
diff changeset
    26
  val restore_naming: local_theory -> local_theory -> local_theory
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
    27
  val standard_morphism: local_theory -> Proof.context -> morphism
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
    28
  val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    29
  val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
    30
  val raw_theory: (theory -> theory) -> local_theory -> local_theory
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38756
diff changeset
    31
  val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38756
diff changeset
    32
  val background_theory: (theory -> theory) -> local_theory -> local_theory
49042
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
    33
  val begin_proofs: local_theory -> local_theory
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
    34
  val register_proofs: thm list -> local_theory -> local_theory
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
    35
  val target_of: local_theory -> Proof.context
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    36
  val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
    37
  val target_morphism: local_theory -> morphism
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 36610
diff changeset
    38
  val propagate_ml_env: generic_theory -> generic_theory
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
    39
  val operations_of: local_theory -> operations
33764
7bcefaab8d41 Local_Theory.define: eliminated slightly odd kind argument -- such low-level definitions should be hardly ever exposed to end-users anyway;
wenzelm
parents: 33724
diff changeset
    40
  val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
    41
    (term * (string * thm)) * local_theory
46992
eeea81b86b70 refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents: 45298
diff changeset
    42
  val define_internal: (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
eeea81b86b70 refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents: 45298
diff changeset
    43
    (term * (string * thm)) * local_theory
33670
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
    44
  val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
    45
  val notes: (Attrib.binding * (thm list * Attrib.src list) list) list ->
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
    46
    local_theory -> (string * thm list) list * local_theory
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
    47
  val notes_kind: string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
    48
    local_theory -> (string * thm list) list * local_theory
38308
0e4649095739 try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents: 37949
diff changeset
    49
  val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
0e4649095739 try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents: 37949
diff changeset
    50
    (term * term) * local_theory
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
    51
  val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory
38308
0e4649095739 try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents: 37949
diff changeset
    52
  val pretty: local_theory -> Pretty.T list
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
    53
  val set_defsort: sort -> local_theory -> local_theory
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 33764
diff changeset
    54
  val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24933
diff changeset
    55
  val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
35738
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
    56
  val class_alias: binding -> class -> local_theory -> local_theory
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
    57
  val type_alias: binding -> string -> local_theory -> local_theory
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
    58
  val const_alias: binding -> string -> local_theory -> local_theory
47061
355317493f34 clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
wenzelm
parents: 46992
diff changeset
    59
  val init: Name_Space.naming -> operations -> Proof.context -> local_theory
21292
11143b6ad87f simplified exit;
wenzelm
parents: 21273
diff changeset
    60
  val exit: local_theory -> Proof.context
28395
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
    61
  val exit_global: local_theory -> theory
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
    62
  val exit_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * Proof.context
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
    63
  val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    64
end;
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    65
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
    66
structure Local_Theory: LOCAL_THEORY =
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    67
struct
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    68
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    69
(** local theory data **)
19059
b4ca3100e818 init/exit no longer change the theory (no naming);
wenzelm
parents: 19032
diff changeset
    70
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    71
(* type lthy *)
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    72
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    73
type operations =
46992
eeea81b86b70 refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents: 45298
diff changeset
    74
 {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
    75
    (term * (string * thm)) * local_theory,
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
    76
  notes: string ->
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    77
    (Attrib.binding * (thm list * Attrib.src list) list) list ->
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 28017
diff changeset
    78
    local_theory -> (string * thm list) list * local_theory,
38308
0e4649095739 try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents: 37949
diff changeset
    79
  abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
0e4649095739 try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents: 37949
diff changeset
    80
    (term * term) * local_theory,
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
    81
  declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory,
38308
0e4649095739 try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents: 37949
diff changeset
    82
  pretty: local_theory -> Pretty.T list,
21292
11143b6ad87f simplified exit;
wenzelm
parents: 21273
diff changeset
    83
  exit: local_theory -> Proof.context};
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    84
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    85
type lthy =
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
    86
 {naming: Name_Space.naming,
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    87
  operations: operations,
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    88
  target: Proof.context};
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    89
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    90
fun make_lthy (naming, operations, target) : lthy =
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    91
  {naming = naming, operations = operations, target = target};
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    92
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    93
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
    94
(* context data *)
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    95
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33456
diff changeset
    96
structure Data = Proof_Data
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
    97
(
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    98
  type T = lthy list;
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
    99
  fun init _ = [];
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   100
);
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   101
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   102
fun assert lthy =
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   103
  if null (Data.get lthy) then error "Missing local theory context" else lthy;
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   104
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   105
val get_last_lthy = List.last o Data.get o assert;
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   106
val get_first_lthy = hd o Data.get o assert;
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   107
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   108
fun map_first_lthy f = assert #>
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   109
  Data.map (fn {naming, operations, target} :: parents =>
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   110
    make_lthy (f (naming, operations, target)) :: parents);
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   111
47273
ea089b484157 better restore to first target, not last target;
wenzelm
parents: 47271
diff changeset
   112
fun restore lthy = #target (get_first_lthy lthy) |> Data.put (Data.get lthy);
ea089b484157 better restore to first target, not last target;
wenzelm
parents: 47271
diff changeset
   113
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   114
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   115
(* nested structure *)
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
   116
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   117
val level = length o Data.get;
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   118
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   119
fun assert_bottom b lthy =
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   120
  let
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   121
    val _ = assert lthy;
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   122
    val b' = level lthy <= 1;
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   123
  in
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   124
    if b andalso not b' then error "Not at bottom of local theory nesting"
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   125
    else if not b andalso b' then error "Already at bottom of local theory nesting"
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   126
    else lthy
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   127
  end;
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   128
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   129
fun open_target naming operations target = assert target
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   130
  |> Data.map (cons (make_lthy (naming, operations, target)));
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   131
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   132
fun close_target lthy =
47281
d6c76b1823fb better restore after close_target;
wenzelm
parents: 47273
diff changeset
   133
  assert_bottom false lthy |> Data.map tl |> restore;
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   134
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   135
fun map_contexts f lthy =
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   136
  let val n = level lthy in
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   137
    lthy |> (Data.map o map_index) (fn (i, {naming, operations, target}) =>
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   138
      make_lthy (naming, operations,
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   139
        target
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   140
        |> Context_Position.set_visible false
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   141
        |> f (n - i - 1)
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   142
        |> Context_Position.restore_visible target))
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   143
    |> f n
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   144
  end;
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   145
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   146
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
   147
(* naming *)
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   148
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   149
val naming_of = #naming o get_first_lthy;
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
   150
val full_name = Name_Space.full_name o naming_of;
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   151
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   152
fun map_naming f =
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   153
  map_first_lthy (fn (naming, operations, target) => (f naming, operations, target));
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   154
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
   155
val conceal = map_naming Name_Space.conceal;
33724
5ee13e0428d2 uniform new_group/reset_group;
wenzelm
parents: 33671
diff changeset
   156
val new_group = map_naming Name_Space.new_group;
5ee13e0428d2 uniform new_group/reset_group;
wenzelm
parents: 33671
diff changeset
   157
val reset_group = map_naming Name_Space.reset_group;
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   158
33276
f2bc8bc6e73d added restore_naming;
wenzelm
parents: 33166
diff changeset
   159
val restore_naming = map_naming o K o naming_of;
f2bc8bc6e73d added restore_naming;
wenzelm
parents: 33166
diff changeset
   160
26131
91024979b9eb maintain group in lthy data, implicit use in operations;
wenzelm
parents: 25984
diff changeset
   161
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   162
(* standard morphisms *)
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   163
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   164
fun standard_morphism lthy ctxt =
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   165
  Proof_Context.norm_export_morphism lthy ctxt $>
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   166
  Morphism.binding_morphism (Name_Space.transform_binding (naming_of lthy));
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   167
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   168
fun standard_form lthy ctxt x =
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   169
  Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x);
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   170
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   171
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   172
(* background theory *)
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   173
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   174
fun raw_theory_result f lthy =
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   175
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   176
    val (res, thy') = f (Proof_Context.theory_of lthy);
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   177
    val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy;
20960
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   178
  in (res, lthy') end;
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   179
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   180
fun raw_theory f = #2 o raw_theory_result (f #> pair ());
f342e82f4e58 added raw_theory(_result);
wenzelm
parents: 20910
diff changeset
   181
28379
0de8a35b866a Theory.checkpoint for main operations, admits concurrent proofs;
wenzelm
parents: 28115
diff changeset
   182
val checkpoint = raw_theory Theory.checkpoint;
0de8a35b866a Theory.checkpoint for main operations, admits concurrent proofs;
wenzelm
parents: 28115
diff changeset
   183
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38756
diff changeset
   184
fun background_theory_result f lthy =
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
   185
  lthy |> raw_theory_result (fn thy =>
33157
56f836b9414f allow name space entries to be "concealed" -- via binding/naming/local_theory;
wenzelm
parents: 33095
diff changeset
   186
    thy
33166
55f250ef9e31 maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents: 33157
diff changeset
   187
    |> Sign.map_naming (K (naming_of lthy))
33157
56f836b9414f allow name space entries to be "concealed" -- via binding/naming/local_theory;
wenzelm
parents: 33095
diff changeset
   188
    |> f
36004
5d79ca55a52b implicit checkpoint in Local_Theory.theory as well -- no longer export Local_Theory.checkpoint;
wenzelm
parents: 35798
diff changeset
   189
    ||> Sign.restore_naming thy
5d79ca55a52b implicit checkpoint in Local_Theory.theory as well -- no longer export Local_Theory.checkpoint;
wenzelm
parents: 35798
diff changeset
   190
    ||> Theory.checkpoint);
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   191
38757
2b3e054ae6fc renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
wenzelm
parents: 38756
diff changeset
   192
fun background_theory f = #2 o background_theory_result (f #> pair ());
19661
baab85f25e0e added syntax interface;
wenzelm
parents: 19630
diff changeset
   193
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   194
49042
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
   195
(* forked proofs *)
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
   196
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
   197
val begin_proofs = background_theory (Context.theory_map Thm.begin_proofs);
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
   198
val register_proofs = background_theory o Context.theory_map o Thm.register_proofs;
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
   199
01041f7bf9b4 more precise register_proofs for local goals;
wenzelm
parents: 47281
diff changeset
   200
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   201
(* target contexts *)
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   202
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   203
val target_of = #target o get_last_lthy;
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   204
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   205
fun target f lthy =
18744
a9a5ee0e43e2 Local theory operations, with optional target locale.
wenzelm
parents:
diff changeset
   206
  let
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   207
    val ctxt = target_of lthy;
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   208
    val ctxt' = ctxt
33383
12d79ece3f7e modernized structure Context_Position;
wenzelm
parents: 33281
diff changeset
   209
      |> Context_Position.set_visible false
28406
daeb21fec18f target update: invisible context position avoids duplication of reports etc.;
wenzelm
parents: 28395
diff changeset
   210
      |> f
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   211
      |> Context_Position.restore_visible ctxt;
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   212
    val thy' = Proof_Context.theory_of ctxt';
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   213
  in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end;
18876
ddb6803da197 added consts_retricted;
wenzelm
parents: 18823
diff changeset
   214
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   215
fun target_morphism lthy = standard_morphism lthy (target_of lthy);
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   216
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 36610
diff changeset
   217
fun propagate_ml_env (context as Context.Proof lthy) =
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   218
      let val inherit = ML_Env.inherit context in
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   219
        lthy
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   220
        |> background_theory (Context.theory_map inherit)
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   221
        |> map_contexts (K (Context.proof_map inherit))
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   222
        |> Context.Proof
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   223
      end
37949
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 36610
diff changeset
   224
  | propagate_ml_env context = context;
48a874444164 moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
wenzelm
parents: 36610
diff changeset
   225
18767
2f064e6bea7e added actual operations;
wenzelm
parents: 18744
diff changeset
   226
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   227
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   228
(** operations **)
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   229
47245
ff1770df59b8 Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents: 47081
diff changeset
   230
val operations_of = #operations o get_first_lthy;
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   231
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   232
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   233
(* primitives *)
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   234
47066
8a6124d09ff5 basic support for nested contexts including bundles;
wenzelm
parents: 47064
diff changeset
   235
fun operation f lthy = f (operations_of lthy) lthy;
47061
355317493f34 clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
wenzelm
parents: 46992
diff changeset
   236
fun operation2 f x y = operation (fn ops => f ops x y);
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   237
20888
0ddd2f297ae7 turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents: 20784
diff changeset
   238
val pretty = operation #pretty;
28379
0de8a35b866a Theory.checkpoint for main operations, admits concurrent proofs;
wenzelm
parents: 28115
diff changeset
   239
val abbrev = apsnd checkpoint ooo operation2 #abbrev;
46992
eeea81b86b70 refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents: 45298
diff changeset
   240
val define = apsnd checkpoint oo operation2 #define false;
eeea81b86b70 refined Local_Theory.define vs. Local_Theory.define_internal, which allows to pass alternative name to the foundational axiom -- expecially important for 'instantiation' or 'overloading', which loose name information due to Long_Name.base_name cooking etc.;
wenzelm
parents: 45298
diff changeset
   241
val define_internal = apsnd checkpoint oo operation2 #define true;
33670
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
   242
val notes_kind = apsnd checkpoint ooo operation2 #notes;
33456
fbd47f9b9b12 allow "pervasive" local theory declarations, which are applied the background theory;
wenzelm
parents: 33383
diff changeset
   243
val declaration = checkpoint ooo operation2 #declaration;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   244
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   245
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   246
(* basic derived operations *)
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   247
33670
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
   248
val notes = notes_kind "";
02b7738aef6a eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents: 33519
diff changeset
   249
fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single;
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   250
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   251
fun set_defsort S =
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   252
  declaration {syntax = true, pervasive = false}
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   253
    (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S)));
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   254
35738
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   255
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   256
(* notation *)
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   257
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 33764
diff changeset
   258
fun type_notation add mode raw_args lthy =
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   259
  let val args = map (apfst (Morphism.typ (target_morphism lthy))) raw_args in
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   260
    declaration {syntax = true, pervasive = false}
47247
23654b331d5c tuned signature;
wenzelm
parents: 47245
diff changeset
   261
      (Proof_Context.generic_type_notation add mode args) lthy
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   262
  end;
35412
b8dead547d9e more uniform treatment of syntax for types vs. consts;
wenzelm
parents: 33764
diff changeset
   263
24949
5f00e3532418 generalized notation interface (add or del);
wenzelm
parents: 24933
diff changeset
   264
fun notation add mode raw_args lthy =
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   265
  let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args in
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   266
    declaration {syntax = true, pervasive = false}
47247
23654b331d5c tuned signature;
wenzelm
parents: 47245
diff changeset
   267
      (Proof_Context.generic_notation add mode args) lthy
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   268
  end;
21743
fe94c6797c09 added notation/abbrev (from term_syntax.ML);
wenzelm
parents: 21700
diff changeset
   269
21664
dd4e89123359 notation/abbreviation: more serious handling of morphisms;
wenzelm
parents: 21614
diff changeset
   270
35738
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   271
(* name space aliases *)
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   272
35798
fd1bb29f8170 replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents: 35738
diff changeset
   273
fun alias global_alias local_alias b name =
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 42360
diff changeset
   274
  declaration {syntax = true, pervasive = false} (fn phi =>
35738
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   275
    let val b' = Morphism.binding phi b
45298
aa35859c8741 uniform treatment of syntax declaration wrt. aux. context (NB: notation avoids duplicate mixfix internally);
wenzelm
parents: 45291
diff changeset
   276
    in Context.mapping (global_alias b' name) (local_alias b' name) end);
35738
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   277
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   278
val class_alias = alias Sign.class_alias Proof_Context.class_alias;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   279
val type_alias = alias Sign.type_alias Proof_Context.type_alias;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   280
val const_alias = alias Sign.const_alias Proof_Context.const_alias;
35738
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   281
98fd035c3fe3 added Local_Theory.alias operations (independent of target);
wenzelm
parents: 35412
diff changeset
   282
36451
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   283
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   284
(** init and exit **)
ddc965e172c4 localized default sort;
wenzelm
parents: 36004
diff changeset
   285
28395
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   286
(* init *)
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   287
47061
355317493f34 clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
wenzelm
parents: 46992
diff changeset
   288
fun init naming operations target =
355317493f34 clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
wenzelm
parents: 46992
diff changeset
   289
  target |> Data.map
47064
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   290
    (fn [] => [make_lthy (naming, operations, target)]
6cd9d6c93276 basic support for nested local theory targets;
wenzelm
parents: 47061
diff changeset
   291
      | _ => error "Local theory already initialized")
47061
355317493f34 clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
wenzelm
parents: 46992
diff changeset
   292
  |> checkpoint;
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   293
28395
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   294
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   295
(* exit *)
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   296
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   297
val exit = Proof_Context.background_theory Theory.checkpoint o operation #exit;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   298
val exit_global = Proof_Context.theory_of o exit;
28395
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   299
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   300
fun exit_result f (x, lthy) =
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   301
  let
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   302
    val ctxt = exit lthy;
33281
223ef9bc399a let naming transform binding beforehand -- covering only the "conceal" flag for now;
wenzelm
parents: 33276
diff changeset
   303
    val phi = standard_morphism lthy ctxt;
28395
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   304
  in (f phi x, ctxt) end;
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   305
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   306
fun exit_result_global f (x, lthy) =
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   307
  let
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   308
    val thy = exit_global lthy;
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38757
diff changeset
   309
    val thy_ctxt = Proof_Context.init_global thy;
33281
223ef9bc399a let naming transform binding beforehand -- covering only the "conceal" flag for now;
wenzelm
parents: 33276
diff changeset
   310
    val phi = standard_morphism lthy thy_ctxt;
28395
984fcc8feb24 added exit_global, exit_result, exit_result_global;
wenzelm
parents: 28379
diff changeset
   311
  in (f phi x, thy) end;
20910
0e129a1df180 added exit;
wenzelm
parents: 20888
diff changeset
   312
18781
ea3b5e22eab5 added constant definition;
wenzelm
parents: 18767
diff changeset
   313
end;