author | blanchet |
Wed, 15 May 2013 17:43:42 +0200 | |
changeset 51998 | f732a674db1b |
parent 51877 | 71052c42edf2 |
child 52110 | 411db77f96f2 |
permissions | -rw-r--r-- |
50263 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_annotate.ML |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
Author: Steffen Juilf Smolka, TU Muenchen |
|
4 |
||
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
50263
diff
changeset
|
5 |
Supplement term with explicit type constraints that show up as |
50263 | 6 |
type annotations when printing the term. |
7 |
*) |
|
8 |
||
50258 | 9 |
signature SLEDGEHAMMER_ANNOTATE = |
10 |
sig |
|
11 |
val annotate_types : Proof.context -> term -> term |
|
12 |
end |
|
13 |
||
14 |
structure Sledgehammer_Annotate : SLEDGEHAMMER_ANNOTATE = |
|
15 |
struct |
|
16 |
||
17 |
(* Util *) |
|
18 |
fun post_traverse_term_type' f _ (t as Const (_, T)) s = f t T s |
|
19 |
| post_traverse_term_type' f _ (t as Free (_, T)) s = f t T s |
|
20 |
| post_traverse_term_type' f _ (t as Var (_, T)) s = f t T s |
|
21 |
| post_traverse_term_type' f env (t as Bound i) s = f t (nth env i) s |
|
22 |
| post_traverse_term_type' f env (Abs (x, T1, b)) s = |
|
23 |
let |
|
24 |
val ((b', s'), T2) = post_traverse_term_type' f (T1 :: env) b s |
|
25 |
in f (Abs (x, T1, b')) (T1 --> T2) s' end |
|
26 |
| post_traverse_term_type' f env (u $ v) s = |
|
27 |
let |
|
28 |
val ((u', s'), Type (_, [_, T])) = post_traverse_term_type' f env u s |
|
29 |
val ((v', s''), _) = post_traverse_term_type' f env v s' |
|
30 |
in f (u' $ v') T s'' end |
|
31 |
||
32 |
fun post_traverse_term_type f s t = |
|
33 |
post_traverse_term_type' (fn t => fn T => fn s => (f t T s, T)) [] t s |> fst |
|
34 |
fun post_fold_term_type f s t = |
|
35 |
post_traverse_term_type (fn t => fn T => fn s => (t, f t T s)) s t |> snd |
|
36 |
||
37 |
(* Data structures, orders *) |
|
38 |
val cost_ord = prod_ord int_ord (prod_ord int_ord int_ord) |
|
39 |
structure Var_Set_Tab = Table( |
|
40 |
type key = indexname list |
|
41 |
val ord = list_ord Term_Ord.fast_indexname_ord) |
|
42 |
||
43 |
(* (1) Generalize types *) |
|
44 |
fun generalize_types ctxt t = |
|
45 |
t |> map_types (fn _ => dummyT) |
|
46 |
|> Syntax.check_term |
|
47 |
(Proof_Context.set_mode Proof_Context.mode_pattern ctxt) |
|
48 |
||
49 |
(* (2) Typing-spot table *) |
|
50 |
local |
|
51 |
fun key_of_atype (TVar (z, _)) = |
|
52 |
Ord_List.insert Term_Ord.fast_indexname_ord z |
|
53 |
| key_of_atype _ = I |
|
54 |
fun key_of_type T = fold_atyps key_of_atype T [] |
|
55 |
fun update_tab t T (tab, pos) = |
|
56 |
(case key_of_type T of |
|
57 |
[] => tab |
|
58 |
| key => |
|
59 |
let val cost = (size_of_typ T, (size_of_term t, pos)) in |
|
60 |
case Var_Set_Tab.lookup tab key of |
|
61 |
NONE => Var_Set_Tab.update_new (key, cost) tab |
|
62 |
| SOME old_cost => |
|
63 |
(case cost_ord (cost, old_cost) of |
|
64 |
LESS => Var_Set_Tab.update (key, cost) tab |
|
65 |
| _ => tab) |
|
66 |
end, |
|
67 |
pos + 1) |
|
68 |
in |
|
69 |
val typing_spot_table = |
|
70 |
post_fold_term_type update_tab (Var_Set_Tab.empty, 0) #> fst |
|
71 |
end |
|
72 |
||
73 |
(* (3) Reverse-greedy *) |
|
74 |
fun reverse_greedy typing_spot_tab = |
|
75 |
let |
|
76 |
fun update_count z = |
|
77 |
fold (fn tvar => fn tab => |
|
78 |
let val c = Vartab.lookup tab tvar |> the_default 0 in |
|
79 |
Vartab.update (tvar, c + z) tab |
|
80 |
end) |
|
81 |
fun superfluous tcount = |
|
82 |
forall (fn tvar => the (Vartab.lookup tcount tvar) > 1) |
|
83 |
fun drop_superfluous (tvars, (_, (_, spot))) (spots, tcount) = |
|
84 |
if superfluous tcount tvars then (spots, update_count ~1 tvars tcount) |
|
85 |
else (spot :: spots, tcount) |
|
86 |
val (typing_spots, tvar_count_tab) = |
|
87 |
Var_Set_Tab.fold |
|
88 |
(fn kv as (k, _) => apfst (cons kv) #> apsnd (update_count 1 k)) |
|
89 |
typing_spot_tab ([], Vartab.empty) |
|
90 |
|>> sort_distinct (rev_order o cost_ord o pairself snd) |
|
91 |
in fold drop_superfluous typing_spots ([], tvar_count_tab) |> fst end |
|
92 |
||
93 |
(* (4) Introduce annotations *) |
|
94 |
fun introduce_annotations ctxt spots t t' = |
|
95 |
let |
|
96 |
val thy = Proof_Context.theory_of ctxt |
|
97 |
val get_types = post_fold_term_type (K cons) [] |
|
98 |
fun match_types tp = |
|
99 |
fold (Sign.typ_match thy) (op ~~ (pairself get_types tp)) Vartab.empty |
|
100 |
fun unica' b x [] = if b then [x] else [] |
|
101 |
| unica' b x (y :: ys) = |
|
102 |
if x = y then unica' false x ys |
|
103 |
else unica' true y ys |> b ? cons x |
|
104 |
fun unica ord xs = |
|
105 |
case sort ord xs of x :: ys => unica' true x ys | [] => [] |
|
106 |
val add_all_tfree_namesT = fold_atyps (fn TFree (x, _) => cons x | _ => I) |
|
107 |
fun erase_unica_tfrees env = |
|
108 |
let |
|
109 |
val unica = |
|
110 |
Vartab.fold (add_all_tfree_namesT o snd o snd) env [] |
|
111 |
|> filter_out (Variable.is_declared ctxt) |
|
112 |
|> unica fast_string_ord |
|
113 |
val erase_unica = map_atyps |
|
114 |
(fn T as TFree (s, _) => |
|
115 |
if Ord_List.member fast_string_ord unica s then dummyT else T |
|
116 |
| T => T) |
|
117 |
in Vartab.map (K (apsnd erase_unica)) env end |
|
118 |
val env = match_types (t', t) |> erase_unica_tfrees |
|
119 |
fun get_annot env (TFree _) = (false, (env, dummyT)) |
|
120 |
| get_annot env (T as TVar (v, S)) = |
|
121 |
let val T' = Envir.subst_type env T in |
|
122 |
if T' = dummyT then (false, (env, dummyT)) |
|
123 |
else (true, (Vartab.update (v, (S, dummyT)) env, T')) |
|
124 |
end |
|
125 |
| get_annot env (Type (S, Ts)) = |
|
126 |
(case fold_rev (fn T => fn (b, (env, Ts)) => |
|
127 |
let |
|
128 |
val (b', (env', T)) = get_annot env T |
|
129 |
in (b orelse b', (env', T :: Ts)) end) |
|
130 |
Ts (false, (env, [])) of |
|
131 |
(true, (env', Ts)) => (true, (env', Type (S, Ts))) |
|
132 |
| (false, (env', _)) => (false, (env', dummyT))) |
|
133 |
fun post1 _ T (env, cp, ps as p :: ps', annots) = |
|
134 |
if p <> cp then |
|
135 |
(env, cp + 1, ps, annots) |
|
136 |
else |
|
51877
71052c42edf2
avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents:
51179
diff
changeset
|
137 |
let val (annot_necessary, (env', T')) = get_annot env T in |
71052c42edf2
avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents:
51179
diff
changeset
|
138 |
(env', cp + 1, ps', annots |> annot_necessary ? cons (p, T')) |
50258 | 139 |
end |
140 |
| post1 _ _ accum = accum |
|
141 |
val (_, _, _, annots) = post_fold_term_type post1 (env, 0, spots, []) t' |
|
142 |
fun post2 t _ (cp, annots as (p, T) :: annots') = |
|
143 |
if p <> cp then (t, (cp + 1, annots)) |
|
144 |
else (Type.constraint T t, (cp + 1, annots')) |
|
145 |
| post2 t _ x = (t, x) |
|
146 |
in post_traverse_term_type post2 (0, rev annots) t |> fst end |
|
147 |
||
148 |
(* (5) Annotate *) |
|
149 |
fun annotate_types ctxt t = |
|
150 |
let |
|
151 |
val t' = generalize_types ctxt t |
|
152 |
val typing_spots = |
|
153 |
t' |> typing_spot_table |
|
154 |
|> reverse_greedy |
|
155 |
|> sort int_ord |
|
156 |
in introduce_annotations ctxt typing_spots t t' end |
|
157 |
||
158 |
end |