| author | huffman | 
| Thu, 05 Apr 2012 15:23:26 +0200 | |
| changeset 47375 | 8e6a45f1bf8f | 
| parent 47281 | d6c76b1823fb | 
| child 49042 | 01041f7bf9b4 | 
| 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 | |
| 18744 | 10 | signature LOCAL_THEORY = | 
| 11 | sig | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 12 | type operations | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 13 | val assert: local_theory -> local_theory | 
| 47273 | 14 | val restore: local_theory -> local_theory | 
| 47064 | 15 | val level: Proof.context -> int | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 16 | val assert_bottom: bool -> local_theory -> local_theory | 
| 47064 | 17 | val open_target: Name_Space.naming -> operations -> local_theory -> local_theory | 
| 18 | val close_target: local_theory -> local_theory | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 19 | val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 20 | val naming_of: local_theory -> Name_Space.naming | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 21 | val full_name: local_theory -> binding -> string | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 22 | val map_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 23 | val conceal: local_theory -> local_theory | 
| 33724 | 24 | val new_group: local_theory -> local_theory | 
| 25 | val reset_group: local_theory -> local_theory | |
| 33276 | 26 | val restore_naming: local_theory -> local_theory -> local_theory | 
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 27 | val standard_morphism: local_theory -> Proof.context -> morphism | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 28 | val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a | 
| 20960 | 29 | val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory | 
| 30 | 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 | 31 | 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 | 32 | 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 | 33 | val target_of: local_theory -> Proof.context | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 34 | 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 | 35 | 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 | 36 | val propagate_ml_env: generic_theory -> generic_theory | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 37 | 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 | 38 | val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 39 | (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 | 40 | 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 | 41 | (term * (string * thm)) * local_theory | 
| 33670 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 42 | val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory | 
| 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 43 | val notes: (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 44 | local_theory -> (string * thm list) list * local_theory | 
| 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 45 | val notes_kind: string -> (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 46 | 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 | 47 | 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 | 48 | (term * term) * local_theory | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 49 |   val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory
 | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
37949diff
changeset | 50 | val pretty: local_theory -> Pretty.T list | 
| 36451 | 51 | val set_defsort: sort -> local_theory -> local_theory | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 52 | val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory | 
| 24949 | 53 | val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory | 
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 54 | val class_alias: binding -> class -> local_theory -> local_theory | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 55 | val type_alias: binding -> string -> local_theory -> local_theory | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 56 | val const_alias: binding -> string -> local_theory -> local_theory | 
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
46992diff
changeset | 57 | val init: Name_Space.naming -> operations -> Proof.context -> local_theory | 
| 21292 | 58 | val exit: local_theory -> Proof.context | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 59 | val exit_global: local_theory -> theory | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 60 | 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 | 61 | val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory | 
| 18744 | 62 | end; | 
| 63 | ||
| 33671 | 64 | structure Local_Theory: LOCAL_THEORY = | 
| 18744 | 65 | struct | 
| 66 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 67 | (** local theory data **) | 
| 19059 
b4ca3100e818
init/exit no longer change the theory (no naming);
 wenzelm parents: 
19032diff
changeset | 68 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 69 | (* type lthy *) | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 70 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 71 | 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 | 72 |  {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
 | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 73 | (term * (string * thm)) * local_theory, | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 74 | notes: string -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 75 | (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 76 | 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 | 77 | 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 | 78 | (term * term) * local_theory, | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 79 |   declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory,
 | 
| 38308 
0e4649095739
try to uniformly follow define/note/abbrev/declaration order as close as possible
 haftmann parents: 
37949diff
changeset | 80 | pretty: local_theory -> Pretty.T list, | 
| 21292 | 81 | exit: local_theory -> Proof.context}; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 82 | |
| 47064 | 83 | type lthy = | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 84 |  {naming: Name_Space.naming,
 | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 85 | operations: operations, | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 86 | target: Proof.context}; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 87 | |
| 47064 | 88 | fun make_lthy (naming, operations, target) : lthy = | 
| 89 |   {naming = naming, operations = operations, target = target};
 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 90 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 91 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 92 | (* context data *) | 
| 18744 | 93 | |
| 33519 | 94 | structure Data = Proof_Data | 
| 18744 | 95 | ( | 
| 47064 | 96 | type T = lthy list; | 
| 97 | fun init _ = []; | |
| 18744 | 98 | ); | 
| 99 | ||
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 100 | fun assert lthy = | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 101 | if null (Data.get lthy) then error "Missing local theory context" else lthy; | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 102 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 103 | val get_last_lthy = List.last o Data.get o assert; | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 104 | val get_first_lthy = hd o Data.get o assert; | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 105 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 106 | fun map_first_lthy f = assert #> | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 107 |   Data.map (fn {naming, operations, target} :: parents =>
 | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 108 | make_lthy (f (naming, operations, target)) :: parents); | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 109 | |
| 47273 | 110 | fun restore lthy = #target (get_first_lthy lthy) |> Data.put (Data.get lthy); | 
| 111 | ||
| 47064 | 112 | |
| 113 | (* nested structure *) | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 114 | |
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 115 | val level = length o Data.get; | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 116 | |
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 117 | fun assert_bottom b lthy = | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 118 | let | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 119 | val _ = assert lthy; | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 120 | val b' = level lthy <= 1; | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 121 | in | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 122 | if b andalso not b' then error "Not at bottom of local theory nesting" | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 123 | else if not b andalso b' then error "Already at bottom of local theory nesting" | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 124 | else lthy | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 125 | end; | 
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 126 | |
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 127 | fun open_target naming operations target = assert target | 
| 47064 | 128 | |> Data.map (cons (make_lthy (naming, operations, target))); | 
| 129 | ||
| 130 | fun close_target lthy = | |
| 47281 | 131 | assert_bottom false lthy |> Data.map tl |> restore; | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 132 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 133 | fun map_contexts f lthy = | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 134 | let val n = level lthy in | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 135 |     lthy |> (Data.map o map_index) (fn (i, {naming, operations, target}) =>
 | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 136 | make_lthy (naming, operations, | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 137 | target | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 138 | |> Context_Position.set_visible false | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 139 | |> f (n - i - 1) | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 140 | |> Context_Position.restore_visible target)) | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 141 | |> f n | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 142 | end; | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 143 | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 144 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 145 | (* naming *) | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 146 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 147 | val naming_of = #naming o get_first_lthy; | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 148 | val full_name = Name_Space.full_name o naming_of; | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 149 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 150 | fun map_naming f = | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 151 | map_first_lthy (fn (naming, operations, target) => (f naming, operations, target)); | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 152 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 153 | val conceal = map_naming Name_Space.conceal; | 
| 33724 | 154 | val new_group = map_naming Name_Space.new_group; | 
| 155 | val reset_group = map_naming Name_Space.reset_group; | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 156 | |
| 33276 | 157 | val restore_naming = map_naming o K o naming_of; | 
| 158 | ||
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 159 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 160 | (* standard morphisms *) | 
| 18767 | 161 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 162 | fun standard_morphism lthy ctxt = | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 163 | Proof_Context.norm_export_morphism lthy ctxt $> | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 164 | Morphism.binding_morphism (Name_Space.transform_binding (naming_of lthy)); | 
| 18767 | 165 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 166 | fun standard_form lthy ctxt x = | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 167 | Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); | 
| 18767 | 168 | |
| 169 | ||
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 170 | (* background theory *) | 
| 19661 | 171 | |
| 20960 | 172 | fun raw_theory_result f lthy = | 
| 173 | let | |
| 42360 | 174 | 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 | 175 | val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; | 
| 20960 | 176 | in (res, lthy') end; | 
| 177 | ||
| 178 | fun raw_theory f = #2 o raw_theory_result (f #> pair ()); | |
| 179 | ||
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 180 | val checkpoint = raw_theory Theory.checkpoint; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 181 | |
| 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 | 182 | fun background_theory_result f lthy = | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 183 | lthy |> raw_theory_result (fn thy => | 
| 33157 
56f836b9414f
allow name space entries to be "concealed" -- via binding/naming/local_theory;
 wenzelm parents: 
33095diff
changeset | 184 | thy | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 185 | |> Sign.map_naming (K (naming_of lthy)) | 
| 33157 
56f836b9414f
allow name space entries to be "concealed" -- via binding/naming/local_theory;
 wenzelm parents: 
33095diff
changeset | 186 | |> f | 
| 36004 
5d79ca55a52b
implicit checkpoint in Local_Theory.theory as well -- no longer export Local_Theory.checkpoint;
 wenzelm parents: 
35798diff
changeset | 187 | ||> Sign.restore_naming thy | 
| 
5d79ca55a52b
implicit checkpoint in Local_Theory.theory as well -- no longer export Local_Theory.checkpoint;
 wenzelm parents: 
35798diff
changeset | 188 | ||> Theory.checkpoint); | 
| 19661 | 189 | |
| 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 | 190 | fun background_theory f = #2 o background_theory_result (f #> pair ()); | 
| 19661 | 191 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 192 | |
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 193 | (* target contexts *) | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 194 | |
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 195 | val target_of = #target o get_last_lthy; | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 196 | |
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 197 | fun target f lthy = | 
| 18744 | 198 | let | 
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 199 | val ctxt = target_of lthy; | 
| 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 200 | val ctxt' = ctxt | 
| 33383 | 201 | |> Context_Position.set_visible false | 
| 28406 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 202 | |> f | 
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 203 | |> Context_Position.restore_visible ctxt; | 
| 42360 | 204 | 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 | 205 | in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; | 
| 18876 | 206 | |
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 207 | fun target_morphism lthy = standard_morphism lthy (target_of lthy); | 
| 18767 | 208 | |
| 37949 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 wenzelm parents: 
36610diff
changeset | 209 | fun propagate_ml_env (context as Context.Proof lthy) = | 
| 47064 | 210 | let val inherit = ML_Env.inherit context in | 
| 211 | lthy | |
| 212 | |> 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 | 213 | |> map_contexts (K (Context.proof_map inherit)) | 
| 47064 | 214 | |> Context.Proof | 
| 215 | 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 | 216 | | 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 | 217 | |
| 18767 | 218 | |
| 47064 | 219 | |
| 220 | (** operations **) | |
| 221 | ||
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 222 | val operations_of = #operations o get_first_lthy; | 
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 223 | |
| 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 224 | |
| 47064 | 225 | (* primitives *) | 
| 18781 | 226 | |
| 47066 
8a6124d09ff5
basic support for nested contexts including bundles;
 wenzelm parents: 
47064diff
changeset | 227 | fun operation f lthy = f (operations_of lthy) lthy; | 
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
46992diff
changeset | 228 | fun operation2 f x y = operation (fn ops => f ops x y); | 
| 18781 | 229 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 230 | val pretty = operation #pretty; | 
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 231 | val abbrev = apsnd checkpoint ooo operation2 #abbrev; | 
| 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 | 232 | val define = apsnd checkpoint oo operation2 #define false; | 
| 
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 | 233 | val define_internal = apsnd checkpoint oo operation2 #define true; | 
| 33670 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 234 | val notes_kind = apsnd checkpoint ooo operation2 #notes; | 
| 33456 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 wenzelm parents: 
33383diff
changeset | 235 | val declaration = checkpoint ooo operation2 #declaration; | 
| 18781 | 236 | |
| 36451 | 237 | |
| 47064 | 238 | (* basic derived operations *) | 
| 36451 | 239 | |
| 33670 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 240 | val notes = notes_kind ""; | 
| 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 241 | fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; | 
| 18781 | 242 | |
| 36451 | 243 | fun set_defsort S = | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 244 |   declaration {syntax = true, pervasive = false}
 | 
| 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 245 | (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); | 
| 36451 | 246 | |
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 247 | |
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 248 | (* notation *) | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 249 | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 250 | fun type_notation add mode raw_args lthy = | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 251 | let val args = map (apfst (Morphism.typ (target_morphism lthy))) raw_args in | 
| 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 252 |     declaration {syntax = true, pervasive = false}
 | 
| 47247 | 253 | (Proof_Context.generic_type_notation add mode args) lthy | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 254 | end; | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 255 | |
| 24949 | 256 | fun notation add mode raw_args lthy = | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 257 | let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args in | 
| 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 258 |     declaration {syntax = true, pervasive = false}
 | 
| 47247 | 259 | (Proof_Context.generic_notation add mode args) lthy | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 260 | end; | 
| 21743 | 261 | |
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
changeset | 262 | |
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 263 | (* name space aliases *) | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 264 | |
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 265 | fun alias global_alias local_alias b name = | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
42360diff
changeset | 266 |   declaration {syntax = true, pervasive = false} (fn phi =>
 | 
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 267 | 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 | 268 | 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 | 269 | |
| 42360 | 270 | val class_alias = alias Sign.class_alias Proof_Context.class_alias; | 
| 271 | val type_alias = alias Sign.type_alias Proof_Context.type_alias; | |
| 272 | val const_alias = alias Sign.const_alias Proof_Context.const_alias; | |
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 273 | |
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 274 | |
| 36451 | 275 | |
| 276 | (** init and exit **) | |
| 277 | ||
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 278 | (* init *) | 
| 20910 | 279 | |
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
46992diff
changeset | 280 | fun init naming operations target = | 
| 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
46992diff
changeset | 281 | target |> Data.map | 
| 47064 | 282 | (fn [] => [make_lthy (naming, operations, target)] | 
| 283 | | _ => error "Local theory already initialized") | |
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
46992diff
changeset | 284 | |> checkpoint; | 
| 20910 | 285 | |
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 286 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 287 | (* exit *) | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 288 | |
| 42360 | 289 | val exit = Proof_Context.background_theory Theory.checkpoint o operation #exit; | 
| 290 | val exit_global = Proof_Context.theory_of o exit; | |
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 291 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 292 | fun exit_result f (x, lthy) = | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 293 | let | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 294 | val ctxt = exit lthy; | 
| 33281 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 295 | val phi = standard_morphism lthy ctxt; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 296 | in (f phi x, ctxt) end; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 297 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 298 | fun exit_result_global f (x, lthy) = | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 299 | let | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 300 | val thy = exit_global lthy; | 
| 42360 | 301 | 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 | 302 | val phi = standard_morphism lthy thy_ctxt; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 303 | in (f phi x, thy) end; | 
| 20910 | 304 | |
| 18781 | 305 | end; |