| author | paulson <lp15@cam.ac.uk> | 
| Thu, 05 May 2022 16:39:48 +0100 | |
| changeset 75449 | 51e05af57455 | 
| parent 74338 | 534b231ce041 | 
| child 78035 | bd5f6cee8001 | 
| permissions | -rw-r--r-- | 
| 18744 | 1  | 
(* Title: Pure/Isar/local_theory.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
4  | 
Local theory operations, with abstract target context.  | 
| 18744 | 5  | 
*)  | 
6  | 
||
| 18951 | 7  | 
type local_theory = Proof.context;  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
8  | 
type generic_theory = Context.generic;  | 
| 18951 | 9  | 
|
| 
57926
 
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
 
wenzelm 
parents: 
57925 
diff
changeset
 | 
10  | 
structure Attrib =  | 
| 
 
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
 
wenzelm 
parents: 
57925 
diff
changeset
 | 
11  | 
struct  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57941 
diff
changeset
 | 
12  | 
type binding = binding * Token.src list;  | 
| 63267 | 13  | 
type thms = (thm list * Token.src list) list;  | 
| 67652 | 14  | 
type fact = binding * thms;  | 
| 
57926
 
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
 
wenzelm 
parents: 
57925 
diff
changeset
 | 
15  | 
end;  | 
| 
 
59b2572e8e93
load local_theory.ML before attrib.ML, with subtle change of semantics due to canonical Local_Theory.map_contexts instead of private Local_Theory.map_top;
 
wenzelm 
parents: 
57925 
diff
changeset
 | 
16  | 
|
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
17  | 
structure Locale =  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
18  | 
struct  | 
| 69058 | 19  | 
  type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism};
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
20  | 
end;  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
21  | 
|
| 18744 | 22  | 
signature LOCAL_THEORY =  | 
23  | 
sig  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
24  | 
type operations  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
25  | 
val assert: local_theory -> local_theory  | 
| 47064 | 26  | 
val level: Proof.context -> int  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
27  | 
val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory  | 
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
28  | 
val background_naming_of: local_theory -> Name_Space.naming  | 
| 
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
29  | 
val map_background_naming: (Name_Space.naming -> Name_Space.naming) ->  | 
| 
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
30  | 
local_theory -> local_theory  | 
| 
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
31  | 
val restore_background_naming: local_theory -> local_theory -> local_theory  | 
| 
33166
 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 
wenzelm 
parents: 
33157 
diff
changeset
 | 
32  | 
val full_name: local_theory -> binding -> string  | 
| 33724 | 33  | 
val new_group: local_theory -> local_theory  | 
34  | 
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
 | 
35  | 
val standard_morphism: local_theory -> Proof.context -> morphism  | 
| 63270 | 36  | 
val standard_morphism_theory: local_theory -> morphism  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
37  | 
val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a  | 
| 20960 | 38  | 
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory  | 
39  | 
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
 | 
40  | 
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
 | 
41  | 
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
 | 
42  | 
val target_of: local_theory -> Proof.context  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
43  | 
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
 | 
44  | 
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
 | 
45  | 
val propagate_ml_env: generic_theory -> generic_theory  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
46  | 
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
 | 
47  | 
val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
48  | 
(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
 | 
49  | 
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
 | 
50  | 
(term * (string * thm)) * local_theory  | 
| 
33670
 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
51  | 
val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory  | 
| 67652 | 52  | 
val notes: Attrib.fact list -> local_theory -> (string * thm list) list * local_theory  | 
53  | 
val notes_kind: string -> Attrib.fact list -> local_theory ->  | 
|
| 63267 | 54  | 
(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
 | 
55  | 
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
 | 
56  | 
(term * term) * local_theory  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
57  | 
  val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
58  | 
val theory_registration: Locale.registration -> local_theory -> local_theory  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
59  | 
val locale_dependency: Locale.registration -> local_theory -> local_theory  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
37949 
diff
changeset
 | 
60  | 
val pretty: local_theory -> Pretty.T list  | 
| 
57929
 
c5063c033a5a
tuned signature -- proper Local_Theory.add_thms_dynamic;
 
wenzelm 
parents: 
57926 
diff
changeset
 | 
61  | 
val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory  | 
| 36451 | 62  | 
val set_defsort: sort -> local_theory -> local_theory  | 
| 74333 | 63  | 
val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list ->  | 
64  | 
local_theory -> local_theory  | 
|
65  | 
val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list ->  | 
|
66  | 
local_theory -> local_theory  | 
|
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33764 
diff
changeset
 | 
67  | 
val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory  | 
| 74338 | 68  | 
val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list ->  | 
69  | 
local_theory -> local_theory  | 
|
| 24949 | 70  | 
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory  | 
| 74338 | 71  | 
val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory  | 
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
72  | 
val type_alias: binding -> string -> local_theory -> local_theory  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
73  | 
val const_alias: binding -> string -> local_theory -> local_theory  | 
| 
72516
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
74  | 
  val init: {background_naming: Name_Space.naming, setup: theory -> Proof.context,
 | 
| 
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
75  | 
conclude: local_theory -> Proof.context} -> operations -> theory -> local_theory  | 
| 21292 | 76  | 
val exit: local_theory -> Proof.context  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
77  | 
val exit_global: local_theory -> theory  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
78  | 
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
 | 
79  | 
val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory  | 
| 72450 | 80  | 
val begin_nested: local_theory -> Binding.scope * local_theory  | 
81  | 
val end_nested: local_theory -> local_theory  | 
|
82  | 
val end_nested_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * local_theory  | 
|
| 18744 | 83  | 
end;  | 
84  | 
||
| 
69708
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
85  | 
signature PRIVATE_LOCAL_THEORY =  | 
| 
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
86  | 
sig  | 
| 
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
87  | 
include LOCAL_THEORY  | 
| 
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
88  | 
val reset: local_theory -> local_theory  | 
| 
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
89  | 
end  | 
| 
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
90  | 
|
| 
 
1c201e4792cb
Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
 
haftmann 
parents: 
69699 
diff
changeset
 | 
91  | 
structure Local_Theory: PRIVATE_LOCAL_THEORY =  | 
| 18744 | 92  | 
struct  | 
93  | 
||
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
94  | 
(** local theory data **)  | 
| 
19059
 
b4ca3100e818
init/exit no longer change the theory (no naming);
 
wenzelm 
parents: 
19032 
diff
changeset
 | 
95  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
96  | 
(* type lthy *)  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
97  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
98  | 
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
 | 
99  | 
 {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
 | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
100  | 
(term * (string * thm)) * local_theory,  | 
| 67652 | 101  | 
notes: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory,  | 
| 
38308
 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 
haftmann 
parents: 
37949 
diff
changeset
 | 
102  | 
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
 | 
103  | 
(term * term) * local_theory,  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
104  | 
  declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory,
 | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
105  | 
theory_registration: Locale.registration -> local_theory -> local_theory,  | 
| 
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
106  | 
locale_dependency: Locale.registration -> local_theory -> local_theory,  | 
| 
66335
 
a849ce33923d
treat exit separate from regular local theory operations
 
haftmann 
parents: 
66259 
diff
changeset
 | 
107  | 
pretty: local_theory -> Pretty.T list};  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
108  | 
|
| 47064 | 109  | 
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
 | 
110  | 
 {background_naming: Name_Space.naming,
 | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
111  | 
operations: operations,  | 
| 72517 | 112  | 
conclude: Proof.context -> Proof.context,  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
113  | 
target: Proof.context};  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
114  | 
|
| 72953 | 115  | 
fun make_lthy (background_naming, operations, conclude, target) : lthy =  | 
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
116  | 
  {background_naming = background_naming, operations = operations,
 | 
| 72953 | 117  | 
conclude = conclude, target = target};  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
118  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
119  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
120  | 
(* context data *)  | 
| 18744 | 121  | 
|
| 33519 | 122  | 
structure Data = Proof_Data  | 
| 18744 | 123  | 
(  | 
| 
52118
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
124  | 
type T = lthy list;  | 
| 
 
2a976115c7c3
mark local theory as brittle also after interpretation inside locales;
 
haftmann 
parents: 
51735 
diff
changeset
 | 
125  | 
fun init _ = [];  | 
| 18744 | 126  | 
);  | 
127  | 
||
| 72153 | 128  | 
|
| 66259 | 129  | 
(* nested structure *)  | 
130  | 
||
| 72153 | 131  | 
val level = length o Data.get; (*1: main target at bottom, >= 2: nested target context*)  | 
| 66259 | 132  | 
|
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
133  | 
fun assert lthy =  | 
| 66259 | 134  | 
if level lthy = 0 then error "Missing local theory context" else lthy;  | 
135  | 
||
136  | 
fun assert_bottom lthy =  | 
|
137  | 
let  | 
|
138  | 
val _ = assert lthy;  | 
|
139  | 
in  | 
|
140  | 
if level lthy > 1 then error "Not at bottom of local theory nesting"  | 
|
141  | 
else lthy  | 
|
142  | 
end;  | 
|
143  | 
||
144  | 
fun assert_not_bottom lthy =  | 
|
145  | 
let  | 
|
146  | 
val _ = assert lthy;  | 
|
147  | 
in  | 
|
148  | 
if level lthy = 1 then error "Already at bottom of local theory nesting"  | 
|
149  | 
else lthy  | 
|
150  | 
end;  | 
|
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
151  | 
|
| 
57194
 
d110b0d1bc12
tuned terminology: emphasize stack-like nature of nested local theories
 
haftmann 
parents: 
56809 
diff
changeset
 | 
152  | 
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
 | 
153  | 
val top_of = hd o Data.get o assert;  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
154  | 
|
| 
57925
 
2bd92d3f92d7
clarified terminology: first is top (amending d110b0d1bc12);
 
wenzelm 
parents: 
57924 
diff
changeset
 | 
155  | 
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
 | 
156  | 
assert #>  | 
| 72953 | 157  | 
  Data.map (fn {background_naming, operations, conclude, target} :: parents =>
 | 
158  | 
make_lthy (f (background_naming, operations, conclude, target)) :: parents);  | 
|
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
159  | 
|
| 62514 | 160  | 
fun reset lthy = #target (top_of lthy) |> Data.put (Data.get lthy);  | 
| 47273 | 161  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
162  | 
fun map_contexts f lthy =  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
163  | 
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
 | 
164  | 
lthy |> (Data.map o map_index)  | 
| 72953 | 165  | 
      (fn (i, {background_naming, operations, conclude, target}) =>
 | 
166  | 
make_lthy (background_naming, operations, conclude,  | 
|
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
167  | 
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
 | 
168  | 
|> 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
 | 
169  | 
|> 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
 | 
170  | 
|> 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
 | 
171  | 
|> f n  | 
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
172  | 
end;  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
173  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
174  | 
|
| 
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
 | 
175  | 
(* naming for background theory *)  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
176  | 
|
| 
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
 | 
177  | 
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
 | 
178  | 
|
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
179  | 
fun map_background_naming f =  | 
| 72953 | 180  | 
map_top (fn (background_naming, operations, conclude, target) =>  | 
181  | 
(f background_naming, operations, conclude, target));  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
182  | 
|
| 
59880
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
183  | 
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
 | 
184  | 
|
| 
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
 | 
185  | 
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
 | 
186  | 
|
| 
 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 
wenzelm 
parents: 
59859 
diff
changeset
 | 
187  | 
val 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
 | 
188  | 
val reset_group = map_background_naming Name_Space.reset_group;  | 
| 33276 | 189  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
190  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
191  | 
(* standard morphisms *)  | 
| 18767 | 192  | 
|
| 
47245
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
193  | 
fun standard_morphism lthy ctxt =  | 
| 
 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 
wenzelm 
parents: 
47081 
diff
changeset
 | 
194  | 
Proof_Context.norm_export_morphism lthy ctxt $>  | 
| 54740 | 195  | 
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
 | 
196  | 
(Name_Space.transform_binding (Proof_Context.naming_of lthy));  | 
| 18767 | 197  | 
|
| 63270 | 198  | 
fun standard_morphism_theory lthy =  | 
199  | 
standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy));  | 
|
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) =  | 
| 
68865
 
dd44e31ca2c6
support multiple inheritance of ML environments, with canonical merge order as in Context.begin_theory;
 
wenzelm 
parents: 
68851 
diff
changeset
 | 
248  | 
let val inherit = ML_Env.inherit [context] in  | 
| 47064 | 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;  | 
| 73846 | 263  | 
fun operation1 f x = operation (fn ops => f ops x);  | 
| 57938 | 264  | 
fun operation2 f x y = operation (fn ops => f ops x y);  | 
265  | 
||
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47064 
diff
changeset
 | 
266  | 
|
| 47064 | 267  | 
(* primitives *)  | 
| 18781 | 268  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
269  | 
val pretty = operation #pretty;  | 
| 52788 | 270  | 
val abbrev = operation2 #abbrev;  | 
271  | 
val define = operation2 #define false;  | 
|
272  | 
val define_internal = operation2 #define true;  | 
|
273  | 
val notes_kind = operation2 #notes;  | 
|
274  | 
val declaration = operation2 #declaration;  | 
|
| 73846 | 275  | 
val theory_registration = operation1 #theory_registration;  | 
| 
68851
 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 
wenzelm 
parents: 
67652 
diff
changeset
 | 
276  | 
fun locale_dependency registration =  | 
| 73846 | 277  | 
assert_bottom #> operation1 #locale_dependency registration;  | 
| 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)  | 
| 66246 | 296  | 
(Proof_Context.alias_fact binding' name)  | 
| 
57929
 
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  | 
|
| 74333 | 307  | 
(* syntax *)  | 
308  | 
||
309  | 
fun gen_syntax prep_type add mode raw_args lthy =  | 
|
310  | 
let  | 
|
311  | 
val args = map (fn (c, T, mx) => (c, prep_type lthy T, mx)) raw_args;  | 
|
312  | 
val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args;  | 
|
313  | 
val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args;  | 
|
314  | 
in  | 
|
315  | 
    declaration {syntax = true, pervasive = false}
 | 
|
316  | 
(fn _ => Proof_Context.generic_syntax add mode args') lthy  | 
|
317  | 
end;  | 
|
318  | 
||
319  | 
val syntax = gen_syntax (K I);  | 
|
320  | 
val syntax_cmd = gen_syntax Proof_Context.read_typ_syntax;  | 
|
321  | 
||
322  | 
||
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
323  | 
(* notation *)  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
324  | 
|
| 74338 | 325  | 
local  | 
326  | 
||
327  | 
fun gen_type_notation prep_type add mode raw_args lthy =  | 
|
| 
50741
 
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
 
wenzelm 
parents: 
50739 
diff
changeset
 | 
328  | 
let  | 
| 74338 | 329  | 
val prepare = prep_type lthy #> Logic.type_map (Assumption.export_term lthy (target_of lthy));  | 
330  | 
val args = map (apfst prepare) raw_args;  | 
|
| 62844 | 331  | 
val args' = map (apsnd Mixfix.reset_pos) args;  | 
332  | 
val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args;  | 
|
| 
50741
 
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
 
wenzelm 
parents: 
50739 
diff
changeset
 | 
333  | 
in  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
334  | 
    declaration {syntax = true, pervasive = false}
 | 
| 62844 | 335  | 
(Proof_Context.generic_type_notation add mode args') lthy  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
336  | 
end;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
33764 
diff
changeset
 | 
337  | 
|
| 74338 | 338  | 
fun gen_notation prep_const add mode raw_args lthy =  | 
| 
50741
 
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
 
wenzelm 
parents: 
50739 
diff
changeset
 | 
339  | 
let  | 
| 74338 | 340  | 
val prepare = prep_const lthy #> Assumption.export_term lthy (target_of lthy);  | 
341  | 
val args = map (apfst prepare) raw_args;  | 
|
| 62844 | 342  | 
val args' = map (apsnd Mixfix.reset_pos) args;  | 
343  | 
val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args;  | 
|
| 
50741
 
20e6e1a92e54
less aggressive Assumption.export_term (similar to Generic_Target.abbrev): merely expand local defs and thus allow notation for local fixes;
 
wenzelm 
parents: 
50739 
diff
changeset
 | 
344  | 
in  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
345  | 
    declaration {syntax = true, pervasive = false}
 | 
| 62844 | 346  | 
(Proof_Context.generic_notation add mode args') lthy  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
347  | 
end;  | 
| 21743 | 348  | 
|
| 74338 | 349  | 
in  | 
350  | 
||
351  | 
val type_notation = gen_type_notation (K I);  | 
|
352  | 
val type_notation_cmd =  | 
|
353  | 
  gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false});
 | 
|
354  | 
||
355  | 
val notation = gen_notation (K I);  | 
|
356  | 
val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false});
 | 
|
357  | 
||
358  | 
end;  | 
|
359  | 
||
| 
21664
 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 
wenzelm 
parents: 
21614 
diff
changeset
 | 
360  | 
|
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
361  | 
(* name space aliases *)  | 
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
362  | 
|
| 57938 | 363  | 
fun syntax_alias global_alias local_alias b name =  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
364  | 
  declaration {syntax = true, pervasive = false} (fn phi =>
 | 
| 
35738
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
365  | 
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
 | 
366  | 
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
 | 
367  | 
|
| 57938 | 368  | 
val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias;  | 
369  | 
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
 | 
370  | 
|
| 
 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
371  | 
|
| 
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
 | 
372  | 
|
| 59890 | 373  | 
(** manage targets **)  | 
| 36451 | 374  | 
|
| 72153 | 375  | 
(* main target *)  | 
| 20910 | 376  | 
|
| 72517 | 377  | 
fun init_target background_naming conclude operations target =  | 
| 72953 | 378  | 
Data.map (K [make_lthy (background_naming, operations, conclude, target)]) target  | 
| 
72516
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
379  | 
|
| 
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
380  | 
fun init {background_naming, setup, conclude} operations thy =
 | 
| 
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
381  | 
thy  | 
| 
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
382  | 
|> Sign.change_begin  | 
| 
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
383  | 
|> setup  | 
| 
 
17dc99589a91
unified Local_Theory.init with Generic_Target.init
 
haftmann 
parents: 
72505 
diff
changeset
 | 
384  | 
|> init_target background_naming (conclude #> target_of #> Sign.change_end_local) operations;  | 
| 20910 | 385  | 
|
| 72517 | 386  | 
val exit_of = #conclude o bottom_of;  | 
| 
66335
 
a849ce33923d
treat exit separate from regular local theory operations
 
haftmann 
parents: 
66259 
diff
changeset
 | 
387  | 
|
| 72502 | 388  | 
fun exit lthy = exit_of lthy (assert_bottom lthy);  | 
| 42360 | 389  | 
val exit_global = Proof_Context.theory_of o exit;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
390  | 
|
| 72153 | 391  | 
fun exit_result decl (x, lthy) =  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
392  | 
let  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
393  | 
val ctxt = exit lthy;  | 
| 
33281
 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 
wenzelm 
parents: 
33276 
diff
changeset
 | 
394  | 
val phi = standard_morphism lthy ctxt;  | 
| 72153 | 395  | 
in (decl phi x, ctxt) end;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
396  | 
|
| 72153 | 397  | 
fun exit_result_global decl (x, lthy) =  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
398  | 
let  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
399  | 
val thy = exit_global lthy;  | 
| 42360 | 400  | 
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
 | 
401  | 
val phi = standard_morphism lthy thy_ctxt;  | 
| 72153 | 402  | 
in (decl phi x, thy) end;  | 
| 20910 | 403  | 
|
| 59890 | 404  | 
|
405  | 
(* nested targets *)  | 
|
406  | 
||
| 72450 | 407  | 
fun begin_nested lthy =  | 
| 59890 | 408  | 
let  | 
409  | 
val _ = assert lthy;  | 
|
410  | 
val (scope, target) = Proof_Context.new_scope lthy;  | 
|
| 
72436
 
d62d84634b06
direct exit to theory when ending nested target on theory target
 
haftmann 
parents: 
72433 
diff
changeset
 | 
411  | 
val entry = make_lthy (background_naming_of lthy, operations_of lthy,  | 
| 72953 | 412  | 
Proof_Context.restore_naming lthy, target);  | 
| 72153 | 413  | 
val lthy' = Data.map (cons entry) target;  | 
| 59890 | 414  | 
in (scope, lthy') end;  | 
415  | 
||
| 72450 | 416  | 
fun end_nested lthy =  | 
| 59890 | 417  | 
let  | 
| 66259 | 418  | 
val _ = assert_not_bottom lthy;  | 
| 72517 | 419  | 
    val ({conclude, ...} :: rest) = Data.get lthy;
 | 
420  | 
in lthy |> Data.put rest |> reset |> conclude end;  | 
|
| 59890 | 421  | 
|
| 72450 | 422  | 
fun end_nested_result decl (x, lthy) =  | 
423  | 
let  | 
|
424  | 
val outer_lthy = end_nested lthy;  | 
|
425  | 
val phi = Proof_Context.export_morphism lthy outer_lthy;  | 
|
426  | 
in (decl phi x, outer_lthy) end;  | 
|
427  | 
||
| 18781 | 428  | 
end;  |