src/HOL/Tools/ATP/atp_translate.ML
changeset 44810 081718c0b0a8
parent 44807 127749bbc639
child 44832 91294d386539
     1.1 --- a/src/HOL/Tools/ATP/atp_translate.ML	Thu Jul 21 08:33:57 2011 +0200
     1.2 +++ b/src/HOL/Tools/ATP/atp_translate.ML	Thu Jul 21 21:29:10 2011 +0200
     1.3 @@ -88,7 +88,6 @@
     1.4    val unmangled_const_name : string -> string
     1.5    val helper_table : ((string * bool) * thm list) list
     1.6    val factsN : string
     1.7 -  val conceal_lambdas : Proof.context -> term -> term
     1.8    val introduce_combinators : Proof.context -> term -> term
     1.9    val prepare_atp_problem :
    1.10      Proof.context -> format -> formula_kind -> formula_kind -> type_enc -> bool
    1.11 @@ -161,8 +160,6 @@
    1.12  (* Freshness almost guaranteed! *)
    1.13  val atp_weak_prefix = "ATP:"
    1.14  
    1.15 -val concealed_lambda_prefix = atp_weak_prefix ^ "lambda_"
    1.16 -
    1.17  (*Escaping of special characters.
    1.18    Alphanumeric characters are left unchanged.
    1.19    The character _ goes to __
    1.20 @@ -926,10 +923,8 @@
    1.21      do_conceal_lambdas Ts t1 $ do_conceal_lambdas Ts t2
    1.22    | do_conceal_lambdas Ts (Abs (_, T, t)) =
    1.23      (* slightly unsound because of hash collisions *)
    1.24 -    Free (concealed_lambda_prefix ^ string_of_int (hash_term t),
    1.25 -          T --> fastype_of1 (Ts, t))
    1.26 +    Free (polymorphic_free_prefix ^ serial_string (), T --> fastype_of1 (Ts, t))
    1.27    | do_conceal_lambdas _ t = t
    1.28 -val conceal_lambdas = simple_translate_lambdas (K do_conceal_lambdas)
    1.29  
    1.30  fun do_introduce_combinators ctxt Ts t =
    1.31    let val thy = Proof_Context.theory_of ctxt in