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