src/HOL/Tools/ATP/atp_util.ML
author blanchet
Mon, 06 Jun 2011 20:36:35 +0200
changeset 44012 37e1431cc213
parent 43926 0a2f5b86bdd7
child 44284 717880e98e6b
permissions -rw-r--r--
gracefully handle the case where a constant is partially or not instantiated at all, as may happen when reconstructing Metis proofs for polymorphic type encodings
blanchet@43926
     1
(*  Title:      HOL/Tools/ATP/atp_util.ML
blanchet@43926
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@43926
     3
blanchet@43926
     4
General-purpose functions used by the ATP module.
blanchet@43926
     5
*)
blanchet@43926
     6
blanchet@43926
     7
signature ATP_UTIL =
blanchet@43926
     8
sig
blanchet@43926
     9
  val timestamp : unit -> string
blanchet@43926
    10
  val hashw : word * word -> word
blanchet@43926
    11
  val hashw_string : string * word -> word
blanchet@43926
    12
  val strip_spaces : bool -> (char -> bool) -> string -> string
blanchet@43926
    13
  val nat_subscript : int -> string
blanchet@43926
    14
  val unyxml : string -> string
blanchet@43926
    15
  val maybe_quote : string -> string
blanchet@43926
    16
  val string_from_ext_time : bool * Time.time -> string
blanchet@43926
    17
  val string_from_time : Time.time -> string
blanchet@43926
    18
  val varify_type : Proof.context -> typ -> typ
blanchet@43926
    19
  val instantiate_type : theory -> typ -> typ -> typ -> typ
blanchet@43926
    20
  val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
blanchet@43926
    21
  val typ_of_dtyp :
blanchet@43926
    22
    Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
blanchet@43926
    23
    -> typ
blanchet@43926
    24
  val is_type_surely_finite : Proof.context -> typ -> bool
blanchet@43926
    25
  val is_type_surely_infinite : Proof.context -> typ list -> typ -> bool
blanchet@43926
    26
  val monomorphic_term : Type.tyenv -> term -> term
blanchet@43926
    27
  val eta_expand : typ list -> term -> int -> term
blanchet@43926
    28
  val transform_elim_prop : term -> term
blanchet@43926
    29
  val specialize_type : theory -> (string * typ) -> term -> term
blanchet@43926
    30
  val strip_subgoal :
blanchet@43926
    31
    Proof.context -> thm -> int -> (string * typ) list * term list * term
blanchet@43926
    32
end;
blanchet@43926
    33
blanchet@43926
    34
structure ATP_Util : ATP_UTIL =
blanchet@43926
    35
struct
blanchet@43926
    36
blanchet@43926
    37
val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
blanchet@43926
    38
blanchet@43926
    39
(* This hash function is recommended in "Compilers: Principles, Techniques, and
blanchet@43926
    40
   Tools" by Aho, Sethi, and Ullman. The "hashpjw" function, which they
blanchet@43926
    41
   particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
blanchet@43926
    42
fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
blanchet@43926
    43
fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
blanchet@43926
    44
fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s
blanchet@43926
    45
blanchet@43926
    46
fun strip_c_style_comment _ [] = []
blanchet@43926
    47
  | strip_c_style_comment is_evil (#"*" :: #"/" :: cs) =
blanchet@43926
    48
    strip_spaces_in_list true is_evil cs
blanchet@43926
    49
  | strip_c_style_comment is_evil (_ :: cs) = strip_c_style_comment is_evil cs
blanchet@43926
    50
and strip_spaces_in_list _ _ [] = []
blanchet@43926
    51
  | strip_spaces_in_list true is_evil (#"%" :: cs) =
blanchet@43926
    52
    strip_spaces_in_list true is_evil
blanchet@43926
    53
                         (cs |> chop_while (not_equal #"\n") |> snd)
blanchet@43926
    54
  | strip_spaces_in_list true is_evil (#"/" :: #"*" :: cs) =
blanchet@43926
    55
    strip_c_style_comment is_evil cs
blanchet@43926
    56
  | strip_spaces_in_list _ _ [c1] = if Char.isSpace c1 then [] else [str c1]
blanchet@43926
    57
  | strip_spaces_in_list skip_comments is_evil [c1, c2] =
blanchet@43926
    58
    strip_spaces_in_list skip_comments is_evil [c1] @
blanchet@43926
    59
    strip_spaces_in_list skip_comments is_evil [c2]
blanchet@43926
    60
  | strip_spaces_in_list skip_comments is_evil (c1 :: c2 :: c3 :: cs) =
blanchet@43926
    61
    if Char.isSpace c1 then
blanchet@43926
    62
      strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
blanchet@43926
    63
    else if Char.isSpace c2 then
blanchet@43926
    64
      if Char.isSpace c3 then
blanchet@43926
    65
        strip_spaces_in_list skip_comments is_evil (c1 :: c3 :: cs)
blanchet@43926
    66
      else
blanchet@43926
    67
        str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
blanchet@43926
    68
        strip_spaces_in_list skip_comments is_evil (c3 :: cs)
blanchet@43926
    69
    else
blanchet@43926
    70
      str c1 :: strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
blanchet@43926
    71
fun strip_spaces skip_comments is_evil =
blanchet@43926
    72
  implode o strip_spaces_in_list skip_comments is_evil o String.explode
blanchet@43926
    73
blanchet@43926
    74
val subscript = implode o map (prefix "\<^isub>") o raw_explode  (* FIXME Symbol.explode (?) *)
blanchet@43926
    75
fun nat_subscript n =
blanchet@43926
    76
  n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
blanchet@43926
    77
blanchet@43926
    78
val unyxml = XML.content_of o YXML.parse_body
blanchet@43926
    79
blanchet@43926
    80
val is_long_identifier = forall Lexicon.is_identifier o space_explode "."
blanchet@43926
    81
fun maybe_quote y =
blanchet@43926
    82
  let val s = unyxml y in
blanchet@43926
    83
    y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
blanchet@43926
    84
           not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
blanchet@43926
    85
           Keyword.is_keyword s) ? quote
blanchet@43926
    86
  end
blanchet@43926
    87
blanchet@43926
    88
fun string_from_ext_time (plus, time) =
blanchet@43926
    89
  let val ms = Time.toMilliseconds time in
blanchet@43926
    90
    (if plus then "> " else "") ^
blanchet@43926
    91
    (if plus andalso ms mod 1000 = 0 then
blanchet@43926
    92
       signed_string_of_int (ms div 1000) ^ " s"
blanchet@43926
    93
     else if ms < 1000 then
blanchet@43926
    94
       signed_string_of_int ms ^ " ms"
blanchet@43926
    95
     else
blanchet@43926
    96
       string_of_real (0.01 * Real.fromInt (ms div 10)) ^ " s")
blanchet@43926
    97
  end
blanchet@43926
    98
blanchet@43926
    99
val string_from_time = string_from_ext_time o pair false
blanchet@43926
   100
blanchet@43926
   101
fun varify_type ctxt T =
blanchet@43926
   102
  Variable.polymorphic_types ctxt [Const (@{const_name undefined}, T)]
blanchet@43926
   103
  |> snd |> the_single |> dest_Const |> snd
blanchet@43926
   104
blanchet@43926
   105
(* TODO: use "Term_Subst.instantiateT" instead? *)
blanchet@43926
   106
fun instantiate_type thy T1 T1' T2 =
blanchet@43926
   107
  Same.commit (Envir.subst_type_same
blanchet@43926
   108
                   (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
blanchet@43926
   109
  handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], [])
blanchet@43926
   110
blanchet@43926
   111
fun varify_and_instantiate_type ctxt T1 T1' T2 =
blanchet@43926
   112
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43926
   113
    instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
blanchet@43926
   114
  end
blanchet@43926
   115
blanchet@43926
   116
fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
blanchet@43926
   117
    the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
blanchet@43926
   118
  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
blanchet@43926
   119
    Type (s, map (typ_of_dtyp descr typ_assoc) Us)
blanchet@43926
   120
  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
blanchet@43926
   121
    let val (s, ds, _) = the (AList.lookup (op =) descr i) in
blanchet@43926
   122
      Type (s, map (typ_of_dtyp descr typ_assoc) ds)
blanchet@43926
   123
    end
blanchet@43926
   124
blanchet@43926
   125
fun datatype_constrs thy (T as Type (s, Ts)) =
blanchet@43926
   126
    (case Datatype.get_info thy s of
blanchet@43926
   127
       SOME {index, descr, ...} =>
blanchet@43926
   128
       let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in
blanchet@43926
   129
         map (apsnd (fn Us => map (typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T))
blanchet@43926
   130
             constrs
blanchet@43926
   131
       end
blanchet@43926
   132
     | NONE => [])
blanchet@43926
   133
  | datatype_constrs _ _ = []
blanchet@43926
   134
blanchet@43926
   135
(* Similar to "Nitpick_HOL.bounded_exact_card_of_type".
blanchet@43926
   136
   0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
blanchet@43926
   137
   cardinality 2 or more. The specified default cardinality is returned if the
blanchet@43926
   138
   cardinality of the type can't be determined. *)
blanchet@43926
   139
fun tiny_card_of_type ctxt default_card assigns T =
blanchet@43926
   140
  let
blanchet@43926
   141
    val thy = Proof_Context.theory_of ctxt
blanchet@43926
   142
    val max = 2 (* 1 would be too small for the "fun" case *)
blanchet@43926
   143
    fun aux slack avoid T =
blanchet@43926
   144
      if member (op =) avoid T then
blanchet@43926
   145
        0
blanchet@43926
   146
      else case AList.lookup (Sign.typ_instance thy o swap) assigns T of
blanchet@43926
   147
        SOME k => k
blanchet@43926
   148
      | NONE =>
blanchet@43926
   149
        case T of
blanchet@43926
   150
          Type (@{type_name fun}, [T1, T2]) =>
blanchet@43926
   151
          (case (aux slack avoid T1, aux slack avoid T2) of
blanchet@43926
   152
             (k, 1) => if slack andalso k = 0 then 0 else 1
blanchet@43926
   153
           | (0, _) => 0
blanchet@43926
   154
           | (_, 0) => 0
blanchet@43926
   155
           | (k1, k2) =>
blanchet@43926
   156
             if k1 >= max orelse k2 >= max then max
blanchet@43926
   157
             else Int.min (max, Integer.pow k2 k1))
blanchet@43926
   158
        | @{typ prop} => 2
blanchet@43926
   159
        | @{typ bool} => 2 (* optimization *)
blanchet@43926
   160
        | @{typ nat} => 0 (* optimization *)
blanchet@43926
   161
        | Type ("Int.int", []) => 0 (* optimization *)
blanchet@43926
   162
        | Type (s, _) =>
blanchet@43926
   163
          (case datatype_constrs thy T of
blanchet@43926
   164
             constrs as _ :: _ =>
blanchet@43926
   165
             let
blanchet@43926
   166
               val constr_cards =
blanchet@43926
   167
                 map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
blanchet@43926
   168
                      o snd) constrs
blanchet@43926
   169
             in
blanchet@43926
   170
               if exists (curry (op =) 0) constr_cards then 0
blanchet@43926
   171
               else Int.min (max, Integer.sum constr_cards)
blanchet@43926
   172
             end
blanchet@43926
   173
           | [] =>
blanchet@43926
   174
             case Typedef.get_info ctxt s of
blanchet@43926
   175
               ({abs_type, rep_type, ...}, _) :: _ =>
blanchet@43926
   176
               (* We cheat here by assuming that typedef types are infinite if
blanchet@43926
   177
                  their underlying type is infinite. This is unsound in general
blanchet@43926
   178
                  but it's hard to think of a realistic example where this would
blanchet@43926
   179
                  not be the case. We are also slack with representation types:
blanchet@43926
   180
                  If a representation type has the form "sigma => tau", we
blanchet@43926
   181
                  consider it enough to check "sigma" for infiniteness. (Look
blanchet@43926
   182
                  for "slack" in this function.) *)
blanchet@43926
   183
               (case varify_and_instantiate_type ctxt
blanchet@43926
   184
                         (Logic.varifyT_global abs_type) T
blanchet@43926
   185
                         (Logic.varifyT_global rep_type)
blanchet@43926
   186
                     |> aux true avoid of
blanchet@43926
   187
                  0 => 0
blanchet@43926
   188
                | 1 => 1
blanchet@43926
   189
                | _ => default_card)
blanchet@43926
   190
             | [] => default_card)
blanchet@43926
   191
          (* Very slightly unsound: Type variables are assumed not to be
blanchet@43926
   192
             constrained to cardinality 1. (In practice, the user would most
blanchet@43926
   193
             likely have used "unit" directly anyway.) *)
blanchet@43926
   194
        | TFree _ => if default_card = 1 then 2 else default_card
blanchet@43926
   195
          (* Schematic type variables that contain only unproblematic sorts
blanchet@43926
   196
             (with no finiteness axiom) can safely be considered infinite. *)
blanchet@43926
   197
        | TVar _ => default_card
blanchet@43926
   198
  in Int.min (max, aux false [] T) end
blanchet@43926
   199
blanchet@43926
   200
fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt 0 [] T <> 0
blanchet@43926
   201
fun is_type_surely_infinite ctxt infinite_Ts T =
blanchet@43926
   202
  tiny_card_of_type ctxt 1 (map (rpair 0) infinite_Ts) T = 0
blanchet@43926
   203
blanchet@44012
   204
fun monomorphic_term subst =
blanchet@43926
   205
  map_types (map_type_tvar (fn v =>
blanchet@43926
   206
      case Type.lookup subst v of
blanchet@43926
   207
        SOME typ => typ
blanchet@44012
   208
      | NONE => TVar v))
blanchet@43926
   209
blanchet@43926
   210
fun eta_expand _ t 0 = t
blanchet@43926
   211
  | eta_expand Ts (Abs (s, T, t')) n =
blanchet@43926
   212
    Abs (s, T, eta_expand (T :: Ts) t' (n - 1))
blanchet@43926
   213
  | eta_expand Ts t n =
blanchet@43926
   214
    fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t'))
blanchet@43926
   215
             (List.take (binder_types (fastype_of1 (Ts, t)), n))
blanchet@43926
   216
             (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0)))
blanchet@43926
   217
blanchet@43926
   218
(* Converts an elim-rule into an equivalent theorem that does not have the
blanchet@43926
   219
   predicate variable. Leaves other theorems unchanged. We simply instantiate
blanchet@43926
   220
   the conclusion variable to False. (Cf. "transform_elim_theorem" in
blanchet@43926
   221
   "Meson_Clausify".) *)
blanchet@43926
   222
fun transform_elim_prop t =
blanchet@43926
   223
  case Logic.strip_imp_concl t of
blanchet@43926
   224
    @{const Trueprop} $ Var (z, @{typ bool}) =>
blanchet@43926
   225
    subst_Vars [(z, @{const False})] t
blanchet@43926
   226
  | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
blanchet@43926
   227
  | _ => t
blanchet@43926
   228
blanchet@43926
   229
fun specialize_type thy (s, T) t =
blanchet@43926
   230
  let
blanchet@43926
   231
    fun subst_for (Const (s', T')) =
blanchet@43926
   232
      if s = s' then
blanchet@43926
   233
        SOME (Sign.typ_match thy (T', T) Vartab.empty)
blanchet@43926
   234
        handle Type.TYPE_MATCH => NONE
blanchet@43926
   235
      else
blanchet@43926
   236
        NONE
blanchet@43926
   237
    | subst_for (t1 $ t2) =
blanchet@43926
   238
      (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
blanchet@43926
   239
    | subst_for (Abs (_, _, t')) = subst_for t'
blanchet@43926
   240
    | subst_for _ = NONE
blanchet@43926
   241
  in
blanchet@43926
   242
    case subst_for t of
blanchet@43926
   243
      SOME subst => monomorphic_term subst t
blanchet@43926
   244
    | NONE => raise Type.TYPE_MATCH
blanchet@43926
   245
  end
blanchet@43926
   246
blanchet@43926
   247
fun strip_subgoal ctxt goal i =
blanchet@43926
   248
  let
blanchet@43926
   249
    val (t, (frees, params)) =
blanchet@43926
   250
      Logic.goal_params (prop_of goal) i
blanchet@43926
   251
      ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
blanchet@43926
   252
    val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
blanchet@43926
   253
    val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
blanchet@43926
   254
  in (rev params, hyp_ts, concl_t) end
blanchet@43926
   255
blanchet@43926
   256
end;