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