| author | haftmann | 
| Wed, 23 Sep 2009 08:25:51 +0200 | |
| changeset 32698 | be4b248616c0 | 
| parent 30578 | 9863745880db | 
| child 33095 | bbd52d2f8696 | 
| 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 | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 12 | val group_of: local_theory -> string | 
| 28017 | 13 | val group_properties_of: local_theory -> Properties.T | 
| 14 | val group_position_of: local_theory -> Properties.T | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 15 | val set_group: string -> local_theory -> local_theory | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 16 | val target_of: local_theory -> Proof.context | 
| 20960 | 17 | val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory | 
| 18 | val raw_theory: (theory -> theory) -> local_theory -> local_theory | |
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 19 | val checkpoint: local_theory -> local_theory | 
| 22240 | 20 | val full_naming: local_theory -> NameSpace.naming | 
| 29581 | 21 | val full_name: local_theory -> binding -> string | 
| 20960 | 22 | val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory | 
| 18951 | 23 | val theory: (theory -> theory) -> local_theory -> local_theory | 
| 20960 | 24 | 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 | 25 | val target: (Proof.context -> Proof.context) -> local_theory -> local_theory | 
| 30578 | 26 | val map_contexts: (Context.generic -> Context.generic) -> local_theory -> local_theory | 
| 21860 | 27 | val affirm: local_theory -> local_theory | 
| 20960 | 28 | val pretty: local_theory -> Pretty.T list | 
| 29581 | 29 | val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> | 
| 25120 | 30 | (term * term) * local_theory | 
| 29581 | 31 | val define: string -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 32 | (term * (string * thm)) * local_theory | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 33 | val note: string -> Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 34 | val notes: string -> (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 35 | local_theory -> (string * thm list) list * local_theory | 
| 24020 | 36 | val type_syntax: declaration -> local_theory -> local_theory | 
| 37 | val term_syntax: declaration -> local_theory -> local_theory | |
| 38 | val declaration: declaration -> local_theory -> local_theory | |
| 24949 | 39 | val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory | 
| 21529 | 40 | val target_morphism: local_theory -> morphism | 
| 41 | val init: string -> operations -> Proof.context -> local_theory | |
| 21273 | 42 | val restore: local_theory -> local_theory | 
| 25289 | 43 | val reinit: local_theory -> local_theory | 
| 21292 | 44 | val exit: local_theory -> Proof.context | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 45 | val exit_global: local_theory -> theory | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 46 | 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 | 47 | val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory | 
| 18744 | 48 | end; | 
| 49 | ||
| 50 | structure LocalTheory: LOCAL_THEORY = | |
| 51 | struct | |
| 52 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 53 | (** local theory data **) | 
| 19059 
b4ca3100e818
init/exit no longer change the theory (no naming);
 wenzelm parents: 
19032diff
changeset | 54 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 55 | (* type lthy *) | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 56 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 57 | type operations = | 
| 20960 | 58 |  {pretty: local_theory -> Pretty.T list,
 | 
| 29581 | 59 | abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 60 | (term * term) * local_theory, | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 61 | define: string -> | 
| 29581 | 62 | (binding * mixfix) * (Attrib.binding * term) -> local_theory -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 63 | (term * (string * thm)) * local_theory, | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 64 | notes: string -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 65 | (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 66 | local_theory -> (string * thm list) list * local_theory, | 
| 24020 | 67 | type_syntax: declaration -> local_theory -> local_theory, | 
| 68 | term_syntax: declaration -> local_theory -> local_theory, | |
| 69 | declaration: declaration -> local_theory -> local_theory, | |
| 25289 | 70 | reinit: local_theory -> local_theory, | 
| 21292 | 71 | exit: local_theory -> Proof.context}; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 72 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 73 | datatype lthy = LThy of | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 74 |  {group: string,
 | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 75 | theory_prefix: string, | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 76 | operations: operations, | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 77 | target: Proof.context}; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 78 | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 79 | fun make_lthy (group, theory_prefix, operations, target) = | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 80 |   LThy {group = group, theory_prefix = theory_prefix, operations = operations, target = target};
 | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 81 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 82 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 83 | (* context data *) | 
| 18744 | 84 | |
| 85 | structure Data = ProofDataFun | |
| 86 | ( | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 87 | type T = lthy option; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 88 | fun init _ = NONE; | 
| 18744 | 89 | ); | 
| 90 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 91 | fun get_lthy lthy = | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 92 | (case Data.get lthy of | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 93 | NONE => error "No local theory context" | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 94 | | SOME (LThy data) => data); | 
| 18951 | 95 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 96 | fun map_lthy f lthy = | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 97 |   let val {group, theory_prefix, operations, target} = get_lthy lthy
 | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 98 | in Data.put (SOME (make_lthy (f (group, theory_prefix, operations, target)))) lthy end; | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 99 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 100 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 101 | (* group *) | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 102 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 103 | val group_of = #group o get_lthy; | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 104 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 105 | fun group_properties_of lthy = | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 106 | (case group_of lthy of | 
| 
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 | | group => [(Markup.groupN, group)]); | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 109 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 110 | fun group_position_of lthy = | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 111 | group_properties_of lthy @ Position.properties_of (Position.thread_data ()); | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 112 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 113 | fun set_group group = map_lthy (fn (_, theory_prefix, operations, target) => | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 114 | (group, theory_prefix, operations, target)); | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 115 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 116 | |
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 117 | (* target *) | 
| 18767 | 118 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 119 | val target_of = #target o get_lthy; | 
| 21860 | 120 | val affirm = tap target_of; | 
| 18767 | 121 | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 122 | fun map_target f = map_lthy (fn (group, theory_prefix, operations, target) => | 
| 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 123 | (group, theory_prefix, operations, f target)); | 
| 18767 | 124 | |
| 125 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 126 | (* substructure mappings *) | 
| 19661 | 127 | |
| 20960 | 128 | fun raw_theory_result f lthy = | 
| 129 | let | |
| 130 | val (res, thy') = f (ProofContext.theory_of lthy); | |
| 131 | val lthy' = lthy | |
| 132 | |> map_target (ProofContext.transfer thy') | |
| 133 | |> ProofContext.transfer thy'; | |
| 134 | in (res, lthy') end; | |
| 135 | ||
| 136 | fun raw_theory f = #2 o raw_theory_result (f #> pair ()); | |
| 137 | ||
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 138 | val checkpoint = raw_theory Theory.checkpoint; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 139 | |
| 22240 | 140 | fun full_naming lthy = | 
| 141 | Sign.naming_of (ProofContext.theory_of lthy) | |
| 30469 | 142 | |> NameSpace.mandatory_path (#theory_prefix (get_lthy lthy)); | 
| 22240 | 143 | |
| 28991 | 144 | fun full_name naming = NameSpace.full_name (full_naming naming); | 
| 21614 | 145 | |
| 21529 | 146 | fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy | 
| 30469 | 147 | |> Sign.mandatory_path (#theory_prefix (get_lthy lthy)) | 
| 21529 | 148 | |> f | 
| 149 | ||> Sign.restore_naming thy); | |
| 19661 | 150 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 151 | fun theory f = #2 o theory_result (f #> pair ()); | 
| 19661 | 152 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 153 | fun target_result f lthy = | 
| 18744 | 154 | let | 
| 28406 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 155 | val (res, ctxt') = target_of lthy | 
| 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 156 | |> ContextPosition.set_visible false | 
| 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 157 | |> f | 
| 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 wenzelm parents: 
28395diff
changeset | 158 | ||> ContextPosition.restore_visible lthy; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 159 | val thy' = ProofContext.theory_of ctxt'; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 160 | val lthy' = lthy | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 161 | |> map_target (K ctxt') | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 162 | |> ProofContext.transfer thy'; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 163 | in (res, lthy') end; | 
| 18876 | 164 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 165 | fun target f = #2 o target_result (f #> pair ()); | 
| 18767 | 166 | |
| 30578 | 167 | fun map_contexts f = | 
| 168 | theory (Context.theory_map f) #> | |
| 169 | target (Context.proof_map f) #> | |
| 170 | Context.proof_map f; | |
| 171 | ||
| 18767 | 172 | |
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
changeset | 173 | (* basic operations *) | 
| 18781 | 174 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 175 | fun operation f lthy = f (#operations (get_lthy lthy)) lthy; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 176 | 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 | 177 | fun operation2 f x y lthy = operation (fn ops => f ops x y) lthy; | 
| 18781 | 178 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 179 | val pretty = operation #pretty; | 
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 180 | val abbrev = apsnd checkpoint ooo operation2 #abbrev; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 181 | val define = apsnd checkpoint ooo operation2 #define; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 182 | val notes = apsnd checkpoint ooo operation2 #notes; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 183 | val type_syntax = checkpoint oo operation1 #type_syntax; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 184 | val term_syntax = checkpoint oo operation1 #term_syntax; | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 185 | val declaration = checkpoint oo operation1 #declaration; | 
| 18781 | 186 | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 187 | fun note kind (a, ths) = notes kind [(a, [(ths, [])])] #>> the_single; | 
| 18781 | 188 | |
| 24933 
5471b164813b
removed LocalTheory.defs/target_morphism operations;
 wenzelm parents: 
24554diff
changeset | 189 | fun target_morphism lthy = | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 190 | ProofContext.norm_export_morphism lthy (target_of lthy); | 
| 24933 
5471b164813b
removed LocalTheory.defs/target_morphism operations;
 wenzelm parents: 
24554diff
changeset | 191 | |
| 24949 | 192 | fun notation add mode raw_args lthy = | 
| 21743 | 193 | let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args | 
| 24949 | 194 | in term_syntax (ProofContext.target_notation add mode args) lthy end; | 
| 21743 | 195 | |
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
changeset | 196 | |
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 197 | (* init *) | 
| 20910 | 198 | |
| 199 | fun init theory_prefix operations target = target |> Data.map | |
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 200 |   (fn NONE => SOME (make_lthy ("", theory_prefix, operations, target))
 | 
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 201 | | SOME _ => error "Local theory already initialized") | 
| 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 202 | |> checkpoint; | 
| 20910 | 203 | |
| 21273 | 204 | fun restore lthy = | 
| 26131 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 wenzelm parents: 
25984diff
changeset | 205 |   let val {group = _, theory_prefix, operations, target} = get_lthy lthy
 | 
| 20910 | 206 | in init theory_prefix operations target end; | 
| 207 | ||
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 208 | val reinit = checkpoint o operation #reinit; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 209 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 210 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 211 | (* exit *) | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 212 | |
| 28379 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 wenzelm parents: 
28115diff
changeset | 213 | val exit = ProofContext.theory Theory.checkpoint o operation #exit; | 
| 28395 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 214 | val exit_global = ProofContext.theory_of o exit; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 215 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 216 | fun exit_result f (x, lthy) = | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 217 | let | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 218 | val ctxt = exit lthy; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 219 | val phi = ProofContext.norm_export_morphism lthy ctxt; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 220 | in (f phi x, ctxt) end; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 221 | |
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 222 | fun exit_result_global f (x, lthy) = | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 223 | let | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 224 | val thy = exit_global lthy; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 225 | val thy_ctxt = ProofContext.init thy; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 226 | val phi = ProofContext.norm_export_morphism lthy thy_ctxt; | 
| 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 wenzelm parents: 
28379diff
changeset | 227 | in (f phi x, thy) end; | 
| 20910 | 228 | |
| 18781 | 229 | end; |