| author | haftmann | 
| Thu, 04 Jan 2007 17:11:09 +0100 | |
| changeset 21993 | 4b802a9e0738 | 
| parent 21860 | c4492c6bf450 | 
| child 22203 | efc0cdc01307 | 
| 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 | |
| 21614 | 16 | val full_name: local_theory -> bstring -> string | 
| 20960 | 17 | val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory | 
| 18951 | 18 | val theory: (theory -> theory) -> local_theory -> local_theory | 
| 20960 | 19 | 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 | 20 | val target: (Proof.context -> Proof.context) -> local_theory -> local_theory | 
| 21860 | 21 | val affirm: local_theory -> local_theory | 
| 20960 | 22 | val pretty: local_theory -> Pretty.T list | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 23 | val consts: (string * typ -> bool) -> | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 24 | ((bstring * typ) * mixfix) list -> local_theory -> (term * thm) list * local_theory | 
| 21433 | 25 | val axioms: string -> ((bstring * Attrib.src list) * term list) list -> local_theory -> | 
| 18951 | 26 | (bstring * thm list) list * local_theory | 
| 21802 | 27 | val abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> | 
| 28 | (term * term) * local_theory | |
| 21433 | 29 | val defs: string -> ((bstring * mixfix) * ((bstring * Attrib.src list) * term)) list -> | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 30 | local_theory -> (term * (bstring * thm)) list * 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 | 
| 21498 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 33 | val type_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory | 
| 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 34 | val term_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory | 
| 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 35 | val declaration: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory | 
| 21433 | 36 | val def: string -> (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> | 
| 18951 | 37 | local_theory -> (term * (bstring * thm)) * local_theory | 
| 21433 | 38 | val note: string -> (bstring * Attrib.src list) * thm list -> | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 39 | local_theory -> (bstring * thm list) * Proof.context | 
| 21743 | 40 | val notation: Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory | 
| 21529 | 41 | val target_morphism: local_theory -> morphism | 
| 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: 
20784diff
changeset | 52 | (** local theory data **) | 
| 19059 
b4ca3100e818
init/exit no longer change the theory (no naming);
 wenzelm parents: 
19032diff
changeset | 53 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 54 | (* type lthy *) | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 55 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
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: 
20784diff
changeset | 58 | consts: (string * typ -> bool) -> ((bstring * typ) * mixfix) list -> local_theory -> | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
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: 
20784diff
changeset | 61 | (bstring * thm list) list * local_theory, | 
| 21802 | 62 | abbrev: Syntax.mode -> (bstring * mixfix) * term -> local_theory -> (term * term) * local_theory, | 
| 21433 | 63 | defs: string -> ((bstring * mixfix) * ((bstring * Attrib.src list) * term)) list -> | 
| 64 | local_theory -> (term * (bstring * thm)) list * local_theory, | |
| 65 | notes: string -> | |
| 66 | ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> | |
| 67 | local_theory -> (bstring * thm list) list * local_theory, | |
| 21498 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 68 | type_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory, | 
| 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 69 | term_syntax: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory, | 
| 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 70 | declaration: (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory, | 
| 21529 | 71 | target_morphism: local_theory -> morphism, | 
| 72 | target_name: local_theory -> bstring -> string, | |
| 21292 | 73 | reinit: local_theory -> theory -> local_theory, | 
| 74 | exit: local_theory -> Proof.context}; | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 75 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 76 | datatype lthy = LThy of | 
| 21529 | 77 |  {theory_prefix: string,
 | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 78 | operations: operations, | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 79 | target: Proof.context}; | 
| 
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 | fun make_lthy (theory_prefix, operations, target) = | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 82 |   LThy {theory_prefix = theory_prefix, operations = operations, target = target};
 | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 83 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 84 | |
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 85 | (* context data *) | 
| 18744 | 86 | |
| 87 | structure Data = ProofDataFun | |
| 88 | ( | |
| 18876 | 89 | val name = "Pure/local_theory"; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 90 | type T = lthy option; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 91 | fun init _ = NONE; | 
| 18744 | 92 | fun print _ _ = (); | 
| 93 | ); | |
| 94 | val _ = Context.add_setup Data.init; | |
| 95 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 96 | fun get_lthy lthy = | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 97 | (case Data.get lthy of | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 98 | NONE => error "No local theory context" | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 99 | | SOME (LThy data) => data); | 
| 18951 | 100 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 101 | fun map_lthy f lthy = | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 102 |   let val {theory_prefix, operations, target} = get_lthy lthy
 | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 103 | in Data.put (SOME (make_lthy (f (theory_prefix, operations, target)))) lthy end; | 
| 18767 | 104 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 105 | val target_of = #target o get_lthy; | 
| 21860 | 106 | val affirm = tap target_of; | 
| 18767 | 107 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 108 | 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 | 109 | (theory_prefix, operations, f target)); | 
| 18767 | 110 | |
| 111 | ||
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 112 | (* substructure mappings *) | 
| 19661 | 113 | |
| 20960 | 114 | fun raw_theory_result f lthy = | 
| 115 | let | |
| 116 | val (res, thy') = f (ProofContext.theory_of lthy); | |
| 117 | val lthy' = lthy | |
| 118 | |> map_target (ProofContext.transfer thy') | |
| 119 | |> ProofContext.transfer thy'; | |
| 120 | in (res, lthy') end; | |
| 121 | ||
| 122 | fun raw_theory f = #2 o raw_theory_result (f #> pair ()); | |
| 123 | ||
| 21614 | 124 | fun full_name lthy = | 
| 125 | let | |
| 126 | val naming = Sign.naming_of (ProofContext.theory_of lthy) | |
| 127 | |> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy)) | |
| 128 | |> NameSpace.qualified_names; | |
| 21685 | 129 | in NameSpace.full naming end; | 
| 21614 | 130 | |
| 21529 | 131 | fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy | 
| 132 | |> Sign.sticky_prefix (#theory_prefix (get_lthy lthy)) | |
| 133 | |> Sign.qualified_names | |
| 134 | |> f | |
| 135 | ||> Sign.restore_naming thy); | |
| 19661 | 136 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 137 | fun theory f = #2 o theory_result (f #> pair ()); | 
| 19661 | 138 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 139 | fun target_result f lthy = | 
| 18744 | 140 | let | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 141 | val (res, ctxt') = f (#target (get_lthy lthy)); | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 142 | val thy' = ProofContext.theory_of ctxt'; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 143 | val lthy' = lthy | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 144 | |> map_target (K ctxt') | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 145 | |> ProofContext.transfer thy'; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 146 | in (res, lthy') end; | 
| 18876 | 147 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 148 | fun target f = #2 o target_result (f #> pair ()); | 
| 18767 | 149 | |
| 150 | ||
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
changeset | 151 | (* basic operations *) | 
| 18781 | 152 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 153 | fun operation f lthy = f (#operations (get_lthy lthy)) lthy; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 154 | fun operation1 f x = operation (fn ops => f ops x); | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 155 | fun operation2 f x y = operation (fn ops => f ops x y); | 
| 18781 | 156 | |
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 157 | val pretty = operation #pretty; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 158 | val consts = operation2 #consts; | 
| 21433 | 159 | val axioms = operation2 #axioms; | 
| 21743 | 160 | val abbrev = operation2 #abbrev; | 
| 21433 | 161 | val defs = operation2 #defs; | 
| 162 | val notes = operation2 #notes; | |
| 21498 
6382c3a1e7cf
uniform interface for type_syntax/term_syntax/declaration, dependent on morphism;
 wenzelm parents: 
21433diff
changeset | 163 | val type_syntax = operation1 #type_syntax; | 
| 20888 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 164 | val term_syntax = operation1 #term_syntax; | 
| 
0ddd2f297ae7
turned into abstract wrapper module, cf. theory_target.ML;
 wenzelm parents: 
20784diff
changeset | 165 | val declaration = operation1 #declaration; | 
| 21529 | 166 | val target_morphism = operation #target_morphism; | 
| 167 | val target_name = operation #target_name; | |
| 18781 | 168 | |
| 21433 | 169 | fun def kind arg lthy = lthy |> defs kind [arg] |>> hd; | 
| 170 | fun note kind (a, ths) lthy = lthy |> notes kind [(a, [(ths, [])])] |>> hd; | |
| 18781 | 171 | |
| 21743 | 172 | fun notation mode raw_args lthy = | 
| 173 | let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args | |
| 174 | in term_syntax (ProofContext.target_notation mode args) lthy end; | |
| 175 | ||
| 21664 
dd4e89123359
notation/abbreviation: more serious handling of morphisms;
 wenzelm parents: 
21614diff
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; |