| author | haftmann |
| Fri, 15 Feb 2008 16:09:12 +0100 | |
| changeset 26072 | f65a7fa2da6c |
| parent 25984 | da0399c9ffcb |
| child 26131 | 91024979b9eb |
| 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 |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
13 |
val target_of: local_theory -> Proof.context |
| 20960 | 14 |
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory |
15 |
val raw_theory: (theory -> theory) -> local_theory -> local_theory |
|
| 22240 | 16 |
val full_naming: local_theory -> NameSpace.naming |
| 21614 | 17 |
val full_name: local_theory -> bstring -> string |
| 20960 | 18 |
val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory |
| 18951 | 19 |
val theory: (theory -> theory) -> local_theory -> local_theory |
| 20960 | 20 |
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
|
21 |
val target: (Proof.context -> Proof.context) -> local_theory -> local_theory |
| 21860 | 22 |
val affirm: local_theory -> local_theory |
| 20960 | 23 |
val pretty: local_theory -> Pretty.T list |
| 24984 | 24 |
val axioms: string -> |
25 |
((bstring * typ) * mixfix) list * ((bstring * Attrib.src list) * term list) list -> local_theory |
|
26 |
-> (term list * (bstring * thm list) list) * local_theory |
|
| 25984 | 27 |
val axioms_grouped: string -> string -> |
28 |
((bstring * typ) * mixfix) list * ((bstring * Attrib.src list) * term list) list -> local_theory |
|
29 |
-> (term list * (bstring * thm list) list) * local_theory |
|
| 21802 | 30 |
val abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> |
| 25120 | 31 |
(term * term) * local_theory |
| 25021 | 32 |
val define: string -> (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> local_theory -> |
33 |
(term * (bstring * thm)) * local_theory |
|
| 25984 | 34 |
val define_grouped: string -> string -> |
35 |
(bstring * mixfix) * ((bstring * Attrib.src list) * term) -> local_theory -> |
|
36 |
(term * (bstring * thm)) * local_theory |
|
37 |
val note: string -> (bstring * Attrib.src list) * thm list -> |
|
38 |
local_theory -> (bstring * thm list) * local_theory |
|
| 21433 | 39 |
val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
| 18951 | 40 |
local_theory -> (bstring * thm list) list * local_theory |
| 25984 | 41 |
val note_grouped: string -> string -> |
42 |
(bstring * Attrib.src list) * thm list -> |
|
43 |
local_theory -> (bstring * thm list) * local_theory |
|
44 |
val notes_grouped: string -> string -> |
|
45 |
((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
|
46 |
local_theory -> (bstring * thm list) list * local_theory |
|
| 24020 | 47 |
val type_syntax: declaration -> local_theory -> local_theory |
48 |
val term_syntax: declaration -> local_theory -> local_theory |
|
49 |
val declaration: declaration -> local_theory -> local_theory |
|
| 24949 | 50 |
val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory |
| 21529 | 51 |
val target_morphism: local_theory -> morphism |
52 |
val init: string -> operations -> Proof.context -> local_theory |
|
| 21273 | 53 |
val restore: local_theory -> local_theory |
| 25289 | 54 |
val reinit: local_theory -> local_theory |
| 21292 | 55 |
val exit: local_theory -> Proof.context |
| 18744 | 56 |
end; |
57 |
||
58 |
structure LocalTheory: LOCAL_THEORY = |
|
59 |
struct |
|
60 |
||
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
61 |
(** local theory data **) |
|
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
62 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
63 |
(* type lthy *) |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
64 |
|
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
65 |
type operations = |
| 20960 | 66 |
{pretty: local_theory -> Pretty.T list,
|
| 25984 | 67 |
axioms: string -> string -> |
| 24984 | 68 |
((bstring * typ) * mixfix) list * ((bstring * Attrib.src list) * term list) list -> local_theory |
69 |
-> (term list * (bstring * thm list) list) * local_theory, |
|
| 25120 | 70 |
abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> (term * term) * local_theory, |
| 25984 | 71 |
define: string -> string -> |
72 |
(bstring * mixfix) * ((bstring * Attrib.src list) * term) -> local_theory -> |
|
| 25021 | 73 |
(term * (bstring * thm)) * local_theory, |
| 25984 | 74 |
notes: string -> string -> |
| 21433 | 75 |
((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
76 |
local_theory -> (bstring * thm list) list * local_theory, |
|
| 24020 | 77 |
type_syntax: declaration -> local_theory -> local_theory, |
78 |
term_syntax: declaration -> local_theory -> local_theory, |
|
79 |
declaration: declaration -> local_theory -> local_theory, |
|
| 25289 | 80 |
reinit: local_theory -> local_theory, |
| 21292 | 81 |
exit: local_theory -> Proof.context}; |
|
20888
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 |
datatype lthy = LThy of |
| 21529 | 84 |
{theory_prefix: string,
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
85 |
operations: operations, |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
86 |
target: Proof.context}; |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
87 |
|
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
88 |
fun make_lthy (theory_prefix, operations, target) = |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
89 |
LThy {theory_prefix = theory_prefix, operations = operations, target = target};
|
|
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 |
|
94 |
structure Data = ProofDataFun |
|
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 = |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
106 |
let val {theory_prefix, operations, target} = get_lthy lthy
|
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
107 |
in Data.put (SOME (make_lthy (f (theory_prefix, operations, target)))) lthy end; |
| 18767 | 108 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
109 |
val target_of = #target o get_lthy; |
| 21860 | 110 |
val affirm = tap target_of; |
| 18767 | 111 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
112 |
fun map_target f = map_lthy (fn (theory_prefix, operations, target) => |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
113 |
(theory_prefix, operations, f target)); |
| 18767 | 114 |
|
115 |
||
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
116 |
(* substructure mappings *) |
| 19661 | 117 |
|
| 20960 | 118 |
fun raw_theory_result f lthy = |
119 |
let |
|
120 |
val (res, thy') = f (ProofContext.theory_of lthy); |
|
121 |
val lthy' = lthy |
|
122 |
|> map_target (ProofContext.transfer thy') |
|
123 |
|> ProofContext.transfer thy'; |
|
124 |
in (res, lthy') end; |
|
125 |
||
126 |
fun raw_theory f = #2 o raw_theory_result (f #> pair ()); |
|
127 |
||
| 22240 | 128 |
fun full_naming lthy = |
129 |
Sign.naming_of (ProofContext.theory_of lthy) |
|
130 |
|> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy)) |
|
131 |
|> NameSpace.qualified_names; |
|
132 |
||
133 |
val full_name = NameSpace.full o full_naming; |
|
| 21614 | 134 |
|
| 21529 | 135 |
fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy |
136 |
|> Sign.sticky_prefix (#theory_prefix (get_lthy lthy)) |
|
137 |
|> Sign.qualified_names |
|
138 |
|> f |
|
139 |
||> Sign.restore_naming thy); |
|
| 19661 | 140 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
141 |
fun theory f = #2 o theory_result (f #> pair ()); |
| 19661 | 142 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
143 |
fun target_result f lthy = |
| 18744 | 144 |
let |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
145 |
val (res, ctxt') = f (#target (get_lthy lthy)); |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
146 |
val thy' = ProofContext.theory_of ctxt'; |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
147 |
val lthy' = lthy |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
148 |
|> map_target (K ctxt') |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
149 |
|> ProofContext.transfer thy'; |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
150 |
in (res, lthy') end; |
| 18876 | 151 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
152 |
fun target f = #2 o target_result (f #> pair ()); |
| 18767 | 153 |
|
154 |
||
|
21664
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
wenzelm
parents:
21614
diff
changeset
|
155 |
(* basic operations *) |
| 18781 | 156 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
157 |
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
|
158 |
fun operation1 f x = operation (fn ops => f ops x); |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
159 |
fun operation2 f x y = operation (fn ops => f ops x y); |
| 25984 | 160 |
fun operation3 f x y z = operation (fn ops => f ops x y z); |
| 18781 | 161 |
|
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
162 |
val pretty = operation #pretty; |
| 25984 | 163 |
val axioms_grouped = operation3 #axioms; |
| 21743 | 164 |
val abbrev = operation2 #abbrev; |
| 25984 | 165 |
val define_grouped = operation3 #define; |
166 |
val notes_grouped = operation3 #notes; |
|
|
21498
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
wenzelm
parents:
21433
diff
changeset
|
167 |
val type_syntax = operation1 #type_syntax; |
|
20888
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
168 |
val term_syntax = operation1 #term_syntax; |
|
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
wenzelm
parents:
20784
diff
changeset
|
169 |
val declaration = operation1 #declaration; |
| 18781 | 170 |
|
| 25984 | 171 |
fun note_grouped kind group (a, ths) lthy = lthy |
172 |
|> notes_grouped kind group [(a, [(ths, [])])] |>> the_single; |
|
173 |
||
174 |
fun axioms kind = axioms_grouped kind ""; |
|
175 |
fun define kind = define_grouped kind ""; |
|
176 |
fun notes kind = notes_grouped kind ""; |
|
177 |
fun note kind = note_grouped kind ""; |
|
| 18781 | 178 |
|
|
24933
5471b164813b
removed LocalTheory.defs/target_morphism operations;
wenzelm
parents:
24554
diff
changeset
|
179 |
fun target_morphism lthy = |
|
5471b164813b
removed LocalTheory.defs/target_morphism operations;
wenzelm
parents:
24554
diff
changeset
|
180 |
ProofContext.export_morphism lthy (target_of lthy) $> |
|
5471b164813b
removed LocalTheory.defs/target_morphism operations;
wenzelm
parents:
24554
diff
changeset
|
181 |
Morphism.thm_morphism Goal.norm_result; |
|
5471b164813b
removed LocalTheory.defs/target_morphism operations;
wenzelm
parents:
24554
diff
changeset
|
182 |
|
| 24949 | 183 |
fun notation add mode raw_args lthy = |
| 21743 | 184 |
let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args |
| 24949 | 185 |
in term_syntax (ProofContext.target_notation add mode args) lthy end; |
| 21743 | 186 |
|
|
21664
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
wenzelm
parents:
21614
diff
changeset
|
187 |
|
| 20910 | 188 |
(* init -- exit *) |
189 |
||
190 |
fun init theory_prefix operations target = target |> Data.map |
|
191 |
(fn NONE => SOME (make_lthy (theory_prefix, operations, target)) |
|
192 |
| SOME _ => error "Local theory already initialized"); |
|
193 |
||
| 21273 | 194 |
fun restore lthy = |
| 20910 | 195 |
let val {theory_prefix, operations, target} = get_lthy lthy
|
196 |
in init theory_prefix operations target end; |
|
197 |
||
| 21292 | 198 |
val reinit = operation #reinit; |
199 |
val exit = operation #exit; |
|
| 20910 | 200 |
|
| 18781 | 201 |
end; |