| author | paulson | 
| Thu, 07 Jan 2016 17:42:01 +0000 | |
| changeset 62088 | 8463e386eaec | 
| parent 61890 | f6ded81f5690 | 
| child 62514 | aae510e9a698 | 
| 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;  | 
| 
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
 | 
13  | 
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
 | 
14  | 
|
| 18744 | 15  | 
signature LOCAL_THEORY =  | 
16  | 
sig  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
17  | 
type operations  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
18  | 
val assert: local_theory -> local_theory  | 
| 47273 | 19  | 
val restore: local_theory -> local_theory  | 
| 47064 | 20  | 
val level: Proof.context -> int  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
21  | 
val assert_bottom: bool -> local_theory -> local_theory  | 
| 
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
 | 
22  | 
val mark_brittle: local_theory -> local_theory  | 
| 51735 | 23  | 
val assert_nonbrittle: local_theory -> local_theory  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
24  | 
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
 | 
25  | 
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
 | 
26  | 
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
 | 
27  | 
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
 | 
28  | 
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
 | 
29  | 
val full_name: local_theory -> binding -> string  | 
| 33724 | 30  | 
val new_group: local_theory -> local_theory  | 
31  | 
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
 | 
32  | 
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
 | 
33  | 
val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a  | 
| 20960 | 34  | 
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory  | 
35  | 
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
 | 
36  | 
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
 | 
37  | 
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
 | 
38  | 
val target_of: local_theory -> Proof.context  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
39  | 
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
 | 
40  | 
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
 | 
41  | 
val propagate_ml_env: generic_theory -> generic_theory  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
42  | 
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
 | 
43  | 
val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
44  | 
(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
 | 
45  | 
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
 | 
46  | 
(term * (string * thm)) * local_theory  | 
| 
33670
 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
47  | 
val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57941 
diff
changeset
 | 
48  | 
val notes: (Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
33670
 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
49  | 
local_theory -> (string * thm list) list * local_theory  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57941 
diff
changeset
 | 
50  | 
val notes_kind: string -> (Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
51  | 
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
 | 
52  | 
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
 | 
53  | 
(term * term) * local_theory  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
54  | 
  val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory
 | 
| 
61890
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
55  | 
val theory_registration: string * morphism -> (morphism * bool) option -> morphism ->  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
56  | 
local_theory -> local_theory  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
57  | 
val locale_dependency: string * morphism -> (morphism * bool) option -> morphism ->  | 
| 
56723
 
a8f71445c265
subscription as target-specific implementation device
 
haftmann 
parents: 
56141 
diff
changeset
 | 
58  | 
local_theory -> local_theory  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
37949 
diff
changeset
 | 
59  | 
val pretty: local_theory -> Pretty.T list  | 
| 
57929
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
60  | 
val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory  | 
| 36451 | 61  | 
val set_defsort: sort -> local_theory -> local_theory  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33764 
diff
changeset
 | 
62  | 
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory  | 
| 24949 | 63  | 
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory  | 
| 57938 | 64  | 
val generic_alias:  | 
65  | 
    (('a Name_Space.table -> 'a Name_Space.table) -> Context.generic -> Context.generic) ->
 | 
|
66  | 
binding -> string -> local_theory -> local_theory  | 
|
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
67  | 
val class_alias: binding -> class -> local_theory -> local_theory  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
68  | 
val type_alias: binding -> string -> local_theory -> local_theory  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
69  | 
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
 | 
70  | 
val init: Name_Space.naming -> operations -> Proof.context -> local_theory  | 
| 21292 | 71  | 
val exit: local_theory -> Proof.context  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
72  | 
val exit_global: local_theory -> theory  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
73  | 
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
 | 
74  | 
val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory  | 
| 59890 | 75  | 
val init_target: Name_Space.naming -> operations -> (local_theory -> local_theory) ->  | 
76  | 
local_theory -> Binding.scope * local_theory  | 
|
77  | 
val open_target: local_theory -> Binding.scope * local_theory  | 
|
78  | 
val close_target: local_theory -> local_theory  | 
|
| 18744 | 79  | 
end;  | 
80  | 
||
| 33671 | 81  | 
structure Local_Theory: LOCAL_THEORY =  | 
| 18744 | 82  | 
struct  | 
83  | 
||
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
84  | 
(** local theory data **)  | 
| 
19059
 
b4ca3100e818
init/exit no longer change the theory (no naming);
 
wenzelm 
parents: 
19032 
diff
changeset
 | 
85  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
86  | 
(* type lthy *)  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
87  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
88  | 
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
 | 
89  | 
 {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
 | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
90  | 
(term * (string * thm)) * local_theory,  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
91  | 
notes: string ->  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57941 
diff
changeset
 | 
92  | 
(Attrib.binding * (thm list * Token.src list) list) list ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
93  | 
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
 | 
94  | 
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
 | 
95  | 
(term * term) * local_theory,  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
96  | 
  declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory,
 | 
| 
61890
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
97  | 
theory_registration: string * morphism -> (morphism * bool) option -> morphism ->  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
98  | 
local_theory -> local_theory,  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
99  | 
locale_dependency: string * morphism -> (morphism * bool) option -> morphism ->  | 
| 
56723
 
a8f71445c265
subscription as target-specific implementation device
 
haftmann 
parents: 
56141 
diff
changeset
 | 
100  | 
local_theory -> local_theory,  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
37949 
diff
changeset
 | 
101  | 
pretty: local_theory -> Pretty.T list,  | 
| 21292 | 102  | 
exit: local_theory -> Proof.context};  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
103  | 
|
| 47064 | 104  | 
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
 | 
105  | 
 {background_naming: Name_Space.naming,
 | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
106  | 
operations: operations,  | 
| 
50739
 
5165d7e6d3b9
more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
 
wenzelm 
parents: 
49062 
diff
changeset
 | 
107  | 
after_close: local_theory -> local_theory,  | 
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
108  | 
brittle: bool,  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
109  | 
target: Proof.context};  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
110  | 
|
| 
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
 | 
111  | 
fun make_lthy (background_naming, operations, after_close, brittle, target) : lthy =  | 
| 
 
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
 | 
112  | 
  {background_naming = background_naming, operations = operations,
 | 
| 
 
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
 | 
113  | 
after_close = after_close, brittle = brittle, target = target};  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
114  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
115  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
116  | 
(* context data *)  | 
| 18744 | 117  | 
|
| 33519 | 118  | 
structure Data = Proof_Data  | 
| 18744 | 119  | 
(  | 
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
120  | 
type T = lthy list;  | 
| 
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
121  | 
fun init _ = [];  | 
| 18744 | 122  | 
);  | 
123  | 
||
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
124  | 
fun assert lthy =  | 
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
125  | 
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
 | 
126  | 
|
| 
57194
 
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
 
haftmann 
parents: 
56809 
diff
changeset
 | 
127  | 
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
 | 
128  | 
val top_of = hd o Data.get o assert;  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
129  | 
|
| 
57925
 
2bd92d3f92d7
clarified terminology: first is top (amending d110b0d1bc12);
 
wenzelm 
parents: 
57924 
diff
changeset
 | 
130  | 
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
 | 
131  | 
assert #>  | 
| 
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
 | 
132  | 
  Data.map (fn {background_naming, operations, after_close, brittle, target} :: parents =>
 | 
| 
 
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
 | 
133  | 
make_lthy (f (background_naming, operations, after_close, brittle, target)) :: parents);  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
134  | 
|
| 
57194
 
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
 
haftmann 
parents: 
56809 
diff
changeset
 | 
135  | 
fun restore lthy = #target (top_of lthy) |> Data.put (Data.get lthy);  | 
| 47273 | 136  | 
|
| 47064 | 137  | 
|
138  | 
(* nested structure *)  | 
|
| 
33166
 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 
wenzelm 
parents: 
33157 
diff
changeset
 | 
139  | 
|
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
140  | 
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
 | 
141  | 
|
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
142  | 
fun assert_bottom b lthy =  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
143  | 
let  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
144  | 
val _ = assert lthy;  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
145  | 
val b' = level lthy <= 1;  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
146  | 
in  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
147  | 
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
 | 
148  | 
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
 | 
149  | 
else lthy  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
150  | 
end;  | 
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
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_contexts f lthy =  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
153  | 
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
 | 
154  | 
lthy |> (Data.map o map_index)  | 
| 
 
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
 | 
155  | 
      (fn (i, {background_naming, operations, after_close, brittle, 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
 | 
156  | 
make_lthy (background_naming, operations, after_close, brittle,  | 
| 
 
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
 | 
157  | 
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
 | 
158  | 
|> 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
 | 
159  | 
|> 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
 | 
160  | 
|> 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
 | 
161  | 
|> f n  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
162  | 
end;  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
163  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
164  | 
|
| 51735 | 165  | 
(* brittle context -- implicit for nested structures *)  | 
166  | 
||
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
167  | 
fun mark_brittle lthy =  | 
| 57924 | 168  | 
if level lthy = 1 then  | 
| 
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
 | 
169  | 
map_top (fn (background_naming, operations, after_close, _, 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
 | 
170  | 
(background_naming, operations, after_close, true, target)) lthy  | 
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
171  | 
else lthy;  | 
| 51735 | 172  | 
|
173  | 
fun assert_nonbrittle lthy =  | 
|
| 57924 | 174  | 
if #brittle (top_of lthy) then error "Brittle local theory context"  | 
| 51735 | 175  | 
else lthy;  | 
176  | 
||
177  | 
||
| 
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
 | 
178  | 
(* naming for background theory *)  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
179  | 
|
| 
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
 | 
180  | 
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
 | 
181  | 
|
| 
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
 | 
182  | 
fun map_background_naming f =  | 
| 
 
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  | 
map_top (fn (background_naming, operations, after_close, brittle, 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
 | 
184  | 
(f background_naming, operations, after_close, brittle, target));  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
185  | 
|
| 
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
 | 
186  | 
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
 | 
187  | 
|
| 
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
 | 
188  | 
val full_name = Name_Space.full_name o background_naming_of;  | 
| 
 
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  | 
|
| 
 
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
 | 
190  | 
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
 | 
191  | 
val reset_group = map_background_naming Name_Space.reset_group;  | 
| 33276 | 192  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
193  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
194  | 
(* standard morphisms *)  | 
| 18767 | 195  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
196  | 
fun standard_morphism lthy ctxt =  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
197  | 
Proof_Context.norm_export_morphism lthy ctxt $>  | 
| 54740 | 198  | 
Morphism.binding_morphism "Local_Theory.standard_binding"  | 
| 
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
 | 
199  | 
(Name_Space.transform_binding (Proof_Context.naming_of lthy));  | 
| 18767 | 200  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
201  | 
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
 | 
202  | 
Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x);  | 
| 18767 | 203  | 
|
204  | 
||
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
205  | 
(* background theory *)  | 
| 19661 | 206  | 
|
| 20960 | 207  | 
fun raw_theory_result f lthy =  | 
208  | 
let  | 
|
| 42360 | 209  | 
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
 | 
210  | 
val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy;  | 
| 20960 | 211  | 
in (res, lthy') end;  | 
212  | 
||
213  | 
fun raw_theory f = #2 o raw_theory_result (f #> pair ());  | 
|
214  | 
||
| 
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
 | 
215  | 
fun background_theory_result f lthy =  | 
| 59887 | 216  | 
let  | 
217  | 
val naming =  | 
|
218  | 
background_naming_of lthy  | 
|
219  | 
|> Name_Space.transform_naming (Proof_Context.naming_of lthy);  | 
|
220  | 
in  | 
|
221  | 
lthy |> raw_theory_result (fn thy =>  | 
|
222  | 
thy  | 
|
223  | 
|> Sign.map_naming (K naming)  | 
|
224  | 
|> f  | 
|
225  | 
||> Sign.restore_naming thy)  | 
|
226  | 
end;  | 
|
| 19661 | 227  | 
|
| 
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
 | 
228  | 
fun background_theory f = #2 o background_theory_result (f #> pair ());  | 
| 19661 | 229  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
230  | 
|
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
231  | 
(* target contexts *)  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
232  | 
|
| 
57194
 
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
 
haftmann 
parents: 
56809 
diff
changeset
 | 
233  | 
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
 | 
234  | 
|
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
235  | 
fun target f lthy =  | 
| 18744 | 236  | 
let  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
237  | 
val ctxt = target_of lthy;  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
238  | 
val ctxt' = ctxt  | 
| 33383 | 239  | 
|> Context_Position.set_visible false  | 
| 
28406
 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 
wenzelm 
parents: 
28395 
diff
changeset
 | 
240  | 
|> f  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
241  | 
|> Context_Position.restore_visible ctxt;  | 
| 42360 | 242  | 
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
 | 
243  | 
in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end;  | 
| 18876 | 244  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
245  | 
fun target_morphism lthy = standard_morphism lthy (target_of lthy);  | 
| 18767 | 246  | 
|
| 
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
 | 
247  | 
fun propagate_ml_env (context as Context.Proof lthy) =  | 
| 47064 | 248  | 
let val inherit = ML_Env.inherit context in  | 
249  | 
lthy  | 
|
250  | 
|> 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
 | 
251  | 
|> map_contexts (K (Context.proof_map inherit))  | 
| 47064 | 252  | 
|> Context.Proof  | 
253  | 
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
 | 
254  | 
| 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
 | 
255  | 
|
| 18767 | 256  | 
|
| 47064 | 257  | 
|
258  | 
(** operations **)  | 
|
259  | 
||
| 
57194
 
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
 
haftmann 
parents: 
56809 
diff
changeset
 | 
260  | 
val operations_of = #operations o top_of;  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
261  | 
|
| 57938 | 262  | 
fun operation f lthy = f (operations_of lthy) lthy;  | 
263  | 
fun operation2 f x y = operation (fn ops => f ops x y);  | 
|
264  | 
||
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
265  | 
|
| 47064 | 266  | 
(* primitives *)  | 
| 18781 | 267  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
268  | 
val pretty = operation #pretty;  | 
| 52788 | 269  | 
val abbrev = operation2 #abbrev;  | 
270  | 
val define = operation2 #define false;  | 
|
271  | 
val define_internal = operation2 #define true;  | 
|
272  | 
val notes_kind = operation2 #notes;  | 
|
273  | 
val declaration = operation2 #declaration;  | 
|
| 
61890
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
274  | 
fun theory_registration dep_morph mixin export =  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
275  | 
assert_bottom true #> operation (fn ops => #theory_registration ops dep_morph mixin export);  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
276  | 
fun locale_dependency dep_morph mixin export =  | 
| 
 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 
haftmann 
parents: 
61111 
diff
changeset
 | 
277  | 
assert_bottom true #> operation (fn ops => #locale_dependency ops dep_morph mixin export);  | 
| 18781 | 278  | 
|
| 36451 | 279  | 
|
| 57938 | 280  | 
(* theorems *)  | 
| 36451 | 281  | 
|
| 
33670
 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
282  | 
val notes = notes_kind "";  | 
| 
 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
283  | 
fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single;  | 
| 18781 | 284  | 
|
| 
57929
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
285  | 
fun add_thms_dynamic (binding, f) lthy =  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
286  | 
lthy  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
287  | 
|> background_theory_result (fn thy => thy  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
288  | 
|> 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
 | 
289  | 
|-> (fn name =>  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
290  | 
map_contexts (fn _ => fn ctxt =>  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
291  | 
Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #>  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
292  | 
    declaration {syntax = false, pervasive = false} (fn phi =>
 | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
293  | 
let val binding' = Morphism.binding phi binding in  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
294  | 
Context.mapping  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
295  | 
(Global_Theory.alias_fact binding' name)  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
296  | 
(Proof_Context.fact_alias binding' name)  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
297  | 
end));  | 
| 
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
298  | 
|
| 57938 | 299  | 
|
300  | 
(* default sort *)  | 
|
301  | 
||
| 36451 | 302  | 
fun set_defsort S =  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
303  | 
  declaration {syntax = true, pervasive = false}
 | 
| 
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
304  | 
(K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S)));  | 
| 36451 | 305  | 
|
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
306  | 
|
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
307  | 
(* notation *)  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
308  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33764 
diff
changeset
 | 
309  | 
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
 | 
310  | 
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
 | 
311  | 
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
 | 
312  | 
in  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
313  | 
    declaration {syntax = true, pervasive = false}
 | 
| 47247 | 314  | 
(Proof_Context.generic_type_notation add mode args) lthy  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
315  | 
end;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33764 
diff
changeset
 | 
316  | 
|
| 24949 | 317  | 
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
 | 
318  | 
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
 | 
319  | 
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
 | 
320  | 
in  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
321  | 
    declaration {syntax = true, pervasive = false}
 | 
| 47247 | 322  | 
(Proof_Context.generic_notation add mode args) lthy  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
323  | 
end;  | 
| 21743 | 324  | 
|
| 
21664
 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 
wenzelm 
parents: 
21614 
diff
changeset
 | 
325  | 
|
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
326  | 
(* name space aliases *)  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
327  | 
|
| 57938 | 328  | 
fun generic_alias app b name =  | 
329  | 
  declaration {syntax = false, pervasive = false} (fn phi => fn context =>
 | 
|
330  | 
let  | 
|
331  | 
val naming = Name_Space.naming_of context;  | 
|
332  | 
val b' = Morphism.binding phi b;  | 
|
333  | 
in app (Name_Space.alias_table naming b' name) context end);  | 
|
334  | 
||
335  | 
fun syntax_alias global_alias local_alias b name =  | 
|
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
336  | 
  declaration {syntax = true, pervasive = false} (fn phi =>
 | 
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
337  | 
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
 | 
338  | 
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
 | 
339  | 
|
| 57938 | 340  | 
val class_alias = syntax_alias Sign.class_alias Proof_Context.class_alias;  | 
341  | 
val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias;  | 
|
342  | 
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
 | 
343  | 
|
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
344  | 
|
| 
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
 | 
345  | 
|
| 59890 | 346  | 
(** manage targets **)  | 
| 36451 | 347  | 
|
| 59890 | 348  | 
(* outermost target *)  | 
| 20910 | 349  | 
|
| 
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
 | 
350  | 
fun init background_naming operations target =  | 
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
351  | 
target |> Data.map  | 
| 
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
 | 
352  | 
(fn [] => [make_lthy (background_naming, operations, I, false, target)]  | 
| 52788 | 353  | 
| _ => error "Local theory already initialized");  | 
| 20910 | 354  | 
|
| 52788 | 355  | 
val exit = operation #exit;  | 
| 42360 | 356  | 
val exit_global = Proof_Context.theory_of o exit;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
357  | 
|
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
358  | 
fun exit_result f (x, lthy) =  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
359  | 
let  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
360  | 
val ctxt = exit lthy;  | 
| 
33281
 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 
wenzelm 
parents: 
33276 
diff
changeset
 | 
361  | 
val phi = standard_morphism lthy ctxt;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
362  | 
in (f phi x, ctxt) end;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
363  | 
|
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
364  | 
fun exit_result_global f (x, lthy) =  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
365  | 
let  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
366  | 
val thy = exit_global lthy;  | 
| 42360 | 367  | 
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
 | 
368  | 
val phi = standard_morphism lthy thy_ctxt;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
369  | 
in (f phi x, thy) end;  | 
| 20910 | 370  | 
|
| 59890 | 371  | 
|
372  | 
(* nested targets *)  | 
|
373  | 
||
374  | 
fun init_target background_naming operations after_close lthy =  | 
|
375  | 
let  | 
|
376  | 
val _ = assert lthy;  | 
|
| 
61111
 
2618e7e3b5ea
proper restore naming after close, which is important for packages that used nested targets internally, e.g. BNF datatype;
 
wenzelm 
parents: 
59890 
diff
changeset
 | 
377  | 
val after_close' = Proof_Context.restore_naming lthy #> after_close;  | 
| 59890 | 378  | 
val (scope, target) = Proof_Context.new_scope lthy;  | 
379  | 
val lthy' =  | 
|
380  | 
target  | 
|
| 
61111
 
2618e7e3b5ea
proper restore naming after close, which is important for packages that used nested targets internally, e.g. BNF datatype;
 
wenzelm 
parents: 
59890 
diff
changeset
 | 
381  | 
|> Data.map (cons (make_lthy (background_naming, operations, after_close', true, target)));  | 
| 59890 | 382  | 
in (scope, lthy') end;  | 
383  | 
||
384  | 
fun open_target lthy =  | 
|
385  | 
init_target (background_naming_of lthy) (operations_of lthy) I lthy;  | 
|
386  | 
||
387  | 
fun close_target lthy =  | 
|
388  | 
let  | 
|
389  | 
val _ = assert_bottom false lthy;  | 
|
390  | 
    val ({after_close, ...} :: rest) = Data.get lthy;
 | 
|
391  | 
in lthy |> Data.put rest |> restore |> after_close end;  | 
|
392  | 
||
| 18781 | 393  | 
end;  |