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