| author | haftmann | 
| Thu, 04 Dec 2008 14:43:33 +0100 | |
| changeset 28965 | 1de908189869 | 
| parent 28406 | daeb21fec18f | 
| child 28991 | 694227dd3e8c | 
| permissions | -rw-r--r-- | 
| 18744 | 1  | 
(* Title: Pure/Isar/local_theory.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Makarius  | 
|
4  | 
||
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
5  | 
Local theory operations, with abstract target context.  | 
| 18744 | 6  | 
*)  | 
7  | 
||
| 18951 | 8  | 
type local_theory = Proof.context;  | 
9  | 
||
| 18744 | 10  | 
signature LOCAL_THEORY =  | 
11  | 
sig  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
12  | 
type operations  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
13  | 
val group_of: local_theory -> string  | 
| 28017 | 14  | 
val group_properties_of: local_theory -> Properties.T  | 
15  | 
val group_position_of: local_theory -> Properties.T  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
16  | 
val set_group: string -> local_theory -> local_theory  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
17  | 
val target_of: local_theory -> Proof.context  | 
| 20960 | 18  | 
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory  | 
19  | 
val raw_theory: (theory -> theory) -> local_theory -> local_theory  | 
|
| 
28379
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
20  | 
val checkpoint: local_theory -> local_theory  | 
| 22240 | 21  | 
val full_naming: local_theory -> NameSpace.naming  | 
| 21614 | 22  | 
val full_name: local_theory -> bstring -> string  | 
| 20960 | 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: 
20784 
diff
changeset
 | 
26  | 
val target: (Proof.context -> Proof.context) -> local_theory -> local_theory  | 
| 21860 | 27  | 
val affirm: local_theory -> local_theory  | 
| 20960 | 28  | 
val pretty: local_theory -> Pretty.T list  | 
| 28965 | 29  | 
val abbrev: Syntax.mode -> (Binding.T * mixfix) * term -> local_theory ->  | 
| 25120 | 30  | 
(term * term) * local_theory  | 
| 28965 | 31  | 
val define: string -> (Binding.T * mixfix) * (Attrib.binding * term) -> local_theory ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
32  | 
(term * (string * thm)) * local_theory  | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
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: 
28083 
diff
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: 
28017 
diff
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: 
28379 
diff
changeset
 | 
45  | 
val exit_global: local_theory -> theory  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
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: 
28379 
diff
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: 
20784 
diff
changeset
 | 
53  | 
(** local theory data **)  | 
| 
19059
 
b4ca3100e818
init/exit no longer change the theory (no naming);
 
wenzelm 
parents: 
19032 
diff
changeset
 | 
54  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
55  | 
(* type lthy *)  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
56  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
57  | 
type operations =  | 
| 20960 | 58  | 
 {pretty: local_theory -> Pretty.T list,
 | 
| 28965 | 59  | 
abbrev: Syntax.mode -> (Binding.T * mixfix) * term -> local_theory ->  | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
60  | 
(term * term) * local_theory,  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
61  | 
define: string ->  | 
| 28965 | 62  | 
(Binding.T * mixfix) * (Attrib.binding * term) -> local_theory ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
63  | 
(term * (string * thm)) * local_theory,  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
64  | 
notes: string ->  | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28083 
diff
changeset
 | 
65  | 
(Attrib.binding * (thm list * Attrib.src list) list) list ->  | 
| 
28083
 
103d9282a946
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
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: 
20784 
diff
changeset
 | 
72  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
73  | 
datatype lthy = LThy of  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
74  | 
 {group: string,
 | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
75  | 
theory_prefix: string,  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
76  | 
operations: operations,  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
77  | 
target: Proof.context};  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
78  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
79  | 
fun make_lthy (group, theory_prefix, operations, target) =  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
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: 
20784 
diff
changeset
 | 
81  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
82  | 
|
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
83  | 
(* context data *)  | 
| 18744 | 84  | 
|
85  | 
structure Data = ProofDataFun  | 
|
86  | 
(  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
87  | 
type T = lthy option;  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
88  | 
fun init _ = NONE;  | 
| 18744 | 89  | 
);  | 
90  | 
||
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
91  | 
fun get_lthy lthy =  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
92  | 
(case Data.get lthy of  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
93  | 
NONE => error "No local theory context"  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
94  | 
| SOME (LThy data) => data);  | 
| 18951 | 95  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
96  | 
fun map_lthy f lthy =  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
97  | 
  let val {group, theory_prefix, operations, target} = get_lthy lthy
 | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
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: 
25984 
diff
changeset
 | 
99  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
100  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
101  | 
(* group *)  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
102  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
103  | 
val group_of = #group o get_lthy;  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
104  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
105  | 
fun group_properties_of lthy =  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
106  | 
(case group_of lthy of  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
107  | 
"" => []  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
108  | 
| group => [(Markup.groupN, group)]);  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
109  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
110  | 
fun group_position_of lthy =  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
111  | 
group_properties_of lthy @ Position.properties_of (Position.thread_data ());  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
112  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
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: 
25984 
diff
changeset
 | 
114  | 
(group, theory_prefix, operations, target));  | 
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
115  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
116  | 
|
| 
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
117  | 
(* target *)  | 
| 18767 | 118  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
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: 
25984 
diff
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: 
25984 
diff
changeset
 | 
123  | 
(group, theory_prefix, operations, f target));  | 
| 18767 | 124  | 
|
125  | 
||
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
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: 
28115 
diff
changeset
 | 
138  | 
val checkpoint = raw_theory Theory.checkpoint;  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
139  | 
|
| 22240 | 140  | 
fun full_naming lthy =  | 
141  | 
Sign.naming_of (ProofContext.theory_of lthy)  | 
|
142  | 
|> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy))  | 
|
143  | 
|> NameSpace.qualified_names;  | 
|
144  | 
||
| 28965 | 145  | 
fun full_name naming = NameSpace.full_name (full_naming naming) o Binding.name;  | 
| 21614 | 146  | 
|
| 21529 | 147  | 
fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy  | 
148  | 
|> Sign.sticky_prefix (#theory_prefix (get_lthy lthy))  | 
|
149  | 
|> Sign.qualified_names  | 
|
150  | 
|> f  | 
|
151  | 
||> Sign.restore_naming thy);  | 
|
| 19661 | 152  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
153  | 
fun theory f = #2 o theory_result (f #> pair ());  | 
| 19661 | 154  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
155  | 
fun target_result f lthy =  | 
| 18744 | 156  | 
let  | 
| 
28406
 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 
wenzelm 
parents: 
28395 
diff
changeset
 | 
157  | 
val (res, ctxt') = target_of lthy  | 
| 
 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 
wenzelm 
parents: 
28395 
diff
changeset
 | 
158  | 
|> ContextPosition.set_visible false  | 
| 
 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 
wenzelm 
parents: 
28395 
diff
changeset
 | 
159  | 
|> f  | 
| 
 
daeb21fec18f
target update: invisible context position avoids duplication of reports etc.;
 
wenzelm 
parents: 
28395 
diff
changeset
 | 
160  | 
||> ContextPosition.restore_visible lthy;  | 
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
161  | 
val thy' = ProofContext.theory_of ctxt';  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
162  | 
val lthy' = lthy  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
163  | 
|> map_target (K ctxt')  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
164  | 
|> ProofContext.transfer thy';  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
165  | 
in (res, lthy') end;  | 
| 18876 | 166  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
167  | 
fun target f = #2 o target_result (f #> pair ());  | 
| 18767 | 168  | 
|
169  | 
||
| 
21664
 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 
wenzelm 
parents: 
21614 
diff
changeset
 | 
170  | 
(* basic operations *)  | 
| 18781 | 171  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
172  | 
fun operation f lthy = f (#operations (get_lthy lthy)) lthy;  | 
| 
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
173  | 
fun operation1 f x = operation (fn ops => f ops x);  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
174  | 
fun operation2 f x y lthy = operation (fn ops => f ops x y) lthy;  | 
| 18781 | 175  | 
|
| 
20888
 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 
wenzelm 
parents: 
20784 
diff
changeset
 | 
176  | 
val pretty = operation #pretty;  | 
| 
28379
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
177  | 
val abbrev = apsnd checkpoint ooo operation2 #abbrev;  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
178  | 
val define = apsnd checkpoint ooo operation2 #define;  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
179  | 
val notes = apsnd checkpoint ooo operation2 #notes;  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
180  | 
val type_syntax = checkpoint oo operation1 #type_syntax;  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
181  | 
val term_syntax = checkpoint oo operation1 #term_syntax;  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
182  | 
val declaration = checkpoint oo operation1 #declaration;  | 
| 18781 | 183  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
184  | 
fun note kind (a, ths) = notes kind [(a, [(ths, [])])] #>> the_single;  | 
| 18781 | 185  | 
|
| 
24933
 
5471b164813b
removed LocalTheory.defs/target_morphism operations;
 
wenzelm 
parents: 
24554 
diff
changeset
 | 
186  | 
fun target_morphism lthy =  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
187  | 
ProofContext.norm_export_morphism lthy (target_of lthy);  | 
| 
24933
 
5471b164813b
removed LocalTheory.defs/target_morphism operations;
 
wenzelm 
parents: 
24554 
diff
changeset
 | 
188  | 
|
| 24949 | 189  | 
fun notation add mode raw_args lthy =  | 
| 21743 | 190  | 
let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args  | 
| 24949 | 191  | 
in term_syntax (ProofContext.target_notation add mode args) lthy end;  | 
| 21743 | 192  | 
|
| 
21664
 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 
wenzelm 
parents: 
21614 
diff
changeset
 | 
193  | 
|
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
194  | 
(* init *)  | 
| 20910 | 195  | 
|
196  | 
fun init theory_prefix operations target = target |> Data.map  | 
|
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
197  | 
  (fn NONE => SOME (make_lthy ("", theory_prefix, operations, target))
 | 
| 
28379
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
198  | 
| SOME _ => error "Local theory already initialized")  | 
| 
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
199  | 
|> checkpoint;  | 
| 20910 | 200  | 
|
| 21273 | 201  | 
fun restore lthy =  | 
| 
26131
 
91024979b9eb
maintain group in lthy data, implicit use in operations;
 
wenzelm 
parents: 
25984 
diff
changeset
 | 
202  | 
  let val {group = _, theory_prefix, operations, target} = get_lthy lthy
 | 
| 20910 | 203  | 
in init theory_prefix operations target end;  | 
204  | 
||
| 
28379
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
205  | 
val reinit = checkpoint o operation #reinit;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
206  | 
|
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
207  | 
|
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
208  | 
(* exit *)  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
209  | 
|
| 
28379
 
0de8a35b866a
Theory.checkpoint for main operations, admits concurrent proofs;
 
wenzelm 
parents: 
28115 
diff
changeset
 | 
210  | 
val exit = ProofContext.theory Theory.checkpoint o operation #exit;  | 
| 
28395
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
211  | 
val exit_global = ProofContext.theory_of o exit;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
212  | 
|
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
213  | 
fun exit_result f (x, lthy) =  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
214  | 
let  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
215  | 
val ctxt = exit lthy;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
216  | 
val phi = ProofContext.norm_export_morphism lthy ctxt;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
217  | 
in (f phi x, ctxt) end;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
218  | 
|
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
219  | 
fun exit_result_global f (x, lthy) =  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
220  | 
let  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
221  | 
val thy = exit_global lthy;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
222  | 
val thy_ctxt = ProofContext.init thy;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
223  | 
val phi = ProofContext.norm_export_morphism lthy thy_ctxt;  | 
| 
 
984fcc8feb24
added exit_global, exit_result, exit_result_global;
 
wenzelm 
parents: 
28379 
diff
changeset
 | 
224  | 
in (f phi x, thy) end;  | 
| 20910 | 225  | 
|
| 18781 | 226  | 
end;  |