src/HOL/Tools/ATP/atp_util.ML
author blanchet
Fri, 18 Jan 2013 14:33:16 +0100
changeset 51983 3686bc0d4df2
parent 51254 fb579401dc26
child 52334 1c6031e5d284
permissions -rw-r--r--
pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
     1 (*  Title:      HOL/Tools/ATP/atp_util.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3 
     4 General-purpose functions used by the ATP module.
     5 *)
     6 
     7 signature ATP_UTIL =
     8 sig
     9   val timestamp : unit -> string
    10   val hash_string : string -> int
    11   val hash_term : term -> int
    12   val chunk_list : int -> 'a list -> 'a list list
    13   val stringN_of_int : int -> int -> string
    14   val strip_spaces : bool -> (char -> bool) -> string -> string
    15   val strip_spaces_except_between_idents : string -> string
    16   val elide_string : int -> string -> string
    17   val nat_subscript : int -> string
    18   val unyxml : string -> string
    19   val maybe_quote : string -> string
    20   val string_from_ext_time : bool * Time.time -> string
    21   val string_from_time : Time.time -> string
    22   val type_instance : theory -> typ -> typ -> bool
    23   val type_generalization : theory -> typ -> typ -> bool
    24   val type_intersect : theory -> typ -> typ -> bool
    25   val type_equiv : theory -> typ * typ -> bool
    26   val varify_type : Proof.context -> typ -> typ
    27   val instantiate_type : theory -> typ -> typ -> typ -> typ
    28   val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
    29   val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
    30   val is_type_surely_finite : Proof.context -> typ -> bool
    31   val is_type_surely_infinite : Proof.context -> bool -> typ list -> typ -> bool
    32   val s_not : term -> term
    33   val s_conj : term * term -> term
    34   val s_disj : term * term -> term
    35   val s_imp : term * term -> term
    36   val s_iff : term * term -> term
    37   val close_form : term -> term
    38   val hol_close_form_prefix : string
    39   val hol_close_form : term -> term
    40   val hol_open_form : (string -> string) -> term -> term
    41   val monomorphic_term : Type.tyenv -> term -> term
    42   val eta_expand : typ list -> term -> int -> term
    43   val cong_extensionalize_term : theory -> term -> term
    44   val abs_extensionalize_term : Proof.context -> term -> term
    45   val unextensionalize_def : term -> term
    46   val is_legitimate_tptp_def : term -> bool
    47   val transform_elim_prop : term -> term
    48   val specialize_type : theory -> (string * typ) -> term -> term
    49   val strip_subgoal :
    50     Proof.context -> thm -> int -> (string * typ) list * term list * term
    51 end;
    52 
    53 structure ATP_Util : ATP_UTIL =
    54 struct
    55 
    56 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
    57 
    58 (* This hash function is recommended in "Compilers: Principles, Techniques, and
    59    Tools" by Aho, Sethi, and Ullman. The "hashpjw" function, which they
    60    particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
    61 fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
    62 fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
    63 fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s
    64 fun hashw_term (t1 $ t2) = hashw (hashw_term t1, hashw_term t2)
    65   | hashw_term (Const (s, _)) = hashw_string (s, 0w0)
    66   | hashw_term (Free (s, _)) = hashw_string (s, 0w0)
    67   | hashw_term _ = 0w0
    68 
    69 fun hash_string s = Word.toInt (hashw_string (s, 0w0))
    70 val hash_term = Word.toInt o hashw_term
    71 
    72 fun chunk_list _ [] = []
    73   | chunk_list k xs =
    74     let val (xs1, xs2) = chop k xs in xs1 :: chunk_list k xs2 end
    75 
    76 fun stringN_of_int 0 _ = ""
    77   | stringN_of_int k n =
    78     stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
    79 
    80 fun strip_spaces skip_comments is_evil =
    81   let
    82     fun strip_c_style_comment [] accum = accum
    83       | strip_c_style_comment (#"*" :: #"/" :: cs) accum =
    84         strip_spaces_in_list true cs accum
    85       | strip_c_style_comment (_ :: cs) accum = strip_c_style_comment cs accum
    86     and strip_spaces_in_list _ [] accum = accum
    87       | strip_spaces_in_list true (#"%" :: cs) accum =
    88         strip_spaces_in_list true (cs |> take_prefix (not_equal #"\n") |> snd)
    89                              accum
    90       | strip_spaces_in_list true (#"/" :: #"*" :: cs) accum =
    91         strip_c_style_comment cs accum
    92       | strip_spaces_in_list _ [c1] accum =
    93         accum |> not (Char.isSpace c1) ? cons c1
    94       | strip_spaces_in_list skip_comments (cs as [_, _]) accum =
    95         accum |> fold (strip_spaces_in_list skip_comments o single) cs
    96       | strip_spaces_in_list skip_comments (c1 :: c2 :: c3 :: cs) accum =
    97         if Char.isSpace c1 then
    98           strip_spaces_in_list skip_comments (c2 :: c3 :: cs) accum
    99         else if Char.isSpace c2 then
   100           if Char.isSpace c3 then
   101             strip_spaces_in_list skip_comments (c1 :: c3 :: cs) accum
   102           else
   103             strip_spaces_in_list skip_comments (c3 :: cs)
   104                 (c1 :: accum |> forall is_evil [c1, c3] ? cons #" ")
   105         else
   106           strip_spaces_in_list skip_comments (c2 :: c3 :: cs) (cons c1 accum)
   107   in
   108     String.explode
   109     #> rpair [] #-> strip_spaces_in_list skip_comments
   110     #> rev #> String.implode
   111   end
   112 
   113 fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
   114 val strip_spaces_except_between_idents = strip_spaces true is_ident_char
   115 
   116 fun elide_string threshold s =
   117   if size s > threshold then
   118     String.extract (s, 0, SOME (threshold div 2 - 5)) ^ " ...... " ^
   119     String.extract (s, size s - (threshold + 1) div 2 + 6, NONE)
   120   else
   121     s
   122 
   123 val subscript = implode o map (prefix "\<^isub>") o raw_explode  (* FIXME Symbol.explode (?) *)
   124 fun nat_subscript n =
   125   n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
   126 
   127 val unyxml = XML.content_of o YXML.parse_body
   128 
   129 val is_long_identifier = forall Symbol_Pos.is_identifier o Long_Name.explode
   130 fun maybe_quote y =
   131   let val s = unyxml y in
   132     y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
   133            not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
   134            Keyword.is_keyword s) ? quote
   135   end
   136 
   137 fun string_from_ext_time (plus, time) =
   138   let val ms = Time.toMilliseconds time in
   139     (if plus then "> " else "") ^
   140     (if plus andalso ms mod 1000 = 0 then
   141        signed_string_of_int (ms div 1000) ^ " s"
   142      else if ms < 1000 then
   143        signed_string_of_int ms ^ " ms"
   144      else
   145        string_of_real (0.01 * Real.fromInt (ms div 10)) ^ " s")
   146   end
   147 
   148 val string_from_time = string_from_ext_time o pair false
   149 
   150 fun type_instance thy T T' = Sign.typ_instance thy (T, T')
   151 fun type_generalization thy T T' = Sign.typ_instance thy (T', T)
   152 
   153 fun type_intersect _ (TVar _) _ = true
   154   | type_intersect _ _ (TVar _) = true
   155   | type_intersect thy T T' =
   156     let
   157       val tvars = Term.add_tvar_namesT T []
   158       val tvars' = Term.add_tvar_namesT T' []
   159       val maxidx' = maxidx_of_typ T'
   160       val T =
   161         T |> exists (member (op =) tvars') tvars ? Logic.incr_tvar (maxidx' + 1)
   162       val maxidx = Integer.max (maxidx_of_typ T) maxidx'
   163     in can (Sign.typ_unify thy (T, T')) (Vartab.empty, maxidx) end
   164 
   165 val type_equiv = Sign.typ_equiv
   166 
   167 fun varify_type ctxt T =
   168   Variable.polymorphic_types ctxt [Const (@{const_name undefined}, T)]
   169   |> snd |> the_single |> dest_Const |> snd
   170 
   171 (* TODO: use "Term_Subst.instantiateT" instead? *)
   172 fun instantiate_type thy T1 T1' T2 =
   173   Same.commit (Envir.subst_type_same
   174                    (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
   175   handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], [])
   176 
   177 fun varify_and_instantiate_type ctxt T1 T1' T2 =
   178   let val thy = Proof_Context.theory_of ctxt in
   179     instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
   180   end
   181 
   182 fun typ_of_dtyp _ typ_assoc (Datatype.DtTFree a) =
   183     the (AList.lookup (op =) typ_assoc (Datatype.DtTFree a))
   184   | typ_of_dtyp descr typ_assoc (Datatype.DtType (s, Us)) =
   185     Type (s, map (typ_of_dtyp descr typ_assoc) Us)
   186   | typ_of_dtyp descr typ_assoc (Datatype.DtRec i) =
   187     let val (s, ds, _) = the (AList.lookup (op =) descr i) in
   188       Type (s, map (typ_of_dtyp descr typ_assoc) ds)
   189     end
   190 
   191 fun datatype_constrs thy (T as Type (s, Ts)) =
   192     (case Datatype.get_info thy s of
   193        SOME {index, descr, ...} =>
   194        let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in
   195          map (apsnd (fn Us => map (typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T))
   196              constrs
   197        end
   198      | NONE => [])
   199   | datatype_constrs _ _ = []
   200 
   201 (* Similar to "Nitpick_HOL.bounded_exact_card_of_type".
   202    0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
   203    cardinality 2 or more. The specified default cardinality is returned if the
   204    cardinality of the type can't be determined. *)
   205 fun tiny_card_of_type ctxt sound assigns default_card T =
   206   let
   207     val thy = Proof_Context.theory_of ctxt
   208     val max = 2 (* 1 would be too small for the "fun" case *)
   209     fun aux slack avoid T =
   210       if member (op =) avoid T then
   211         0
   212       else case AList.lookup (type_equiv thy) assigns T of
   213         SOME k => k
   214       | NONE =>
   215         case T of
   216           Type (@{type_name fun}, [T1, T2]) =>
   217           (case (aux slack avoid T1, aux slack avoid T2) of
   218              (k, 1) => if slack andalso k = 0 then 0 else 1
   219            | (0, _) => 0
   220            | (_, 0) => 0
   221            | (k1, k2) =>
   222              if k1 >= max orelse k2 >= max then max
   223              else Int.min (max, Integer.pow k2 k1))
   224         | Type (@{type_name set}, [T']) => aux slack avoid (T' --> @{typ bool})
   225         | @{typ prop} => 2
   226         | @{typ bool} => 2 (* optimization *)
   227         | @{typ nat} => 0 (* optimization *)
   228         | Type ("Int.int", []) => 0 (* optimization *)
   229         | Type (s, _) =>
   230           (case datatype_constrs thy T of
   231              constrs as _ :: _ =>
   232              let
   233                val constr_cards =
   234                  map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
   235                       o snd) constrs
   236              in
   237                if exists (curry (op =) 0) constr_cards then 0
   238                else Int.min (max, Integer.sum constr_cards)
   239              end
   240            | [] =>
   241              case Typedef.get_info ctxt s of
   242                ({abs_type, rep_type, ...}, _) :: _ =>
   243                if not sound then
   244                  (* We cheat here by assuming that typedef types are infinite if
   245                     their underlying type is infinite. This is unsound in
   246                     general but it's hard to think of a realistic example where
   247                     this would not be the case. We are also slack with
   248                     representation types: If a representation type has the form
   249                     "sigma => tau", we consider it enough to check "sigma" for
   250                     infiniteness. *)
   251                  (case varify_and_instantiate_type ctxt
   252                            (Logic.varifyT_global abs_type) T
   253                            (Logic.varifyT_global rep_type)
   254                        |> aux true avoid of
   255                     0 => 0
   256                   | 1 => 1
   257                   | _ => default_card)
   258                else
   259                  default_card
   260              | [] => default_card)
   261           (* Very slightly unsound: Type variables are assumed not to be
   262              constrained to cardinality 1. (In practice, the user would most
   263              likely have used "unit" directly anyway.) *)
   264         | TFree _ =>
   265           if not sound andalso default_card = 1 then 2 else default_card
   266         | TVar _ => default_card
   267   in Int.min (max, aux false [] T) end
   268 
   269 fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt true [] 0 T <> 0
   270 fun is_type_surely_infinite ctxt sound infinite_Ts T =
   271   tiny_card_of_type ctxt sound (map (rpair 0) infinite_Ts) 1 T = 0
   272 
   273 (* Simple simplifications to ensure that sort annotations don't leave a trail of
   274    spurious "True"s. *)
   275 fun s_not (Const (@{const_name All}, T) $ Abs (s, T', t')) =
   276     Const (@{const_name Ex}, T) $ Abs (s, T', s_not t')
   277   | s_not (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
   278     Const (@{const_name All}, T) $ Abs (s, T', s_not t')
   279   | s_not (@{const HOL.implies} $ t1 $ t2) = @{const HOL.conj} $ t1 $ s_not t2
   280   | s_not (@{const HOL.conj} $ t1 $ t2) =
   281     @{const HOL.disj} $ s_not t1 $ s_not t2
   282   | s_not (@{const HOL.disj} $ t1 $ t2) =
   283     @{const HOL.conj} $ s_not t1 $ s_not t2
   284   | s_not (@{const False}) = @{const True}
   285   | s_not (@{const True}) = @{const False}
   286   | s_not (@{const Not} $ t) = t
   287   | s_not t = @{const Not} $ t
   288 fun s_conj (@{const True}, t2) = t2
   289   | s_conj (t1, @{const True}) = t1
   290   | s_conj p = HOLogic.mk_conj p
   291 fun s_disj (@{const False}, t2) = t2
   292   | s_disj (t1, @{const False}) = t1
   293   | s_disj p = HOLogic.mk_disj p
   294 fun s_imp (@{const True}, t2) = t2
   295   | s_imp (t1, @{const False}) = s_not t1
   296   | s_imp p = HOLogic.mk_imp p
   297 fun s_iff (@{const True}, t2) = t2
   298   | s_iff (t1, @{const True}) = t1
   299   | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
   300 
   301 (* cf. "close_form" in "refute.ML" *)
   302 fun close_form t =
   303   fold (fn ((s, i), T) => fn t' =>
   304            Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t')))
   305        (Term.add_vars t []) t
   306 
   307 val hol_close_form_prefix = "ATP.close_form."
   308 
   309 fun hol_close_form t =
   310   fold (fn ((s, i), T) => fn t' =>
   311            HOLogic.all_const T
   312            $ Abs (hol_close_form_prefix ^ s, T,
   313                   abstract_over (Var ((s, i), T), t')))
   314        (Term.add_vars t []) t
   315 
   316 fun hol_open_form unprefix
   317       (t as Const (@{const_name All}, _) $ Abs (s, T, t')) =
   318     (case try unprefix s of
   319        SOME s =>
   320        let
   321          val names = Name.make_context (map fst (Term.add_var_names t' []))
   322          val (s, _) = Name.variant s names
   323        in hol_open_form unprefix (subst_bound (Var ((s, 0), T), t')) end
   324      | NONE => t)
   325   | hol_open_form _ t = t
   326 
   327 fun monomorphic_term subst =
   328   map_types (map_type_tvar (fn v =>
   329       case Type.lookup subst v of
   330         SOME typ => typ
   331       | NONE => TVar v))
   332 
   333 fun eta_expand _ t 0 = t
   334   | eta_expand Ts (Abs (s, T, t')) n =
   335     Abs (s, T, eta_expand (T :: Ts) t' (n - 1))
   336   | eta_expand Ts t n =
   337     fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t'))
   338              (List.take (binder_types (fastype_of1 (Ts, t)), n))
   339              (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0)))
   340 
   341 fun cong_extensionalize_term thy t =
   342   if exists_Const (fn (s, _) => s = @{const_name Not}) t then
   343     t |> Skip_Proof.make_thm thy
   344       |> Meson.cong_extensionalize_thm thy
   345       |> prop_of
   346   else
   347     t
   348 
   349 fun is_fun_equality (@{const_name HOL.eq},
   350                      Type (_, [Type (@{type_name fun}, _), _])) = true
   351   | is_fun_equality _ = false
   352 
   353 fun abs_extensionalize_term ctxt t =
   354   if exists_Const is_fun_equality t then
   355     let val thy = Proof_Context.theory_of ctxt in
   356       t |> cterm_of thy |> Meson.abs_extensionalize_conv ctxt
   357         |> prop_of |> Logic.dest_equals |> snd
   358     end
   359   else
   360     t
   361 
   362 fun unextensionalize_def t =
   363   case t of
   364     @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ lhs $ rhs) =>
   365     (case strip_comb lhs of
   366        (c as Const (_, T), args) =>
   367        if forall is_Var args andalso not (has_duplicates (op =) args) then
   368          @{const Trueprop}
   369          $ (Const (@{const_name HOL.eq}, T --> T --> @{typ bool})
   370             $ c $ fold_rev lambda args rhs)
   371        else
   372          t
   373      | _ => t)
   374   | _ => t
   375 
   376 fun is_legitimate_tptp_def (@{const Trueprop} $ t) = is_legitimate_tptp_def t
   377   | is_legitimate_tptp_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
   378     (is_Const t orelse is_Free t) andalso
   379     not (exists_subterm (curry (op =) t) u)
   380   | is_legitimate_tptp_def _ = false
   381 
   382 (* Converts an elim-rule into an equivalent theorem that does not have the
   383    predicate variable. Leaves other theorems unchanged. We simply instantiate
   384    the conclusion variable to "False". (Cf. "transform_elim_theorem" in
   385    "Meson_Clausify".) *)
   386 fun transform_elim_prop t =
   387   case Logic.strip_imp_concl t of
   388     @{const Trueprop} $ Var (z, @{typ bool}) =>
   389     subst_Vars [(z, @{const False})] t
   390   | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
   391   | _ => t
   392 
   393 fun specialize_type thy (s, T) t =
   394   let
   395     fun subst_for (Const (s', T')) =
   396       if s = s' then
   397         SOME (Sign.typ_match thy (T', T) Vartab.empty)
   398         handle Type.TYPE_MATCH => NONE
   399       else
   400         NONE
   401     | subst_for (t1 $ t2) =
   402       (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
   403     | subst_for (Abs (_, _, t')) = subst_for t'
   404     | subst_for _ = NONE
   405   in
   406     case subst_for t of
   407       SOME subst => monomorphic_term subst t
   408     | NONE => raise Type.TYPE_MATCH
   409   end
   410 
   411 fun strip_subgoal ctxt goal i =
   412   let
   413     val (t, (frees, params)) =
   414       Logic.goal_params (prop_of goal) i
   415       ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
   416     val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
   417     val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
   418   in (rev params, hyp_ts, concl_t) end
   419 
   420 end;