src/HOL/Tools/SMT/smt_builtin.ML
changeset 42361 23f352990944
parent 41473 3717fc42ebe9
child 46042 ab32a87ba01a
equal deleted inserted replaced
42360:da8817d01e7c 42361:23f352990944
    71 
    71 
    72 fun merge_ttab ttabp =
    72 fun merge_ttab ttabp =
    73   SMT_Utils.dict_merge (Ord_List.merge typ_ord) ttabp
    73   SMT_Utils.dict_merge (Ord_List.merge typ_ord) ttabp
    74 
    74 
    75 fun lookup_ttab ctxt ttab T =
    75 fun lookup_ttab ctxt ttab T =
    76   let fun match (U, _) = Sign.typ_instance (ProofContext.theory_of ctxt) (T, U)
    76   let fun match (U, _) = Sign.typ_instance (Proof_Context.theory_of ctxt) (T, U)
    77   in
    77   in
    78     get_first (find_first match)
    78     get_first (find_first match)
    79       (SMT_Utils.dict_lookup ttab (SMT_Config.solver_class_of ctxt))
    79       (SMT_Utils.dict_lookup ttab (SMT_Config.solver_class_of ctxt))
    80   end
    80   end
    81 
    81