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