author | blanchet |
Fri, 01 Aug 2014 14:43:57 +0200 | |
changeset 57743 | 0af2d5dfb0ac |
parent 57194 | d110b0d1bc12 |
child 57924 | a3360da1d2f0 |
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 |
|
18744 | 10 |
signature LOCAL_THEORY = |
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 | 14 |
val restore: local_theory -> local_theory |
47064 | 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 |
51735 | 17 |
val assert_nonbrittle: local_theory -> local_theory |
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
18 |
val open_target: Name_Space.naming -> operations -> (local_theory -> local_theory) -> |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
19 |
local_theory -> local_theory |
47064 | 20 |
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
|
21 |
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
|
22 |
val naming_of: local_theory -> Name_Space.naming |
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
23 |
val full_name: local_theory -> binding -> string |
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
24 |
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
|
25 |
val conceal: local_theory -> local_theory |
33724 | 26 |
val new_group: local_theory -> local_theory |
27 |
val reset_group: local_theory -> local_theory |
|
33276 | 28 |
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
|
29 |
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
|
30 |
val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a |
20960 | 31 |
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory |
32 |
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
|
33 |
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
|
34 |
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
|
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 |
56723
a8f71445c265
subscription as target-specific implementation device
haftmann
parents:
56141
diff
changeset
|
52 |
val subscription: string * morphism -> (morphism * bool) option -> morphism -> |
a8f71445c265
subscription as target-specific implementation device
haftmann
parents:
56141
diff
changeset
|
53 |
local_theory -> local_theory |
38308
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents:
37949
diff
changeset
|
54 |
val pretty: local_theory -> Pretty.T list |
36451 | 55 |
val set_defsort: sort -> local_theory -> local_theory |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
33764
diff
changeset
|
56 |
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory |
24949 | 57 |
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
|
58 |
val class_alias: binding -> class -> local_theory -> local_theory |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
59 |
val type_alias: binding -> string -> local_theory -> local_theory |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
60 |
val const_alias: binding -> string -> local_theory -> local_theory |
52140 | 61 |
val activate: string * morphism -> (morphism * bool) option -> morphism -> |
62 |
local_theory -> local_theory |
|
52153
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
haftmann
parents:
52140
diff
changeset
|
63 |
val activate_nonbrittle: string * morphism -> (morphism * bool) option -> morphism -> |
52140 | 64 |
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
|
65 |
val init: Name_Space.naming -> operations -> Proof.context -> local_theory |
21292 | 66 |
val exit: local_theory -> Proof.context |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
67 |
val exit_global: local_theory -> theory |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
68 |
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
|
69 |
val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory |
18744 | 70 |
end; |
71 |
||
33671 | 72 |
structure Local_Theory: LOCAL_THEORY = |
18744 | 73 |
struct |
74 |
||
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
75 |
(** local theory data **) |
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
76 |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
77 |
(* type lthy *) |
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
78 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
79 |
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
|
80 |
{define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
81 |
(term * (string * thm)) * local_theory, |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
82 |
notes: string -> |
28084
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents:
28083
diff
changeset
|
83 |
(Attrib.binding * (thm list * Attrib.src list) list) list -> |
28083
103d9282a946
explicit type Name.binding for higher-specification elements;
wenzelm
parents:
28017
diff
changeset
|
84 |
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
|
85 |
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
|
86 |
(term * term) * local_theory, |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
87 |
declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory, |
56723
a8f71445c265
subscription as target-specific implementation device
haftmann
parents:
56141
diff
changeset
|
88 |
subscription: string * morphism -> (morphism * bool) option -> morphism -> |
a8f71445c265
subscription as target-specific implementation device
haftmann
parents:
56141
diff
changeset
|
89 |
local_theory -> local_theory, |
38308
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
haftmann
parents:
37949
diff
changeset
|
90 |
pretty: local_theory -> Pretty.T list, |
21292 | 91 |
exit: local_theory -> Proof.context}; |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
92 |
|
47064 | 93 |
type lthy = |
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
94 |
{naming: Name_Space.naming, |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
95 |
operations: operations, |
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
96 |
after_close: local_theory -> local_theory, |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
97 |
brittle: bool, |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
98 |
target: Proof.context}; |
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
99 |
|
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
100 |
fun make_lthy (naming, operations, after_close, brittle, target) : lthy = |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
101 |
{naming = naming, operations = operations, after_close = after_close, brittle = brittle, target = target}; |
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
102 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
103 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
104 |
(* context data *) |
18744 | 105 |
|
33519 | 106 |
structure Data = Proof_Data |
18744 | 107 |
( |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
108 |
type T = lthy list; |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
109 |
fun init _ = []; |
18744 | 110 |
); |
111 |
||
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
112 |
fun assert lthy = |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
113 |
if null (Data.get lthy) then error "Missing local theory context" else lthy; |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
114 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
115 |
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
|
116 |
val top_of = hd o Data.get o assert; |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
117 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
118 |
fun map_bottom f = |
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
119 |
assert #> |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
120 |
Data.map (fn {naming, operations, after_close, brittle, target} :: parents => |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
121 |
make_lthy (f (naming, operations, after_close, brittle, target)) :: parents); |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
122 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
123 |
fun restore lthy = #target (top_of lthy) |> Data.put (Data.get lthy); |
47273 | 124 |
|
47064 | 125 |
|
126 |
(* nested structure *) |
|
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
127 |
|
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
128 |
val level = length o Data.get; (*1: main target at bottom, >= 2: nested context*) |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
129 |
|
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
130 |
fun assert_bottom b lthy = |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
131 |
let |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
132 |
val _ = assert lthy; |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
133 |
val b' = level lthy <= 1; |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
134 |
in |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
135 |
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
|
136 |
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
|
137 |
else lthy |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
138 |
end; |
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
139 |
|
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
140 |
fun open_target naming operations after_close target = |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
141 |
assert target |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
142 |
|> Data.map (cons (make_lthy (naming, operations, after_close, true, target))); |
47064 | 143 |
|
144 |
fun close_target lthy = |
|
50739
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
145 |
let |
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents:
49062
diff
changeset
|
146 |
val _ = assert_bottom false lthy; |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
147 |
val ({after_close, ...} :: rest) = Data.get lthy; |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
148 |
in lthy |> Data.put rest |> restore |> after_close end; |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
149 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
150 |
fun map_contexts f lthy = |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
151 |
let val n = level lthy in |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
152 |
lthy |> (Data.map o map_index) (fn (i, {naming, operations, after_close, brittle, target}) => |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
153 |
make_lthy (naming, operations, after_close, brittle, |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
154 |
target |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
155 |
|> Context_Position.set_visible false |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
156 |
|> f (n - i - 1) |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
157 |
|> Context_Position.restore_visible target)) |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
158 |
|> f n |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
159 |
end; |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
160 |
|
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
161 |
|
51735 | 162 |
(* brittle context -- implicit for nested structures *) |
163 |
||
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
164 |
fun mark_brittle lthy = |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
165 |
if level lthy = 1 |
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
166 |
then map_bottom (fn (naming, operations, after_close, brittle, target) => |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
167 |
(naming, operations, after_close, true, target)) lthy |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
168 |
else lthy; |
51735 | 169 |
|
170 |
fun assert_nonbrittle lthy = |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
171 |
if #brittle (top_of lthy) |
51735 | 172 |
then error "Brittle local theory context" |
173 |
else lthy; |
|
174 |
||
175 |
||
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
176 |
(* naming *) |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
177 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
178 |
val naming_of = #naming o top_of; |
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
179 |
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
|
180 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
181 |
fun map_naming f = |
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
182 |
map_bottom (fn (naming, operations, after_close, brittle, target) => |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
183 |
(f naming, operations, after_close, brittle, target)); |
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
184 |
|
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
185 |
val conceal = map_naming Name_Space.conceal; |
33724 | 186 |
val new_group = map_naming Name_Space.new_group; |
187 |
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
|
188 |
|
33276 | 189 |
val restore_naming = map_naming o K o naming_of; |
190 |
||
26131
91024979b9eb
maintain group in lthy data, implicit use in operations;
wenzelm
parents:
25984
diff
changeset
|
191 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
192 |
(* standard morphisms *) |
18767 | 193 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
194 |
fun standard_morphism lthy ctxt = |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
195 |
Proof_Context.norm_export_morphism lthy ctxt $> |
54740 | 196 |
Morphism.binding_morphism "Local_Theory.standard_binding" |
197 |
(Name_Space.transform_binding (naming_of lthy)); |
|
18767 | 198 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
199 |
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
|
200 |
Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); |
18767 | 201 |
|
202 |
||
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
203 |
(* background theory *) |
19661 | 204 |
|
20960 | 205 |
fun raw_theory_result f lthy = |
206 |
let |
|
42360 | 207 |
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
|
208 |
val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; |
20960 | 209 |
in (res, lthy') end; |
210 |
||
211 |
fun raw_theory f = #2 o raw_theory_result (f #> pair ()); |
|
212 |
||
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
|
213 |
fun background_theory_result f lthy = |
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
214 |
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
|
215 |
thy |
33166
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
wenzelm
parents:
33157
diff
changeset
|
216 |
|> 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
|
217 |
|> f |
52788 | 218 |
||> Sign.restore_naming thy); |
19661 | 219 |
|
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
|
220 |
fun background_theory f = #2 o background_theory_result (f #> pair ()); |
19661 | 221 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
222 |
|
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
223 |
(* target contexts *) |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
224 |
|
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
225 |
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
|
226 |
|
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
227 |
fun target f lthy = |
18744 | 228 |
let |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
229 |
val ctxt = target_of lthy; |
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
230 |
val ctxt' = ctxt |
33383 | 231 |
|> Context_Position.set_visible false |
28406
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
wenzelm
parents:
28395
diff
changeset
|
232 |
|> f |
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
233 |
|> Context_Position.restore_visible ctxt; |
42360 | 234 |
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
|
235 |
in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; |
18876 | 236 |
|
47245
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
wenzelm
parents:
47081
diff
changeset
|
237 |
fun target_morphism lthy = standard_morphism lthy (target_of lthy); |
18767 | 238 |
|
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
|
239 |
fun propagate_ml_env (context as Context.Proof lthy) = |
47064 | 240 |
let val inherit = ML_Env.inherit context in |
241 |
lthy |
|
242 |
|> 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
|
243 |
|> map_contexts (K (Context.proof_map inherit)) |
47064 | 244 |
|> Context.Proof |
245 |
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
|
246 |
| 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
|
247 |
|
18767 | 248 |
|
47064 | 249 |
|
250 |
(** operations **) |
|
251 |
||
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
252 |
val operations_of = #operations o top_of; |
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
253 |
|
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
254 |
|
47064 | 255 |
(* primitives *) |
18781 | 256 |
|
47066
8a6124d09ff5
basic support for nested contexts including bundles;
wenzelm
parents:
47064
diff
changeset
|
257 |
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
|
258 |
fun operation2 f x y = operation (fn ops => f ops x y); |
18781 | 259 |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
260 |
val pretty = operation #pretty; |
52788 | 261 |
val abbrev = operation2 #abbrev; |
262 |
val define = operation2 #define false; |
|
263 |
val define_internal = operation2 #define true; |
|
264 |
val notes_kind = operation2 #notes; |
|
265 |
val declaration = operation2 #declaration; |
|
56809
b60009672a65
prevent subscription in nested contexts explicitly -- at foundational and user level
haftmann
parents:
56723
diff
changeset
|
266 |
fun subscription dep_morph mixin export = |
b60009672a65
prevent subscription in nested contexts explicitly -- at foundational and user level
haftmann
parents:
56723
diff
changeset
|
267 |
assert_bottom true #> operation (fn ops => #subscription ops dep_morph mixin export); |
18781 | 268 |
|
36451 | 269 |
|
47064 | 270 |
(* basic derived operations *) |
36451 | 271 |
|
33670
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33519
diff
changeset
|
272 |
val notes = notes_kind ""; |
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
wenzelm
parents:
33519
diff
changeset
|
273 |
fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; |
18781 | 274 |
|
36451 | 275 |
fun set_defsort S = |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
276 |
declaration {syntax = true, pervasive = false} |
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
277 |
(K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); |
36451 | 278 |
|
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
279 |
|
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
280 |
(* notation *) |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
281 |
|
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
33764
diff
changeset
|
282 |
fun type_notation 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
|
283 |
let |
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
|
284 |
val args = map (apfst (Logic.type_map (Assumption.export_term lthy (target_of lthy)))) raw_args; |
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
|
285 |
in |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
286 |
declaration {syntax = true, pervasive = false} |
47247 | 287 |
(Proof_Context.generic_type_notation add mode args) lthy |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
288 |
end; |
35412
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
wenzelm
parents:
33764
diff
changeset
|
289 |
|
24949 | 290 |
fun notation 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
|
291 |
let |
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
|
292 |
val args = map (apfst (Assumption.export_term lthy (target_of lthy))) raw_args |
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
|
293 |
in |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
294 |
declaration {syntax = true, pervasive = false} |
47247 | 295 |
(Proof_Context.generic_notation add mode args) lthy |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
296 |
end; |
21743 | 297 |
|
21664
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
wenzelm
parents:
21614
diff
changeset
|
298 |
|
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
299 |
(* name space aliases *) |
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
300 |
|
35798
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
wenzelm
parents:
35738
diff
changeset
|
301 |
fun alias global_alias local_alias b name = |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
302 |
declaration {syntax = true, pervasive = false} (fn phi => |
35738
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
303 |
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
|
304 |
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
|
305 |
|
42360 | 306 |
val class_alias = alias Sign.class_alias Proof_Context.class_alias; |
307 |
val type_alias = alias Sign.type_alias Proof_Context.type_alias; |
|
308 |
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
|
309 |
|
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
wenzelm
parents:
35412
diff
changeset
|
310 |
|
52140 | 311 |
(* activation of locale fragments *) |
312 |
||
52153
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
haftmann
parents:
52140
diff
changeset
|
313 |
fun activate_nonbrittle dep_morph mixin export = |
57194
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
haftmann
parents:
56809
diff
changeset
|
314 |
map_bottom (fn (naming, operations, after_close, brittle, target) => |
52140 | 315 |
(naming, operations, after_close, brittle, |
316 |
(Context.proof_map ooo Locale.add_registration) dep_morph mixin export target)); |
|
317 |
||
318 |
fun activate dep_morph mixin export = |
|
52153
f5773a46cf05
more specific structure for registration into theory and dependency onto locale
haftmann
parents:
52140
diff
changeset
|
319 |
mark_brittle #> activate_nonbrittle dep_morph mixin export; |
52140 | 320 |
|
36451 | 321 |
|
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
|
322 |
|
36451 | 323 |
(** init and exit **) |
324 |
||
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
325 |
(* init *) |
20910 | 326 |
|
47061
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
wenzelm
parents:
46992
diff
changeset
|
327 |
fun init naming operations target = |
52118
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
328 |
target |> Data.map |
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
haftmann
parents:
51735
diff
changeset
|
329 |
(fn [] => [make_lthy (naming, operations, I, false, target)] |
52788 | 330 |
| _ => error "Local theory already initialized"); |
20910 | 331 |
|
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
332 |
|
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
333 |
(* exit *) |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
334 |
|
52788 | 335 |
val exit = operation #exit; |
42360 | 336 |
val exit_global = Proof_Context.theory_of o exit; |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
337 |
|
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
338 |
fun exit_result f (x, lthy) = |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
339 |
let |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
340 |
val ctxt = exit lthy; |
33281
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
wenzelm
parents:
33276
diff
changeset
|
341 |
val phi = standard_morphism lthy ctxt; |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
342 |
in (f phi x, ctxt) end; |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
343 |
|
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
344 |
fun exit_result_global f (x, lthy) = |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
345 |
let |
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
346 |
val thy = exit_global lthy; |
42360 | 347 |
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
|
348 |
val phi = standard_morphism lthy thy_ctxt; |
28395
984fcc8feb24
added exit_global, exit_result, exit_result_global;
wenzelm
parents:
28379
diff
changeset
|
349 |
in (f phi x, thy) end; |
20910 | 350 |
|
18781 | 351 |
end; |