| author | blanchet | 
| Mon, 07 Mar 2016 23:20:11 +0100 | |
| changeset 62535 | cb262f03ac12 | 
| parent 59058 | a78612c67ec0 | 
| child 64429 | 582f54f6b29b | 
| permissions | -rw-r--r-- | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
54821diff
changeset | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_isar_annotate.ML | 
| 55212 | 2 | Author: Steffen Juilf Smolka, TU Muenchen | 
| 50263 | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 4 | ||
| 55212 | 5 | Supplements term with a locally minmal, complete set of type constraints. Complete: The constraints | 
| 6 | suffice to infer the term's types. Minimal: Reducing the set of constraints further will make it | |
| 7 | incomplete. | |
| 52369 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 8 | |
| 55212 | 9 | When configuring the pretty printer appropriately, the constraints will show up as type annotations | 
| 10 | when printing the term. This allows the term to be printed and reparsed without a change of types. | |
| 52369 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 11 | |
| 55213 | 12 | Note: Terms should be unchecked before calling "annotate_types_in_term" to avoid awkward syntax. | 
| 50263 | 13 | *) | 
| 14 | ||
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
54821diff
changeset | 15 | signature SLEDGEHAMMER_ISAR_ANNOTATE = | 
| 50258 | 16 | sig | 
| 55213 | 17 | val annotate_types_in_term : Proof.context -> term -> term | 
| 54504 | 18 | end; | 
| 50258 | 19 | |
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
54821diff
changeset | 20 | structure Sledgehammer_Isar_Annotate : SLEDGEHAMMER_ISAR_ANNOTATE = | 
| 50258 | 21 | struct | 
| 22 | ||
| 23 | fun post_traverse_term_type' f _ (t as Const (_, T)) s = f t T s | |
| 24 | | post_traverse_term_type' f _ (t as Free (_, T)) s = f t T s | |
| 25 | | post_traverse_term_type' f _ (t as Var (_, T)) s = f t T s | |
| 26 | | post_traverse_term_type' f env (t as Bound i) s = f t (nth env i) s | |
| 27 | | post_traverse_term_type' f env (Abs (x, T1, b)) s = | |
| 55243 | 28 | let val ((b', s'), T2) = post_traverse_term_type' f (T1 :: env) b s in | 
| 29 | f (Abs (x, T1, b')) (T1 --> T2) s' | |
| 30 | end | |
| 50258 | 31 | | post_traverse_term_type' f env (u $ v) s = | 
| 32 | let | |
| 33 | val ((u', s'), Type (_, [_, T])) = post_traverse_term_type' f env u s | |
| 34 | val ((v', s''), _) = post_traverse_term_type' f env v s' | |
| 35 | in f (u' $ v') T s'' end | |
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
54821diff
changeset | 36 | handle Bind => raise Fail "Sledgehammer_Isar_Annotate: post_traverse_term_type'" | 
| 50258 | 37 | |
| 38 | fun post_traverse_term_type f s t = | |
| 39 | post_traverse_term_type' (fn t => fn T => fn s => (f t T s, T)) [] t s |> fst | |
| 40 | fun post_fold_term_type f s t = | |
| 41 | post_traverse_term_type (fn t => fn T => fn s => (t, f t T s)) s t |> snd | |
| 42 | ||
| 52452 | 43 | fun fold_map_atypes f T s = | 
| 55286 | 44 | (case T of | 
| 52452 | 45 | Type (name, Ts) => | 
| 55286 | 46 | let val (Ts, s) = fold_map (fold_map_atypes f) Ts s in | 
| 47 | (Type (name, Ts), s) | |
| 48 | end | |
| 49 | | _ => f T s) | |
| 52452 | 50 | |
| 51 | val indexname_ord = Term_Ord.fast_indexname_ord | |
| 50258 | 52 | val cost_ord = prod_ord int_ord (prod_ord int_ord int_ord) | 
| 55243 | 53 | |
| 50258 | 54 | structure Var_Set_Tab = Table( | 
| 55 | type key = indexname list | |
| 52452 | 56 | val ord = list_ord indexname_ord) | 
| 50258 | 57 | |
| 58 | fun generalize_types ctxt t = | |
| 52369 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 59 | let | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 60 | val erase_types = map_types (fn _ => dummyT) | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 61 | (* use schematic type variables *) | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 62 | val ctxt = ctxt |> Proof_Context.set_mode Proof_Context.mode_pattern | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 63 | val infer_types = singleton (Type_Infer_Context.infer_types ctxt) | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 64 | in | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 65 | t |> erase_types |> infer_types | 
| 
0b395800fdf0
uncheck terms before annotation to avoid awkward syntax
 smolkas parents: 
52366diff
changeset | 66 | end | 
| 50258 | 67 | |
| 52452 | 68 | fun match_types ctxt t1 t2 = | 
| 69 | let | |
| 70 | val thy = Proof_Context.theory_of ctxt | |
| 71 | val get_types = post_fold_term_type (K cons) [] | |
| 72 | in | |
| 57467 
03345dad8430
robustness in the face of ill-typed "unchecked" terms (e.g. case expressions)
 blanchet parents: 
55286diff
changeset | 73 | fold (perhaps o try o Sign.typ_match thy) (get_types t1 ~~ get_types t2) Vartab.empty | 
| 52452 | 74 | end | 
| 75 | ||
| 57467 
03345dad8430
robustness in the face of ill-typed "unchecked" terms (e.g. case expressions)
 blanchet parents: 
55286diff
changeset | 76 | fun handle_trivial_tfrees ctxt t' subst = | 
| 52452 | 77 | let | 
| 55243 | 78 | val add_tfree_names = snd #> snd #> fold_atyps (fn TFree (x, _) => cons x | _ => I) | 
| 52452 | 79 | |
| 80 | val trivial_tfree_names = | |
| 81 | Vartab.fold add_tfree_names subst [] | |
| 82 | |> filter_out (Variable.is_declared ctxt) | |
| 55243 | 83 | |> distinct (op =) | 
| 52452 | 84 | val tfree_name_trivial = Ord_List.member fast_string_ord trivial_tfree_names | 
| 85 | ||
| 86 | val trivial_tvar_names = | |
| 87 | Vartab.fold | |
| 88 | (fn (tvar_name, (_, TFree (tfree_name, _))) => | |
| 89 | tfree_name_trivial tfree_name ? cons tvar_name | |
| 90 | | _ => I) | |
| 91 | subst | |
| 92 | [] | |
| 93 | |> sort indexname_ord | |
| 94 | val tvar_name_trivial = Ord_List.member indexname_ord trivial_tvar_names | |
| 95 | ||
| 96 | val t' = | |
| 97 | t' |> map_types | |
| 98 | (map_type_tvar | |
| 99 | (fn (idxn, sort) => | |
| 100 | if tvar_name_trivial idxn then dummyT else TVar (idxn, sort))) | |
| 101 | ||
| 102 | val subst = | |
| 103 | subst |> fold Vartab.delete trivial_tvar_names | |
| 104 | |> Vartab.map | |
| 105 | (K (apsnd (map_type_tfree | |
| 106 | (fn (name, sort) => | |
| 107 | if tfree_name_trivial name then dummyT | |
| 108 | else TFree (name, sort))))) | |
| 109 | in | |
| 110 | (t', subst) | |
| 111 | end | |
| 112 | ||
| 54821 | 113 | fun key_of_atype (TVar (z, _)) = Ord_List.insert indexname_ord z | 
| 50258 | 114 | | key_of_atype _ = I | 
| 115 | fun key_of_type T = fold_atyps key_of_atype T [] | |
| 55243 | 116 | |
| 50258 | 117 | fun update_tab t T (tab, pos) = | 
| 55243 | 118 | ((case key_of_type T of | 
| 50258 | 119 | [] => tab | 
| 120 | | key => | |
| 121 | let val cost = (size_of_typ T, (size_of_term t, pos)) in | |
| 55243 | 122 | (case Var_Set_Tab.lookup tab key of | 
| 50258 | 123 | NONE => Var_Set_Tab.update_new (key, cost) tab | 
| 124 | | SOME old_cost => | |
| 125 | (case cost_ord (cost, old_cost) of | |
| 55243 | 126 | LESS => Var_Set_Tab.update (key, cost) tab | 
| 127 | | _ => tab)) | |
| 128 | end), | |
| 50258 | 129 | pos + 1) | 
| 130 | ||
| 55243 | 131 | val typing_spot_table = post_fold_term_type update_tab (Var_Set_Tab.empty, 0) #> fst | 
| 132 | ||
| 50258 | 133 | fun reverse_greedy typing_spot_tab = | 
| 134 | let | |
| 135 | fun update_count z = | |
| 136 | fold (fn tvar => fn tab => | |
| 137 | let val c = Vartab.lookup tab tvar |> the_default 0 in | |
| 138 | Vartab.update (tvar, c + z) tab | |
| 139 | end) | |
| 55243 | 140 | fun superfluous tcount = forall (fn tvar => the (Vartab.lookup tcount tvar) > 1) | 
| 50258 | 141 | fun drop_superfluous (tvars, (_, (_, spot))) (spots, tcount) = | 
| 142 | if superfluous tcount tvars then (spots, update_count ~1 tvars tcount) | |
| 143 | else (spot :: spots, tcount) | |
| 55243 | 144 | |
| 50258 | 145 | val (typing_spots, tvar_count_tab) = | 
| 55243 | 146 | Var_Set_Tab.fold (fn kv as (k, _) => apfst (cons kv) #> apsnd (update_count 1 k)) | 
| 50258 | 147 | typing_spot_tab ([], Vartab.empty) | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57467diff
changeset | 148 | |>> sort_distinct (rev_order o cost_ord o apply2 snd) | 
| 55243 | 149 | in | 
| 150 | fold drop_superfluous typing_spots ([], tvar_count_tab) |> fst | |
| 151 | end | |
| 50258 | 152 | |
| 52452 | 153 | fun introduce_annotations subst spots t t' = | 
| 50258 | 154 | let | 
| 52452 | 155 | fun subst_atype (T as TVar (idxn, S)) subst = | 
| 54821 | 156 | (Envir.subst_type subst T, Vartab.update (idxn, (S, dummyT)) subst) | 
| 52452 | 157 | | subst_atype T subst = (T, subst) | 
| 55243 | 158 | |
| 52452 | 159 | val subst_type = fold_map_atypes subst_atype | 
| 55243 | 160 | |
| 52452 | 161 | fun collect_annot _ T (subst, cp, ps as p :: ps', annots) = | 
| 50258 | 162 | if p <> cp then | 
| 52452 | 163 | (subst, cp + 1, ps, annots) | 
| 50258 | 164 | else | 
| 52452 | 165 | let val (T, subst) = subst_type T subst in | 
| 55243 | 166 | (subst, cp + 1, ps', (p, T) :: annots) | 
| 50258 | 167 | end | 
| 52452 | 168 | | collect_annot _ _ x = x | 
| 55243 | 169 | |
| 170 | val (_, _, _, annots) = post_fold_term_type collect_annot (subst, 0, spots, []) t' | |
| 171 | ||
| 52452 | 172 | fun insert_annot t _ (cp, annots as (p, T) :: annots') = | 
| 54821 | 173 | if p <> cp then (t, (cp + 1, annots)) else (Type.constraint T t, (cp + 1, annots')) | 
| 52452 | 174 | | insert_annot t _ x = (t, x) | 
| 52110 
411db77f96f2
prevent pretty printer from automatically annotating numerals
 smolkas parents: 
51877diff
changeset | 175 | in | 
| 55243 | 176 | t |> post_traverse_term_type insert_annot (0, rev annots) |> fst | 
| 52110 
411db77f96f2
prevent pretty printer from automatically annotating numerals
 smolkas parents: 
51877diff
changeset | 177 | end | 
| 50258 | 178 | |
| 55213 | 179 | fun annotate_types_in_term ctxt t = | 
| 50258 | 180 | let | 
| 181 | val t' = generalize_types ctxt t | |
| 52452 | 182 | val subst = match_types ctxt t' t | 
| 57467 
03345dad8430
robustness in the face of ill-typed "unchecked" terms (e.g. case expressions)
 blanchet parents: 
55286diff
changeset | 183 | val (t'', subst') = handle_trivial_tfrees ctxt t' subst | 
| 
03345dad8430
robustness in the face of ill-typed "unchecked" terms (e.g. case expressions)
 blanchet parents: 
55286diff
changeset | 184 | val typing_spots = t'' |> typing_spot_table |> reverse_greedy |> sort int_ord | 
| 55243 | 185 | in | 
| 57467 
03345dad8430
robustness in the face of ill-typed "unchecked" terms (e.g. case expressions)
 blanchet parents: 
55286diff
changeset | 186 | introduce_annotations subst' typing_spots t t'' | 
| 55243 | 187 | end | 
| 50258 | 188 | |
| 54504 | 189 | end; |