author | wenzelm |
Wed, 02 Oct 2024 11:27:19 +0200 | |
changeset 81100 | 6ae3d0b2b8ad |
parent 79368 | a2d8ecb13d3f |
child 81591 | d570d215e380 |
permissions | -rw-r--r-- |
18744 | 1 |
(* Title: Pure/Isar/local_theory.ML |
2 |
Author: Makarius |
|
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 | 5 |
*) |
6 |
||
18951 | 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 | 9 |
|
57926
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
wenzelm
parents:
57925
diff
changeset
|
10 |
structure Attrib = |
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
wenzelm
parents:
57925
diff
changeset
|
11 |
struct |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57941
diff
changeset
|
12 |
type binding = binding * Token.src list; |
63267 | 13 |
type thms = (thm list * Token.src list) list; |
67652 | 14 |
type fact = binding * thms; |
57926
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
wenzelm
parents:
57925
diff
changeset
|
15 |
end; |
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
wenzelm
parents:
57925
diff
changeset
|
16 |
|
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
17 |
structure Locale = |
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
18 |
struct |
69058 | 19 |
type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism}; |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
20 |
end; |
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
21 |
|
18744 | 22 |
signature LOCAL_THEORY = |
23 |
sig |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
24 |
type operations |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
25 |
val assert: local_theory -> local_theory |
47064 | 26 |
val level: Proof.context -> int |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
27 |
val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory |
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
28 |
val background_naming_of: local_theory -> Name_Space.naming |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
29 |
val map_background_naming: (Name_Space.naming -> Name_Space.naming) -> |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
30 |
local_theory -> local_theory |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
31 |
val restore_background_naming: local_theory -> local_theory -> local_theory |
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
32 |
val full_name: local_theory -> binding -> string |
79368 | 33 |
val full_name_pos: local_theory -> binding -> string * Position.T |
33724 | 34 |
val new_group: local_theory -> local_theory |
35 |
val reset_group: local_theory -> local_theory |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
36 |
val standard_morphism: local_theory -> Proof.context -> morphism |
63270 | 37 |
val standard_morphism_theory: local_theory -> morphism |
78072 | 38 |
val standard_form: local_theory -> Proof.context -> 'a Morphism.entity -> 'a |
20960 | 39 |
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory |
40 |
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
|
41 |
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
|
42 |
val background_theory: (theory -> 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
|
43 |
val target_of: local_theory -> Proof.context |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
44 |
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
|
45 |
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
|
46 |
val propagate_ml_env: generic_theory -> generic_theory |
78035 | 47 |
val touch_ml_env: generic_theory -> generic_theory |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
48 |
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
|
49 |
val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
50 |
(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
|
51 |
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
|
52 |
(term * (string * thm)) * local_theory |
33670
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33519
diff
changeset
|
53 |
val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory |
67652 | 54 |
val notes: Attrib.fact list -> local_theory -> (string * thm list) list * local_theory |
55 |
val notes_kind: string -> Attrib.fact list -> local_theory -> |
|
63267 | 56 |
(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
|
57 |
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
|
58 |
(term * term) * local_theory |
78095 | 59 |
val declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration -> |
78072 | 60 |
local_theory -> local_theory |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
61 |
val theory_registration: Locale.registration -> local_theory -> local_theory |
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
62 |
val locale_dependency: Locale.registration -> local_theory -> local_theory |
38308
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents:
37949
diff
changeset
|
63 |
val pretty: local_theory -> Pretty.T list |
57929
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
64 |
val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory |
36451 | 65 |
val set_defsort: sort -> local_theory -> local_theory |
74333 | 66 |
val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> |
67 |
local_theory -> local_theory |
|
68 |
val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list -> |
|
69 |
local_theory -> local_theory |
|
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
33764
diff
changeset
|
70 |
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory |
74338 | 71 |
val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> |
72 |
local_theory -> local_theory |
|
24949 | 73 |
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory |
74338 | 74 |
val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory |
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
75 |
val type_alias: binding -> string -> local_theory -> local_theory |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
76 |
val const_alias: binding -> string -> local_theory -> local_theory |
72516
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
77 |
val init: {background_naming: Name_Space.naming, setup: theory -> Proof.context, |
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
78 |
conclude: local_theory -> Proof.context} -> operations -> theory -> local_theory |
21292 | 79 |
val exit: local_theory -> Proof.context |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
80 |
val exit_global: local_theory -> theory |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
81 |
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
|
82 |
val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory |
72450 | 83 |
val begin_nested: local_theory -> Binding.scope * local_theory |
84 |
val end_nested: local_theory -> local_theory |
|
85 |
val end_nested_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * local_theory |
|
18744 | 86 |
end; |
87 |
||
69708
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
88 |
signature PRIVATE_LOCAL_THEORY = |
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
89 |
sig |
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
90 |
include LOCAL_THEORY |
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
91 |
val reset: local_theory -> local_theory |
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
92 |
end |
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
93 |
|
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents:
69699
diff
changeset
|
94 |
structure Local_Theory: PRIVATE_LOCAL_THEORY = |
18744 | 95 |
struct |
96 |
||
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
97 |
(** local theory data **) |
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
98 |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
99 |
(* type lthy *) |
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
100 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
101 |
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
|
102 |
{define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
103 |
(term * (string * thm)) * local_theory, |
67652 | 104 |
notes: string -> Attrib.fact list -> 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
|
105 |
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
|
106 |
(term * term) * local_theory, |
78095 | 107 |
declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration_entity -> |
78072 | 108 |
local_theory -> local_theory, |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
109 |
theory_registration: Locale.registration -> local_theory -> local_theory, |
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
110 |
locale_dependency: Locale.registration -> local_theory -> local_theory, |
66335
a849ce33923d
treat exit separate from regular local theory operations
haftmann
parents:
66259
diff
changeset
|
111 |
pretty: local_theory -> Pretty.T list}; |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
112 |
|
47064 | 113 |
type lthy = |
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
114 |
{background_naming: Name_Space.naming, |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
115 |
operations: operations, |
72517 | 116 |
conclude: Proof.context -> Proof.context, |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
117 |
target: Proof.context}; |
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
118 |
|
72953 | 119 |
fun make_lthy (background_naming, operations, conclude, target) : lthy = |
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
120 |
{background_naming = background_naming, operations = operations, |
72953 | 121 |
conclude = conclude, target = target}; |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
122 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
123 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
124 |
(* context data *) |
18744 | 125 |
|
33519 | 126 |
structure Data = Proof_Data |
18744 | 127 |
( |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
128 |
type T = lthy list; |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
129 |
fun init _ = []; |
18744 | 130 |
); |
131 |
||
72153 | 132 |
|
66259 | 133 |
(* nested structure *) |
134 |
||
72153 | 135 |
val level = length o Data.get; (*1: main target at bottom, >= 2: nested target context*) |
66259 | 136 |
|
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
137 |
fun assert lthy = |
66259 | 138 |
if level lthy = 0 then error "Missing local theory context" else lthy; |
139 |
||
140 |
fun assert_bottom lthy = |
|
141 |
let |
|
142 |
val _ = assert lthy; |
|
143 |
in |
|
144 |
if level lthy > 1 then error "Not at bottom of local theory nesting" |
|
145 |
else lthy |
|
146 |
end; |
|
147 |
||
148 |
fun assert_not_bottom lthy = |
|
149 |
let |
|
150 |
val _ = assert lthy; |
|
151 |
in |
|
152 |
if level lthy = 1 then error "Already at bottom of local theory nesting" |
|
153 |
else lthy |
|
154 |
end; |
|
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
155 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
156 |
val bottom_of = List.last o Data.get o assert; |
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
157 |
val top_of = hd o Data.get o assert; |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
158 |
|
57925
2bd92d3f92d7
clarified terminology: first is top (amending d110b0d1bc12);
wenzelm
parents:
57924
diff
changeset
|
159 |
fun map_top f = |
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
160 |
assert #> |
72953 | 161 |
Data.map (fn {background_naming, operations, conclude, target} :: parents => |
162 |
make_lthy (f (background_naming, operations, conclude, target)) :: parents); |
|
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
163 |
|
62514 | 164 |
fun reset lthy = #target (top_of lthy) |> Data.put (Data.get lthy); |
47273 | 165 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
166 |
fun map_contexts f lthy = |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
167 |
let val n = level lthy in |
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
168 |
lthy |> (Data.map o map_index) |
72953 | 169 |
(fn (i, {background_naming, operations, conclude, target}) => |
170 |
make_lthy (background_naming, operations, conclude, |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
171 |
target |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
172 |
|> Context_Position.set_visible false |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
173 |
|> f (n - i - 1) |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
174 |
|> Context_Position.restore_visible target)) |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
175 |
|> f n |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
176 |
end; |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
177 |
|
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
178 |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
179 |
(* naming for background theory *) |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
180 |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
181 |
val background_naming_of = #background_naming o top_of; |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
182 |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
183 |
fun map_background_naming f = |
72953 | 184 |
map_top (fn (background_naming, operations, conclude, target) => |
185 |
(f background_naming, operations, conclude, target)); |
|
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
186 |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
187 |
val restore_background_naming = map_background_naming o K o background_naming_of; |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
188 |
|
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
189 |
val full_name = Name_Space.full_name o background_naming_of; |
79368 | 190 |
fun full_name_pos lthy b = (full_name lthy b, Binding.default_pos_of b); |
59880
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
191 |
|
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
192 |
val new_group = map_background_naming Name_Space.new_group; |
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents:
59859
diff
changeset
|
193 |
val reset_group = map_background_naming Name_Space.reset_group; |
33276 | 194 |
|
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
195 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
196 |
(* standard morphisms *) |
18767 | 197 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
198 |
fun standard_morphism lthy ctxt = |
78064
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78035
diff
changeset
|
199 |
Morphism.set_context' lthy |
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78072
diff
changeset
|
200 |
(Proof_Context.export_morphism lthy ctxt $> |
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78072
diff
changeset
|
201 |
Morphism.thm_morphism' "Local_Theory.standard" (Goal.norm_result o Proof_Context.init_global) $> |
78064
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78035
diff
changeset
|
202 |
Morphism.binding_morphism "Local_Theory.standard_binding" |
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
wenzelm
parents:
78035
diff
changeset
|
203 |
(Name_Space.transform_binding (Proof_Context.naming_of lthy))); |
18767 | 204 |
|
63270 | 205 |
fun standard_morphism_theory lthy = |
206 |
standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy)); |
|
207 |
||
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
208 |
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
|
209 |
Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); |
18767 | 210 |
|
211 |
||
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
212 |
(* background theory *) |
19661 | 213 |
|
20960 | 214 |
fun raw_theory_result f lthy = |
215 |
let |
|
42360 | 216 |
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
|
217 |
val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; |
20960 | 218 |
in (res, lthy') end; |
219 |
||
220 |
fun raw_theory f = #2 o raw_theory_result (f #> pair ()); |
|
221 |
||
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
|
222 |
fun background_theory_result f lthy = |
59887 | 223 |
let |
224 |
val naming = |
|
225 |
background_naming_of lthy |
|
226 |
|> Name_Space.transform_naming (Proof_Context.naming_of lthy); |
|
227 |
in |
|
228 |
lthy |> raw_theory_result (fn thy => |
|
229 |
thy |
|
230 |
|> Sign.map_naming (K naming) |
|
231 |
|> f |
|
232 |
||> Sign.restore_naming thy) |
|
233 |
end; |
|
19661 | 234 |
|
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
|
235 |
fun background_theory f = #2 o background_theory_result (f #> pair ()); |
19661 | 236 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
237 |
|
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
238 |
(* target contexts *) |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
239 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
240 |
val target_of = #target o bottom_of; |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
241 |
|
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
242 |
fun target f lthy = |
18744 | 243 |
let |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
244 |
val ctxt = target_of lthy; |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
245 |
val ctxt' = ctxt |
33383 | 246 |
|> Context_Position.set_visible false |
28406
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
wenzelm
parents:
28395
diff
changeset
|
247 |
|> f |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
248 |
|> Context_Position.restore_visible ctxt; |
42360 | 249 |
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
|
250 |
in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; |
18876 | 251 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
252 |
fun target_morphism lthy = standard_morphism lthy (target_of lthy); |
18767 | 253 |
|
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
|
254 |
fun propagate_ml_env (context as Context.Proof lthy) = |
68865
dd44e31ca2c6
support multiple inheritance of ML environments, with canonical merge order as in Context.begin_theory;
wenzelm
parents:
68851
diff
changeset
|
255 |
let val inherit = ML_Env.inherit [context] in |
47064 | 256 |
lthy |
257 |
|> 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
|
258 |
|> map_contexts (K (Context.proof_map inherit)) |
47064 | 259 |
|> Context.Proof |
260 |
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
|
261 |
| 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
|
262 |
|
78035 | 263 |
fun touch_ml_env context = |
264 |
if Context.enabled_tracing () then |
|
265 |
(case context of |
|
266 |
Context.Theory _ => ML_Env.touch context |
|
267 |
| Context.Proof _ => context) |
|
268 |
else context; |
|
269 |
||
18767 | 270 |
|
47064 | 271 |
|
272 |
(** operations **) |
|
273 |
||
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
274 |
val operations_of = #operations o top_of; |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
275 |
|
57938 | 276 |
fun operation f lthy = f (operations_of lthy) lthy; |
73846 | 277 |
fun operation1 f x = operation (fn ops => f ops x); |
57938 | 278 |
fun operation2 f x y = operation (fn ops => f ops x y); |
279 |
||
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
280 |
|
47064 | 281 |
(* primitives *) |
18781 | 282 |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
283 |
val pretty = operation #pretty; |
52788 | 284 |
val abbrev = operation2 #abbrev; |
285 |
val define = operation2 #define false; |
|
286 |
val define_internal = operation2 #define true; |
|
287 |
val notes_kind = operation2 #notes; |
|
78072 | 288 |
fun declaration args = operation2 #declaration args o Morphism.entity; |
73846 | 289 |
val theory_registration = operation1 #theory_registration; |
68851
6c9825c1e26b
clarified signature: explicit type Locale.registration;
wenzelm
parents:
67652
diff
changeset
|
290 |
fun locale_dependency registration = |
73846 | 291 |
assert_bottom #> operation1 #locale_dependency registration; |
18781 | 292 |
|
36451 | 293 |
|
57938 | 294 |
(* theorems *) |
36451 | 295 |
|
33670
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33519
diff
changeset
|
296 |
val notes = notes_kind ""; |
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33519
diff
changeset
|
297 |
fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; |
18781 | 298 |
|
57929
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
299 |
fun add_thms_dynamic (binding, f) lthy = |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
300 |
lthy |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
301 |
|> background_theory_result (fn thy => thy |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
302 |
|> Global_Theory.add_thms_dynamic' (Sign.inherit_naming thy lthy) (binding, f)) |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
303 |
|-> (fn name => |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
304 |
map_contexts (fn _ => fn ctxt => |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
305 |
Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #> |
78095 | 306 |
declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} (fn phi => |
57929
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
307 |
let val binding' = Morphism.binding phi binding in |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
308 |
Context.mapping |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
309 |
(Global_Theory.alias_fact binding' name) |
66246 | 310 |
(Proof_Context.alias_fact binding' name) |
57929
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
311 |
end)); |
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
wenzelm
parents:
57926
diff
changeset
|
312 |
|
57938 | 313 |
|
314 |
(* default sort *) |
|
315 |
||
36451 | 316 |
fun set_defsort S = |
78095 | 317 |
declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
318 |
(K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); |
36451 | 319 |
|
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
320 |
|
74333 | 321 |
(* syntax *) |
322 |
||
323 |
fun gen_syntax prep_type add mode raw_args lthy = |
|
324 |
let |
|
325 |
val args = map (fn (c, T, mx) => (c, prep_type lthy T, mx)) raw_args; |
|
326 |
val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args; |
|
327 |
val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args; |
|
328 |
in |
|
78095 | 329 |
declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} |
74333 | 330 |
(fn _ => Proof_Context.generic_syntax add mode args') lthy |
331 |
end; |
|
332 |
||
333 |
val syntax = gen_syntax (K I); |
|
334 |
val syntax_cmd = gen_syntax Proof_Context.read_typ_syntax; |
|
335 |
||
336 |
||
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
337 |
(* notation *) |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
338 |
|
74338 | 339 |
local |
340 |
||
341 |
fun gen_type_notation prep_type add mode raw_args lthy = |
|
50741
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
wenzelm
parents:
50739
diff
changeset
|
342 |
let |
74338 | 343 |
val prepare = prep_type lthy #> Logic.type_map (Assumption.export_term lthy (target_of lthy)); |
344 |
val args = map (apfst prepare) raw_args; |
|
62844 | 345 |
val args' = map (apsnd Mixfix.reset_pos) args; |
346 |
val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args; |
|
50741
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
wenzelm
parents:
50739
diff
changeset
|
347 |
in |
78095 | 348 |
declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} |
62844 | 349 |
(Proof_Context.generic_type_notation add mode args') lthy |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
350 |
end; |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
33764
diff
changeset
|
351 |
|
74338 | 352 |
fun gen_notation prep_const add mode raw_args lthy = |
50741
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
wenzelm
parents:
50739
diff
changeset
|
353 |
let |
74338 | 354 |
val prepare = prep_const lthy #> Assumption.export_term lthy (target_of lthy); |
355 |
val args = map (apfst prepare) raw_args; |
|
62844 | 356 |
val args' = map (apsnd Mixfix.reset_pos) args; |
357 |
val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args; |
|
50741
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
wenzelm
parents:
50739
diff
changeset
|
358 |
in |
78095 | 359 |
declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} |
62844 | 360 |
(Proof_Context.generic_notation add mode args') lthy |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
361 |
end; |
21743 | 362 |
|
74338 | 363 |
in |
364 |
||
365 |
val type_notation = gen_type_notation (K I); |
|
366 |
val type_notation_cmd = |
|
367 |
gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); |
|
368 |
||
369 |
val notation = gen_notation (K I); |
|
370 |
val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); |
|
371 |
||
372 |
end; |
|
373 |
||
21664
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
wenzelm
parents:
21614
diff
changeset
|
374 |
|
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
375 |
(* name space aliases *) |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
376 |
|
57938 | 377 |
fun syntax_alias global_alias local_alias b name = |
78095 | 378 |
declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (fn phi => |
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
379 |
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
|
380 |
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
|
381 |
|
57938 | 382 |
val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias; |
383 |
val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias; |
|
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
384 |
|
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
385 |
|
56055
8fe7414f00b1
slightly more rubust (and opportunistic) exit for old-fashioned theory_to_proof, which is used by global 'sublocale' with Named_Target.init but without proper exit;
wenzelm
parents:
54740
diff
changeset
|
386 |
|
59890 | 387 |
(** manage targets **) |
36451 | 388 |
|
72153 | 389 |
(* main target *) |
20910 | 390 |
|
72517 | 391 |
fun init_target background_naming conclude operations target = |
72953 | 392 |
Data.map (K [make_lthy (background_naming, operations, conclude, target)]) target |
72516
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
393 |
|
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
394 |
fun init {background_naming, setup, conclude} operations thy = |
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
395 |
thy |
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
396 |
|> Sign.change_begin |
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
397 |
|> setup |
17dc99589a91
unified Local_Theory.init with Generic_Target.init
haftmann
parents:
72505
diff
changeset
|
398 |
|> init_target background_naming (conclude #> target_of #> Sign.change_end_local) operations; |
20910 | 399 |
|
72517 | 400 |
val exit_of = #conclude o bottom_of; |
66335
a849ce33923d
treat exit separate from regular local theory operations
haftmann
parents:
66259
diff
changeset
|
401 |
|
72502 | 402 |
fun exit lthy = exit_of lthy (assert_bottom lthy); |
42360 | 403 |
val exit_global = Proof_Context.theory_of o exit; |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
404 |
|
72153 | 405 |
fun exit_result decl (x, lthy) = |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
406 |
let |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
407 |
val ctxt = exit lthy; |
33281
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
wenzelm
parents:
33276
diff
changeset
|
408 |
val phi = standard_morphism lthy ctxt; |
72153 | 409 |
in (decl phi x, ctxt) end; |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
410 |
|
72153 | 411 |
fun exit_result_global decl (x, lthy) = |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
412 |
let |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
413 |
val thy = exit_global lthy; |
42360 | 414 |
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
|
415 |
val phi = standard_morphism lthy thy_ctxt; |
72153 | 416 |
in (decl phi x, thy) end; |
20910 | 417 |
|
59890 | 418 |
|
419 |
(* nested targets *) |
|
420 |
||
72450 | 421 |
fun begin_nested lthy = |
59890 | 422 |
let |
423 |
val _ = assert lthy; |
|
424 |
val (scope, target) = Proof_Context.new_scope lthy; |
|
72436
d62d84634b06
direct exit to theory when ending nested target on theory target
haftmann
parents:
72433
diff
changeset
|
425 |
val entry = make_lthy (background_naming_of lthy, operations_of lthy, |
72953 | 426 |
Proof_Context.restore_naming lthy, target); |
72153 | 427 |
val lthy' = Data.map (cons entry) target; |
59890 | 428 |
in (scope, lthy') end; |
429 |
||
72450 | 430 |
fun end_nested lthy = |
59890 | 431 |
let |
66259 | 432 |
val _ = assert_not_bottom lthy; |
72517 | 433 |
val ({conclude, ...} :: rest) = Data.get lthy; |
434 |
in lthy |> Data.put rest |> reset |> conclude end; |
|
59890 | 435 |
|
72450 | 436 |
fun end_nested_result decl (x, lthy) = |
437 |
let |
|
438 |
val outer_lthy = end_nested lthy; |
|
439 |
val phi = Proof_Context.export_morphism lthy outer_lthy; |
|
440 |
in (decl phi x, outer_lthy) end; |
|
441 |
||
18781 | 442 |
end; |