author | wenzelm |
Thu, 17 Aug 2000 10:33:37 +0200 | |
changeset 9619 | 6125cc9efc18 |
parent 9566 | 0874bf3a909d |
child 9733 | 99fda46926cc |
permissions | -rw-r--r-- |
5819 | 1 |
(* Title: Pure/Isar/proof_context.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
8807 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
5819 | 5 |
|
6 |
Proof context information. |
|
7 |
*) |
|
8 |
||
9 |
signature PROOF_CONTEXT = |
|
10 |
sig |
|
11 |
type context |
|
9470 | 12 |
type exporter |
5819 | 13 |
exception CONTEXT of string * context |
14 |
val theory_of: context -> theory |
|
15 |
val sign_of: context -> Sign.sg |
|
7557 | 16 |
val prems_of: context -> thm list |
6985 | 17 |
val show_hyps: bool ref |
18 |
val pretty_thm: thm -> Pretty.T |
|
6528 | 19 |
val verbose: bool ref |
9515 | 20 |
val setmp_verbose: ('a -> 'b) -> 'a -> 'b |
5819 | 21 |
val print_binds: context -> unit |
22 |
val print_thms: context -> unit |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
23 |
val print_cases: context -> unit |
8462 | 24 |
val pretty_prems: context -> Pretty.T list |
25 |
val pretty_context: context -> Pretty.T list |
|
5819 | 26 |
val print_proof_data: theory -> unit |
5874 | 27 |
val init: theory -> context |
9470 | 28 |
val assumptions: context -> (cterm list * exporter) list |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
29 |
val fixed_names: context -> string list |
5819 | 30 |
val read_typ: context -> string -> typ |
9504 | 31 |
val read_typ_no_norm: context -> string -> typ |
5819 | 32 |
val cert_typ: context -> typ -> typ |
9504 | 33 |
val cert_typ_no_norm: context -> typ -> typ |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
34 |
val intern_skolem: context -> term -> term |
9133 | 35 |
val extern_skolem: context -> term -> term |
5874 | 36 |
val read_termTs: context -> (string * typ) list -> term list * (indexname * typ) list |
5819 | 37 |
val read_term: context -> string -> term |
38 |
val read_prop: context -> string -> term |
|
8096 | 39 |
val read_termT_pats: context -> (string * typ) list -> term list |
40 |
val read_term_pats: typ -> context -> string list -> term list |
|
41 |
val read_prop_pats: context -> string list -> term list |
|
5819 | 42 |
val cert_term: context -> term -> term |
43 |
val cert_prop: context -> term -> term |
|
8096 | 44 |
val cert_term_pats: typ -> context -> term list -> term list |
45 |
val cert_prop_pats: context -> term list -> term list |
|
5819 | 46 |
val declare_term: term -> context -> context |
47 |
val declare_terms: term list -> context -> context |
|
7925 | 48 |
val warn_extra_tfrees: context -> context -> context |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
49 |
val generalizeT: context -> context -> typ -> typ |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
50 |
val generalize: context -> context -> term -> term |
9470 | 51 |
val find_free: term -> string -> term option |
9553 | 52 |
val norm_hhf: thm -> thm |
9470 | 53 |
val export_wrt: bool -> context -> context option -> (thm -> thm Seq.seq) * (int -> tactic) list |
7925 | 54 |
val auto_bind_goal: term -> context -> context |
55 |
val auto_bind_facts: string -> term list -> context -> context |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
56 |
val match_bind: bool -> (string list * string) list -> context -> context |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
57 |
val match_bind_i: bool -> (term list * term) list -> context -> context |
8096 | 58 |
val read_propp: context * (string * (string list * string list)) |
59 |
-> context * (term * (term list * term list)) |
|
60 |
val cert_propp: context * (term * (term list * term list)) |
|
61 |
-> context * (term * (term list * term list)) |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
62 |
val bind_propp: context * (string * (string list * string list)) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
63 |
-> context * (term * (context -> context)) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
64 |
val bind_propp_i: context * (term * (term list * term list)) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
65 |
-> context * (term * (context -> context)) |
6091 | 66 |
val get_thm: context -> string -> thm |
9566 | 67 |
val get_thm_closure: context -> string -> thm |
6091 | 68 |
val get_thms: context -> string -> thm list |
9566 | 69 |
val get_thms_closure: context -> string -> thm list |
6091 | 70 |
val put_thm: string * thm -> context -> context |
71 |
val put_thms: string * thm list -> context -> context |
|
72 |
val put_thmss: (string * thm list) list -> context -> context |
|
7606 | 73 |
val reset_thms: string -> context -> context |
9196 | 74 |
val have_thmss: |
75 |
((string * context attribute list) * (thm list * context attribute list) list) list -> |
|
76 |
context -> context * (string * thm list) list |
|
9470 | 77 |
val assume: exporter |
7270 | 78 |
-> (string * context attribute list * (string * (string list * string list)) list) list |
79 |
-> context -> context * ((string * thm list) list * thm list) |
|
9470 | 80 |
val assume_i: exporter |
7270 | 81 |
-> (string * context attribute list * (term * (term list * term list)) list) list |
82 |
-> context -> context * ((string * thm list) list * thm list) |
|
8096 | 83 |
val read_vars: context * (string list * string option) -> context * (string list * typ option) |
84 |
val cert_vars: context * (string list * typ option) -> context * (string list * typ option) |
|
7411 | 85 |
val fix: (string list * string option) list -> context -> context |
7663 | 86 |
val fix_i: (string list * typ option) list -> context -> context |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
87 |
val bind_skolem: context -> string list -> term -> term |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
88 |
val get_case: context -> string -> RuleCases.T |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
89 |
val add_cases: (string * RuleCases.T) list -> context -> context |
5819 | 90 |
val setup: (theory -> theory) list |
91 |
end; |
|
92 |
||
8151 | 93 |
signature PRIVATE_PROOF_CONTEXT = |
5819 | 94 |
sig |
95 |
include PROOF_CONTEXT |
|
96 |
val init_data: Object.kind -> (theory -> Object.T) * (context -> Object.T -> unit) |
|
97 |
-> theory -> theory |
|
98 |
val print_data: Object.kind -> context -> unit |
|
99 |
val get_data: Object.kind -> (Object.T -> 'a) -> context -> 'a |
|
100 |
val put_data: Object.kind -> ('a -> Object.T) -> 'a -> context -> context |
|
101 |
end; |
|
102 |
||
8151 | 103 |
structure ProofContext: PRIVATE_PROOF_CONTEXT = |
5819 | 104 |
struct |
105 |
||
106 |
||
107 |
(** datatype context **) |
|
108 |
||
9470 | 109 |
type exporter = |
110 |
(cterm list -> thm -> thm Seq.seq) * (bool -> int -> (int -> tactic) list); |
|
111 |
||
5819 | 112 |
datatype context = |
113 |
Context of |
|
6847
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
114 |
{thy: theory, (*current theory*) |
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
115 |
data: Object.T Symtab.table, (*user data*) |
5819 | 116 |
asms: |
9470 | 117 |
((cterm list * exporter) list * (*assumes: A ==> _*) |
6931 | 118 |
(string * thm list) list) * |
6847
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
119 |
((string * string) list * string list), (*fixes: !!x. _*) |
7606 | 120 |
binds: (term * typ) option Vartab.table, (*term bindings*) |
121 |
thms: thm list option Symtab.table, (*local thms*) |
|
8426 | 122 |
cases: (string * RuleCases.T) list, (*local contexts*) |
5819 | 123 |
defs: |
6847
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
124 |
typ Vartab.table * (*type constraints*) |
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
125 |
sort Vartab.table * (*default sorts*) |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
126 |
(string list * term list Symtab.table)}; (*used type variables*) |
5819 | 127 |
|
128 |
exception CONTEXT of string * context; |
|
129 |
||
130 |
||
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
131 |
fun make_context (thy, data, asms, binds, thms, cases, defs) = |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
132 |
Context {thy = thy, data = data, asms = asms, binds = binds, thms = thms, |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
133 |
cases = cases, defs = defs}; |
5819 | 134 |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
135 |
fun map_context f (Context {thy, data, asms, binds, thms, cases, defs}) = |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
136 |
make_context (f (thy, data, asms, binds, thms, cases, defs)); |
5819 | 137 |
|
138 |
fun theory_of (Context {thy, ...}) = thy; |
|
139 |
val sign_of = Theory.sign_of o theory_of; |
|
140 |
||
7270 | 141 |
fun prems_of (Context {asms = ((_, asms), _), ...}) = flat (map #2 asms); |
142 |
||
5819 | 143 |
|
144 |
||
145 |
(** print context information **) |
|
146 |
||
6985 | 147 |
val show_hyps = ref false; |
7575 | 148 |
|
149 |
fun pretty_thm thm = |
|
150 |
if ! show_hyps then setmp Display.show_hyps true Display.pretty_thm_no_quote thm |
|
151 |
else Display.pretty_cterm (#prop (Thm.crep_thm thm)); |
|
6985 | 152 |
|
6528 | 153 |
val verbose = ref false; |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
154 |
fun verb f x = if ! verbose then f (x ()) else []; |
8462 | 155 |
fun verb_single x = verb Library.single x; |
5819 | 156 |
|
9515 | 157 |
fun setmp_verbose f x = Library.setmp verbose true f x; |
158 |
||
8462 | 159 |
fun pretty_items prt name items = |
5819 | 160 |
let |
8462 | 161 |
fun prt_itms (name, [x]) = Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, prt x] |
162 |
| prt_itms (name, xs) = Pretty.big_list (name ^ ":") (map prt xs); |
|
6721 | 163 |
in |
6847
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
164 |
if null items andalso not (! verbose) then [] |
8462 | 165 |
else [Pretty.big_list name (map prt_itms items)] |
6721 | 166 |
end; |
5819 | 167 |
|
168 |
||
169 |
(* term bindings *) |
|
170 |
||
7606 | 171 |
val smash_option = fn (_, None) => None | (xi, Some b) => Some (xi, b); |
172 |
||
8462 | 173 |
fun pretty_binds (ctxt as Context {binds, ...}) = |
5819 | 174 |
let |
7575 | 175 |
val prt_term = Sign.pretty_term (sign_of ctxt); |
8462 | 176 |
fun prt_bind (xi, (t, T)) = prt_term (Logic.mk_equals (Var (xi, T), t)); |
7606 | 177 |
val bs = mapfilter smash_option (Vartab.dest binds); |
6721 | 178 |
in |
7606 | 179 |
if null bs andalso not (! verbose) then [] |
8462 | 180 |
else [Pretty.big_list "term bindings:" (map prt_bind bs)] |
6721 | 181 |
end; |
5819 | 182 |
|
8462 | 183 |
val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds; |
6847
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
184 |
|
5819 | 185 |
|
186 |
(* local theorems *) |
|
187 |
||
8462 | 188 |
fun pretty_thms (Context {thms, ...}) = |
9274 | 189 |
pretty_items pretty_thm "facts:" (mapfilter smash_option (Symtab.dest thms)); |
6847
f175f56c57a6
tuned output: print_context replaced by strings_of_context;
wenzelm
parents:
6797
diff
changeset
|
190 |
|
8462 | 191 |
val print_thms = Pretty.writeln o Pretty.chunks o pretty_thms; |
5819 | 192 |
|
193 |
||
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
194 |
(* local contexts *) |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
195 |
|
8462 | 196 |
fun pretty_cases (ctxt as Context {cases, ...}) = |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
197 |
let |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
198 |
val prt_term = Sign.pretty_term (sign_of ctxt); |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
199 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
200 |
fun prt_sect _ _ [] = [] |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
201 |
| prt_sect s prt xs = [Pretty.block (Pretty.breaks (Pretty.str s :: map prt xs))]; |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
202 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
203 |
fun prt_case (name, (xs, ts)) = Pretty.block (Pretty.breaks |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
204 |
(Pretty.str (name ^ ":") :: |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
205 |
prt_sect "fix" (prt_term o Free) xs @ |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
206 |
prt_sect "assume" (Pretty.quote o prt_term) ts)); |
8426 | 207 |
|
208 |
val cases' = rev (Library.gen_distinct Library.eq_fst cases); |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
209 |
in |
8426 | 210 |
if null cases andalso not (! verbose) then [] |
8462 | 211 |
else [Pretty.big_list "cases:" (map prt_case cases')] |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
212 |
end; |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
213 |
|
8462 | 214 |
val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases; |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
215 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
216 |
|
5819 | 217 |
(* main context *) |
218 |
||
8462 | 219 |
fun pretty_prems ctxt = |
7575 | 220 |
(case prems_of ctxt of |
221 |
[] => [] |
|
8462 | 222 |
| prems => [Pretty.big_list "prems:" (map pretty_thm prems)]); |
7575 | 223 |
|
8462 | 224 |
fun pretty_context (ctxt as Context {asms = (_, (fixes, _)), cases, |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
225 |
defs = (types, sorts, (used, _)), ...}) = |
5819 | 226 |
let |
7575 | 227 |
val sign = sign_of ctxt; |
5819 | 228 |
|
229 |
val prt_term = Sign.pretty_term sign; |
|
230 |
val prt_typ = Sign.pretty_typ sign; |
|
231 |
val prt_sort = Sign.pretty_sort sign; |
|
232 |
||
233 |
(*theory*) |
|
234 |
val pretty_thy = Pretty.block [Pretty.str "Theory:", Pretty.brk 1, Sign.pretty_sg sign]; |
|
235 |
||
236 |
(*fixes*) |
|
7606 | 237 |
fun prt_fix (x, x') = Pretty.block |
238 |
[prt_term (Syntax.free x), Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')]; |
|
239 |
||
9515 | 240 |
fun prt_fixes [] = [] |
241 |
| prt_fixes xs = [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 :: |
|
242 |
Pretty.commas (map prt_fix xs))]; |
|
7200 | 243 |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
244 |
(*defaults*) |
5819 | 245 |
fun prt_atom prt prtT (x, X) = Pretty.block |
246 |
[prt x, Pretty.str " ::", Pretty.brk 1, prtT X]; |
|
247 |
||
248 |
fun prt_var (x, ~1) = prt_term (Syntax.free x) |
|
249 |
| prt_var xi = prt_term (Syntax.var xi); |
|
250 |
||
251 |
fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) |
|
252 |
| prt_varT xi = prt_typ (TVar (xi, [])); |
|
253 |
||
254 |
val prt_defT = prt_atom prt_var prt_typ; |
|
255 |
val prt_defS = prt_atom prt_varT prt_sort; |
|
256 |
in |
|
8462 | 257 |
verb_single (K pretty_thy) @ |
9515 | 258 |
prt_fixes (rev (filter_out (can Syntax.dest_internal o #1) fixes)) @ |
8462 | 259 |
pretty_prems ctxt @ |
260 |
verb pretty_binds (K ctxt) @ |
|
261 |
verb pretty_thms (K ctxt) @ |
|
262 |
verb pretty_cases (K ctxt) @ |
|
263 |
verb_single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @ |
|
264 |
verb_single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts))) @ |
|
265 |
verb_single (fn () => Pretty.strs ("used type variable names:" :: used)) |
|
5819 | 266 |
end; |
267 |
||
268 |
||
269 |
||
270 |
(** user data **) |
|
271 |
||
272 |
(* errors *) |
|
273 |
||
274 |
fun of_theory thy = "\nof theory " ^ Sign.str_of_sg (Theory.sign_of thy); |
|
275 |
||
276 |
fun err_inconsistent kinds = |
|
277 |
error ("Attempt to merge different versions of " ^ commas_quote kinds ^ " proof data"); |
|
278 |
||
279 |
fun err_dup_init thy kind = |
|
280 |
error ("Duplicate initialization of " ^ quote kind ^ " proof data" ^ of_theory thy); |
|
281 |
||
282 |
fun err_undef ctxt kind = |
|
283 |
raise CONTEXT ("Tried to access undefined " ^ quote kind ^ " proof data", ctxt); |
|
284 |
||
285 |
fun err_uninit ctxt kind = |
|
286 |
raise CONTEXT ("Tried to access uninitialized " ^ quote kind ^ " proof data" ^ |
|
287 |
of_theory (theory_of ctxt), ctxt); |
|
288 |
||
289 |
fun err_access ctxt kind = |
|
290 |
raise CONTEXT ("Unauthorized access to " ^ quote kind ^ " proof data" ^ |
|
291 |
of_theory (theory_of ctxt), ctxt); |
|
292 |
||
293 |
||
294 |
(* data kind 'Isar/proof_data' *) |
|
295 |
||
296 |
structure ProofDataDataArgs = |
|
297 |
struct |
|
298 |
val name = "Isar/proof_data"; |
|
299 |
type T = (Object.kind * ((theory -> Object.T) * (context -> Object.T -> unit))) Symtab.table; |
|
300 |
||
301 |
val empty = Symtab.empty; |
|
6550 | 302 |
val copy = I; |
5819 | 303 |
val prep_ext = I; |
304 |
fun merge tabs = Symtab.merge (Object.eq_kind o pairself fst) tabs |
|
305 |
handle Symtab.DUPS kinds => err_inconsistent kinds; |
|
306 |
fun print _ tab = Pretty.writeln (Pretty.strs (map #1 (Symtab.dest tab))); |
|
307 |
end; |
|
308 |
||
309 |
structure ProofDataData = TheoryDataFun(ProofDataDataArgs); |
|
310 |
val print_proof_data = ProofDataData.print; |
|
311 |
||
312 |
||
313 |
(* init proof data *) |
|
314 |
||
315 |
fun init_data kind meths thy = |
|
316 |
let |
|
317 |
val name = Object.name_of_kind kind; |
|
318 |
val tab = Symtab.update_new ((name, (kind, meths)), ProofDataData.get thy) |
|
319 |
handle Symtab.DUP _ => err_dup_init thy name; |
|
320 |
in thy |> ProofDataData.put tab end; |
|
321 |
||
322 |
||
323 |
(* access data *) |
|
324 |
||
325 |
fun lookup_data (ctxt as Context {data, ...}) kind = |
|
326 |
let |
|
327 |
val thy = theory_of ctxt; |
|
328 |
val name = Object.name_of_kind kind; |
|
329 |
in |
|
330 |
(case Symtab.lookup (ProofDataData.get thy, name) of |
|
331 |
Some (k, meths) => |
|
332 |
if Object.eq_kind (kind, k) then |
|
333 |
(case Symtab.lookup (data, name) of |
|
334 |
Some x => (x, meths) |
|
335 |
| None => err_undef ctxt name) |
|
336 |
else err_access ctxt name |
|
337 |
| None => err_uninit ctxt name) |
|
338 |
end; |
|
339 |
||
340 |
fun get_data kind f ctxt = |
|
341 |
let val (x, _) = lookup_data ctxt kind |
|
342 |
in f x handle Match => Object.kind_error kind end; |
|
343 |
||
344 |
fun print_data kind ctxt = |
|
345 |
let val (x, (_, prt)) = lookup_data ctxt kind |
|
346 |
in prt ctxt x end; |
|
347 |
||
348 |
fun put_data kind f x ctxt = |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
349 |
(lookup_data ctxt kind; |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
350 |
ctxt |> map_context (fn (thy, data, asms, binds, thms, cases, defs) => |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
351 |
(thy, Symtab.update ((Object.name_of_kind kind, f x), data), asms, binds, thms, cases, defs))); |
5819 | 352 |
|
353 |
||
354 |
(* init context *) |
|
355 |
||
5874 | 356 |
fun init thy = |
357 |
let val data = Symtab.map (fn (_, (f, _)) => f thy) (ProofDataData.get thy) in |
|
8426 | 358 |
make_context (thy, data, (([], []), ([], [])), Vartab.empty, Symtab.empty, [], |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
359 |
(Vartab.empty, Vartab.empty, ([], Symtab.empty))) |
5819 | 360 |
end; |
361 |
||
362 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
363 |
(* get assumptions *) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
364 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
365 |
fun assumptions (Context {asms = ((asms, _), _), ...}) = asms; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
366 |
fun fixed_names (Context {asms = (_, (fixes, _)), ...}) = map #2 fixes; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
367 |
fun used_table (Context {defs = (_, _, (_, tab)), ...}) = tab; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
368 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
369 |
|
8096 | 370 |
|
7663 | 371 |
(** default sorts and types **) |
372 |
||
7670 | 373 |
fun def_sort (Context {defs = (_, sorts, _), ...}) xi = Vartab.lookup (sorts, xi); |
7663 | 374 |
|
7670 | 375 |
fun def_type (Context {binds, defs = (types, _, _), ...}) is_pat xi = |
7663 | 376 |
(case Vartab.lookup (types, xi) of |
377 |
None => |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
378 |
if is_pat then None else |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
379 |
(case Vartab.lookup (binds, xi) of |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
380 |
Some (Some (_, T)) => Some (TypeInfer.polymorphicT T) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
381 |
| _ => None) |
7663 | 382 |
| some => some); |
383 |
||
384 |
||
5819 | 385 |
|
386 |
(** prepare types **) |
|
387 |
||
9504 | 388 |
local |
389 |
||
390 |
fun read_typ_aux read ctxt s = |
|
391 |
transform_error (read (sign_of ctxt, def_sort ctxt)) s |
|
7663 | 392 |
handle ERROR_MESSAGE msg => raise CONTEXT (msg, ctxt); |
5819 | 393 |
|
9504 | 394 |
fun cert_typ_aux cert ctxt raw_T = cert (sign_of ctxt) raw_T |
395 |
handle TYPE (msg, _, _) => raise CONTEXT (msg, ctxt); |
|
396 |
||
397 |
in |
|
398 |
||
399 |
val read_typ = read_typ_aux Sign.read_typ; |
|
400 |
val read_typ_no_norm = read_typ_aux Sign.read_typ_no_norm; |
|
401 |
val cert_typ = cert_typ_aux Sign.certify_typ; |
|
402 |
val cert_typ_no_norm = cert_typ_aux Sign.certify_typ_no_norm; |
|
403 |
||
404 |
end; |
|
405 |
||
5819 | 406 |
|
407 |
||
7679 | 408 |
(* internalize Skolem constants *) |
409 |
||
9133 | 410 |
fun fixes_of (Context {asms = (_, (fixes, _)), ...}) = fixes; |
411 |
fun get_skolem ctxt x = assoc (fixes_of ctxt, x); |
|
7679 | 412 |
|
9325 | 413 |
fun no_skolem no_internal ctxt x = |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
414 |
if can Syntax.dest_skolem x then |
7679 | 415 |
raise CONTEXT ("Illegal reference to internal Skolem constant: " ^ quote x, ctxt) |
9325 | 416 |
else if no_internal andalso can Syntax.dest_internal x then |
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
417 |
raise CONTEXT ("Illegal reference to internal variable: " ^ quote x, ctxt) |
7679 | 418 |
else x; |
419 |
||
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
420 |
fun intern_skolem ctxt = |
7679 | 421 |
let |
422 |
fun intern (t as Free (x, T)) = |
|
9325 | 423 |
(case get_skolem ctxt (no_skolem true ctxt x) of |
7679 | 424 |
Some x' => Free (x', T) |
425 |
| None => t) |
|
426 |
| intern (t $ u) = intern t $ intern u |
|
427 |
| intern (Abs (x, T, t)) = Abs (x, T, intern t) |
|
428 |
| intern a = a; |
|
429 |
in intern end; |
|
430 |
||
431 |
||
9133 | 432 |
(* externalize Skolem constants -- for printing purposes only *) |
433 |
||
434 |
fun extern_skolem ctxt = |
|
435 |
let |
|
436 |
val rev_fixes = map Library.swap (fixes_of ctxt); |
|
437 |
||
438 |
fun extern (t as Free (x, T)) = |
|
439 |
(case assoc (rev_fixes, x) of |
|
440 |
Some x' => Free (if get_skolem ctxt x' = Some x then x' else NameSpace.hidden x', T) |
|
441 |
| None => t) |
|
442 |
| extern (t $ u) = extern t $ extern u |
|
443 |
| extern (Abs (x, T, t)) = Abs (x, T, extern t) |
|
444 |
| extern a = a; |
|
445 |
in extern end |
|
446 |
||
8096 | 447 |
|
5819 | 448 |
(** prepare terms and propositions **) |
449 |
||
450 |
(* |
|
451 |
(1) read / certify wrt. signature of context |
|
452 |
(2) intern Skolem constants |
|
453 |
(3) expand term bindings |
|
454 |
*) |
|
455 |
||
456 |
||
457 |
(* read / certify wrt. signature *) (*exception ERROR*) (*exception TERM*) |
|
458 |
||
5874 | 459 |
fun read_def_termTs freeze sg (types, sorts, used) sTs = |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
460 |
Sign.read_def_terms (sg, types, sorts) used freeze sTs; |
5874 | 461 |
|
462 |
fun read_def_termT freeze sg defs sT = apfst hd (read_def_termTs freeze sg defs [sT]); |
|
463 |
||
8096 | 464 |
fun read_term_sg freeze sg defs s = #1 (read_def_termT freeze sg defs (s, TypeInfer.logicT)); |
465 |
fun read_prop_sg freeze sg defs s = #1 (read_def_termT freeze sg defs (s, propT)); |
|
5819 | 466 |
|
467 |
||
468 |
fun cert_term_sg sg t = Thm.term_of (Thm.cterm_of sg t); |
|
469 |
||
470 |
fun cert_prop_sg sg tm = |
|
471 |
let |
|
472 |
val ctm = Thm.cterm_of sg tm; |
|
473 |
val {t, T, ...} = Thm.rep_cterm ctm; |
|
474 |
in |
|
475 |
if T = propT then t |
|
476 |
else raise TERM ("Term not of type prop", [t]) |
|
477 |
end; |
|
478 |
||
479 |
||
480 |
(* norm_term *) |
|
481 |
||
482 |
(*beta normal form for terms (not eta normal form), chase variables in |
|
483 |
bindings environment (code taken from Pure/envir.ML)*) |
|
484 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
485 |
fun unifyT ctxt (T, U) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
486 |
let val maxidx = Int.max (Term.maxidx_of_typ T, Term.maxidx_of_typ U) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
487 |
in #1 (Type.unify (Sign.tsig_of (sign_of ctxt)) maxidx Vartab.empty (T, U)) end; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
488 |
|
5819 | 489 |
fun norm_term (ctxt as Context {binds, ...}) = |
490 |
let |
|
491 |
(*raised when norm has no effect on a term, to do sharing instead of copying*) |
|
492 |
exception SAME; |
|
493 |
||
494 |
fun norm (t as Var (xi, T)) = |
|
495 |
(case Vartab.lookup (binds, xi) of |
|
7606 | 496 |
Some (Some (u, U)) => |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
497 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
498 |
val env = unifyT ctxt (T, U) handle Type.TUNIFY => |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
499 |
raise TYPE ("norm_term: ill-typed variable assigment", [T, U], [t, u]); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
500 |
val u' = Term.subst_TVars_Vartab env u; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
501 |
in norm u' handle SAME => u' end |
7606 | 502 |
| _ => raise CONTEXT ("Unbound schematic variable: " ^ Syntax.string_of_vname xi, ctxt)) |
5819 | 503 |
| norm (Abs (a, T, body)) = Abs (a, T, norm body) |
504 |
| norm (Abs (_, _, body) $ t) = normh (subst_bound (t, body)) |
|
505 |
| norm (f $ t) = |
|
506 |
((case norm f of |
|
507 |
Abs (_, _, body) => normh (subst_bound (t, body)) |
|
508 |
| nf => nf $ (norm t handle SAME => t)) handle SAME => f $ norm t) |
|
509 |
| norm _ = raise SAME |
|
510 |
and normh t = norm t handle SAME => t |
|
511 |
in normh end; |
|
512 |
||
513 |
||
6550 | 514 |
(* dummy patterns *) |
515 |
||
9540 | 516 |
fun prepare_dummies t = #2 (Term.replace_dummy_patterns (1, t)); |
6762 | 517 |
|
9540 | 518 |
fun reject_dummies ctxt t = Term.no_dummy_patterns t |
519 |
handle TERM _ => raise CONTEXT ("Illegal dummy pattern(s) in term", ctxt); |
|
6550 | 520 |
|
521 |
||
5819 | 522 |
(* read terms *) |
523 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
524 |
fun gen_read read app is_pat (ctxt as Context {defs = (_, _, (used, _)), ...}) s = |
7663 | 525 |
(transform_error (read (sign_of ctxt) (def_type ctxt is_pat, def_sort ctxt, used)) s |
526 |
handle TERM (msg, _) => raise CONTEXT (msg, ctxt) |
|
527 |
| ERROR_MESSAGE msg => raise CONTEXT (msg, ctxt)) |
|
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
528 |
|> app (intern_skolem ctxt) |
7663 | 529 |
|> app (if is_pat then I else norm_term ctxt) |
530 |
|> app (if is_pat then prepare_dummies else (reject_dummies ctxt)); |
|
5819 | 531 |
|
5874 | 532 |
val read_termTs = gen_read (read_def_termTs false) (apfst o map) false; |
8096 | 533 |
val read_termT_pats = #1 oo gen_read (read_def_termTs false) (apfst o map) true; |
534 |
||
535 |
fun read_term_pats T ctxt pats = read_termT_pats ctxt (map (rpair T) pats); |
|
536 |
val read_prop_pats = read_term_pats propT; |
|
537 |
||
6762 | 538 |
val read_term = gen_read (read_term_sg true) I false; |
539 |
val read_prop = gen_read (read_prop_sg true) I false; |
|
5819 | 540 |
|
541 |
||
542 |
(* certify terms *) |
|
543 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
544 |
fun gen_cert cert is_pat ctxt t = t |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
545 |
|> (if is_pat then I else norm_term ctxt) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
546 |
|> (fn t' => cert (sign_of ctxt) t' handle TERM (msg, _) => raise CONTEXT (msg, ctxt)); |
5819 | 547 |
|
548 |
val cert_term = gen_cert cert_term_sg false; |
|
549 |
val cert_prop = gen_cert cert_prop_sg false; |
|
8096 | 550 |
|
551 |
fun cert_term_pats _ = map o gen_cert cert_term_sg true; |
|
552 |
val cert_prop_pats = map o gen_cert cert_prop_sg true; |
|
5819 | 553 |
|
554 |
||
555 |
(* declare terms *) |
|
556 |
||
557 |
val ins_types = foldl_aterms |
|
558 |
(fn (types, Free (x, T)) => Vartab.update (((x, ~1), T), types) |
|
559 |
| (types, Var v) => Vartab.update (v, types) |
|
560 |
| (types, _) => types); |
|
561 |
||
562 |
val ins_sorts = foldl_types (foldl_atyps |
|
563 |
(fn (sorts, TFree (x, S)) => Vartab.update (((x, ~1), S), sorts) |
|
564 |
| (sorts, TVar v) => Vartab.update (v, sorts) |
|
565 |
| (sorts, _) => sorts)); |
|
566 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
567 |
val ins_used = foldl_term_types (fn t => foldl_atyps |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
568 |
(fn ((used, tab), TFree (x, _)) => (x ins used, Symtab.update_multi ((x, t), tab)) |
5819 | 569 |
| (used, _) => used)); |
570 |
||
7663 | 571 |
fun ins_skolem def_ty = foldr |
5994 | 572 |
(fn ((x, x'), types) => |
7663 | 573 |
(case def_ty x' of |
5994 | 574 |
Some T => Vartab.update (((x, ~1), T), types) |
575 |
| None => types)); |
|
576 |
||
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
577 |
fun map_defaults f = map_context (fn (thy, data, asms, binds, thms, cases, defs) => |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
578 |
(thy, data, asms, binds, thms, cases, f defs)); |
5819 | 579 |
|
5994 | 580 |
fun declare (ctxt as Context {asms = (_, (fixes, _)), ...}, t) = |
5819 | 581 |
ctxt |
7670 | 582 |
|> map_defaults (fn (types, sorts, used) => (ins_types (types, t), sorts, used)) |
583 |
|> map_defaults (fn (types, sorts, used) => (types, ins_sorts (sorts, t), used)) |
|
584 |
|> map_defaults (fn (types, sorts, used) => (types, sorts, ins_used (used, t))) |
|
585 |
|> map_defaults (fn (types, sorts, used) => |
|
586 |
(ins_skolem (fn x => Vartab.lookup (types, (x, ~1))) (fixes, types), sorts, used)); |
|
5819 | 587 |
|
588 |
||
589 |
fun declare_term t ctxt = declare (ctxt, t); |
|
590 |
fun declare_terms ts ctxt = foldl declare (ctxt, ts); |
|
591 |
||
592 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
593 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
594 |
(** Hindley-Milner polymorphism **) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
595 |
|
7925 | 596 |
(* warn_extra_tfrees *) |
597 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
598 |
local |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
599 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
600 |
fun used_free (a, ts) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
601 |
(case mapfilter (fn Free (x, _) => Some x | _ => None) ts of |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
602 |
[] => None |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
603 |
| xs => Some (a, xs)); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
604 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
605 |
fun warn_extra (names1: string list, tab1) (names2, tab2) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
606 |
if names1 = names2 then () |
7925 | 607 |
else |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
608 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
609 |
val extra = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
610 |
Library.gen_rems Library.eq_fst (Symtab.dest tab2, Symtab.dest tab1) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
611 |
|> mapfilter used_free; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
612 |
val tfrees = map #1 extra; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
613 |
val frees = Library.sort_strings (Library.distinct (flat (map #2 extra))); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
614 |
in |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
615 |
if null extra then () |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
616 |
else warning ("Danger! Just introduced free type variable(s): " ^ commas tfrees ^ " in " ^ |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
617 |
space_implode " or " frees) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
618 |
end; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
619 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
620 |
in |
7925 | 621 |
|
622 |
fun warn_extra_tfrees (ctxt1 as Context {defs = (_, _, used1), ...}) |
|
623 |
(ctxt2 as Context {defs = (_, _, used2), ...}) = (warn_extra used1 used2; ctxt2); |
|
624 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
625 |
end; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
626 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
627 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
628 |
(* generalize type variables *) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
629 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
630 |
fun gen_tfrees inner opt_outer = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
631 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
632 |
val inner_used = used_table inner; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
633 |
val inner_fixes = fixed_names inner; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
634 |
val (outer_used, outer_fixes) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
635 |
(case opt_outer of |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
636 |
None => (Symtab.empty, []) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
637 |
| Some ctxt => (used_table ctxt, fixed_names ctxt)); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
638 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
639 |
val extra_fixes = inner_fixes \\ outer_fixes; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
640 |
fun still_fixed (Free (x, _)) = not (x mem_string extra_fixes) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
641 |
| still_fixed _ = false; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
642 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
643 |
fun add (gen, (a, xs)) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
644 |
if is_some (Symtab.lookup (outer_used, a)) orelse exists still_fixed xs |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
645 |
then gen else a :: gen; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
646 |
in Symtab.foldl add ([], inner_used) end; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
647 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
648 |
fun generalizeT inner outer = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
649 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
650 |
val tfrees = gen_tfrees inner (Some outer); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
651 |
fun gen (x, S) = if x mem_string tfrees then TVar ((x, 0), S) else TFree (x, S); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
652 |
in Term.map_type_tfree gen end; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
653 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
654 |
val generalize = Term.map_term_types oo generalizeT; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
655 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
656 |
|
9553 | 657 |
|
658 |
(** export theorems **) |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
659 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
660 |
fun get_free x (None, t as Free (y, _)) = if x = y then Some t else None |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
661 |
| get_free _ (opt, _) = opt; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
662 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
663 |
fun find_free t x = foldl_aterms (get_free x) (None, t); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
664 |
|
9553 | 665 |
val norm_hhf = |
666 |
Drule.forall_elim_vars_safe o Tactic.rewrite_rule [Drule.norm_hhf_eq]; |
|
667 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
668 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
669 |
local |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
670 |
|
9470 | 671 |
fun export tfrees fixes goal_asms thm = |
672 |
thm |
|
9553 | 673 |
|> norm_hhf |
9470 | 674 |
|> Seq.EVERY (rev (map op |> goal_asms)) |
675 |
|> Seq.map (fn rule => |
|
676 |
let |
|
677 |
val {sign, prop, maxidx, ...} = Thm.rep_thm rule; |
|
678 |
val frees = map (Thm.cterm_of sign) (mapfilter (find_free prop) fixes); |
|
679 |
in |
|
680 |
rule |
|
681 |
|> Drule.forall_intr_list frees |
|
9553 | 682 |
|> norm_hhf |
9470 | 683 |
|> Drule.tvars_intr_list tfrees |
684 |
end); |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
685 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
686 |
fun diff_context inner None = (gen_tfrees inner None, fixed_names inner, assumptions inner) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
687 |
| diff_context inner (Some outer) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
688 |
(gen_tfrees inner (Some outer), |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
689 |
fixed_names inner \\ fixed_names outer, |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
690 |
Library.drop (length (assumptions outer), assumptions inner)); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
691 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
692 |
in |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
693 |
|
9470 | 694 |
fun export_wrt is_goal inner opt_outer = |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
695 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
696 |
val (tfrees, fixes, asms) = diff_context inner opt_outer; |
9470 | 697 |
val goal_asms = map (fn (cprops, (exp, _)) => (map Drule.mk_cgoal cprops, exp)) asms; |
698 |
val tacs = flat (map (fn (cprops, (_, f)) => f is_goal (length cprops)) asms); |
|
699 |
in (export tfrees fixes goal_asms, tacs) end; |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
700 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
701 |
end; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
702 |
|
7925 | 703 |
|
5819 | 704 |
|
705 |
(** bindings **) |
|
706 |
||
707 |
(* update_binds *) |
|
708 |
||
7606 | 709 |
fun del_bind (ctxt, xi) = |
710 |
ctxt |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
711 |
|> map_context (fn (thy, data, asms, binds, thms, cases, defs) => |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
712 |
(thy, data, asms, Vartab.update ((xi, None), binds), thms, cases, defs)); |
7606 | 713 |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
714 |
fun upd_bind (ctxt, ((x, i), t)) = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
715 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
716 |
val T = Term.fastype_of t; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
717 |
val t' = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
718 |
if null (Term.term_tvars t \\ Term.typ_tvars T) then t |
8637 | 719 |
else Var ((x ^ "_has_extra_type_vars_on_rhs", i), T); |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
720 |
in |
5819 | 721 |
ctxt |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
722 |
|> declare_term t' |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
723 |
|> map_context (fn (thy, data, asms, binds, thms, cases, defs) => |
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
724 |
(thy, data, asms, Vartab.update (((x, i), Some (t', T)), binds), thms, cases, defs)) |
5819 | 725 |
end; |
726 |
||
7606 | 727 |
fun del_upd_bind (ctxt, (xi, None)) = del_bind (ctxt, xi) |
728 |
| del_upd_bind (ctxt, (xi, Some t)) = upd_bind (ctxt, (xi, t)); |
|
729 |
||
5819 | 730 |
fun update_binds bs ctxt = foldl upd_bind (ctxt, bs); |
7606 | 731 |
fun delete_update_binds bs ctxt = foldl del_upd_bind (ctxt, bs); |
732 |
||
5819 | 733 |
|
8096 | 734 |
(* simult_matches *) |
735 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
736 |
fun simult_matches ctxt [] = [] |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
737 |
| simult_matches ctxt pairs = |
8096 | 738 |
let |
739 |
val maxidx = foldl (fn (i, (t1, t2)) => |
|
740 |
Int.max (i, Int.max (Term.maxidx_of_term t1, Term.maxidx_of_term t2))) (~1, pairs); |
|
741 |
val envs = Unify.smash_unifiers (sign_of ctxt, Envir.empty maxidx, pairs); |
|
742 |
val env = |
|
743 |
(case Seq.pull envs of |
|
744 |
None => raise CONTEXT ("Pattern match failed!", ctxt) (* FIXME improve msg (!?) *) |
|
745 |
| Some (env, _) => env); |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
746 |
val binds = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
747 |
(*Note: Envir.norm_term ensures proper type instantiation*) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
748 |
map (apsnd (Envir.norm_term env)) (Envir.alist_of env); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
749 |
in binds end; |
8096 | 750 |
|
751 |
||
752 |
(* add_binds(_i) *) |
|
5819 | 753 |
|
7925 | 754 |
local |
755 |
||
5819 | 756 |
fun gen_bind prep (ctxt, (xi as (x, _), raw_t)) = |
7606 | 757 |
ctxt |> delete_update_binds [(xi, apsome (prep ctxt) raw_t)]; |
5819 | 758 |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
759 |
fun gen_binds prep binds ctxt = foldl (gen_bind prep) (ctxt, binds); |
7925 | 760 |
|
761 |
in |
|
5819 | 762 |
|
763 |
val add_binds = gen_binds read_term; |
|
764 |
val add_binds_i = gen_binds cert_term; |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
765 |
|
7925 | 766 |
val auto_bind_goal = add_binds_i o AutoBind.goal; |
767 |
val auto_bind_facts = add_binds_i oo AutoBind.facts; |
|
768 |
||
769 |
end; |
|
5819 | 770 |
|
771 |
||
8096 | 772 |
(* match_bind(_i) *) |
5819 | 773 |
|
8096 | 774 |
local |
775 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
776 |
fun prep_bind (prep, prep_pats) (ctxt, (raw_pats, raw_t)) = |
5819 | 777 |
let |
5935 | 778 |
val t = prep ctxt raw_t; |
8096 | 779 |
val ctxt' = declare_term t ctxt; |
780 |
val pats = prep_pats (fastype_of t) ctxt' raw_pats; |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
781 |
val binds = simult_matches ctxt' (map (rpair t) pats); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
782 |
in (ctxt', binds) end; |
7670 | 783 |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
784 |
fun gen_binds prepp gen raw_binds ctxt = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
785 |
let |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
786 |
val (ctxt', binds) = apsnd flat (foldl_map (prep_bind prepp) (ctxt, raw_binds)); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
787 |
val binds' = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
788 |
if gen then map (apsnd (generalize ctxt' ctxt)) binds |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
789 |
else binds; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
790 |
val binds'' = map (apsnd Some) binds'; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
791 |
in |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
792 |
warn_extra_tfrees ctxt |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
793 |
(if gen then ctxt (*sic!*) |> declare_terms (map #2 binds') |> add_binds_i binds'' |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
794 |
else ctxt' |> add_binds_i binds'') |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
795 |
end; |
8096 | 796 |
|
797 |
in |
|
5935 | 798 |
|
8096 | 799 |
val match_bind = gen_binds (read_term, read_term_pats); |
800 |
val match_bind_i = gen_binds (cert_term, cert_term_pats); |
|
801 |
||
802 |
end; |
|
5935 | 803 |
|
804 |
||
8096 | 805 |
(* proposition patterns *) |
5935 | 806 |
|
8096 | 807 |
fun prep_propp prep_prop prep_pats (ctxt, (raw_prop, (raw_pats1, raw_pats2))) = |
6931 | 808 |
let |
8096 | 809 |
val prop = prep_prop ctxt raw_prop; |
810 |
val ctxt' = declare_term prop ctxt; |
|
811 |
||
812 |
val pats = prep_pats ctxt' (raw_pats1 @ raw_pats2); (*simultaneous type inference!*) |
|
813 |
val len1 = length raw_pats1; |
|
814 |
in (ctxt', (prop, (take (len1, pats), drop (len1, pats)))) end; |
|
815 |
||
816 |
val read_propp = prep_propp read_prop read_prop_pats; |
|
817 |
val cert_propp = prep_propp cert_prop cert_prop_pats; |
|
5935 | 818 |
|
8096 | 819 |
|
820 |
fun gen_bind_propp prepp (ctxt, propp) = |
|
821 |
let |
|
822 |
val (ctxt', (prop, (pats1, pats2))) = prepp (ctxt, propp); |
|
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
823 |
val pairs = |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
824 |
map (rpair prop) pats1 @ |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
825 |
map (rpair (Logic.strip_imp_concl prop)) pats2; (* FIXME handle params!? *) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
826 |
val binds = simult_matches ctxt' pairs; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
827 |
|
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
828 |
(*note: context evaluated now, binds added later (lazy)*) |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
829 |
val gen = generalize ctxt' ctxt; |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
830 |
fun gen_binds ct = ct |> add_binds_i (map (apsnd (Some o gen)) binds); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
831 |
in (ctxt' |> add_binds_i (map (apsnd Some) binds), (prop, gen_binds)) end; |
8096 | 832 |
|
833 |
val bind_propp = gen_bind_propp read_propp; |
|
834 |
val bind_propp_i = gen_bind_propp cert_propp; |
|
6789 | 835 |
|
836 |
||
5819 | 837 |
|
838 |
(** theorems **) |
|
839 |
||
6091 | 840 |
(* get_thm(s) *) |
5819 | 841 |
|
9566 | 842 |
(*beware of proper order of evaluation!*) |
5819 | 843 |
|
9566 | 844 |
fun retrieve_thms f g (ctxt as Context {thy, thms, ...}) = |
845 |
let |
|
846 |
val sg_ref = Sign.self_ref (Theory.sign_of thy); |
|
847 |
val get_from_thy = f thy; |
|
848 |
in |
|
849 |
fn name => |
|
850 |
(case Symtab.lookup (thms, name) of |
|
851 |
Some (Some ths) => map (Thm.transfer_sg (Sign.deref sg_ref)) ths |
|
852 |
| _ => get_from_thy name) |> g name |
|
853 |
end; |
|
5819 | 854 |
|
9566 | 855 |
val get_thm = retrieve_thms PureThy.get_thms PureThy.single_thm; |
856 |
val get_thm_closure = retrieve_thms PureThy.get_thms_closure PureThy.single_thm; |
|
857 |
val get_thms = retrieve_thms PureThy.get_thms (K I); |
|
858 |
val get_thms_closure = retrieve_thms PureThy.get_thms_closure (K I); |
|
5819 | 859 |
|
860 |
||
6091 | 861 |
(* put_thm(s) *) |
5819 | 862 |
|
7479 | 863 |
fun put_thms ("", _) = I |
864 |
| put_thms (name, ths) = map_context |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
865 |
(fn (thy, data, asms, binds, thms, cases, defs) => |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
866 |
(thy, data, asms, binds, Symtab.update ((name, Some ths), thms), cases, defs)); |
5819 | 867 |
|
6091 | 868 |
fun put_thm (name, th) = put_thms (name, [th]); |
5819 | 869 |
|
6091 | 870 |
fun put_thmss [] ctxt = ctxt |
871 |
| put_thmss (th :: ths) ctxt = ctxt |> put_thms th |> put_thmss ths; |
|
5819 | 872 |
|
873 |
||
7606 | 874 |
(* reset_thms *) |
875 |
||
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
876 |
fun reset_thms name = map_context (fn (thy, data, asms, binds, thms, cases, defs) => |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
877 |
(thy, data, asms, binds, Symtab.update ((name, None), thms), cases, defs)); |
7606 | 878 |
|
879 |
||
6091 | 880 |
(* have_thmss *) |
5819 | 881 |
|
9196 | 882 |
fun have_thss (ctxt, ((name, more_attrs), ths_attrs)) = |
5819 | 883 |
let |
884 |
fun app ((ct, ths), (th, attrs)) = |
|
6091 | 885 |
let val (ct', th') = Thm.applys_attributes ((ct, th), attrs @ more_attrs) |
5819 | 886 |
in (ct', th' :: ths) end |
887 |
val (ctxt', rev_thms) = foldl app ((ctxt, []), ths_attrs); |
|
9196 | 888 |
val thms = flat (rev rev_thms); |
6091 | 889 |
in (ctxt' |> put_thms (name, thms), (name, thms)) end; |
5819 | 890 |
|
9196 | 891 |
fun have_thmss args ctxt = foldl_map have_thss (ctxt, args); |
892 |
||
5819 | 893 |
|
894 |
||
895 |
(** assumptions **) |
|
896 |
||
897 |
(* assume *) |
|
898 |
||
7270 | 899 |
local |
6797 | 900 |
|
9470 | 901 |
fun gen_assm prepp (ctxt, (name, attrs, raw_prop_pats)) = |
5819 | 902 |
let |
8096 | 903 |
val (ctxt', props) = foldl_map prepp (ctxt, raw_prop_pats); |
5819 | 904 |
|
7575 | 905 |
val cprops = map (Thm.cterm_of (sign_of ctxt')) props; |
9553 | 906 |
val asms = map (norm_hhf o Drule.assume_goal) cprops; |
5919 | 907 |
|
908 |
val ths = map (fn th => ([th], [])) asms; |
|
9196 | 909 |
val (ctxt'', [(_, thms)]) = |
5819 | 910 |
ctxt' |
6797 | 911 |
|> auto_bind_facts name props |
9196 | 912 |
|> have_thmss [((name, attrs @ [Drule.tag_assumption]), ths)]; |
9470 | 913 |
in (ctxt'', (cprops, (name, asms), (name, thms))) end; |
5819 | 914 |
|
9470 | 915 |
fun gen_assms prepp exp args ctxt = |
916 |
let |
|
917 |
val (ctxt', results) = foldl_map (gen_assm prepp) (ctxt, args); |
|
918 |
val cprops = flat (map #1 results); |
|
919 |
val asmss = map #2 results; |
|
920 |
val thmss = map #3 results; |
|
921 |
val ctxt'' = |
|
922 |
ctxt' |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
923 |
|> map_context (fn (thy, data, ((asms_ct, asms_th), fixes), binds, thms, cases, defs) => |
9470 | 924 |
(thy, data, ((asms_ct @ [(cprops, exp)], asms_th @ asmss), fixes), binds, thms, |
925 |
cases, defs)); |
|
926 |
in (warn_extra_tfrees ctxt ctxt'', (thmss, prems_of ctxt'')) end; |
|
5819 | 927 |
|
7270 | 928 |
in |
929 |
||
8616
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
930 |
val assume = gen_assms (apsnd #1 o bind_propp); |
90d2fed59be1
support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents:
8462
diff
changeset
|
931 |
val assume_i = gen_assms (apsnd #1 o bind_propp_i); |
7270 | 932 |
|
933 |
end; |
|
5819 | 934 |
|
935 |
||
8096 | 936 |
(* variables *) |
937 |
||
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
938 |
fun prep_vars prep_typ (ctxt, (xs, raw_T)) = |
8096 | 939 |
let |
9325 | 940 |
val _ = (case filter (not o Syntax.is_identifier) (map (no_skolem false ctxt) xs) of |
8096 | 941 |
[] => () | bads => raise CONTEXT ("Bad variable name(s): " ^ commas_quote bads, ctxt)); |
942 |
||
943 |
val opt_T = apsome (prep_typ ctxt) raw_T; |
|
944 |
val T = if_none opt_T TypeInfer.logicT; |
|
945 |
val ctxt' = ctxt |> declare_terms (map (fn x => Free (x, T)) xs); |
|
946 |
in (ctxt', (xs, opt_T)) end; |
|
947 |
||
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
948 |
val read_vars = prep_vars read_typ; |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
949 |
val cert_vars = prep_vars cert_typ; |
8096 | 950 |
|
951 |
||
5819 | 952 |
(* fix *) |
953 |
||
8096 | 954 |
local |
955 |
||
956 |
fun add_vars xs (fixes, names) = |
|
957 |
let |
|
958 |
val xs' = variantlist (xs, names); |
|
959 |
val fixes' = (xs ~~ map Syntax.skolem xs') @ fixes; |
|
960 |
val names' = xs' @ names; |
|
961 |
in (fixes', names') end; |
|
962 |
||
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
963 |
fun map_vars f = map_context (fn (thy, data, (assumes, vars), binds, thms, cases, defs) => |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
964 |
(thy, data, (assumes, f vars), binds, thms, cases, defs)); |
5819 | 965 |
|
8096 | 966 |
fun gen_fix prep raw_vars ctxt = |
967 |
let |
|
968 |
val (ctxt', varss) = foldl_map prep (ctxt, raw_vars); |
|
969 |
val xs = flat (map fst varss); |
|
970 |
in |
|
971 |
(case Library.duplicates xs of |
|
972 |
[] => () |
|
973 |
| dups => raise CONTEXT ("Duplicate variable name(s): " ^ commas_quote dups, ctxt)); |
|
974 |
ctxt' |> map_vars (add_vars xs) |
|
975 |
end; |
|
5819 | 976 |
|
8096 | 977 |
in |
7679 | 978 |
|
8096 | 979 |
val fix = gen_fix read_vars; |
980 |
val fix_i = gen_fix cert_vars; |
|
981 |
||
982 |
end; |
|
5819 | 983 |
|
6895 | 984 |
|
9291
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
985 |
(*Note: improper use may result in variable capture / dynamic scoping!*) |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
986 |
fun bind_skolem ctxt xs = |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
987 |
let |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
988 |
val ctxt' = ctxt |> fix_i [(xs, None)]; |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
989 |
fun bind (t as Free (x, T)) = |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
990 |
if x mem_string xs then |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
991 |
(case get_skolem ctxt' x of Some x' => Free (x', T) | None => t) |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
992 |
else t |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
993 |
| bind (t $ u) = bind t $ bind u |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
994 |
| bind (Abs (x, T, t)) = Abs (x, T, bind t) |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
995 |
| bind a = a; |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
996 |
in bind end; |
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
997 |
|
23705d14be8f
"_i" arguments now expected to have skolems already internalized;
wenzelm
parents:
9274
diff
changeset
|
998 |
|
5819 | 999 |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1000 |
(** cases **) |
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1001 |
|
8403 | 1002 |
fun check_case ctxt name (xs, ts) = |
1003 |
if null (foldr Term.add_typ_tvars (map snd xs, [])) andalso |
|
1004 |
null (foldr Term.add_term_vars (ts, [])) then () |
|
1005 |
else raise CONTEXT ("Illegal schematic variable(s) in case " ^ quote name, ctxt); |
|
1006 |
||
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1007 |
fun get_case (ctxt as Context {cases, ...}) name = |
8426 | 1008 |
(case assoc (cases, name) of |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1009 |
None => raise CONTEXT ("Unknown case: " ^ quote name, ctxt) |
8403 | 1010 |
| Some c => (check_case ctxt name c; c)); |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1011 |
|
8384 | 1012 |
|
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1013 |
fun add_cases xs = map_context (fn (thy, data, asms, binds, thms, cases, defs) => |
8426 | 1014 |
(thy, data, asms, binds, thms, rev (filter_out (equal "" o #1) xs) @ cases, defs)); |
8373
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1015 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1016 |
|
e7237c8fe29e
handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents:
8186
diff
changeset
|
1017 |
|
5819 | 1018 |
(** theory setup **) |
1019 |
||
1020 |
val setup = [ProofDataData.init]; |
|
1021 |
||
1022 |
||
1023 |
end; |