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