src/HOL/Tools/Sledgehammer/sledgehammer_util.ML
author blanchet
Fri, 27 May 2011 10:30:07 +0200
changeset 43845 20e9caff1f86
parent 43802 f30ae82cb62e
child 43846 c96f06bffd90
permissions -rw-r--r--
fix soundness bug in Sledgehammer: distinguish params in goals from fixed variables in context
blanchet@36062
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_util.ML
blanchet@35961
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@35961
     3
blanchet@35961
     4
General-purpose functions used by the Sledgehammer modules.
blanchet@35961
     5
*)
blanchet@35961
     6
blanchet@35961
     7
signature SLEDGEHAMMER_UTIL =
blanchet@35961
     8
sig
blanchet@36140
     9
  val plural_s : int -> string
blanchet@35961
    10
  val serial_commas : string -> string list -> string list
blanchet@38977
    11
  val simplify_spaces : string -> string
blanchet@35961
    12
  val parse_bool_option : bool -> string -> string -> bool option
blanchet@35961
    13
  val parse_time_option : string -> string -> Time.time option
blanchet@40582
    14
  val string_from_time : Time.time -> string
blanchet@36482
    15
  val nat_subscript : int -> string
blanchet@36474
    16
  val unyxml : string -> string
blanchet@36474
    17
  val maybe_quote : string -> string
blanchet@43550
    18
  val typ_of_dtyp :
blanchet@43550
    19
    Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
blanchet@43550
    20
    -> typ
blanchet@43567
    21
  val varify_type : Proof.context -> typ -> typ
blanchet@43567
    22
  val instantiate_type : theory -> typ -> typ -> typ -> typ
blanchet@43567
    23
  val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
blanchet@43595
    24
  val is_type_surely_finite : Proof.context -> typ -> bool
blanchet@43755
    25
  val is_type_surely_infinite : Proof.context -> typ list -> typ -> bool
blanchet@36553
    26
  val monomorphic_term : Type.tyenv -> term -> term
blanchet@38890
    27
  val eta_expand : typ list -> term -> int -> term
blanchet@43785
    28
  val transform_elim_prop : term -> term
blanchet@36553
    29
  val specialize_type : theory -> (string * typ) -> term -> term
blanchet@38290
    30
  val subgoal_count : Proof.state -> int
blanchet@43845
    31
  val strip_subgoal :
blanchet@43845
    32
    Proof.context -> thm -> int -> (string * typ) list * term list * term
blanchet@38935
    33
  val reserved_isar_keyword_table : unit -> unit Symtab.table
blanchet@35961
    34
end;
blanchet@39564
    35
blanchet@35961
    36
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
blanchet@35961
    37
struct
blanchet@35961
    38
blanchet@36140
    39
fun plural_s n = if n = 1 then "" else "s"
blanchet@36062
    40
blanchet@35961
    41
fun serial_commas _ [] = ["??"]
blanchet@35961
    42
  | serial_commas _ [s] = [s]
blanchet@35961
    43
  | serial_commas conj [s1, s2] = [s1, conj, s2]
blanchet@35961
    44
  | serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3]
blanchet@35961
    45
  | serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss
blanchet@35961
    46
blanchet@43802
    47
val simplify_spaces = ATP_Proof.strip_spaces false (K true)
blanchet@38199
    48
blanchet@35961
    49
fun parse_bool_option option name s =
blanchet@35961
    50
  (case s of
blanchet@35961
    51
     "smart" => if option then NONE else raise Option
blanchet@35961
    52
   | "false" => SOME false
blanchet@35961
    53
   | "true" => SOME true
blanchet@35961
    54
   | "" => SOME true
blanchet@35961
    55
   | _ => raise Option)
blanchet@35961
    56
  handle Option.Option =>
blanchet@35961
    57
         let val ss = map quote ((option ? cons "smart") ["true", "false"]) in
blanchet@35961
    58
           error ("Parameter " ^ quote name ^ " must be assigned " ^
blanchet@35961
    59
                  space_implode " " (serial_commas "or" ss) ^ ".")
blanchet@35961
    60
         end
blanchet@35961
    61
blanchet@40582
    62
val has_junk =
wenzelm@40875
    63
  exists (fn s => not (Symbol.is_digit s) andalso s <> ".") o raw_explode (* FIXME Symbol.explode (?) *)
blanchet@40582
    64
blanchet@35961
    65
fun parse_time_option _ "none" = NONE
blanchet@35961
    66
  | parse_time_option name s =
blanchet@40582
    67
    let val secs = if has_junk s then NONE else Real.fromString s in
blanchet@40582
    68
      if is_none secs orelse Real.<= (the secs, 0.0) then
blanchet@40582
    69
        error ("Parameter " ^ quote name ^ " must be assigned a positive \
blanchet@40582
    70
               \number of seconds (e.g., \"60\", \"0.5\") or \"none\".")
blanchet@35961
    71
      else
blanchet@40582
    72
        SOME (seconds (the secs))
blanchet@35961
    73
    end
blanchet@35961
    74
blanchet@40582
    75
fun string_from_time t =
blanchet@40582
    76
  string_of_real (0.01 * Real.fromInt (Time.toMilliseconds t div 10)) ^ " s"
blanchet@40582
    77
wenzelm@40875
    78
val subscript = implode o map (prefix "\<^isub>") o raw_explode  (* FIXME Symbol.explode (?) *)
blanchet@37318
    79
fun nat_subscript n =
blanchet@37318
    80
  n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
blanchet@36482
    81
wenzelm@39812
    82
val unyxml = XML.content_of o YXML.parse_body
blanchet@36474
    83
wenzelm@43162
    84
val is_long_identifier = forall Lexicon.is_identifier o space_explode "."
blanchet@36474
    85
fun maybe_quote y =
blanchet@36474
    86
  let val s = unyxml y in
blanchet@36474
    87
    y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
blanchet@36474
    88
           not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
wenzelm@36970
    89
           Keyword.is_keyword s) ? quote
blanchet@36474
    90
  end
blanchet@36474
    91
blanchet@43550
    92
fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
blanchet@43550
    93
    the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
blanchet@43550
    94
  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
blanchet@43550
    95
    Type (s, map (typ_of_dtyp descr typ_assoc) Us)
blanchet@43550
    96
  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
blanchet@43550
    97
    let val (s, ds, _) = the (AList.lookup (op =) descr i) in
blanchet@43550
    98
      Type (s, map (typ_of_dtyp descr typ_assoc) ds)
blanchet@43550
    99
    end
blanchet@43550
   100
blanchet@43567
   101
fun varify_type ctxt T =
blanchet@43567
   102
  Variable.polymorphic_types ctxt [Const (@{const_name undefined}, T)]
blanchet@43567
   103
  |> snd |> the_single |> dest_Const |> snd
blanchet@43567
   104
blanchet@43567
   105
(* TODO: use "Term_Subst.instantiateT" instead? *)
blanchet@43567
   106
fun instantiate_type thy T1 T1' T2 =
blanchet@43567
   107
  Same.commit (Envir.subst_type_same
blanchet@43567
   108
                   (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
blanchet@43567
   109
  handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], [])
blanchet@43567
   110
blanchet@43567
   111
fun varify_and_instantiate_type ctxt T1 T1' T2 =
blanchet@43567
   112
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43567
   113
    instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
blanchet@43567
   114
  end
blanchet@43567
   115
blanchet@43595
   116
fun datatype_constrs thy (T as Type (s, Ts)) =
blanchet@43595
   117
    (case Datatype.get_info thy s of
blanchet@43595
   118
       SOME {index, descr, ...} =>
blanchet@43595
   119
       let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in
blanchet@43595
   120
         map (apsnd (fn Us => map (typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T))
blanchet@43595
   121
             constrs
blanchet@43595
   122
       end
blanchet@43595
   123
     | NONE => [])
blanchet@43595
   124
  | datatype_constrs _ _ = []
blanchet@43595
   125
blanchet@43595
   126
(* Similar to "Nitpick_HOL.bounded_exact_card_of_type".
blanchet@43595
   127
   0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
blanchet@43595
   128
   cardinality 2 or more. The specified default cardinality is returned if the
blanchet@43595
   129
   cardinality of the type can't be determined. *)
blanchet@43755
   130
fun tiny_card_of_type ctxt default_card assigns T =
blanchet@43595
   131
  let
blanchet@43755
   132
    val thy = Proof_Context.theory_of ctxt
blanchet@43595
   133
    val max = 2 (* 1 would be too small for the "fun" case *)
blanchet@43595
   134
    fun aux slack avoid T =
blanchet@43754
   135
      if member (op =) avoid T then
blanchet@43754
   136
        0
blanchet@43755
   137
      else case AList.lookup (Sign.typ_instance thy o swap) assigns T of
blanchet@43755
   138
        SOME k => k
blanchet@43755
   139
      | NONE =>
blanchet@43755
   140
        case T of
blanchet@43755
   141
          Type (@{type_name fun}, [T1, T2]) =>
blanchet@43755
   142
          (case (aux slack avoid T1, aux slack avoid T2) of
blanchet@43755
   143
             (k, 1) => if slack andalso k = 0 then 0 else 1
blanchet@43755
   144
           | (0, _) => 0
blanchet@43755
   145
           | (_, 0) => 0
blanchet@43755
   146
           | (k1, k2) =>
blanchet@43755
   147
             if k1 >= max orelse k2 >= max then max
blanchet@43755
   148
             else Int.min (max, Integer.pow k2 k1))
blanchet@43755
   149
        | @{typ prop} => 2
blanchet@43755
   150
        | @{typ bool} => 2 (* optimization *)
blanchet@43755
   151
        | @{typ nat} => 0 (* optimization *)
blanchet@43755
   152
        | @{typ int} => 0 (* optimization *)
blanchet@43755
   153
        | Type (s, _) =>
blanchet@43755
   154
          (case datatype_constrs thy T of
blanchet@43755
   155
             constrs as _ :: _ =>
blanchet@43755
   156
             let
blanchet@43755
   157
               val constr_cards =
blanchet@43755
   158
                 map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
blanchet@43755
   159
                      o snd) constrs
blanchet@43755
   160
             in
blanchet@43755
   161
               if exists (curry (op =) 0) constr_cards then 0
blanchet@43755
   162
               else Int.min (max, Integer.sum constr_cards)
blanchet@43755
   163
             end
blanchet@43755
   164
           | [] =>
blanchet@43755
   165
             case Typedef.get_info ctxt s of
blanchet@43755
   166
               ({abs_type, rep_type, ...}, _) :: _ =>
blanchet@43755
   167
               (* We cheat here by assuming that typedef types are infinite if
blanchet@43755
   168
                  their underlying type is infinite. This is unsound in general
blanchet@43755
   169
                  but it's hard to think of a realistic example where this would
blanchet@43755
   170
                  not be the case. We are also slack with representation types:
blanchet@43755
   171
                  If a representation type has the form "sigma => tau", we
blanchet@43755
   172
                  consider it enough to check "sigma" for infiniteness. (Look
blanchet@43755
   173
                  for "slack" in this function.) *)
blanchet@43755
   174
               (case varify_and_instantiate_type ctxt
blanchet@43755
   175
                         (Logic.varifyT_global abs_type) T
blanchet@43755
   176
                         (Logic.varifyT_global rep_type)
blanchet@43755
   177
                     |> aux true avoid of
blanchet@43755
   178
                  0 => 0
blanchet@43755
   179
                | 1 => 1
blanchet@43755
   180
                | _ => default_card)
blanchet@43755
   181
             | [] => default_card)
blanchet@43755
   182
          (* Very slightly unsound: Type variables are assumed not to be
blanchet@43755
   183
             constrained to cardinality 1. (In practice, the user would most
blanchet@43755
   184
             likely have used "unit" directly anyway.) *)
blanchet@43755
   185
        | TFree _ => if default_card = 1 then 2 else default_card
blanchet@43755
   186
          (* Schematic type variables that contain only unproblematic sorts
blanchet@43755
   187
             (with no finiteness axiom) can safely be considered infinite. *)
blanchet@43755
   188
        | TVar _ => default_card
blanchet@43595
   189
  in Int.min (max, aux false [] T) end
blanchet@43595
   190
blanchet@43755
   191
fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt 0 [] T <> 0
blanchet@43755
   192
fun is_type_surely_infinite ctxt infinite_Ts T =
blanchet@43755
   193
  tiny_card_of_type ctxt 1 (map (rpair 0) infinite_Ts) T = 0
blanchet@43595
   194
blanchet@36553
   195
fun monomorphic_term subst t =
blanchet@36553
   196
  map_types (map_type_tvar (fn v =>
blanchet@36553
   197
      case Type.lookup subst v of
blanchet@36553
   198
        SOME typ => typ
blanchet@36553
   199
      | NONE => raise TERM ("monomorphic_term: uninstanitated schematic type \
blanchet@36553
   200
                            \variable", [t]))) t
blanchet@36553
   201
blanchet@38890
   202
fun eta_expand _ t 0 = t
blanchet@38890
   203
  | eta_expand Ts (Abs (s, T, t')) n =
blanchet@38890
   204
    Abs (s, T, eta_expand (T :: Ts) t' (n - 1))
blanchet@38890
   205
  | eta_expand Ts t n =
blanchet@38890
   206
    fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t'))
blanchet@38890
   207
             (List.take (binder_types (fastype_of1 (Ts, t)), n))
blanchet@38890
   208
             (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0)))
blanchet@38890
   209
blanchet@38890
   210
(* Converts an elim-rule into an equivalent theorem that does not have the
blanchet@38890
   211
   predicate variable. Leaves other theorems unchanged. We simply instantiate
blanchet@38890
   212
   the conclusion variable to False. (Cf. "transform_elim_theorem" in
blanchet@40071
   213
   "Meson_Clausify".) *)
blanchet@43785
   214
fun transform_elim_prop t =
blanchet@38890
   215
  case Logic.strip_imp_concl t of
blanchet@38890
   216
    @{const Trueprop} $ Var (z, @{typ bool}) =>
blanchet@38890
   217
    subst_Vars [(z, @{const False})] t
blanchet@38890
   218
  | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
blanchet@38890
   219
  | _ => t
blanchet@38890
   220
blanchet@36553
   221
fun specialize_type thy (s, T) t =
blanchet@36553
   222
  let
blanchet@36553
   223
    fun subst_for (Const (s', T')) =
blanchet@36553
   224
      if s = s' then
blanchet@36553
   225
        SOME (Sign.typ_match thy (T', T) Vartab.empty)
blanchet@36553
   226
        handle Type.TYPE_MATCH => NONE
blanchet@36553
   227
      else
blanchet@36553
   228
        NONE
blanchet@36553
   229
    | subst_for (t1 $ t2) =
blanchet@36553
   230
      (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
blanchet@36553
   231
    | subst_for (Abs (_, _, t')) = subst_for t'
blanchet@36553
   232
    | subst_for _ = NONE
blanchet@36553
   233
  in
blanchet@36553
   234
    case subst_for t of
blanchet@36553
   235
      SOME subst => monomorphic_term subst t
blanchet@36553
   236
    | NONE => raise Type.TYPE_MATCH
blanchet@36553
   237
  end
blanchet@36553
   238
blanchet@38290
   239
val subgoal_count = Logic.count_prems o prop_of o #goal o Proof.goal
blanchet@38290
   240
blanchet@43845
   241
fun strip_subgoal ctxt goal i =
blanchet@38229
   242
  let
blanchet@43845
   243
    val (t, (frees, params)) =
blanchet@43845
   244
      Logic.goal_params (prop_of goal) i
blanchet@43845
   245
      ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
blanchet@38229
   246
    val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
blanchet@38229
   247
    val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
blanchet@43845
   248
  in (rev params, hyp_ts, concl_t) end
blanchet@36553
   249
blanchet@38935
   250
fun reserved_isar_keyword_table () =
blanchet@38935
   251
  union (op =) (Keyword.dest_keywords ()) (Keyword.dest_commands ())
blanchet@38935
   252
  |> map (rpair ()) |> Symtab.make
blanchet@38935
   253
blanchet@35961
   254
end;