author | wenzelm |
Tue, 31 Jan 2006 18:19:31 +0100 | |
changeset 18876 | ddb6803da197 |
parent 18823 | 916c493b7f0c |
child 18951 | 4f5f6c632127 |
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 |
||
8 |
signature LOCAL_THEORY = |
|
9 |
sig |
|
10 |
val params_of: Proof.context -> (string * typ) list |
|
18767 | 11 |
val hyps_of: Proof.context -> term list |
18781 | 12 |
val standard: Proof.context -> thm -> thm |
18876 | 13 |
val pretty_consts: Proof.context -> (string * typ -> bool) -> (string * typ) list -> Pretty.T |
14 |
val print_consts: Proof.context -> (string * typ -> bool) -> (string * typ) list -> unit |
|
18781 | 15 |
val theory: (theory -> theory) -> Proof.context -> Proof.context |
16 |
val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context |
|
17 |
val init: xstring option -> theory -> Proof.context |
|
18 |
val init_i: string option -> theory -> Proof.context |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
19 |
val exit: Proof.context -> Proof.context * theory |
18823 | 20 |
val consts: ((bstring * typ) * mixfix) list -> Proof.context -> term list * Proof.context |
18876 | 21 |
val consts_restricted: (string * typ -> bool) -> |
22 |
((bstring * typ) * mixfix) list -> Proof.context -> term list * Proof.context |
|
18823 | 23 |
val axioms: ((bstring * Attrib.src list) * term list) list -> Proof.context -> |
18781 | 24 |
(bstring * thm list) list * Proof.context |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
25 |
val axioms_finish: (Proof.context -> thm -> thm) -> |
18823 | 26 |
((bstring * Attrib.src list) * term list) list -> Proof.context -> |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
27 |
(bstring * thm list) list * Proof.context |
18767 | 28 |
val notes: string -> ((bstring * Attrib.src list) * (thm list * Attrib.src list) list) list -> |
29 |
Proof.context -> (bstring * thm list) list * Proof.context |
|
18781 | 30 |
val note: (bstring * Attrib.src list) * thm list -> Proof.context -> |
18767 | 31 |
(bstring * thm list) * Proof.context |
18823 | 32 |
val def: (bstring * mixfix) * ((bstring * Attrib.src list) * term) -> |
18781 | 33 |
Proof.context -> (term * (bstring * thm)) * Proof.context |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
34 |
val def_finish: (Proof.context -> term -> thm -> thm) -> |
18823 | 35 |
(bstring * mixfix) * ((bstring * Attrib.src list) * term) -> |
18781 | 36 |
Proof.context -> (term * (bstring * thm)) * Proof.context |
18744 | 37 |
end; |
38 |
||
39 |
structure LocalTheory: LOCAL_THEORY = |
|
40 |
struct |
|
41 |
||
18767 | 42 |
(** local context data **) |
18744 | 43 |
|
44 |
structure Data = ProofDataFun |
|
45 |
( |
|
18876 | 46 |
val name = "Pure/local_theory"; |
18744 | 47 |
type T = |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
48 |
{locale: (string * (cterm list * Proof.context)) option, |
18744 | 49 |
params: (string * typ) list, |
50 |
hyps: term list, |
|
51 |
exit: theory -> theory}; |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
52 |
fun init thy = {locale = NONE, params = [], hyps = [], exit = I}; |
18744 | 53 |
fun print _ _ = (); |
54 |
); |
|
55 |
||
56 |
val _ = Context.add_setup Data.init; |
|
57 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
58 |
fun map_locale f = Data.map (fn {locale, params, hyps, exit} => |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
59 |
{locale = Option.map (fn (loc, (view, ctxt)) => (loc, (view, f ctxt))) locale, |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
60 |
params = params, hyps = hyps, exit = exit}); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
61 |
|
18767 | 62 |
val locale_of = #locale o Data.get; |
63 |
val params_of = #params o Data.get; |
|
64 |
val hyps_of = #hyps o Data.get; |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
65 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
66 |
fun standard ctxt = |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
67 |
(case #locale (Data.get ctxt) of |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
68 |
NONE => Drule.standard |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
69 |
| SOME (_, (_, loc_ctxt)) => ProofContext.export_standard ctxt loc_ctxt); |
18744 | 70 |
|
18767 | 71 |
|
18823 | 72 |
(* print consts *) |
18767 | 73 |
|
74 |
local |
|
75 |
||
76 |
fun pretty_var ctxt (x, T) = |
|
77 |
Pretty.block [Pretty.str x, Pretty.str " ::", Pretty.brk 1, |
|
78 |
Pretty.quote (ProofContext.pretty_typ ctxt T)]; |
|
18744 | 79 |
|
18767 | 80 |
fun pretty_vars ctxt kind vs = Pretty.big_list kind (map (pretty_var ctxt) vs); |
81 |
||
82 |
in |
|
83 |
||
18876 | 84 |
fun pretty_consts ctxt pred cs = |
85 |
(case filter pred (params_of ctxt) of |
|
18767 | 86 |
[] => pretty_vars ctxt "constants" cs |
87 |
| ps => Pretty.chunks [pretty_vars ctxt "parameters" ps, pretty_vars ctxt "constants" cs]); |
|
88 |
||
18876 | 89 |
fun print_consts _ _ [] = () |
90 |
| print_consts ctxt pred cs = Pretty.writeln (pretty_consts ctxt pred cs); |
|
18823 | 91 |
|
18767 | 92 |
end; |
93 |
||
94 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
95 |
(* theory/locale substructures *) |
18781 | 96 |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
97 |
fun transfer thy = |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
98 |
ProofContext.transfer thy #> map_locale (ProofContext.transfer thy); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
99 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
100 |
fun theory f ctxt = transfer (f (ProofContext.theory_of ctxt)) ctxt; |
18781 | 101 |
|
102 |
fun theory_result f ctxt = |
|
103 |
let val (res, thy') = f (ProofContext.theory_of ctxt) |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
104 |
in (res, transfer thy' ctxt) end; |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
105 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
106 |
fun locale_result f ctxt = |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
107 |
(case locale_of ctxt of |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
108 |
NONE => error "Local theory: no locale context" |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
109 |
| SOME (loc, (view, loc_ctxt)) => |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
110 |
let |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
111 |
val (res, loc_ctxt') = f loc_ctxt; |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
112 |
val thy' = ProofContext.theory_of loc_ctxt'; |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
113 |
in (res, ctxt |> map_locale (K loc_ctxt') |> transfer thy') end); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
114 |
|
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 |
(* init -- exit *) |
18781 | 117 |
|
118 |
fun init_i NONE thy = ProofContext.init thy |
|
119 |
| init_i (SOME loc) thy = |
|
120 |
thy |
|
121 |
|> Locale.init loc |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
122 |
||>> ProofContext.inferred_fixes |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
123 |
|> (fn ((view, params), ctxt) => Data.put |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
124 |
{locale = SOME (loc, (view, ctxt)), |
18781 | 125 |
params = params, |
126 |
hyps = ProofContext.assms_of ctxt, |
|
127 |
exit = Sign.restore_naming thy} ctxt) |
|
128 |
|> theory (Sign.add_path (Sign.base_name loc) #> Sign.no_base_names); |
|
129 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
130 |
fun init loc thy = init_i (Option.map (Locale.intern thy) loc) thy; |
18781 | 131 |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
132 |
fun exit ctxt = (ctxt, #exit (Data.get ctxt) (ProofContext.theory_of ctxt)); |
18781 | 133 |
|
134 |
||
135 |
||
136 |
(** local theory **) |
|
137 |
||
18767 | 138 |
(* consts *) |
18744 | 139 |
|
18876 | 140 |
fun consts_restricted pred decls ctxt = |
18744 | 141 |
let |
142 |
val thy = ProofContext.theory_of ctxt; |
|
18876 | 143 |
val params = filter pred (params_of ctxt); |
18744 | 144 |
val ps = map Free params; |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
145 |
val Ps = map #2 params; |
18744 | 146 |
in |
147 |
ctxt |
|
18781 | 148 |
|> theory (Sign.add_consts_i (decls |> map (fn ((c, T), mx) => (c, Ps ---> T, mx)))) |
18767 | 149 |
|> pair (decls |> map (fn ((c, T), _) => |
150 |
Term.list_comb (Const (Sign.full_name thy c, Ps ---> T), ps))) |
|
151 |
end; |
|
152 |
||
18876 | 153 |
val consts = consts_restricted (K true); |
154 |
||
18767 | 155 |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
156 |
(* fact definitions *) |
18767 | 157 |
|
158 |
fun notes kind facts ctxt = |
|
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
159 |
let |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
160 |
val attrib = Attrib.attribute_i (ProofContext.theory_of ctxt); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
161 |
val facts' = map (apsnd (map (apfst (map (standard ctxt))))) facts; |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
162 |
in |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
163 |
ctxt |> |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
164 |
(case locale_of ctxt of |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
165 |
NONE => theory_result (PureThy.note_thmss_i kind (Attrib.map_facts attrib facts')) |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
166 |
| SOME (loc, view_ctxt) => |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
167 |
locale_result (K (apfst #1 (Locale.add_thmss kind loc facts' view_ctxt)))) |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
168 |
||> (#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
|
169 |
end; |
18767 | 170 |
|
18781 | 171 |
fun note (a, ths) ctxt = |
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
172 |
ctxt |> notes PureThy.theoremK [(a, [(ths, [])])] |>> hd; |
18767 | 173 |
|
174 |
||
175 |
(* axioms *) |
|
176 |
||
18781 | 177 |
local |
178 |
||
179 |
fun add_axiom hyps (name, prop) thy = |
|
18767 | 180 |
let |
18781 | 181 |
val name' = if name = "" then "axiom_" ^ string_of_int (serial ()) else name; |
18767 | 182 |
val frees = rev ([] |> Term.add_frees prop |> fold Term.add_frees hyps); |
183 |
val prop' = Term.list_all_free (frees, Logic.list_implies (hyps, prop)); |
|
184 |
val thy' = thy |> Theory.add_axioms_i [(name', prop')]; |
|
18781 | 185 |
|
186 |
val cert = Thm.cterm_of thy'; |
|
187 |
fun all_polymorphic (x, T) th = |
|
188 |
let val var = cert (Var ((x, 0), #1 (Logic.dest_all (Thm.prop_of th)))) |
|
189 |
in ((var, cert (Free (x, T))), Thm.forall_elim var th) end; |
|
190 |
fun implies_polymorphic hyp th = Thm.assume (cert hyp) COMP th; |
|
18767 | 191 |
val th = |
192 |
Thm.get_axiom_i thy' (Sign.full_name thy' name') |
|
18781 | 193 |
|> fold_map all_polymorphic frees |-> Drule.cterm_instantiate |
194 |
|> fold implies_polymorphic hyps |
|
18767 | 195 |
in (th, thy') end; |
196 |
||
18781 | 197 |
in |
198 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
199 |
fun axioms_finish finish = fold_map (fn ((name, atts), props) => |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
200 |
fold ProofContext.fix_frees props |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
201 |
#> (fn ctxt => ctxt |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
202 |
|> 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
|
203 |
|-> (fn ths => note ((name, atts), map (finish ctxt) ths)))); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
204 |
|
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
205 |
val axioms = axioms_finish (K I); |
18744 | 206 |
|
207 |
end; |
|
18781 | 208 |
|
209 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
210 |
(* constant definitions *) |
18781 | 211 |
|
212 |
local |
|
213 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
214 |
fun add_def (name, prop) = |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
215 |
Theory.add_defs_i false [(name, prop)] #> (fn thy => |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
216 |
let |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
217 |
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
|
218 |
val cert = Thm.cterm_of thy; |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
219 |
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
|
220 |
(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
|
221 |
in (Drule.cterm_instantiate subst th, thy) end); |
18781 | 222 |
|
223 |
in |
|
224 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
225 |
fun def_finish finish (var, spec) ctxt = |
18781 | 226 |
let |
227 |
val (x, mx) = var; |
|
228 |
val ((name, atts), rhs) = spec; |
|
229 |
val name' = if name = "" then Thm.def_name x else name; |
|
18876 | 230 |
val rhs_frees = Term.add_frees rhs []; |
18781 | 231 |
in |
232 |
ctxt |
|
18876 | 233 |
|> consts_restricted (member (op =) rhs_frees) [((x, Term.fastype_of rhs), mx)] |
18781 | 234 |
|-> (fn [lhs] => |
235 |
theory_result (add_def (name', Logic.mk_equals (lhs, rhs))) |
|
236 |
#-> (fn th => fn ctxt' => note ((name', atts), [finish ctxt' lhs th]) ctxt') |
|
237 |
#> apfst (fn (b, [th]) => (lhs, (b, th)))) |
|
238 |
end; |
|
239 |
||
18805
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
240 |
val def = def_finish (K (K I)); |
18863b33c831
improved 'notes', including proper treatment of locale results;
wenzelm
parents:
18781
diff
changeset
|
241 |
|
18781 | 242 |
end; |
243 |
||
244 |
end; |