| author | haftmann | 
| Sun, 11 Apr 2010 16:51:06 +0200 | |
| changeset 36110 | 4ab91a42666a | 
| parent 36004 | 5d79ca55a52b | 
| child 36451 | ddc965e172c4 | 
| 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; | 
| 8 | ||
| 18744 | 9 | signature LOCAL_THEORY = | 
| 10 | sig | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 11 | type operations | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 12 | val affirm: local_theory -> local_theory | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 13 | val naming_of: local_theory -> Name_Space.naming | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 14 | val full_name: local_theory -> binding -> string | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 15 | 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 | 16 | val conceal: local_theory -> local_theory | 
| 33724 | 17 | val new_group: local_theory -> local_theory | 
| 18 | val reset_group: local_theory -> local_theory | |
| 33276 | 19 | val restore_naming: local_theory -> local_theory -> local_theory | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 20 | val target_of: local_theory -> Proof.context | 
| 20960 | 21 | val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory | 
| 22 | val raw_theory: (theory -> theory) -> local_theory -> local_theory | |
| 23 | val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory | |
| 18951 | 24 | val theory: (theory -> theory) -> local_theory -> local_theory | 
| 20960 | 25 | val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 26 | val target: (Proof.context -> Proof.context) -> local_theory -> local_theory | 
| 30578 | 27 | val map_contexts: (Context.generic -> Context.generic) -> local_theory -> local_theory | 
| 33281 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 28 | val standard_morphism: local_theory -> Proof.context -> morphism | 
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 29 | val target_morphism: local_theory -> morphism | 
| 33456 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 wenzelm parents: 
33383diff
changeset | 30 | val global_morphism: local_theory -> morphism | 
| 20960 | 31 | val pretty: local_theory -> Pretty.T list | 
| 29581 | 32 | val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> | 
| 25120 | 33 | (term * term) * local_theory | 
| 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 | 34 | val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 35 | (term * (string * thm)) * local_theory | 
| 33670 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 36 | 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 | 37 | 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 | 38 | local_theory -> (string * thm list) list * local_theory | 
| 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 39 | 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 | 40 | local_theory -> (string * thm list) list * local_theory | 
| 33456 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 wenzelm parents: 
33383diff
changeset | 41 | val declaration: bool -> declaration -> local_theory -> local_theory | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 42 | val syntax_declaration: bool -> declaration -> local_theory -> local_theory | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 43 | val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory | 
| 24949 | 44 | 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 | 45 | val class_alias: binding -> class -> local_theory -> local_theory | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 46 | val type_alias: binding -> string -> local_theory -> local_theory | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 47 | val const_alias: binding -> string -> local_theory -> local_theory | 
| 33724 | 48 | val init: serial option -> string -> operations -> Proof.context -> local_theory | 
| 21273 | 49 | val restore: local_theory -> local_theory | 
| 25289 | 50 | val reinit: local_theory -> local_theory | 
| 21292 | 51 | val exit: local_theory -> Proof.context | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 52 | val exit_global: local_theory -> theory | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 53 | 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 | 54 | val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory | 
| 18744 | 55 | end; | 
| 56 | ||
| 33671 | 57 | structure Local_Theory: LOCAL_THEORY = | 
| 18744 | 58 | struct | 
| 59 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 60 | (** local theory data **) | 
| 19059 
b4ca3100e818
init/exit no longer change the theory (no naming);
 wenzelm parents: 
19032diff
changeset | 61 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 62 | (* type lthy *) | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 63 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 64 | type operations = | 
| 20960 | 65 |  {pretty: local_theory -> Pretty.T list,
 | 
| 29581 | 66 | abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 67 | (term * term) * local_theory, | 
| 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 | 68 | define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 69 | (term * (string * thm)) * local_theory, | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 70 | notes: string -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 71 | (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 72 | local_theory -> (string * thm list) list * local_theory, | 
| 33456 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 wenzelm parents: 
33383diff
changeset | 73 | declaration: bool -> declaration -> local_theory -> local_theory, | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 74 | syntax_declaration: bool -> declaration -> local_theory -> local_theory, | 
| 25289 | 75 | reinit: local_theory -> local_theory, | 
| 21292 | 76 | exit: local_theory -> Proof.context}; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 77 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 78 | datatype lthy = LThy of | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 79 |  {naming: Name_Space.naming,
 | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 80 | theory_prefix: string, | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 81 | operations: operations, | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 82 | target: Proof.context}; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 83 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 84 | fun make_lthy (naming, theory_prefix, operations, target) = | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 85 |   LThy {naming = naming, theory_prefix = theory_prefix,
 | 
| 33157 
56f836b9414f
allow name space entries to be "concealed" -- via binding/naming/local_theory;
 wenzelm parents: 
33095diff
changeset | 86 | operations = operations, target = target}; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 87 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 88 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 89 | (* context data *) | 
| 18744 | 90 | |
| 33519 | 91 | structure Data = Proof_Data | 
| 18744 | 92 | ( | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 93 | type T = lthy option; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 94 | fun init _ = NONE; | 
| 18744 | 95 | ); | 
| 96 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 97 | fun get_lthy lthy = | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 98 | (case Data.get lthy of | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 99 | NONE => error "No local theory context" | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 100 | | SOME (LThy data) => data); | 
| 18951 | 101 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 102 | fun map_lthy f lthy = | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 103 |   let val {naming, theory_prefix, operations, target} = get_lthy lthy
 | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 104 | in Data.put (SOME (make_lthy (f (naming, theory_prefix, operations, target)))) lthy end; | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 105 | |
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 106 | val affirm = tap get_lthy; | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 107 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 108 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 109 | (* naming *) | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 110 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 111 | val naming_of = #naming o get_lthy; | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 112 | 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 | 113 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 114 | fun map_naming f = map_lthy (fn (naming, theory_prefix, operations, target) => | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 115 | (f naming, theory_prefix, operations, target)); | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 116 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 117 | val conceal = map_naming Name_Space.conceal; | 
| 33724 | 118 | val new_group = map_naming Name_Space.new_group; | 
| 119 | val reset_group = map_naming Name_Space.reset_group; | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 120 | |
| 33276 | 121 | val restore_naming = map_naming o K o naming_of; | 
| 122 | ||
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 123 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 124 | (* target *) | 
| 18767 | 125 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 126 | val target_of = #target o get_lthy; | 
| 18767 | 127 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 128 | fun map_target f = map_lthy (fn (naming, theory_prefix, operations, target) => | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 129 | (naming, theory_prefix, operations, f target)); | 
| 18767 | 130 | |
| 131 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 132 | (* substructure mappings *) | 
| 19661 | 133 | |
| 20960 | 134 | fun raw_theory_result f lthy = | 
| 135 | let | |
| 136 | val (res, thy') = f (ProofContext.theory_of lthy); | |
| 137 | val lthy' = lthy | |
| 138 | |> map_target (ProofContext.transfer thy') | |
| 139 | |> ProofContext.transfer thy'; | |
| 140 | in (res, lthy') end; | |
| 141 | ||
| 142 | fun raw_theory f = #2 o raw_theory_result (f #> pair ()); | |
| 143 | ||
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 144 | val checkpoint = raw_theory Theory.checkpoint; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 145 | |
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 146 | fun theory_result f lthy = | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 147 | lthy |> raw_theory_result (fn thy => | 
| 33157 
56f836b9414f
allow name space entries to be "concealed" -- via binding/naming/local_theory;
 wenzelm parents: 
33095diff
changeset | 148 | thy | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 149 | |> 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 | 150 | |> f | 
| 36004 
5d79ca55a52b
implicit checkpoint in Local_Theory.theory as well -- no longer export Local_Theory.checkpoint;
 wenzelm parents: 
35798diff
changeset | 151 | ||> Sign.restore_naming thy | 
| 
5d79ca55a52b
implicit checkpoint in Local_Theory.theory as well -- no longer export Local_Theory.checkpoint;
 wenzelm parents: 
35798diff
changeset | 152 | ||> Theory.checkpoint); | 
| 19661 | 153 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 154 | fun theory f = #2 o theory_result (f #> pair ()); | 
| 19661 | 155 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 156 | fun target_result f lthy = | 
| 18744 | 157 | let | 
| 28406 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 158 | val (res, ctxt') = target_of lthy | 
| 33383 | 159 | |> Context_Position.set_visible false | 
| 28406 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 160 | |> f | 
| 33383 | 161 | ||> Context_Position.restore_visible lthy; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 162 | val thy' = ProofContext.theory_of ctxt'; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 163 | val lthy' = lthy | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 164 | |> map_target (K ctxt') | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 165 | |> ProofContext.transfer thy'; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 166 | in (res, lthy') end; | 
| 18876 | 167 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 168 | fun target f = #2 o target_result (f #> pair ()); | 
| 18767 | 169 | |
| 30578 | 170 | fun map_contexts f = | 
| 171 | theory (Context.theory_map f) #> | |
| 172 | target (Context.proof_map f) #> | |
| 173 | Context.proof_map f; | |
| 174 | ||
| 18767 | 175 | |
| 33281 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 176 | (* morphisms *) | 
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 177 | |
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 178 | fun standard_morphism lthy ctxt = | 
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 179 | ProofContext.norm_export_morphism lthy ctxt $> | 
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 180 | Morphism.binding_morphism (Name_Space.transform_binding (naming_of lthy)); | 
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 181 | |
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 182 | fun target_morphism lthy = standard_morphism lthy (target_of lthy); | 
| 33456 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 wenzelm parents: 
33383diff
changeset | 183 | fun global_morphism lthy = standard_morphism lthy (ProofContext.init (ProofContext.theory_of lthy)); | 
| 33281 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 184 | |
| 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 185 | |
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
changeset | 186 | (* basic operations *) | 
| 18781 | 187 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 188 | fun operation f lthy = f (#operations (get_lthy lthy)) lthy; | 
| 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 | 189 | fun operation1 f x = operation (fn ops => f ops x); | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 190 | fun operation2 f x y lthy = operation (fn ops => f ops x y) lthy; | 
| 18781 | 191 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 192 | val pretty = operation #pretty; | 
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 193 | val abbrev = apsnd checkpoint ooo operation2 #abbrev; | 
| 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 | 194 | val define = apsnd checkpoint oo operation1 #define; | 
| 33670 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 195 | 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 | 196 | val declaration = checkpoint ooo operation2 #declaration; | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 197 | val syntax_declaration = checkpoint ooo operation2 #syntax_declaration; | 
| 18781 | 198 | |
| 33670 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 199 | val notes = notes_kind ""; | 
| 
02b7738aef6a
eliminated slightly odd kind argument of LocalTheory.note(s);
 wenzelm parents: 
33519diff
changeset | 200 | fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; | 
| 18781 | 201 | |
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 202 | |
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 203 | (* notation *) | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 204 | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 205 | fun type_notation add mode raw_args lthy = | 
| 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 206 | let val args = map (apfst (Morphism.typ (target_morphism lthy))) raw_args | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 207 | in syntax_declaration false (ProofContext.target_type_notation add mode args) lthy end; | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
33764diff
changeset | 208 | |
| 24949 | 209 | fun notation add mode raw_args lthy = | 
| 21743 | 210 | let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args | 
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 211 | in syntax_declaration false (ProofContext.target_notation add mode args) lthy end; | 
| 21743 | 212 | |
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
changeset | 213 | |
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 214 | (* name space aliases *) | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 215 | |
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 216 | fun alias global_alias local_alias b name = | 
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 217 | syntax_declaration false (fn phi => | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 218 | let val b' = Morphism.binding phi b | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 219 | in Context.mapping (global_alias b' name) (local_alias b' name) end) | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 220 | #> local_alias b name; | 
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 221 | |
| 35798 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 222 | val class_alias = alias Sign.class_alias ProofContext.class_alias; | 
| 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 223 | val type_alias = alias Sign.type_alias ProofContext.type_alias; | 
| 
fd1bb29f8170
replaced type_syntax/term_syntax by uniform syntax_declaration;
 wenzelm parents: 
35738diff
changeset | 224 | val const_alias = alias Sign.const_alias ProofContext.const_alias; | 
| 35738 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 225 | |
| 
98fd035c3fe3
added Local_Theory.alias operations (independent of target);
 wenzelm parents: 
35412diff
changeset | 226 | |
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 227 | (* init *) | 
| 20910 | 228 | |
| 33724 | 229 | fun init group theory_prefix operations target = | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 230 | let val naming = | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 231 | Sign.naming_of (ProofContext.theory_of target) | 
| 33724 | 232 | |> Name_Space.set_group group | 
| 33166 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 233 | |> Name_Space.mandatory_path theory_prefix; | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 234 | in | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 235 | target |> Data.map | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 236 | (fn NONE => SOME (make_lthy (naming, theory_prefix, operations, target)) | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 237 | | SOME _ => error "Local theory already initialized") | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 238 | |> checkpoint | 
| 
55f250ef9e31
maintain proper Name_Space.naming, with conceal and set_group;
 wenzelm parents: 
33157diff
changeset | 239 | end; | 
| 20910 | 240 | |
| 21273 | 241 | fun restore lthy = | 
| 33724 | 242 |   let val {naming, theory_prefix, operations, target} = get_lthy lthy
 | 
| 243 | in init (Name_Space.get_group naming) theory_prefix operations target end; | |
| 20910 | 244 | |
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 245 | val reinit = checkpoint o operation #reinit; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 246 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 247 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 248 | (* exit *) | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 249 | |
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 250 | val exit = ProofContext.theory Theory.checkpoint o operation #exit; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 251 | val exit_global = ProofContext.theory_of o exit; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 252 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 253 | fun exit_result f (x, lthy) = | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 254 | let | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 255 | val ctxt = exit lthy; | 
| 33281 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 256 | val phi = standard_morphism lthy ctxt; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 257 | in (f phi x, ctxt) end; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 258 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 259 | fun exit_result_global f (x, lthy) = | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 260 | let | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 261 | val thy = exit_global lthy; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 262 | val thy_ctxt = ProofContext.init thy; | 
| 33281 
223ef9bc399a
let naming transform binding beforehand -- covering only the "conceal" flag for now;
 wenzelm parents: 
33276diff
changeset | 263 | val phi = standard_morphism lthy thy_ctxt; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 264 | in (f phi x, thy) end; | 
| 20910 | 265 | |
| 18781 | 266 | end; |