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