author | wenzelm |
Thu, 06 Jul 2006 23:36:40 +0200 | |
changeset 20032 | 2087e5634598 |
parent 19991 | 0c9650664d47 |
child 20236 | 54e15870444b |
permissions | -rw-r--r-- |
18744 | 1 |
(* Title: Pure/Isar/local_theory.ML |
2 |
ID: $Id$ |
|
3 |
Author: Makarius |
|
4 |
||
5 |
Local theory operations, with optional target locale. |
|
6 |
*) |
|
7 |
||
18951 | 8 |
type local_theory = Proof.context; |
9 |
||
18744 | 10 |
signature LOCAL_THEORY = |
11 |
sig |
|
18951 | 12 |
val params_of: local_theory -> (string * typ) list |
13 |
val hyps_of: local_theory -> term list |
|
19913 | 14 |
val standard: local_theory -> thm list -> thm list |
18951 | 15 |
val pretty_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> Pretty.T |
16 |
val quiet_mode: bool ref |
|
17 |
val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit |
|
18 |
val theory: (theory -> theory) -> local_theory -> local_theory |
|
19 |
val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory |
|
20 |
val init: xstring option -> theory -> local_theory |
|
21 |
val init_i: string option -> theory -> local_theory |
|
19017 | 22 |
val restore: local_theory -> local_theory |
18951 | 23 |
val exit: local_theory -> local_theory * theory |
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
24 |
val restore_naming: local_theory -> local_theory |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
25 |
val naming: local_theory -> local_theory |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
26 |
val mapping: xstring option -> (local_theory -> local_theory) -> theory -> local_theory * theory |
19680 | 27 |
val term_syntax: (Context.generic -> Context.generic) -> local_theory -> local_theory |
19661 | 28 |
val syntax: string * bool -> (string * mixfix) list -> local_theory -> local_theory |
29 |
val abbrevs: string * bool -> ((bstring * mixfix) * term) list -> local_theory -> local_theory |
|
18951 | 30 |
val consts: ((bstring * typ) * mixfix) list -> local_theory -> term list * local_theory |
18876 | 31 |
val consts_restricted: (string * typ -> bool) -> |
18951 | 32 |
((bstring * typ) * mixfix) list -> local_theory -> term list * local_theory |
33 |
val axioms: ((bstring * Attrib.src list) * term list) list -> local_theory -> |
|
34 |
(bstring * thm list) list * local_theory |
|
35 |
val axioms_finish: (local_theory -> thm -> thm) -> |
|
36 |
((bstring * Attrib.src list) * term list) list -> local_theory -> |
|
37 |
(bstring * thm list) list * local_theory |
|
18767 | 38 |
val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
18951 | 39 |
local_theory -> (bstring * thm list) list * local_theory |
40 |
val note: (bstring * Attrib.src list) * thm list -> local_theory -> |
|
41 |
(bstring * thm list) * local_theory |
|
18823 | 42 |
val def: (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> |
18951 | 43 |
local_theory -> (term * (bstring * thm)) * local_theory |
44 |
val def_finish: (local_theory -> term -> thm -> thm) -> |
|
18823 | 45 |
(bstring * mixfix) * ((bstring * Attrib.src list) * term) -> |
18951 | 46 |
local_theory -> (term * (bstring * thm)) * local_theory |
18744 | 47 |
end; |
48 |
||
49 |
structure LocalTheory: LOCAL_THEORY = |
|
50 |
struct |
|
51 |
||
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
52 |
|
18767 | 53 |
(** local context data **) |
18744 | 54 |
|
55 |
structure Data = ProofDataFun |
|
56 |
( |
|
18876 | 57 |
val name = "Pure/local_theory"; |
18744 | 58 |
type T = |
20032 | 59 |
{locale: (string * Proof.context) option, |
18744 | 60 |
params: (string * typ) list, |
61 |
hyps: term list, |
|
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
62 |
restore_naming: theory -> theory}; |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
63 |
fun init thy = {locale = NONE, params = [], hyps = [], restore_naming = I}; |
18744 | 64 |
fun print _ _ = (); |
65 |
); |
|
66 |
val _ = Context.add_setup Data.init; |
|
67 |
||
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
68 |
fun map_locale f = Data.map (fn {locale, params, hyps, restore_naming} => |
20032 | 69 |
{locale = Option.map (fn (loc, ctxt) => (loc, f ctxt)) locale, |
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
70 |
params = params, hyps = hyps, restore_naming = restore_naming}); |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
71 |
|
18767 | 72 |
val locale_of = #locale o Data.get; |
73 |
val params_of = #params o Data.get; |
|
74 |
val hyps_of = #hyps o Data.get; |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
75 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
76 |
fun standard ctxt = |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
77 |
(case #locale (Data.get ctxt) of |
19913 | 78 |
NONE => map Drule.standard |
20032 | 79 |
| SOME (_, loc_ctxt) => ProofContext.export_standard ctxt loc_ctxt); |
18744 | 80 |
|
18767 | 81 |
|
18823 | 82 |
(* print consts *) |
18767 | 83 |
|
18951 | 84 |
val quiet_mode = ref false; |
85 |
||
18767 | 86 |
local |
87 |
||
88 |
fun pretty_var ctxt (x, T) = |
|
89 |
Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1, |
|
90 |
Pretty.quote (ProofContext.pretty_typ ctxt T)]; |
|
18744 | 91 |
|
18767 | 92 |
fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs); |
93 |
||
94 |
in |
|
95 |
||
18876 | 96 |
fun pretty_consts ctxt pred cs = |
97 |
(case filter pred (params_of ctxt) of |
|
18767 | 98 |
[] => pretty_vars ctxt "constants" cs |
99 |
| ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]); |
|
100 |
||
18876 | 101 |
fun print_consts _ _ [] = () |
18951 | 102 |
| print_consts ctxt pred cs = |
103 |
if ! quiet_mode then () else Pretty.writeln (pretty_consts ctxt pred cs); |
|
18823 | 104 |
|
18767 | 105 |
end; |
106 |
||
107 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
108 |
(* theory/locale substructures *) |
18781 | 109 |
|
110 |
fun theory_result f ctxt = |
|
111 |
let val (res, thy') = f (ProofContext.theory_of ctxt) |
|
19032 | 112 |
in (res, ctxt |> map_locale (ProofContext.transfer thy') |> ProofContext.transfer thy') end; |
113 |
||
114 |
fun theory f ctxt = #2 (theory_result (f #> pair ()) ctxt); |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
115 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
116 |
fun locale_result f ctxt = |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
117 |
(case locale_of ctxt of |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
118 |
NONE => error "Local theory: no locale context" |
20032 | 119 |
| SOME (_, loc_ctxt) => |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
120 |
let |
20032 | 121 |
val (res, loc_ctxt') = f loc_ctxt; |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
122 |
val thy' = ProofContext.theory_of loc_ctxt'; |
19032 | 123 |
in (res, ctxt |> map_locale (K loc_ctxt') |> ProofContext.transfer thy') end); |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
124 |
|
19017 | 125 |
fun locale f ctxt = |
126 |
if is_none (locale_of ctxt) then ctxt |
|
127 |
else #2 (locale_result (f #> pair ()) ctxt); |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
128 |
|
19017 | 129 |
|
130 |
(* init -- restore -- exit *) |
|
18781 | 131 |
|
132 |
fun init_i NONE thy = ProofContext.init thy |
|
133 |
| init_i (SOME loc) thy = |
|
134 |
thy |
|
20032 | 135 |
|> Locale.init loc |
136 |
|> ProofContext.inferred_fixes |
|
137 |
|> (fn (params, ctxt) => Data.put |
|
138 |
{locale = SOME (loc, ctxt), |
|
18781 | 139 |
params = params, |
140 |
hyps = ProofContext.assms_of ctxt, |
|
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
141 |
restore_naming = Sign.restore_naming thy} ctxt); |
18781 | 142 |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
143 |
fun init loc thy = init_i (Option.map (Locale.intern thy) loc) thy; |
18781 | 144 |
|
19017 | 145 |
fun restore ctxt = |
146 |
(case locale_of ctxt of |
|
147 |
NONE => ctxt |
|
20032 | 148 |
| SOME (_, loc_ctxt) => loc_ctxt |> Data.put (Data.get ctxt)); |
19017 | 149 |
|
19059
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
150 |
fun exit ctxt = (ctxt, ProofContext.theory_of ctxt); |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
151 |
|
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
152 |
|
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
153 |
(* local naming *) |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
154 |
|
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
155 |
fun naming ctxt = |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
156 |
(case locale_of ctxt of |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
157 |
NONE => ctxt |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
158 |
| SOME (loc, _) => ctxt |> theory (Sign.sticky_prefix (Sign.base_name loc))); |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
159 |
|
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
160 |
fun restore_naming ctxt = theory (#restore_naming (Data.get ctxt)) ctxt; |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
161 |
|
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
162 |
fun mapping loc f = |
b4ca3100e818
init/exit no longer change the theory (no naming);
wenzelm
parents:
19032
diff
changeset
|
163 |
init loc #> naming #> f #> restore_naming #> exit; |
18781 | 164 |
|
165 |
||
166 |
||
167 |
(** local theory **) |
|
168 |
||
19661 | 169 |
(* term syntax *) |
170 |
||
19680 | 171 |
fun term_syntax f ctxt = ctxt |> |
19661 | 172 |
(case locale_of ctxt of |
19680 | 173 |
NONE => theory (Context.theory_map f) |
19661 | 174 |
| SOME (loc, _) => |
19680 | 175 |
locale (Locale.add_term_syntax loc (Context.proof_map f)) #> |
176 |
Context.proof_map f); |
|
19661 | 177 |
|
19680 | 178 |
fun syntax x y = |
179 |
term_syntax (Context.mapping (Sign.add_const_syntax x y) (ProofContext.add_const_syntax x y)); |
|
19661 | 180 |
|
19680 | 181 |
fun abbrevs x y = |
182 |
term_syntax (Context.mapping (Sign.add_abbrevs x y) (ProofContext.add_abbrevs x y)); |
|
19661 | 183 |
|
184 |
||
18767 | 185 |
(* consts *) |
18744 | 186 |
|
18876 | 187 |
fun consts_restricted pred decls ctxt = |
18744 | 188 |
let |
189 |
val thy = ProofContext.theory_of ctxt; |
|
18876 | 190 |
val params = filter pred (params_of ctxt); |
18744 | 191 |
val ps = map Free params; |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
192 |
val Ps = map #2 params; |
19661 | 193 |
val abbrs = decls |> map (fn ((c, T), mx) => |
194 |
((c, mx), Term.list_comb (Const (Sign.full_name thy c, Ps ---> T), ps))); |
|
18744 | 195 |
in |
19017 | 196 |
ctxt |> |
197 |
(case locale_of ctxt of |
|
19661 | 198 |
NONE => theory (Sign.add_consts_authentic (map (fn ((c, T), mx) => (c, T, mx)) decls)) |
19017 | 199 |
| SOME (loc, _) => |
200 |
theory (Sign.add_consts_i (map (fn ((c, T), _) => (c, Ps ---> T, NoSyn)) decls)) #> |
|
19661 | 201 |
abbrevs Syntax.default_mode abbrs) |
202 |
|> pair (map #2 abbrs) |
|
18767 | 203 |
end; |
204 |
||
18876 | 205 |
val consts = consts_restricted (K true); |
206 |
||
18767 | 207 |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
208 |
(* fact definitions *) |
18767 | 209 |
|
210 |
fun notes kind facts ctxt = |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
211 |
let |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
212 |
val attrib = Attrib.attribute_i (ProofContext.theory_of ctxt); |
19913 | 213 |
val facts' = map (apsnd (map (apfst (standard ctxt)))) facts; (* FIXME burrow standard *) |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
214 |
in |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
215 |
ctxt |> |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
216 |
(case locale_of ctxt of |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
217 |
NONE => theory_result (PureThy.note_thmss_i kind (Attrib.map_facts attrib facts')) |
19017 | 218 |
| SOME (loc, _) => |
20032 | 219 |
locale_result (apfst #1 o Locale.add_thmss kind loc facts')) |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
220 |
||> (#2 o ProofContext.note_thmss_i (Attrib.map_facts attrib facts)) |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
221 |
end; |
18767 | 222 |
|
18781 | 223 |
fun note (a, ths) ctxt = |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
224 |
ctxt |> notes PureThy.theoremK [(a, [(ths, [])])] |>> hd; |
18767 | 225 |
|
226 |
||
227 |
(* axioms *) |
|
228 |
||
18781 | 229 |
local |
230 |
||
231 |
fun add_axiom hyps (name, prop) thy = |
|
18767 | 232 |
let |
19517 | 233 |
val name' = if name = "" then "axiom_" ^ serial_string () else name; |
18767 | 234 |
val frees = rev ([] |> Term.add_frees prop |> fold Term.add_frees hyps); |
235 |
val prop' = Term.list_all_free (frees, Logic.list_implies (hyps, prop)); |
|
236 |
val thy' = thy |> Theory.add_axioms_i [(name', prop')]; |
|
18781 | 237 |
|
238 |
val cert = Thm.cterm_of thy'; |
|
239 |
fun all_polymorphic (x, T) th = |
|
240 |
let val var = cert (Var ((x, 0), #1 (Logic.dest_all (Thm.prop_of th)))) |
|
241 |
in ((var, cert (Free (x, T))), Thm.forall_elim var th) end; |
|
242 |
fun implies_polymorphic hyp th = Thm.assume (cert hyp) COMP th; |
|
18767 | 243 |
val th = |
244 |
Thm.get_axiom_i thy' (Sign.full_name thy' name') |
|
18781 | 245 |
|> fold_map all_polymorphic frees |-> Drule.cterm_instantiate |
246 |
|> fold implies_polymorphic hyps |
|
18767 | 247 |
in (th, thy') end; |
248 |
||
18781 | 249 |
in |
250 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
251 |
fun axioms_finish finish = fold_map (fn ((name, atts), props) => |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
252 |
fold ProofContext.fix_frees props |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
253 |
#> (fn ctxt => ctxt |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
254 |
|> theory_result (fold_map (add_axiom (hyps_of ctxt)) (PureThy.name_multi name props)) |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
255 |
|-> (fn ths => note ((name, atts), map (finish ctxt) ths)))); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
256 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
257 |
val axioms = axioms_finish (K I); |
18744 | 258 |
|
259 |
end; |
|
18781 | 260 |
|
261 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
262 |
(* constant definitions *) |
18781 | 263 |
|
264 |
local |
|
265 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
266 |
fun add_def (name, prop) = |
19630 | 267 |
Theory.add_defs_i false false [(name, prop)] #> (fn thy => |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
268 |
let |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
269 |
val th = Thm.get_axiom_i thy (Sign.full_name thy name); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
270 |
val cert = Thm.cterm_of thy; |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
271 |
val subst = map2 (fn var => fn free => (cert (Var var), cert (Free free))) |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
272 |
(Term.add_vars (Thm.prop_of th) []) (Term.add_frees prop []); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
273 |
in (Drule.cterm_instantiate subst th, thy) end); |
18781 | 274 |
|
275 |
in |
|
276 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
277 |
fun def_finish finish (var, spec) ctxt = |
18781 | 278 |
let |
279 |
val (x, mx) = var; |
|
280 |
val ((name, atts), rhs) = spec; |
|
281 |
val name' = if name = "" then Thm.def_name x else name; |
|
18876 | 282 |
val rhs_frees = Term.add_frees rhs []; |
18781 | 283 |
in |
284 |
ctxt |
|
18876 | 285 |
|> consts_restricted (member (op =) rhs_frees) [((x, Term.fastype_of rhs), mx)] |
18781 | 286 |
|-> (fn [lhs] => |
287 |
theory_result (add_def (name', Logic.mk_equals (lhs, rhs))) |
|
288 |
#-> (fn th => fn ctxt' => note ((name', atts), [finish ctxt' lhs th]) ctxt') |
|
289 |
#> apfst (fn (b, [th]) => (lhs, (b, th)))) |
|
290 |
end; |
|
291 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
292 |
val def = def_finish (K (K I)); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
293 |
|
18781 | 294 |
end; |
295 |
||
296 |
end; |