src/HOL/Tools/ATP/atp_util.ML
author blanchet
Wed, 07 Sep 2011 13:50:17 +0200
changeset 45652 c9a081ef441d
parent 45355 dbd98b549597
child 45730 237ba63d6041
permissions -rw-r--r--
tuning
     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 strip_spaces : bool -> (char -> bool) -> string -> string
    13   val strip_spaces_except_between_idents : string -> string
    14   val nat_subscript : int -> string
    15   val unyxml : string -> string
    16   val maybe_quote : string -> string
    17   val string_from_ext_time : bool * Time.time -> string
    18   val string_from_time : Time.time -> string
    19   val type_instance : Proof.context -> typ -> typ -> bool
    20   val type_generalization : Proof.context -> typ -> typ -> bool
    21   val type_intersect : Proof.context -> typ -> typ -> bool
    22   val type_aconv : Proof.context -> typ * typ -> bool
    23   val varify_type : Proof.context -> typ -> typ
    24   val instantiate_type : theory -> typ -> typ -> typ -> typ
    25   val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
    26   val typ_of_dtyp :
    27     Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
    28     -> typ
    29   val is_type_surely_finite : Proof.context -> typ -> bool
    30   val is_type_surely_infinite : Proof.context -> bool -> typ list -> typ -> bool
    31   val s_not : term -> term
    32   val s_conj : term * term -> term
    33   val s_disj : term * term -> term
    34   val s_imp : term * term -> term
    35   val s_iff : term * term -> term
    36   val close_form : term -> term
    37   val monomorphic_term : Type.tyenv -> term -> term
    38   val eta_expand : typ list -> term -> int -> term
    39   val transform_elim_prop : term -> term
    40   val specialize_type : theory -> (string * typ) -> term -> term
    41   val strip_subgoal :
    42     Proof.context -> thm -> int -> (string * typ) list * term list * term
    43 end;
    44 
    45 structure ATP_Util : ATP_UTIL =
    46 struct
    47 
    48 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
    49 
    50 (* This hash function is recommended in "Compilers: Principles, Techniques, and
    51    Tools" by Aho, Sethi, and Ullman. The "hashpjw" function, which they
    52    particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
    53 fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
    54 fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
    55 fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s
    56 fun hashw_term (t1 $ t2) = hashw (hashw_term t1, hashw_term t2)
    57   | hashw_term (Const (s, _)) = hashw_string (s, 0w0)
    58   | hashw_term (Free (s, _)) = hashw_string (s, 0w0)
    59   | hashw_term _ = 0w0
    60 
    61 fun hash_string s = Word.toInt (hashw_string (s, 0w0))
    62 val hash_term = Word.toInt o hashw_term
    63 
    64 fun strip_c_style_comment _ [] = []
    65   | strip_c_style_comment is_evil (#"*" :: #"/" :: cs) =
    66     strip_spaces_in_list true is_evil cs
    67   | strip_c_style_comment is_evil (_ :: cs) = strip_c_style_comment is_evil cs
    68 and strip_spaces_in_list _ _ [] = []
    69   | strip_spaces_in_list true is_evil (#"%" :: cs) =
    70     strip_spaces_in_list true is_evil
    71                          (cs |> chop_while (not_equal #"\n") |> snd)
    72   | strip_spaces_in_list true is_evil (#"/" :: #"*" :: cs) =
    73     strip_c_style_comment is_evil cs
    74   | strip_spaces_in_list _ _ [c1] = if Char.isSpace c1 then [] else [str c1]
    75   | strip_spaces_in_list skip_comments is_evil [c1, c2] =
    76     strip_spaces_in_list skip_comments is_evil [c1] @
    77     strip_spaces_in_list skip_comments is_evil [c2]
    78   | strip_spaces_in_list skip_comments is_evil (c1 :: c2 :: c3 :: cs) =
    79     if Char.isSpace c1 then
    80       strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
    81     else if Char.isSpace c2 then
    82       if Char.isSpace c3 then
    83         strip_spaces_in_list skip_comments is_evil (c1 :: c3 :: cs)
    84       else
    85         str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
    86         strip_spaces_in_list skip_comments is_evil (c3 :: cs)
    87     else
    88       str c1 :: strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
    89 fun strip_spaces skip_comments is_evil =
    90   implode o strip_spaces_in_list skip_comments is_evil o String.explode
    91 
    92 fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
    93 val strip_spaces_except_between_idents = strip_spaces true is_ident_char
    94 
    95 val subscript = implode o map (prefix "\<^isub>") o raw_explode  (* FIXME Symbol.explode (?) *)
    96 fun nat_subscript n =
    97   n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
    98 
    99 val unyxml = XML.content_of o YXML.parse_body
   100 
   101 val is_long_identifier = forall Lexicon.is_identifier o space_explode "."
   102 fun maybe_quote y =
   103   let val s = unyxml y in
   104     y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
   105            not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
   106            Keyword.is_keyword s) ? quote
   107   end
   108 
   109 fun string_from_ext_time (plus, time) =
   110   let val ms = Time.toMilliseconds time in
   111     (if plus then "> " else "") ^
   112     (if plus andalso ms mod 1000 = 0 then
   113        signed_string_of_int (ms div 1000) ^ " s"
   114      else if ms < 1000 then
   115        signed_string_of_int ms ^ " ms"
   116      else
   117        string_of_real (0.01 * Real.fromInt (ms div 10)) ^ " s")
   118   end
   119 
   120 val string_from_time = string_from_ext_time o pair false
   121 
   122 fun type_instance ctxt T T' =
   123   Sign.typ_instance (Proof_Context.theory_of ctxt) (T, T')
   124 fun type_generalization ctxt T T' = type_instance ctxt T' T
   125 fun type_intersect ctxt T T' =
   126   type_instance ctxt T T' orelse type_generalization ctxt T T'
   127 fun type_aconv ctxt (T, T') =
   128   type_instance ctxt T T' andalso type_instance ctxt T' T
   129 
   130 fun varify_type ctxt T =
   131   Variable.polymorphic_types ctxt [Const (@{const_name undefined}, T)]
   132   |> snd |> the_single |> dest_Const |> snd
   133 
   134 (* TODO: use "Term_Subst.instantiateT" instead? *)
   135 fun instantiate_type thy T1 T1' T2 =
   136   Same.commit (Envir.subst_type_same
   137                    (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
   138   handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], [])
   139 
   140 fun varify_and_instantiate_type ctxt T1 T1' T2 =
   141   let val thy = Proof_Context.theory_of ctxt in
   142     instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
   143   end
   144 
   145 fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
   146     the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
   147   | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
   148     Type (s, map (typ_of_dtyp descr typ_assoc) Us)
   149   | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
   150     let val (s, ds, _) = the (AList.lookup (op =) descr i) in
   151       Type (s, map (typ_of_dtyp descr typ_assoc) ds)
   152     end
   153 
   154 fun datatype_constrs thy (T as Type (s, Ts)) =
   155     (case Datatype.get_info thy s of
   156        SOME {index, descr, ...} =>
   157        let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in
   158          map (apsnd (fn Us => map (typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T))
   159              constrs
   160        end
   161      | NONE => [])
   162   | datatype_constrs _ _ = []
   163 
   164 (* Similar to "Nitpick_HOL.bounded_exact_card_of_type".
   165    0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
   166    cardinality 2 or more. The specified default cardinality is returned if the
   167    cardinality of the type can't be determined. *)
   168 fun tiny_card_of_type ctxt sound assigns default_card T =
   169   let
   170     val thy = Proof_Context.theory_of ctxt
   171     val max = 2 (* 1 would be too small for the "fun" case *)
   172     fun aux slack avoid T =
   173       if member (op =) avoid T then
   174         0
   175       else case AList.lookup (type_aconv ctxt) assigns T of
   176         SOME k => k
   177       | NONE =>
   178         case T of
   179           Type (@{type_name fun}, [T1, T2]) =>
   180           (case (aux slack avoid T1, aux slack avoid T2) of
   181              (k, 1) => if slack andalso k = 0 then 0 else 1
   182            | (0, _) => 0
   183            | (_, 0) => 0
   184            | (k1, k2) =>
   185              if k1 >= max orelse k2 >= max then max
   186              else Int.min (max, Integer.pow k2 k1))
   187         | @{typ prop} => 2
   188         | @{typ bool} => 2 (* optimization *)
   189         | @{typ nat} => 0 (* optimization *)
   190         | Type ("Int.int", []) => 0 (* optimization *)
   191         | Type (s, _) =>
   192           (case datatype_constrs thy T of
   193              constrs as _ :: _ =>
   194              let
   195                val constr_cards =
   196                  map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
   197                       o snd) constrs
   198              in
   199                if exists (curry (op =) 0) constr_cards then 0
   200                else Int.min (max, Integer.sum constr_cards)
   201              end
   202            | [] =>
   203              case Typedef.get_info ctxt s of
   204                ({abs_type, rep_type, ...}, _) :: _ =>
   205                (* We cheat here by assuming that typedef types are infinite if
   206                   their underlying type is infinite. This is unsound in general
   207                   but it's hard to think of a realistic example where this would
   208                   not be the case. We are also slack with representation types:
   209                   If a representation type has the form "sigma => tau", we
   210                   consider it enough to check "sigma" for infiniteness. (Look
   211                   for "slack" in this function.) *)
   212                (case varify_and_instantiate_type ctxt
   213                          (Logic.varifyT_global abs_type) T
   214                          (Logic.varifyT_global rep_type)
   215                      |> aux true avoid of
   216                   0 => 0
   217                 | 1 => 1
   218                 | _ => default_card)
   219              | [] => default_card)
   220           (* Very slightly unsound: Type variables are assumed not to be
   221              constrained to cardinality 1. (In practice, the user would most
   222              likely have used "unit" directly anyway.) *)
   223         | TFree _ =>
   224           if not sound andalso default_card = 1 then 2 else default_card
   225         | TVar _ => default_card
   226   in Int.min (max, aux false [] T) end
   227 
   228 fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt true [] 0 T <> 0
   229 fun is_type_surely_infinite ctxt sound infinite_Ts T =
   230   tiny_card_of_type ctxt sound (map (rpair 0) infinite_Ts) 1 T = 0
   231 
   232 (* Simple simplifications to ensure that sort annotations don't leave a trail of
   233    spurious "True"s. *)
   234 fun s_not (Const (@{const_name All}, T) $ Abs (s, T', t')) =
   235     Const (@{const_name Ex}, T) $ Abs (s, T', s_not t')
   236   | s_not (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
   237     Const (@{const_name All}, T) $ Abs (s, T', s_not t')
   238   | s_not (@{const HOL.implies} $ t1 $ t2) = @{const HOL.conj} $ t1 $ s_not t2
   239   | s_not (@{const HOL.conj} $ t1 $ t2) =
   240     @{const HOL.disj} $ s_not t1 $ s_not t2
   241   | s_not (@{const HOL.disj} $ t1 $ t2) =
   242     @{const HOL.conj} $ s_not t1 $ s_not t2
   243   | s_not (@{const False}) = @{const True}
   244   | s_not (@{const True}) = @{const False}
   245   | s_not (@{const Not} $ t) = t
   246   | s_not t = @{const Not} $ t
   247 fun s_conj (@{const True}, t2) = t2
   248   | s_conj (t1, @{const True}) = t1
   249   | s_conj p = HOLogic.mk_conj p
   250 fun s_disj (@{const False}, t2) = t2
   251   | s_disj (t1, @{const False}) = t1
   252   | s_disj p = HOLogic.mk_disj p
   253 fun s_imp (@{const True}, t2) = t2
   254   | s_imp (t1, @{const False}) = s_not t1
   255   | s_imp p = HOLogic.mk_imp p
   256 fun s_iff (@{const True}, t2) = t2
   257   | s_iff (t1, @{const True}) = t1
   258   | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
   259 
   260 fun close_form t =
   261   fold (fn ((x, i), T) => fn t' =>
   262            HOLogic.all_const T $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
   263        (Term.add_vars t []) t
   264 
   265 fun monomorphic_term subst =
   266   map_types (map_type_tvar (fn v =>
   267       case Type.lookup subst v of
   268         SOME typ => typ
   269       | NONE => TVar v))
   270 
   271 fun eta_expand _ t 0 = t
   272   | eta_expand Ts (Abs (s, T, t')) n =
   273     Abs (s, T, eta_expand (T :: Ts) t' (n - 1))
   274   | eta_expand Ts t n =
   275     fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t'))
   276              (List.take (binder_types (fastype_of1 (Ts, t)), n))
   277              (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0)))
   278 
   279 (* Converts an elim-rule into an equivalent theorem that does not have the
   280    predicate variable. Leaves other theorems unchanged. We simply instantiate
   281    the conclusion variable to "False". (Cf. "transform_elim_theorem" in
   282    "Meson_Clausify".) *)
   283 fun transform_elim_prop t =
   284   case Logic.strip_imp_concl t of
   285     @{const Trueprop} $ Var (z, @{typ bool}) =>
   286     subst_Vars [(z, @{const False})] t
   287   | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
   288   | _ => t
   289 
   290 fun specialize_type thy (s, T) t =
   291   let
   292     fun subst_for (Const (s', T')) =
   293       if s = s' then
   294         SOME (Sign.typ_match thy (T', T) Vartab.empty)
   295         handle Type.TYPE_MATCH => NONE
   296       else
   297         NONE
   298     | subst_for (t1 $ t2) =
   299       (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
   300     | subst_for (Abs (_, _, t')) = subst_for t'
   301     | subst_for _ = NONE
   302   in
   303     case subst_for t of
   304       SOME subst => monomorphic_term subst t
   305     | NONE => raise Type.TYPE_MATCH
   306   end
   307 
   308 fun strip_subgoal ctxt goal i =
   309   let
   310     val (t, (frees, params)) =
   311       Logic.goal_params (prop_of goal) i
   312       ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
   313     val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
   314     val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
   315   in (rev params, hyp_ts, concl_t) end
   316 
   317 end;