src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
author blanchet
Mon, 30 May 2011 17:00:38 +0200
changeset 43905 b6e61d22fa61
parent 43880 b967219cec78
permissions -rw-r--r--
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet@40358
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
blanchet@38506
     2
    Author:     Fabian Immler, TU Muenchen
blanchet@38506
     3
    Author:     Makarius
blanchet@38506
     4
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@38506
     5
blanchet@39734
     6
Translation of HOL to FOL for Sledgehammer.
blanchet@38506
     7
*)
blanchet@38506
     8
blanchet@40249
     9
signature SLEDGEHAMMER_ATP_TRANSLATE =
blanchet@38506
    10
sig
blanchet@43088
    11
  type 'a fo_term = 'a ATP_Problem.fo_term
blanchet@43780
    12
  type format = ATP_Problem.format
blanchet@43580
    13
  type formula_kind = ATP_Problem.formula_kind
blanchet@38506
    14
  type 'a problem = 'a ATP_Problem.problem
blanchet@43511
    15
  type locality = Sledgehammer_Filter.locality
blanchet@43484
    16
blanchet@43484
    17
  datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
blanchet@43484
    18
  datatype type_level =
blanchet@43484
    19
    All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types
blanchet@43708
    20
  datatype type_heaviness = Heavy | Light
blanchet@43484
    21
blanchet@43484
    22
  datatype type_system =
blanchet@43587
    23
    Simple_Types of type_level |
blanchet@43708
    24
    Preds of polymorphism * type_level * type_heaviness |
blanchet@43708
    25
    Tags of polymorphism * type_level * type_heaviness
blanchet@43484
    26
blanchet@40358
    27
  type translated_formula
blanchet@38506
    28
blanchet@43517
    29
  val readable_names : bool Config.T
blanchet@40445
    30
  val fact_prefix : string
blanchet@38506
    31
  val conjecture_prefix : string
blanchet@43750
    32
  val helper_prefix : string
blanchet@43750
    33
  val typed_helper_suffix : string
blanchet@43807
    34
  val predicator_name : string
blanchet@43807
    35
  val app_op_name : string
blanchet@43807
    36
  val type_pred_name : string
blanchet@43803
    37
  val simple_type_prefix : string
blanchet@43484
    38
  val type_sys_from_string : string -> type_system
blanchet@43484
    39
  val polymorphism_of_type_sys : type_system -> polymorphism
blanchet@43484
    40
  val level_of_type_sys : type_system -> type_level
blanchet@43484
    41
  val is_type_sys_virtually_sound : type_system -> bool
blanchet@43484
    42
  val is_type_sys_fairly_sound : type_system -> bool
blanchet@43413
    43
  val unmangled_const : string -> string * string fo_term list
blanchet@41336
    44
  val translate_atp_fact :
blanchet@43835
    45
    Proof.context -> format -> type_system -> bool -> (string * locality) * thm
blanchet@43511
    46
    -> translated_formula option * ((string * locality) * thm)
blanchet@40240
    47
  val prepare_atp_problem :
blanchet@43780
    48
    Proof.context -> format -> formula_kind -> formula_kind -> type_system
blanchet@43905
    49
    -> bool option -> term list -> term
blanchet@41339
    50
    -> (translated_formula option * ((string * 'a) * thm)) list
blanchet@43412
    51
    -> string problem * string Symtab.table * int * int
blanchet@43750
    52
       * (string * 'a) list vector * int list * int Symtab.table
blanchet@41561
    53
  val atp_problem_weights : string problem -> (string * real) list
blanchet@38506
    54
end;
blanchet@38506
    55
blanchet@41388
    56
structure Sledgehammer_ATP_Translate : SLEDGEHAMMER_ATP_TRANSLATE =
blanchet@38506
    57
struct
blanchet@38506
    58
blanchet@38506
    59
open ATP_Problem
blanchet@39734
    60
open Metis_Translate
blanchet@38506
    61
open Sledgehammer_Util
blanchet@43511
    62
open Sledgehammer_Filter
blanchet@43511
    63
blanchet@43511
    64
(* experimental *)
blanchet@43511
    65
val generate_useful_info = false
blanchet@38506
    66
blanchet@43748
    67
fun useful_isabelle_info s =
blanchet@43748
    68
  if generate_useful_info then
blanchet@43748
    69
    SOME (ATerm ("[]", [ATerm ("isabelle_" ^ s, [])]))
blanchet@43748
    70
  else
blanchet@43748
    71
    NONE
blanchet@43748
    72
blanchet@43748
    73
val intro_info = useful_isabelle_info "intro"
blanchet@43748
    74
val elim_info = useful_isabelle_info "elim"
blanchet@43748
    75
val simp_info = useful_isabelle_info "simp"
blanchet@43748
    76
blanchet@43439
    77
(* Readable names are often much shorter, especially if types are mangled in
blanchet@43460
    78
   names. Also, the logic for generating legal SNARK sort names is only
blanchet@43460
    79
   implemented for readable names. Finally, readable names are, well, more
blanchet@43460
    80
   readable. For these reason, they are enabled by default. *)
blanchet@43517
    81
val readable_names =
blanchet@43517
    82
  Attrib.setup_config_bool @{binding sledgehammer_atp_readable_names} (K true)
blanchet@43439
    83
blanchet@43839
    84
val type_decl_prefix = "ty_"
blanchet@43839
    85
val sym_decl_prefix = "sy_"
blanchet@43839
    86
val sym_formula_prefix = "sym_"
blanchet@40445
    87
val fact_prefix = "fact_"
blanchet@38506
    88
val conjecture_prefix = "conj_"
blanchet@38506
    89
val helper_prefix = "help_"
blanchet@43414
    90
val class_rel_clause_prefix = "crel_";
blanchet@38506
    91
val arity_clause_prefix = "arity_"
blanchet@40156
    92
val tfree_prefix = "tfree_"
blanchet@38506
    93
blanchet@43750
    94
val typed_helper_suffix = "_T"
blanchet@43750
    95
val untyped_helper_suffix = "_U"
blanchet@43750
    96
blanchet@43807
    97
val predicator_name = "hBOOL"
blanchet@43807
    98
val app_op_name = "hAPP"
blanchet@43807
    99
val type_pred_name = "is"
blanchet@43803
   100
val simple_type_prefix = "ty_"
blanchet@43402
   101
blanchet@43803
   102
fun make_simple_type s =
blanchet@43835
   103
  if s = tptp_bool_type orelse s = tptp_fun_type orelse
blanchet@43835
   104
     s = tptp_individual_type then
blanchet@43835
   105
    s
blanchet@43835
   106
  else
blanchet@43835
   107
    simple_type_prefix ^ ascii_of s
blanchet@43402
   108
blanchet@38506
   109
(* Freshness almost guaranteed! *)
blanchet@38506
   110
val sledgehammer_weak_prefix = "Sledgehammer:"
blanchet@38506
   111
blanchet@43484
   112
datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
blanchet@43484
   113
datatype type_level =
blanchet@43484
   114
  All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types
blanchet@43708
   115
datatype type_heaviness = Heavy | Light
blanchet@43484
   116
blanchet@43484
   117
datatype type_system =
blanchet@43587
   118
  Simple_Types of type_level |
blanchet@43708
   119
  Preds of polymorphism * type_level * type_heaviness |
blanchet@43708
   120
  Tags of polymorphism * type_level * type_heaviness
blanchet@43484
   121
blanchet@43559
   122
fun try_unsuffixes ss s =
blanchet@43559
   123
  fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
blanchet@43559
   124
blanchet@43484
   125
fun type_sys_from_string s =
blanchet@43587
   126
  (case try (unprefix "poly_") s of
blanchet@43587
   127
     SOME s => (SOME Polymorphic, s)
blanchet@43484
   128
   | NONE =>
blanchet@43484
   129
     case try (unprefix "mono_") s of
blanchet@43587
   130
       SOME s => (SOME Monomorphic, s)
blanchet@43587
   131
     | NONE =>
blanchet@43587
   132
       case try (unprefix "mangled_") s of
blanchet@43587
   133
         SOME s => (SOME Mangled_Monomorphic, s)
blanchet@43587
   134
       | NONE => (NONE, s))
blanchet@43484
   135
  ||> (fn s =>
blanchet@43559
   136
          (* "_query" and "_bang" are for the ASCII-challenged Mirabelle. *)
blanchet@43559
   137
          case try_unsuffixes ["?", "_query"] s of
blanchet@43484
   138
            SOME s => (Nonmonotonic_Types, s)
blanchet@43484
   139
          | NONE =>
blanchet@43559
   140
            case try_unsuffixes ["!", "_bang"] s of
blanchet@43484
   141
              SOME s => (Finite_Types, s)
blanchet@43484
   142
            | NONE => (All_Types, s))
blanchet@43699
   143
  ||> apsnd (fn s =>
blanchet@43708
   144
                case try (unsuffix "_heavy") s of
blanchet@43723
   145
                  SOME s => (Heavy, s)
blanchet@43723
   146
                | NONE => (Light, s))
blanchet@43708
   147
  |> (fn (poly, (level, (heaviness, core))) =>
blanchet@43708
   148
         case (core, (poly, level, heaviness)) of
blanchet@43724
   149
           ("simple", (NONE, _, Light)) => Simple_Types level
blanchet@43723
   150
         | ("preds", (SOME poly, _, _)) => Preds (poly, level, heaviness)
blanchet@43720
   151
         | ("tags", (SOME Polymorphic, All_Types, _)) =>
blanchet@43723
   152
           Tags (Polymorphic, All_Types, heaviness)
blanchet@43755
   153
         | ("tags", (SOME Polymorphic, _, _)) =>
blanchet@43755
   154
           (* The actual light encoding is very unsound. *)
blanchet@43755
   155
           Tags (Polymorphic, level, Heavy)
blanchet@43723
   156
         | ("tags", (SOME poly, _, _)) => Tags (poly, level, heaviness)
blanchet@43723
   157
         | ("args", (SOME poly, All_Types (* naja *), Light)) =>
blanchet@43708
   158
           Preds (poly, Const_Arg_Types, Light)
blanchet@43723
   159
         | ("erased", (NONE, All_Types (* naja *), Light)) =>
blanchet@43708
   160
           Preds (Polymorphic, No_Types, Light)
blanchet@43618
   161
         | _ => raise Same.SAME)
blanchet@43618
   162
  handle Same.SAME => error ("Unknown type system: " ^ quote s ^ ".")
blanchet@43484
   163
blanchet@43587
   164
fun polymorphism_of_type_sys (Simple_Types _) = Mangled_Monomorphic
blanchet@43699
   165
  | polymorphism_of_type_sys (Preds (poly, _, _)) = poly
blanchet@43699
   166
  | polymorphism_of_type_sys (Tags (poly, _, _)) = poly
blanchet@43484
   167
blanchet@43587
   168
fun level_of_type_sys (Simple_Types level) = level
blanchet@43699
   169
  | level_of_type_sys (Preds (_, level, _)) = level
blanchet@43699
   170
  | level_of_type_sys (Tags (_, level, _)) = level
blanchet@43699
   171
blanchet@43708
   172
fun heaviness_of_type_sys (Simple_Types _) = Heavy
blanchet@43708
   173
  | heaviness_of_type_sys (Preds (_, _, heaviness)) = heaviness
blanchet@43708
   174
  | heaviness_of_type_sys (Tags (_, _, heaviness)) = heaviness
blanchet@43702
   175
blanchet@43557
   176
fun is_type_level_virtually_sound level =
blanchet@43557
   177
  level = All_Types orelse level = Nonmonotonic_Types
blanchet@43484
   178
val is_type_sys_virtually_sound =
blanchet@43484
   179
  is_type_level_virtually_sound o level_of_type_sys
blanchet@43484
   180
blanchet@43484
   181
fun is_type_level_fairly_sound level =
blanchet@43484
   182
  is_type_level_virtually_sound level orelse level = Finite_Types
blanchet@43484
   183
val is_type_sys_fairly_sound = is_type_level_fairly_sound o level_of_type_sys
blanchet@43484
   184
blanchet@43835
   185
fun is_setting_higher_order THF (Simple_Types _) = true
blanchet@43835
   186
  | is_setting_higher_order _ _ = false
blanchet@43835
   187
blanchet@40358
   188
type translated_formula =
blanchet@38991
   189
  {name: string,
blanchet@43511
   190
   locality: locality,
blanchet@43396
   191
   kind: formula_kind,
blanchet@43433
   192
   combformula: (name, typ, combterm) formula,
blanchet@43433
   193
   atomic_types: typ list}
blanchet@38506
   194
blanchet@43511
   195
fun update_combformula f ({name, locality, kind, combformula, atomic_types}
blanchet@43511
   196
                          : translated_formula) =
blanchet@43511
   197
  {name = name, locality = locality, kind = kind, combformula = f combformula,
blanchet@43433
   198
   atomic_types = atomic_types} : translated_formula
blanchet@43413
   199
blanchet@43429
   200
fun fact_lift f ({combformula, ...} : translated_formula) = f combformula
blanchet@43429
   201
blanchet@43905
   202
val type_instance = Sign.typ_instance o Proof_Context.theory_of
blanchet@43905
   203
blanchet@43905
   204
fun insert_type ctxt get_T x xs =
blanchet@43905
   205
  let val T = get_T x in
blanchet@43905
   206
    if exists (curry (type_instance ctxt) T o get_T) xs then xs
blanchet@43905
   207
    else x :: filter_out (curry (type_instance ctxt o swap) T o get_T) xs
blanchet@43905
   208
  end
blanchet@43547
   209
blanchet@43618
   210
(* The Booleans indicate whether all type arguments should be kept. *)
blanchet@43618
   211
datatype type_arg_policy =
blanchet@43618
   212
  Explicit_Type_Args of bool |
blanchet@43618
   213
  Mangled_Type_Args of bool |
blanchet@43618
   214
  No_Type_Args
blanchet@41384
   215
blanchet@43707
   216
fun should_drop_arg_type_args (Simple_Types _) =
blanchet@43707
   217
    false (* since TFF doesn't support overloading *)
blanchet@43707
   218
  | should_drop_arg_type_args type_sys =
blanchet@43707
   219
    level_of_type_sys type_sys = All_Types andalso
blanchet@43708
   220
    heaviness_of_type_sys type_sys = Heavy
blanchet@43702
   221
blanchet@43460
   222
fun general_type_arg_policy type_sys =
blanchet@43460
   223
  if level_of_type_sys type_sys = No_Types then
blanchet@43460
   224
    No_Type_Args
blanchet@43460
   225
  else if polymorphism_of_type_sys type_sys = Mangled_Monomorphic then
blanchet@43702
   226
    Mangled_Type_Args (should_drop_arg_type_args type_sys)
blanchet@43460
   227
  else
blanchet@43702
   228
    Explicit_Type_Args (should_drop_arg_type_args type_sys)
blanchet@43434
   229
blanchet@43792
   230
fun type_arg_policy type_sys s =
blanchet@43792
   231
  if s = @{const_name HOL.eq} orelse
blanchet@43807
   232
     (s = app_op_name andalso level_of_type_sys type_sys = Const_Arg_Types) then
blanchet@43792
   233
    No_Type_Args
blanchet@43792
   234
  else
blanchet@43792
   235
    general_type_arg_policy type_sys
blanchet@43088
   236
blanchet@43797
   237
fun atp_type_literals_for_types format type_sys kind Ts =
blanchet@43797
   238
  if level_of_type_sys type_sys = No_Types orelse format = CNF_UEQ then
blanchet@43224
   239
    []
blanchet@43224
   240
  else
blanchet@43224
   241
    Ts |> type_literals_for_types
blanchet@43224
   242
       |> filter (fn TyLitVar _ => kind <> Conjecture
blanchet@43224
   243
                   | TyLitFree _ => kind = Conjecture)
blanchet@41385
   244
blanchet@43405
   245
fun mk_aconns c phis =
blanchet@43405
   246
  let val (phis', phi') = split_last phis in
blanchet@43405
   247
    fold_rev (mk_aconn c) phis' phi'
blanchet@43405
   248
  end
blanchet@38506
   249
fun mk_ahorn [] phi = phi
blanchet@43405
   250
  | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
blanchet@43393
   251
fun mk_aquant _ [] phi = phi
blanchet@43393
   252
  | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
blanchet@43393
   253
    if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
blanchet@43393
   254
  | mk_aquant q xs phi = AQuant (q, xs, phi)
blanchet@38506
   255
blanchet@43393
   256
fun close_universally atom_vars phi =
blanchet@41393
   257
  let
blanchet@41393
   258
    fun formula_vars bounds (AQuant (_, xs, phi)) =
blanchet@43397
   259
        formula_vars (map fst xs @ bounds) phi
blanchet@41393
   260
      | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
blanchet@43393
   261
      | formula_vars bounds (AAtom tm) =
blanchet@43397
   262
        union (op =) (atom_vars tm []
blanchet@43397
   263
                      |> filter_out (member (op =) bounds o fst))
blanchet@43393
   264
  in mk_aquant AForall (formula_vars [] phi []) phi end
blanchet@43393
   265
blanchet@43402
   266
fun combterm_vars (CombApp (tm1, tm2)) = fold combterm_vars [tm1, tm2]
blanchet@43393
   267
  | combterm_vars (CombConst _) = I
blanchet@43445
   268
  | combterm_vars (CombVar (name, T)) = insert (op =) (name, SOME T)
blanchet@43545
   269
fun close_combformula_universally phi = close_universally combterm_vars phi
blanchet@43393
   270
blanchet@43393
   271
fun term_vars (ATerm (name as (s, _), tms)) =
blanchet@43839
   272
  is_tptp_variable s ? insert (op =) (name, NONE) #> fold term_vars tms
blanchet@43545
   273
fun close_formula_universally phi = close_universally term_vars phi
blanchet@41393
   274
blanchet@43835
   275
val homo_infinite_type_name = @{type_name ind} (* any infinite type *)
blanchet@43835
   276
val homo_infinite_type = Type (homo_infinite_type_name, [])
blanchet@43835
   277
blanchet@43835
   278
fun fo_term_from_typ higher_order =
blanchet@43835
   279
  let
blanchet@43835
   280
    fun term (Type (s, Ts)) =
blanchet@43835
   281
      ATerm (case (higher_order, s) of
blanchet@43835
   282
               (true, @{type_name bool}) => `I tptp_bool_type
blanchet@43835
   283
             | (true, @{type_name fun}) => `I tptp_fun_type
blanchet@43835
   284
             | _ => if s = homo_infinite_type_name then `I tptp_individual_type
blanchet@43835
   285
                    else `make_fixed_type_const s,
blanchet@43835
   286
             map term Ts)
blanchet@43835
   287
    | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
blanchet@43835
   288
    | term (TVar ((x as (s, _)), _)) =
blanchet@43835
   289
      ATerm ((make_schematic_type_var x, s), [])
blanchet@43835
   290
  in term end
blanchet@43433
   291
blanchet@43433
   292
(* This shouldn't clash with anything else. *)
blanchet@43413
   293
val mangled_type_sep = "\000"
blanchet@43413
   294
blanchet@43433
   295
fun generic_mangled_type_name f (ATerm (name, [])) = f name
blanchet@43433
   296
  | generic_mangled_type_name f (ATerm (name, tys)) =
blanchet@43626
   297
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
blanchet@43626
   298
    ^ ")"
blanchet@43413
   299
blanchet@43839
   300
val bool_atype = AType (`I tptp_bool_type)
blanchet@43839
   301
blanchet@43835
   302
fun ho_type_from_fo_term higher_order pred_sym ary =
blanchet@43804
   303
  let
blanchet@43804
   304
    fun to_atype ty =
blanchet@43804
   305
      AType ((make_simple_type (generic_mangled_type_name fst ty),
blanchet@43804
   306
              generic_mangled_type_name snd ty))
blanchet@43804
   307
    fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
blanchet@43839
   308
    fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
blanchet@43835
   309
      | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
blanchet@43835
   310
    fun to_ho (ty as ATerm ((s, _), tys)) =
blanchet@43835
   311
      if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
blanchet@43835
   312
  in if higher_order then to_ho else to_fo ary end
blanchet@43804
   313
blanchet@43835
   314
fun mangled_type higher_order pred_sym ary =
blanchet@43835
   315
  ho_type_from_fo_term higher_order pred_sym ary o fo_term_from_typ higher_order
blanchet@43804
   316
blanchet@43835
   317
fun mangled_const_name T_args (s, s') =
blanchet@43804
   318
  let
blanchet@43835
   319
    val ty_args = map (fo_term_from_typ false) T_args
blanchet@43804
   320
    fun type_suffix f g =
blanchet@43804
   321
      fold_rev (curry (op ^) o g o prefix mangled_type_sep
blanchet@43804
   322
                o generic_mangled_type_name f) ty_args ""
blanchet@43804
   323
  in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
blanchet@43413
   324
blanchet@43413
   325
val parse_mangled_ident =
blanchet@43413
   326
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
blanchet@43413
   327
blanchet@43413
   328
fun parse_mangled_type x =
blanchet@43413
   329
  (parse_mangled_ident
blanchet@43413
   330
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
blanchet@43413
   331
                    [] >> ATerm) x
blanchet@43413
   332
and parse_mangled_types x =
blanchet@43413
   333
  (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
blanchet@43413
   334
blanchet@43413
   335
fun unmangled_type s =
blanchet@43413
   336
  s |> suffix ")" |> raw_explode
blanchet@43413
   337
    |> Scan.finite Symbol.stopper
blanchet@43413
   338
           (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
blanchet@43413
   339
                                                quote s)) parse_mangled_type))
blanchet@43413
   340
    |> fst
blanchet@43413
   341
blanchet@43432
   342
val unmangled_const_name = space_explode mangled_type_sep #> hd
blanchet@43413
   343
fun unmangled_const s =
blanchet@43413
   344
  let val ss = space_explode mangled_type_sep s in
blanchet@43413
   345
    (hd ss, map unmangled_type (tl ss))
blanchet@43413
   346
  end
blanchet@43413
   347
blanchet@43858
   348
fun introduce_proxies format type_sys =
blanchet@43439
   349
  let
blanchet@43858
   350
    fun intro top_level (CombApp (tm1, tm2)) =
blanchet@43858
   351
        CombApp (intro top_level tm1, intro false tm2)
blanchet@43858
   352
      | intro top_level (CombConst (name as (s, _), T, T_args)) =
blanchet@43441
   353
        (case proxify_const s of
blanchet@43858
   354
           SOME (_, proxy_base) =>
blanchet@43841
   355
           if top_level orelse is_setting_higher_order format type_sys then
blanchet@43841
   356
             case (top_level, s) of
blanchet@43841
   357
               (_, "c_False") => (`I tptp_false, [])
blanchet@43841
   358
             | (_, "c_True") => (`I tptp_true, [])
blanchet@43841
   359
             | (false, "c_Not") => (`I tptp_not, [])
blanchet@43841
   360
             | (false, "c_conj") => (`I tptp_and, [])
blanchet@43841
   361
             | (false, "c_disj") => (`I tptp_or, [])
blanchet@43841
   362
             | (false, "c_implies") => (`I tptp_implies, [])
blanchet@43841
   363
             | (false, s) =>
blanchet@43858
   364
               if is_tptp_equal s then (`I tptp_equal, [])
blanchet@43858
   365
               else (proxy_base |>> prefix const_prefix, T_args)
blanchet@43841
   366
             | _ => (name, [])
blanchet@43440
   367
           else
blanchet@43445
   368
             (proxy_base |>> prefix const_prefix, T_args)
blanchet@43445
   369
          | NONE => (name, T_args))
blanchet@43445
   370
        |> (fn (name, T_args) => CombConst (name, T, T_args))
blanchet@43858
   371
      | intro _ tm = tm
blanchet@43858
   372
  in intro true end
blanchet@43439
   373
blanchet@43835
   374
fun combformula_from_prop thy format type_sys eq_as_iff =
blanchet@38506
   375
  let
blanchet@43439
   376
    fun do_term bs t atomic_types =
blanchet@41388
   377
      combterm_from_term thy bs (Envir.eta_contract t)
blanchet@43835
   378
      |>> (introduce_proxies format type_sys #> AAtom)
blanchet@43439
   379
      ||> union (op =) atomic_types
blanchet@38506
   380
    fun do_quant bs q s T t' =
blanchet@38743
   381
      let val s = Name.variant (map fst bs) s in
blanchet@38743
   382
        do_formula ((s, T) :: bs) t'
blanchet@43433
   383
        #>> mk_aquant q [(`make_bound_var s, SOME T)]
blanchet@38743
   384
      end
blanchet@38506
   385
    and do_conn bs c t1 t2 =
blanchet@38506
   386
      do_formula bs t1 ##>> do_formula bs t2
blanchet@43402
   387
      #>> uncurry (mk_aconn c)
blanchet@38506
   388
    and do_formula bs t =
blanchet@38506
   389
      case t of
blanchet@43402
   390
        @{const Not} $ t1 => do_formula bs t1 #>> mk_anot
blanchet@38506
   391
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@38506
   392
        do_quant bs AForall s T t'
blanchet@38506
   393
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@38506
   394
        do_quant bs AExists s T t'
haftmann@39028
   395
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
haftmann@39028
   396
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
haftmann@39019
   397
      | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
haftmann@39093
   398
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@41388
   399
        if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t
blanchet@41388
   400
      | _ => do_term bs t
blanchet@38506
   401
  in do_formula [] end
blanchet@38506
   402
blanchet@43615
   403
fun presimplify_term ctxt =
blanchet@43615
   404
  Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
blanchet@43615
   405
  #> Meson.presimplify ctxt
blanchet@43615
   406
  #> prop_of
blanchet@38506
   407
wenzelm@41739
   408
fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j
blanchet@38506
   409
fun conceal_bounds Ts t =
blanchet@38506
   410
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@38506
   411
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@38506
   412
fun reveal_bounds Ts =
blanchet@38506
   413
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@38506
   414
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@38506
   415
blanchet@43612
   416
fun extensionalize_term ctxt t =
blanchet@43612
   417
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43612
   418
    t |> cterm_of thy |> Meson.extensionalize_conv ctxt
blanchet@43612
   419
      |> prop_of |> Logic.dest_equals |> snd
blanchet@43612
   420
  end
blanchet@38831
   421
blanchet@38506
   422
fun introduce_combinators_in_term ctxt kind t =
wenzelm@43232
   423
  let val thy = Proof_Context.theory_of ctxt in
blanchet@38716
   424
    if Meson.is_fol_term thy t then
blanchet@38716
   425
      t
blanchet@38716
   426
    else
blanchet@38716
   427
      let
blanchet@38716
   428
        fun aux Ts t =
blanchet@38716
   429
          case t of
blanchet@38716
   430
            @{const Not} $ t1 => @{const Not} $ aux Ts t1
blanchet@38716
   431
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@38716
   432
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   433
          | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@38890
   434
            aux Ts (t0 $ eta_expand Ts t1 1)
blanchet@38716
   435
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@38716
   436
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   437
          | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@38890
   438
            aux Ts (t0 $ eta_expand Ts t1 1)
haftmann@39028
   439
          | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39028
   440
          | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39019
   441
          | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39093
   442
          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
blanchet@38716
   443
              $ t1 $ t2 =>
blanchet@38716
   444
            t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   445
          | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
blanchet@38716
   446
                   t
blanchet@38716
   447
                 else
blanchet@38716
   448
                   t |> conceal_bounds Ts
blanchet@38716
   449
                     |> Envir.eta_contract
blanchet@38716
   450
                     |> cterm_of thy
blanchet@40071
   451
                     |> Meson_Clausify.introduce_combinators_in_cterm
blanchet@38716
   452
                     |> prop_of |> Logic.dest_equals |> snd
blanchet@38716
   453
                     |> reveal_bounds Ts
blanchet@39616
   454
        val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
blanchet@38716
   455
      in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@38716
   456
      handle THM _ =>
blanchet@38716
   457
             (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@38836
   458
             if kind = Conjecture then HOLogic.false_const
blanchet@38836
   459
             else HOLogic.true_const
blanchet@38716
   460
  end
blanchet@38506
   461
blanchet@38506
   462
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@43224
   463
   same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
blanchet@38506
   464
fun freeze_term t =
blanchet@38506
   465
  let
blanchet@38506
   466
    fun aux (t $ u) = aux t $ aux u
blanchet@38506
   467
      | aux (Abs (s, T, t)) = Abs (s, T, aux t)
blanchet@38506
   468
      | aux (Var ((s, i), T)) =
blanchet@38506
   469
        Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@38506
   470
      | aux t = t
blanchet@38506
   471
  in t |> exists_subterm is_Var t ? aux end
blanchet@38506
   472
blanchet@40445
   473
(* making fact and conjecture formulas *)
blanchet@43835
   474
fun make_formula ctxt format type_sys eq_as_iff presimp name loc kind t =
blanchet@38506
   475
  let
wenzelm@43232
   476
    val thy = Proof_Context.theory_of ctxt
blanchet@38831
   477
    val t = t |> Envir.beta_eta_contract
blanchet@43785
   478
              |> transform_elim_prop
blanchet@41459
   479
              |> Object_Logic.atomize_term thy
blanchet@43434
   480
    val need_trueprop = (fastype_of t = @{typ bool})
blanchet@38890
   481
    val t = t |> need_trueprop ? HOLogic.mk_Trueprop
blanchet@43607
   482
              |> Raw_Simplifier.rewrite_term thy
blanchet@43607
   483
                     (Meson.unfold_set_const_simps ctxt) []
blanchet@43612
   484
              |> extensionalize_term ctxt
blanchet@43615
   485
              |> presimp ? presimplify_term ctxt
blanchet@38506
   486
              |> perhaps (try (HOLogic.dest_Trueprop))
blanchet@38506
   487
              |> introduce_combinators_in_term ctxt kind
blanchet@38836
   488
              |> kind <> Axiom ? freeze_term
blanchet@43803
   489
    val (combformula, atomic_types) =
blanchet@43835
   490
      combformula_from_prop thy format type_sys eq_as_iff t []
blanchet@38506
   491
  in
blanchet@43511
   492
    {name = name, locality = loc, kind = kind, combformula = combformula,
blanchet@43433
   493
     atomic_types = atomic_types}
blanchet@38506
   494
  end
blanchet@38506
   495
blanchet@43835
   496
fun make_fact ctxt format type_sys keep_trivial eq_as_iff presimp
blanchet@43835
   497
              ((name, loc), t) =
blanchet@43803
   498
  case (keep_trivial,
blanchet@43835
   499
        make_formula ctxt format type_sys eq_as_iff presimp name loc Axiom t) of
blanchet@43842
   500
    (false, formula as {combformula = AAtom (CombConst ((s, _), _, _)), ...}) =>
blanchet@43842
   501
    if s = tptp_true then NONE else SOME formula
blanchet@42861
   502
  | (_, formula) => SOME formula
blanchet@43432
   503
blanchet@43835
   504
fun make_conjecture ctxt format prem_kind type_sys ts =
blanchet@38836
   505
  let val last = length ts - 1 in
blanchet@43580
   506
    map2 (fn j => fn t =>
blanchet@43580
   507
             let
blanchet@43580
   508
               val (kind, maybe_negate) =
blanchet@43580
   509
                 if j = last then
blanchet@43580
   510
                   (Conjecture, I)
blanchet@43580
   511
                 else
blanchet@43580
   512
                   (prem_kind,
blanchet@43580
   513
                    if prem_kind = Conjecture then update_combformula mk_anot
blanchet@43580
   514
                    else I)
blanchet@43580
   515
              in
blanchet@43835
   516
                t |> make_formula ctxt format type_sys true true
blanchet@43835
   517
                                  (string_of_int j) General kind
blanchet@43803
   518
                  |> maybe_negate
blanchet@43580
   519
              end)
blanchet@38836
   520
         (0 upto last) ts
blanchet@38836
   521
  end
blanchet@38506
   522
blanchet@43552
   523
(** Finite and infinite type inference **)
blanchet@43552
   524
blanchet@43755
   525
fun deep_freeze_atyp (TVar (_, S)) = TFree ("v", S)
blanchet@43755
   526
  | deep_freeze_atyp T = T
blanchet@43755
   527
val deep_freeze_type = map_atyps deep_freeze_atyp
blanchet@43755
   528
blanchet@43552
   529
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
blanchet@43552
   530
   dangerous because their "exhaust" properties can easily lead to unsound ATP
blanchet@43552
   531
   proofs. On the other hand, all HOL infinite types can be given the same
blanchet@43552
   532
   models in first-order logic (via Löwenheim-Skolem). *)
blanchet@43552
   533
blanchet@43755
   534
fun should_encode_type ctxt (nonmono_Ts as _ :: _) _ T =
blanchet@43755
   535
    exists (curry (type_instance ctxt) (deep_freeze_type T)) nonmono_Ts
blanchet@43707
   536
  | should_encode_type _ _ All_Types _ = true
blanchet@43552
   537
  | should_encode_type ctxt _ Finite_Types T = is_type_surely_finite ctxt T
blanchet@43552
   538
  | should_encode_type _ _ _ _ = false
blanchet@43552
   539
blanchet@43708
   540
fun should_predicate_on_type ctxt nonmono_Ts (Preds (_, level, heaviness))
blanchet@43705
   541
                             should_predicate_on_var T =
blanchet@43747
   542
    (heaviness = Heavy orelse should_predicate_on_var ()) andalso
blanchet@43747
   543
    should_encode_type ctxt nonmono_Ts level T
blanchet@43705
   544
  | should_predicate_on_type _ _ _ _ _ = false
blanchet@43552
   545
blanchet@43707
   546
fun is_var_or_bound_var (CombConst ((s, _), _, _)) =
blanchet@43707
   547
    String.isPrefix bound_var_prefix s
blanchet@43707
   548
  | is_var_or_bound_var (CombVar _) = true
blanchet@43707
   549
  | is_var_or_bound_var _ = false
blanchet@43707
   550
blanchet@43700
   551
datatype tag_site = Top_Level | Eq_Arg | Elsewhere
blanchet@43700
   552
blanchet@43700
   553
fun should_tag_with_type _ _ _ Top_Level _ _ = false
blanchet@43708
   554
  | should_tag_with_type ctxt nonmono_Ts (Tags (_, level, heaviness)) site u T =
blanchet@43708
   555
    (case heaviness of
blanchet@43708
   556
       Heavy => should_encode_type ctxt nonmono_Ts level T
blanchet@43708
   557
     | Light =>
blanchet@43707
   558
       case (site, is_var_or_bound_var u) of
blanchet@43707
   559
         (Eq_Arg, true) => should_encode_type ctxt nonmono_Ts level T
blanchet@43700
   560
       | _ => false)
blanchet@43700
   561
  | should_tag_with_type _ _ _ _ _ _ = false
blanchet@43552
   562
blanchet@43835
   563
fun homogenized_type ctxt nonmono_Ts level =
blanchet@43835
   564
  let
blanchet@43835
   565
    val should_encode = should_encode_type ctxt nonmono_Ts level
blanchet@43835
   566
    fun homo 0 T = if should_encode T then T else homo_infinite_type
blanchet@43835
   567
      | homo ary (Type (@{type_name fun}, [T1, T2])) =
blanchet@43835
   568
        homo 0 T1 --> homo (ary - 1) T2
blanchet@43835
   569
      | homo _ _ = raise Fail "expected function type"
blanchet@43835
   570
  in homo end
blanchet@43552
   571
blanchet@43444
   572
(** "hBOOL" and "hAPP" **)
blanchet@41561
   573
blanchet@43445
   574
type sym_info =
blanchet@43905
   575
  {pred_sym : bool, min_ary : int, max_ary : int, types : typ list}
blanchet@43434
   576
blanchet@43905
   577
fun add_combterm_syms_to_table ctxt explicit_apply =
blanchet@43429
   578
  let
blanchet@43905
   579
    fun consider_var_arity const_T var_T max_ary =
blanchet@43905
   580
      let
blanchet@43905
   581
        fun iter ary T =
blanchet@43905
   582
          if ary = max_ary orelse type_instance ctxt (var_T, T) then ary
blanchet@43905
   583
          else iter (ary + 1) (range_type T)
blanchet@43905
   584
      in iter 0 const_T end
blanchet@43905
   585
    fun add top_level tm (accum as (ho_var_Ts, sym_tab)) =
blanchet@43429
   586
      let val (head, args) = strip_combterm_comb tm in
blanchet@43429
   587
        (case head of
blanchet@43434
   588
           CombConst ((s, _), T, _) =>
blanchet@43429
   589
           if String.isPrefix bound_var_prefix s then
blanchet@43905
   590
             if explicit_apply = NONE andalso can dest_funT T then
blanchet@43905
   591
               let
blanchet@43905
   592
                 fun repair_min_arity {pred_sym, min_ary, max_ary, types} =
blanchet@43905
   593
                   {pred_sym = pred_sym,
blanchet@43905
   594
                    min_ary =
blanchet@43905
   595
                      fold (fn T' => consider_var_arity T' T) types min_ary,
blanchet@43905
   596
                    max_ary = max_ary, types = types}
blanchet@43905
   597
                 val ho_var_Ts' = ho_var_Ts |> insert_type ctxt I T
blanchet@43905
   598
               in
blanchet@43905
   599
                 if pointer_eq (ho_var_Ts', ho_var_Ts) then accum
blanchet@43905
   600
                 else (ho_var_Ts', Symtab.map (K repair_min_arity) sym_tab)
blanchet@43905
   601
               end
blanchet@43905
   602
             else
blanchet@43905
   603
               accum
blanchet@43429
   604
           else
blanchet@43905
   605
             let
blanchet@43905
   606
               val ary = length args
blanchet@43905
   607
             in
blanchet@43905
   608
               (ho_var_Ts,
blanchet@43905
   609
                case Symtab.lookup sym_tab s of
blanchet@43905
   610
                  SOME {pred_sym, min_ary, max_ary, types} =>
blanchet@43905
   611
                  let
blanchet@43905
   612
                    val types' = types |> insert_type ctxt I T
blanchet@43905
   613
                    val min_ary =
blanchet@43905
   614
                      if is_some explicit_apply orelse
blanchet@43905
   615
                         pointer_eq (types', types) then
blanchet@43905
   616
                        min_ary
blanchet@43905
   617
                      else
blanchet@43905
   618
                        fold (consider_var_arity T) ho_var_Ts min_ary
blanchet@43905
   619
                  in
blanchet@43905
   620
                    Symtab.update (s, {pred_sym = pred_sym andalso top_level,
blanchet@43905
   621
                                       min_ary = Int.min (ary, min_ary),
blanchet@43905
   622
                                       max_ary = Int.max (ary, max_ary),
blanchet@43905
   623
                                       types = types'})
blanchet@43905
   624
                                  sym_tab
blanchet@43905
   625
                  end
blanchet@43905
   626
                | NONE =>
blanchet@43905
   627
                  let
blanchet@43905
   628
                    val min_ary =
blanchet@43905
   629
                      case explicit_apply of
blanchet@43905
   630
                        SOME true => 0
blanchet@43905
   631
                      | SOME false => ary
blanchet@43905
   632
                      | NONE => fold (consider_var_arity T) ho_var_Ts ary
blanchet@43905
   633
                  in
blanchet@43905
   634
                    Symtab.update_new (s, {pred_sym = top_level,
blanchet@43905
   635
                                           min_ary = min_ary, max_ary = ary,
blanchet@43905
   636
                                           types = [T]})
blanchet@43905
   637
                                      sym_tab
blanchet@43905
   638
                  end)
blanchet@43905
   639
             end
blanchet@43905
   640
         | _ => accum)
blanchet@43905
   641
        |> fold (add false) args
blanchet@43429
   642
      end
blanchet@43905
   643
  in add true end
blanchet@43905
   644
fun add_fact_syms_to_table ctxt explicit_apply =
blanchet@43905
   645
  fact_lift (formula_fold NONE
blanchet@43905
   646
                          (K (add_combterm_syms_to_table ctxt explicit_apply)))
blanchet@38506
   647
blanchet@43546
   648
val default_sym_table_entries : (string * sym_info) list =
blanchet@43905
   649
  [(tptp_equal, {pred_sym = true, min_ary = 2, max_ary = 2, types = []}),
blanchet@43905
   650
   (tptp_old_equal, {pred_sym = true, min_ary = 2, max_ary = 2, types = []}),
blanchet@43807
   651
   (make_fixed_const predicator_name,
blanchet@43905
   652
    {pred_sym = true, min_ary = 1, max_ary = 1, types = []})] @
blanchet@43439
   653
  ([tptp_false, tptp_true]
blanchet@43905
   654
   |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = []}))
blanchet@41388
   655
blanchet@43905
   656
fun sym_table_for_facts ctxt explicit_apply facts =
blanchet@43905
   657
  Symtab.empty
blanchet@43905
   658
  |> fold Symtab.default default_sym_table_entries
blanchet@43905
   659
  |> pair [] |> fold (add_fact_syms_to_table ctxt explicit_apply) facts |> snd
blanchet@38506
   660
blanchet@43429
   661
fun min_arity_of sym_tab s =
blanchet@43429
   662
  case Symtab.lookup sym_tab s of
blanchet@43445
   663
    SOME ({min_ary, ...} : sym_info) => min_ary
blanchet@43429
   664
  | NONE =>
blanchet@43429
   665
    case strip_prefix_and_unascii const_prefix s of
blanchet@43418
   666
      SOME s =>
blanchet@43441
   667
      let val s = s |> unmangled_const_name |> invert_const in
blanchet@43807
   668
        if s = predicator_name then 1
blanchet@43807
   669
        else if s = app_op_name then 2
blanchet@43807
   670
        else if s = type_pred_name then 1
blanchet@43428
   671
        else 0
blanchet@43418
   672
      end
blanchet@38506
   673
    | NONE => 0
blanchet@38506
   674
blanchet@38506
   675
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
   676
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
   677
   type instantiations). If false, the constant always receives all of its
blanchet@38506
   678
   arguments and is used as a predicate. *)
blanchet@43429
   679
fun is_pred_sym sym_tab s =
blanchet@43429
   680
  case Symtab.lookup sym_tab s of
blanchet@43445
   681
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
blanchet@43445
   682
    pred_sym andalso min_ary = max_ary
blanchet@43429
   683
  | NONE => false
blanchet@38506
   684
blanchet@43439
   685
val predicator_combconst =
blanchet@43807
   686
  CombConst (`make_fixed_const predicator_name, @{typ "bool => bool"}, [])
blanchet@43439
   687
fun predicator tm = CombApp (predicator_combconst, tm)
blanchet@38506
   688
blanchet@43439
   689
fun introduce_predicators_in_combterm sym_tab tm =
blanchet@43413
   690
  case strip_combterm_comb tm of
blanchet@43413
   691
    (CombConst ((s, _), _, _), _) =>
blanchet@43439
   692
    if is_pred_sym sym_tab s then tm else predicator tm
blanchet@43439
   693
  | _ => predicator tm
blanchet@38506
   694
blanchet@43415
   695
fun list_app head args = fold (curry (CombApp o swap)) args head
blanchet@38506
   696
blanchet@43415
   697
fun explicit_app arg head =
blanchet@43415
   698
  let
blanchet@43433
   699
    val head_T = combtyp_of head
blanchet@43563
   700
    val (arg_T, res_T) = dest_funT head_T
blanchet@43415
   701
    val explicit_app =
blanchet@43807
   702
      CombConst (`make_fixed_const app_op_name, head_T --> head_T,
blanchet@43563
   703
                 [arg_T, res_T])
blanchet@43415
   704
  in list_app explicit_app [head, arg] end
blanchet@43415
   705
fun list_explicit_app head args = fold explicit_app args head
blanchet@43415
   706
blanchet@43436
   707
fun introduce_explicit_apps_in_combterm sym_tab =
blanchet@43415
   708
  let
blanchet@43415
   709
    fun aux tm =
blanchet@43415
   710
      case strip_combterm_comb tm of
blanchet@43415
   711
        (head as CombConst ((s, _), _, _), args) =>
blanchet@43415
   712
        args |> map aux
blanchet@43428
   713
             |> chop (min_arity_of sym_tab s)
blanchet@43415
   714
             |>> list_app head
blanchet@43415
   715
             |-> list_explicit_app
blanchet@43415
   716
      | (head, args) => list_explicit_app head (map aux args)
blanchet@43415
   717
  in aux end
blanchet@43415
   718
blanchet@43618
   719
fun chop_fun 0 T = ([], T)
blanchet@43618
   720
  | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
blanchet@43618
   721
    chop_fun (n - 1) ran_T |>> cons dom_T
blanchet@43618
   722
  | chop_fun _ _ = raise Fail "unexpected non-function"
blanchet@43618
   723
blanchet@43651
   724
fun filter_type_args _ _ _ [] = []
blanchet@43651
   725
  | filter_type_args thy s arity T_args =
blanchet@43705
   726
    let
blanchet@43705
   727
      (* will throw "TYPE" for pseudo-constants *)
blanchet@43807
   728
      val U = if s = app_op_name then
blanchet@43705
   729
                @{typ "('a => 'b) => 'a => 'b"} |> Logic.varifyT_global
blanchet@43705
   730
              else
blanchet@43705
   731
                s |> Sign.the_const_type thy
blanchet@43705
   732
    in
blanchet@43652
   733
      case Term.add_tvarsT (U |> chop_fun arity |> snd) [] of
blanchet@43652
   734
        [] => []
blanchet@43652
   735
      | res_U_vars =>
blanchet@43652
   736
        let val U_args = (s, U) |> Sign.const_typargs thy in
blanchet@43652
   737
          U_args ~~ T_args
blanchet@43652
   738
          |> map_filter (fn (U, T) =>
blanchet@43652
   739
                            if member (op =) res_U_vars (dest_TVar U) then
blanchet@43652
   740
                              SOME T
blanchet@43652
   741
                            else
blanchet@43652
   742
                              NONE)
blanchet@43652
   743
        end
blanchet@43651
   744
    end
blanchet@43651
   745
    handle TYPE _ => T_args
blanchet@43618
   746
blanchet@43835
   747
fun enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys =
blanchet@43618
   748
  let
blanchet@43618
   749
    val thy = Proof_Context.theory_of ctxt
blanchet@43618
   750
    fun aux arity (CombApp (tm1, tm2)) =
blanchet@43618
   751
        CombApp (aux (arity + 1) tm1, aux 0 tm2)
blanchet@43618
   752
      | aux arity (CombConst (name as (s, _), T, T_args)) =
blanchet@43571
   753
        let
blanchet@43571
   754
          val level = level_of_type_sys type_sys
blanchet@43571
   755
          val (T, T_args) =
blanchet@43571
   756
            (* Aggressively merge most "hAPPs" if the type system is unsound
blanchet@43571
   757
               anyway, by distinguishing overloads only on the homogenized
blanchet@43708
   758
               result type. Don't do it for lightweight type systems, though,
blanchet@43708
   759
               since it leads to too many unsound proofs. *)
blanchet@43807
   760
            if s = const_prefix ^ app_op_name andalso
blanchet@43591
   761
               length T_args = 2 andalso
blanchet@43707
   762
               not (is_type_sys_virtually_sound type_sys) andalso
blanchet@43708
   763
               heaviness_of_type_sys type_sys = Heavy then
blanchet@43835
   764
              T_args |> map (homogenized_type ctxt nonmono_Ts level 0)
blanchet@43571
   765
                     |> (fn Ts => let val T = hd Ts --> nth Ts 1 in
blanchet@43702
   766
                                    (T --> T, tl Ts)
blanchet@43571
   767
                                  end)
blanchet@43571
   768
            else
blanchet@43571
   769
              (T, T_args)
blanchet@43571
   770
        in
blanchet@43571
   771
          (case strip_prefix_and_unascii const_prefix s of
blanchet@43571
   772
             NONE => (name, T_args)
blanchet@43571
   773
           | SOME s'' =>
blanchet@43618
   774
             let
blanchet@43618
   775
               val s'' = invert_const s''
blanchet@43702
   776
               fun filtered_T_args false = T_args
blanchet@43702
   777
                 | filtered_T_args true = filter_type_args thy s'' arity T_args
blanchet@43618
   778
             in
blanchet@43571
   779
               case type_arg_policy type_sys s'' of
blanchet@43702
   780
                 Explicit_Type_Args drop_args =>
blanchet@43702
   781
                 (name, filtered_T_args drop_args)
blanchet@43702
   782
               | Mangled_Type_Args drop_args =>
blanchet@43835
   783
                 (mangled_const_name (filtered_T_args drop_args) name, [])
blanchet@43618
   784
               | No_Type_Args => (name, [])
blanchet@43571
   785
             end)
blanchet@43571
   786
          |> (fn (name, T_args) => CombConst (name, T, T_args))
blanchet@43571
   787
        end
blanchet@43618
   788
      | aux _ tm = tm
blanchet@43618
   789
  in aux 0 end
blanchet@43444
   790
blanchet@43803
   791
fun repair_combterm ctxt format nonmono_Ts type_sys sym_tab =
blanchet@43835
   792
  not (is_setting_higher_order format type_sys)
blanchet@43835
   793
  ? (introduce_explicit_apps_in_combterm sym_tab
blanchet@43835
   794
     #> introduce_predicators_in_combterm sym_tab)
blanchet@43835
   795
  #> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43803
   796
fun repair_fact ctxt format nonmono_Ts type_sys sym_tab =
blanchet@43571
   797
  update_combformula (formula_map
blanchet@43803
   798
      (repair_combterm ctxt format nonmono_Ts type_sys sym_tab))
blanchet@43444
   799
blanchet@43444
   800
(** Helper facts **)
blanchet@43444
   801
blanchet@43444
   802
fun ti_ti_helper_fact () =
blanchet@43444
   803
  let
blanchet@43444
   804
    fun var s = ATerm (`I s, [])
blanchet@43460
   805
    fun tag tm = ATerm (`make_fixed_const type_tag_name, [var "X", tm])
blanchet@43444
   806
  in
blanchet@43483
   807
    Formula (helper_prefix ^ "ti_ti", Axiom,
blanchet@43841
   808
             AAtom (ATerm (`I tptp_equal, [tag (tag (var "Y")), tag (var "Y")]))
blanchet@43748
   809
             |> close_formula_universally, simp_info, NONE)
blanchet@43444
   810
  end
blanchet@43444
   811
blanchet@43905
   812
fun helper_facts_for_sym ctxt format type_sys (s, {types, ...} : sym_info) =
blanchet@43444
   813
  case strip_prefix_and_unascii const_prefix s of
blanchet@43444
   814
    SOME mangled_s =>
blanchet@43444
   815
    let
blanchet@43444
   816
      val thy = Proof_Context.theory_of ctxt
blanchet@43444
   817
      val unmangled_s = mangled_s |> unmangled_const_name
blanchet@43764
   818
      fun dub_and_inst c needs_fairly_sound (th, j) =
blanchet@43750
   819
        ((c ^ "_" ^ string_of_int j ^
blanchet@43764
   820
          (if needs_fairly_sound then typed_helper_suffix
blanchet@43750
   821
           else untyped_helper_suffix),
blanchet@43750
   822
          General),
blanchet@43444
   823
         let val t = th |> prop_of in
blanchet@43618
   824
           t |> ((case general_type_arg_policy type_sys of
blanchet@43618
   825
                    Mangled_Type_Args _ => true
blanchet@43618
   826
                  | _ => false) andalso
blanchet@43444
   827
                 not (null (Term.hidden_polymorphism t)))
blanchet@43905
   828
                ? (case types of
blanchet@43905
   829
                     [T] => specialize_type thy (invert_const unmangled_s, T)
blanchet@43905
   830
                   | _ => I)
blanchet@43444
   831
         end)
blanchet@43444
   832
      fun make_facts eq_as_iff =
blanchet@43835
   833
        map_filter (make_fact ctxt format type_sys false eq_as_iff false)
blanchet@43764
   834
      val fairly_sound = is_type_sys_fairly_sound type_sys
blanchet@43444
   835
    in
blanchet@43444
   836
      metis_helpers
blanchet@43765
   837
      |> maps (fn (metis_s, (needs_fairly_sound, ths)) =>
blanchet@43444
   838
                  if metis_s <> unmangled_s orelse
blanchet@43765
   839
                     (needs_fairly_sound andalso not fairly_sound) then
blanchet@43444
   840
                    []
blanchet@43444
   841
                  else
blanchet@43444
   842
                    ths ~~ (1 upto length ths)
blanchet@43764
   843
                    |> map (dub_and_inst mangled_s needs_fairly_sound)
blanchet@43764
   844
                    |> make_facts (not needs_fairly_sound))
blanchet@43444
   845
    end
blanchet@43444
   846
  | NONE => []
blanchet@43803
   847
fun helper_facts_for_sym_table ctxt format type_sys sym_tab =
blanchet@43803
   848
  Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_sys) sym_tab
blanchet@43803
   849
                  []
blanchet@43444
   850
blanchet@43835
   851
fun translate_atp_fact ctxt format type_sys keep_trivial =
blanchet@43835
   852
  `(make_fact ctxt format type_sys keep_trivial true true o apsnd prop_of)
blanchet@43444
   853
blanchet@43803
   854
fun translate_formulas ctxt format prem_kind type_sys hyp_ts concl_t
blanchet@43803
   855
                       rich_facts =
blanchet@43444
   856
  let
blanchet@43444
   857
    val thy = Proof_Context.theory_of ctxt
blanchet@43444
   858
    val fact_ts = map (prop_of o snd o snd) rich_facts
blanchet@43444
   859
    val (facts, fact_names) =
blanchet@43444
   860
      rich_facts
blanchet@43444
   861
      |> map_filter (fn (NONE, _) => NONE
blanchet@43444
   862
                      | (SOME fact, (name, _)) => SOME (fact, name))
blanchet@43444
   863
      |> ListPair.unzip
blanchet@43444
   864
    (* Remove existing facts from the conjecture, as this can dramatically
blanchet@43444
   865
       boost an ATP's performance (for some reason). *)
blanchet@43444
   866
    val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts)
blanchet@43444
   867
    val goal_t = Logic.list_implies (hyp_ts, concl_t)
blanchet@43444
   868
    val all_ts = goal_t :: fact_ts
blanchet@43444
   869
    val subs = tfree_classes_of_terms all_ts
blanchet@43444
   870
    val supers = tvar_classes_of_terms all_ts
blanchet@43444
   871
    val tycons = type_consts_of_terms thy all_ts
blanchet@43835
   872
    val conjs =
blanchet@43835
   873
      hyp_ts @ [concl_t] |> make_conjecture ctxt format prem_kind type_sys
blanchet@43444
   874
    val (supers', arity_clauses) =
blanchet@43460
   875
      if level_of_type_sys type_sys = No_Types then ([], [])
blanchet@43444
   876
      else make_arity_clauses thy tycons supers
blanchet@43444
   877
    val class_rel_clauses = make_class_rel_clauses thy subs supers'
blanchet@43444
   878
  in
blanchet@43444
   879
    (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses))
blanchet@43444
   880
  end
blanchet@43444
   881
blanchet@43444
   882
fun fo_literal_from_type_literal (TyLitVar (class, name)) =
blanchet@43444
   883
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@43444
   884
  | fo_literal_from_type_literal (TyLitFree (class, name)) =
blanchet@43444
   885
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@43444
   886
blanchet@43444
   887
fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
blanchet@43444
   888
blanchet@43835
   889
fun type_pred_combterm ctxt nonmono_Ts type_sys T tm =
blanchet@43807
   890
  CombApp (CombConst (`make_fixed_const type_pred_name, T --> @{typ bool}, [T])
blanchet@43835
   891
           |> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys,
blanchet@43444
   892
           tm)
blanchet@43444
   893
blanchet@43747
   894
fun var_occurs_positively_naked_in_term _ (SOME false) _ accum = accum
blanchet@43747
   895
  | var_occurs_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
blanchet@43841
   896
    accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
blanchet@43747
   897
fun is_var_nonmonotonic_in_formula _ _ (SOME false) _ = false
blanchet@43747
   898
  | is_var_nonmonotonic_in_formula pos phi _ name =
blanchet@43747
   899
    formula_fold pos (var_occurs_positively_naked_in_term name) phi false
blanchet@43705
   900
blanchet@43835
   901
fun mk_const_aterm x T_args args =
blanchet@43835
   902
  ATerm (x, map (fo_term_from_typ false) T_args @ args)
blanchet@43835
   903
blanchet@43803
   904
fun tag_with_type ctxt format nonmono_Ts type_sys T tm =
blanchet@43700
   905
  CombConst (`make_fixed_const type_tag_name, T --> T, [T])
blanchet@43835
   906
  |> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43803
   907
  |> term_from_combterm ctxt format nonmono_Ts type_sys Top_Level
blanchet@43700
   908
  |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm]))
blanchet@43803
   909
and term_from_combterm ctxt format nonmono_Ts type_sys =
blanchet@43444
   910
  let
blanchet@43803
   911
    fun aux site u =
blanchet@43803
   912
      let
blanchet@43803
   913
        val (head, args) = strip_combterm_comb u
blanchet@43803
   914
        val (x as (s, _), T_args) =
blanchet@43803
   915
          case head of
blanchet@43803
   916
            CombConst (name, _, T_args) => (name, T_args)
blanchet@43803
   917
          | CombVar (name, _) => (name, [])
blanchet@43803
   918
          | CombApp _ => raise Fail "impossible \"CombApp\""
blanchet@43841
   919
        val arg_site = if site = Top_Level andalso is_tptp_equal s then Eq_Arg
blanchet@43803
   920
                       else Elsewhere
blanchet@43835
   921
        val t = mk_const_aterm x T_args (map (aux arg_site) args)
blanchet@43803
   922
        val T = combtyp_of u
blanchet@43803
   923
      in
blanchet@43803
   924
        t |> (if should_tag_with_type ctxt nonmono_Ts type_sys site u T then
blanchet@43803
   925
                tag_with_type ctxt format nonmono_Ts type_sys T
blanchet@43803
   926
              else
blanchet@43803
   927
                I)
blanchet@43803
   928
      end
blanchet@43803
   929
  in aux end
blanchet@43803
   930
and formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43803
   931
                             should_predicate_on_var =
blanchet@43700
   932
  let
blanchet@43835
   933
    val higher_order = is_setting_higher_order format type_sys
blanchet@43803
   934
    val do_term = term_from_combterm ctxt format nonmono_Ts type_sys Top_Level
blanchet@43444
   935
    val do_bound_type =
blanchet@43552
   936
      case type_sys of
blanchet@43587
   937
        Simple_Types level =>
blanchet@43835
   938
        homogenized_type ctxt nonmono_Ts level 0
blanchet@43835
   939
        #> mangled_type higher_order false 0 #> SOME
blanchet@43552
   940
      | _ => K NONE
blanchet@43747
   941
    fun do_out_of_bound_type pos phi universal (name, T) =
blanchet@43705
   942
      if should_predicate_on_type ctxt nonmono_Ts type_sys
blanchet@43747
   943
             (fn () => should_predicate_on_var pos phi universal name) T then
blanchet@43705
   944
        CombVar (name, T)
blanchet@43835
   945
        |> type_pred_combterm ctxt nonmono_Ts type_sys T
blanchet@43747
   946
        |> do_term |> AAtom |> SOME
blanchet@43444
   947
      else
blanchet@43444
   948
        NONE
blanchet@43747
   949
    fun do_formula pos (AQuant (q, xs, phi)) =
blanchet@43747
   950
        let
blanchet@43747
   951
          val phi = phi |> do_formula pos
blanchet@43747
   952
          val universal = Option.map (q = AExists ? not) pos
blanchet@43747
   953
        in
blanchet@43705
   954
          AQuant (q, xs |> map (apsnd (fn NONE => NONE
blanchet@43705
   955
                                        | SOME T => do_bound_type T)),
blanchet@43705
   956
                  (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
blanchet@43705
   957
                      (map_filter
blanchet@43705
   958
                           (fn (_, NONE) => NONE
blanchet@43705
   959
                             | (s, SOME T) =>
blanchet@43747
   960
                               do_out_of_bound_type pos phi universal (s, T))
blanchet@43747
   961
                           xs)
blanchet@43705
   962
                      phi)
blanchet@43705
   963
        end
blanchet@43747
   964
      | do_formula pos (AConn conn) = aconn_map pos do_formula conn
blanchet@43747
   965
      | do_formula _ (AAtom tm) = AAtom (do_term tm)
blanchet@43747
   966
  in do_formula o SOME end
blanchet@43444
   967
blanchet@43797
   968
fun bound_atomic_types format type_sys Ts =
blanchet@43592
   969
  mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal)
blanchet@43797
   970
                (atp_type_literals_for_types format type_sys Axiom Ts))
blanchet@43592
   971
blanchet@43797
   972
fun formula_for_fact ctxt format nonmono_Ts type_sys
blanchet@43444
   973
                     ({combformula, atomic_types, ...} : translated_formula) =
blanchet@43592
   974
  combformula
blanchet@43592
   975
  |> close_combformula_universally
blanchet@43803
   976
  |> formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43747
   977
                              is_var_nonmonotonic_in_formula true
blanchet@43797
   978
  |> bound_atomic_types format type_sys atomic_types
blanchet@43444
   979
  |> close_formula_universally
blanchet@43444
   980
blanchet@43444
   981
(* Each fact is given a unique fact number to avoid name clashes (e.g., because
blanchet@43444
   982
   of monomorphization). The TPTP explicitly forbids name clashes, and some of
blanchet@43444
   983
   the remote provers might care. *)
blanchet@43797
   984
fun formula_line_for_fact ctxt format prefix nonmono_Ts type_sys
blanchet@43511
   985
                          (j, formula as {name, locality, kind, ...}) =
blanchet@43550
   986
  Formula (prefix ^ (if polymorphism_of_type_sys type_sys = Polymorphic then ""
blanchet@43550
   987
                     else string_of_int j ^ "_") ^
blanchet@43518
   988
           ascii_of name,
blanchet@43797
   989
           kind, formula_for_fact ctxt format nonmono_Ts type_sys formula, NONE,
blanchet@43748
   990
           case locality of
blanchet@43748
   991
             Intro => intro_info
blanchet@43748
   992
           | Elim => elim_info
blanchet@43748
   993
           | Simp => simp_info
blanchet@43748
   994
           | _ => NONE)
blanchet@43444
   995
blanchet@43780
   996
fun formula_line_for_class_rel_clause
blanchet@43780
   997
        (ClassRelClause {name, subclass, superclass, ...}) =
blanchet@43444
   998
  let val ty_arg = ATerm (`I "T", []) in
blanchet@43448
   999
    Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
blanchet@43444
  1000
             AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
blanchet@43444
  1001
                               AAtom (ATerm (superclass, [ty_arg]))])
blanchet@43748
  1002
             |> close_formula_universally, intro_info, NONE)
blanchet@43444
  1003
  end
blanchet@43444
  1004
blanchet@43444
  1005
fun fo_literal_from_arity_literal (TConsLit (c, t, args)) =
blanchet@43444
  1006
    (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
blanchet@43444
  1007
  | fo_literal_from_arity_literal (TVarLit (c, sort)) =
blanchet@43444
  1008
    (false, ATerm (c, [ATerm (sort, [])]))
blanchet@43444
  1009
blanchet@43780
  1010
fun formula_line_for_arity_clause
blanchet@43780
  1011
        (ArityClause {name, prem_lits, concl_lits, ...}) =
blanchet@43448
  1012
  Formula (arity_clause_prefix ^ ascii_of name, Axiom,
blanchet@43444
  1013
           mk_ahorn (map (formula_from_fo_literal o apfst not
blanchet@43766
  1014
                          o fo_literal_from_arity_literal) prem_lits)
blanchet@43444
  1015
                    (formula_from_fo_literal
blanchet@43766
  1016
                         (fo_literal_from_arity_literal concl_lits))
blanchet@43748
  1017
           |> close_formula_universally, intro_info, NONE)
blanchet@43444
  1018
blanchet@43803
  1019
fun formula_line_for_conjecture ctxt format nonmono_Ts type_sys
blanchet@43444
  1020
        ({name, kind, combformula, ...} : translated_formula) =
blanchet@43448
  1021
  Formula (conjecture_prefix ^ name, kind,
blanchet@43803
  1022
           formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43780
  1023
               is_var_nonmonotonic_in_formula false
blanchet@43780
  1024
               (close_combformula_universally combformula)
blanchet@43444
  1025
           |> close_formula_universally, NONE, NONE)
blanchet@43444
  1026
blanchet@43797
  1027
fun free_type_literals format type_sys
blanchet@43797
  1028
                       ({atomic_types, ...} : translated_formula) =
blanchet@43797
  1029
  atomic_types |> atp_type_literals_for_types format type_sys Conjecture
blanchet@43444
  1030
               |> map fo_literal_from_type_literal
blanchet@43444
  1031
blanchet@43444
  1032
fun formula_line_for_free_type j lit =
blanchet@43448
  1033
  Formula (tfree_prefix ^ string_of_int j, Hypothesis,
blanchet@43444
  1034
           formula_from_fo_literal lit, NONE, NONE)
blanchet@43797
  1035
fun formula_lines_for_free_types format type_sys facts =
blanchet@43444
  1036
  let
blanchet@43797
  1037
    val litss = map (free_type_literals format type_sys) facts
blanchet@43444
  1038
    val lits = fold (union (op =)) litss []
blanchet@43444
  1039
  in map2 formula_line_for_free_type (0 upto length lits - 1) lits end
blanchet@43444
  1040
blanchet@43444
  1041
(** Symbol declarations **)
blanchet@43415
  1042
blanchet@43445
  1043
fun should_declare_sym type_sys pred_sym s =
blanchet@43839
  1044
  is_tptp_user_symbol s andalso not (String.isPrefix bound_var_prefix s) andalso
blanchet@43765
  1045
  (case type_sys of
blanchet@43765
  1046
     Simple_Types _ => true
blanchet@43765
  1047
   | Tags (_, _, Light) => true
blanchet@43765
  1048
   | _ => not pred_sym)
blanchet@43413
  1049
blanchet@43755
  1050
fun sym_decl_table_for_facts ctxt type_sys repaired_sym_tab (conjs, facts) =
blanchet@43445
  1051
  let
blanchet@43568
  1052
    fun add_combterm in_conj tm =
blanchet@43445
  1053
      let val (head, args) = strip_combterm_comb tm in
blanchet@43445
  1054
        (case head of
blanchet@43445
  1055
           CombConst ((s, s'), T, T_args) =>
blanchet@43445
  1056
           let val pred_sym = is_pred_sym repaired_sym_tab s in
blanchet@43445
  1057
             if should_declare_sym type_sys pred_sym s then
blanchet@43447
  1058
               Symtab.map_default (s, [])
blanchet@43755
  1059
                   (insert_type ctxt #3 (s', T_args, T, pred_sym, length args,
blanchet@43755
  1060
                                         in_conj))
blanchet@43445
  1061
             else
blanchet@43445
  1062
               I
blanchet@43445
  1063
           end
blanchet@43445
  1064
         | _ => I)
blanchet@43568
  1065
        #> fold (add_combterm in_conj) args
blanchet@43445
  1066
      end
blanchet@43568
  1067
    fun add_fact in_conj =
blanchet@43705
  1068
      fact_lift (formula_fold NONE (K (add_combterm in_conj)))
blanchet@43568
  1069
  in
blanchet@43568
  1070
    Symtab.empty
blanchet@43568
  1071
    |> is_type_sys_fairly_sound type_sys
blanchet@43568
  1072
       ? (fold (add_fact true) conjs #> fold (add_fact false) facts)
blanchet@43568
  1073
  end
blanchet@43445
  1074
blanchet@43755
  1075
(* These types witness that the type classes they belong to allow infinite
blanchet@43755
  1076
   models and hence that any types with these type classes is monotonic. *)
blanchet@43755
  1077
val known_infinite_types = [@{typ nat}, @{typ int}, @{typ "nat => bool"}]
blanchet@43755
  1078
blanchet@43555
  1079
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
blanchet@43555
  1080
   out with monotonicity" paper presented at CADE 2011. *)
blanchet@43755
  1081
fun add_combterm_nonmonotonic_types _ _ (SOME false) _ = I
blanchet@43700
  1082
  | add_combterm_nonmonotonic_types ctxt level _
blanchet@43841
  1083
        (CombApp (CombApp (CombConst ((s, _), Type (_, [T, _]), _), tm1), tm2)) =
blanchet@43841
  1084
    (is_tptp_equal s andalso exists is_var_or_bound_var [tm1, tm2] andalso
blanchet@43700
  1085
     (case level of
blanchet@43755
  1086
        Nonmonotonic_Types =>
blanchet@43755
  1087
        not (is_type_surely_infinite ctxt known_infinite_types T)
blanchet@43700
  1088
      | Finite_Types => is_type_surely_finite ctxt T
blanchet@43755
  1089
      | _ => true)) ? insert_type ctxt I (deep_freeze_type T)
blanchet@43700
  1090
  | add_combterm_nonmonotonic_types _ _ _ _ = I
blanchet@43700
  1091
fun add_fact_nonmonotonic_types ctxt level ({kind, combformula, ...}
blanchet@43700
  1092
                                            : translated_formula) =
blanchet@43705
  1093
  formula_fold (SOME (kind <> Conjecture))
blanchet@43700
  1094
               (add_combterm_nonmonotonic_types ctxt level) combformula
blanchet@43755
  1095
fun nonmonotonic_types_for_facts ctxt type_sys facts =
blanchet@43700
  1096
  let val level = level_of_type_sys type_sys in
blanchet@43755
  1097
    if level = Nonmonotonic_Types orelse level = Finite_Types then
blanchet@43755
  1098
      [] |> fold (add_fact_nonmonotonic_types ctxt level) facts
blanchet@43755
  1099
         (* We must add "bool" in case the helper "True_or_False" is added
blanchet@43755
  1100
            later. In addition, several places in the code rely on the list of
blanchet@43755
  1101
            nonmonotonic types not being empty. *)
blanchet@43755
  1102
         |> insert_type ctxt I @{typ bool}
blanchet@43755
  1103
    else
blanchet@43755
  1104
      []
blanchet@43700
  1105
  end
blanchet@43547
  1106
blanchet@43835
  1107
fun decl_line_for_sym ctxt format nonmono_Ts type_sys s
blanchet@43835
  1108
                      (s', T_args, T, pred_sym, ary, _) =
blanchet@43835
  1109
  let
blanchet@43835
  1110
    val (higher_order, T_arg_Ts, level) =
blanchet@43835
  1111
      case type_sys of
blanchet@43835
  1112
        Simple_Types level => (format = THF, [], level)
blanchet@43835
  1113
      | _ => (false, replicate (length T_args) homo_infinite_type, No_Types)
blanchet@43835
  1114
  in
blanchet@43839
  1115
    Decl (sym_decl_prefix ^ s, (s, s'),
blanchet@43835
  1116
          (T_arg_Ts ---> (T |> homogenized_type ctxt nonmono_Ts level ary))
blanchet@43835
  1117
          |> mangled_type higher_order pred_sym (length T_arg_Ts + ary))
blanchet@43835
  1118
  end
blanchet@43450
  1119
blanchet@43463
  1120
fun is_polymorphic_type T = fold_atyps (fn TVar _ => K true | _ => I) T false
blanchet@43463
  1121
blanchet@43797
  1122
fun formula_line_for_pred_sym_decl ctxt format conj_sym_kind nonmono_Ts type_sys
blanchet@43797
  1123
                                   n s j (s', T_args, T, _, ary, in_conj) =
blanchet@43450
  1124
  let
blanchet@43580
  1125
    val (kind, maybe_negate) =
blanchet@43580
  1126
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43580
  1127
      else (Axiom, I)
blanchet@43618
  1128
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@43450
  1129
    val bound_names =
blanchet@43450
  1130
      1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  1131
    val bounds =
blanchet@43450
  1132
      bound_names ~~ arg_Ts |> map (fn (name, T) => CombConst (name, T, []))
blanchet@43450
  1133
    val bound_Ts =
blanchet@43463
  1134
      arg_Ts |> map (fn T => if n > 1 orelse is_polymorphic_type T then SOME T
blanchet@43463
  1135
                             else NONE)
blanchet@43450
  1136
  in
blanchet@43839
  1137
    Formula (sym_formula_prefix ^ s ^
blanchet@43580
  1138
             (if n > 1 then "_" ^ string_of_int j else ""), kind,
blanchet@43450
  1139
             CombConst ((s, s'), T, T_args)
blanchet@43700
  1140
             |> fold (curry (CombApp o swap)) bounds
blanchet@43835
  1141
             |> type_pred_combterm ctxt nonmono_Ts type_sys res_T
blanchet@43804
  1142
             |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
blanchet@43803
  1143
             |> formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43747
  1144
                                         (K (K (K (K true)))) true
blanchet@43797
  1145
             |> n > 1 ? bound_atomic_types format type_sys (atyps_of T)
blanchet@43580
  1146
             |> close_formula_universally
blanchet@43580
  1147
             |> maybe_negate,
blanchet@43748
  1148
             intro_info, NONE)
blanchet@43450
  1149
  end
blanchet@43450
  1150
blanchet@43797
  1151
fun formula_lines_for_tag_sym_decl ctxt format conj_sym_kind nonmono_Ts type_sys
blanchet@43797
  1152
        n s (j, (s', T_args, T, pred_sym, ary, in_conj)) =
blanchet@43700
  1153
  let
blanchet@43700
  1154
    val ident_base =
blanchet@43839
  1155
      sym_formula_prefix ^ s ^ (if n > 1 then "_" ^ string_of_int j else "")
blanchet@43721
  1156
    val (kind, maybe_negate) =
blanchet@43721
  1157
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43721
  1158
      else (Axiom, I)
blanchet@43700
  1159
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@43700
  1160
    val bound_names =
blanchet@43700
  1161
      1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  1162
    val bounds = bound_names |> map (fn name => ATerm (name, []))
blanchet@43835
  1163
    val cst = mk_const_aterm (s, s') T_args
blanchet@43701
  1164
    val atomic_Ts = atyps_of T
blanchet@43705
  1165
    fun eq tms =
blanchet@43705
  1166
      (if pred_sym then AConn (AIff, map AAtom tms)
blanchet@43841
  1167
       else AAtom (ATerm (`I tptp_equal, tms)))
blanchet@43797
  1168
      |> bound_atomic_types format type_sys atomic_Ts
blanchet@43701
  1169
      |> close_formula_universally
blanchet@43721
  1170
      |> maybe_negate
blanchet@43707
  1171
    val should_encode = should_encode_type ctxt nonmono_Ts All_Types
blanchet@43803
  1172
    val tag_with = tag_with_type ctxt format nonmono_Ts type_sys
blanchet@43700
  1173
    val add_formula_for_res =
blanchet@43700
  1174
      if should_encode res_T then
blanchet@43721
  1175
        cons (Formula (ident_base ^ "_res", kind,
blanchet@43835
  1176
                       eq [tag_with res_T (cst bounds), cst bounds],
blanchet@43748
  1177
                       simp_info, NONE))
blanchet@43700
  1178
      else
blanchet@43700
  1179
        I
blanchet@43700
  1180
    fun add_formula_for_arg k =
blanchet@43700
  1181
      let val arg_T = nth arg_Ts k in
blanchet@43700
  1182
        if should_encode arg_T then
blanchet@43700
  1183
          case chop k bounds of
blanchet@43700
  1184
            (bounds1, bound :: bounds2) =>
blanchet@43721
  1185
            cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,
blanchet@43835
  1186
                           eq [cst (bounds1 @ tag_with arg_T bound :: bounds2),
blanchet@43835
  1187
                               cst bounds],
blanchet@43748
  1188
                           simp_info, NONE))
blanchet@43700
  1189
          | _ => raise Fail "expected nonempty tail"
blanchet@43700
  1190
        else
blanchet@43700
  1191
          I
blanchet@43700
  1192
      end
blanchet@43700
  1193
  in
blanchet@43705
  1194
    [] |> not pred_sym ? add_formula_for_res
blanchet@43700
  1195
       |> fold add_formula_for_arg (ary - 1 downto 0)
blanchet@43700
  1196
  end
blanchet@43700
  1197
blanchet@43707
  1198
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
blanchet@43707
  1199
blanchet@43797
  1200
fun problem_lines_for_sym_decls ctxt format conj_sym_kind nonmono_Ts type_sys
blanchet@43580
  1201
                                (s, decls) =
blanchet@43839
  1202
  case type_sys of
blanchet@43839
  1203
    Simple_Types _ =>
blanchet@43839
  1204
    decls |> map (decl_line_for_sym ctxt format nonmono_Ts type_sys s)
blanchet@43839
  1205
  | Preds _ =>
blanchet@43839
  1206
    let
blanchet@43839
  1207
      val decls =
blanchet@43839
  1208
        case decls of
blanchet@43839
  1209
          decl :: (decls' as _ :: _) =>
blanchet@43839
  1210
          let val T = result_type_of_decl decl in
blanchet@43839
  1211
            if forall (curry (type_instance ctxt o swap) T
blanchet@43839
  1212
                       o result_type_of_decl) decls' then
blanchet@43839
  1213
              [decl]
blanchet@43839
  1214
            else
blanchet@43839
  1215
              decls
blanchet@43839
  1216
          end
blanchet@43839
  1217
        | _ => decls
blanchet@43839
  1218
      val n = length decls
blanchet@43839
  1219
      val decls =
blanchet@43839
  1220
        decls
blanchet@43839
  1221
        |> filter (should_predicate_on_type ctxt nonmono_Ts type_sys (K true)
blanchet@43839
  1222
                   o result_type_of_decl)
blanchet@43839
  1223
    in
blanchet@43839
  1224
      (0 upto length decls - 1, decls)
blanchet@43839
  1225
      |-> map2 (formula_line_for_pred_sym_decl ctxt format conj_sym_kind
blanchet@43839
  1226
                                               nonmono_Ts type_sys n s)
blanchet@43839
  1227
    end
blanchet@43839
  1228
  | Tags (_, _, heaviness) =>
blanchet@43839
  1229
    (case heaviness of
blanchet@43839
  1230
       Heavy => []
blanchet@43839
  1231
     | Light =>
blanchet@43839
  1232
       let val n = length decls in
blanchet@43839
  1233
         (0 upto n - 1 ~~ decls)
blanchet@43839
  1234
         |> maps (formula_lines_for_tag_sym_decl ctxt format conj_sym_kind
blanchet@43839
  1235
                                                 nonmono_Ts type_sys n s)
blanchet@43839
  1236
       end)
blanchet@43450
  1237
blanchet@43797
  1238
fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
blanchet@43797
  1239
                                     type_sys sym_decl_tab =
blanchet@43839
  1240
  sym_decl_tab
blanchet@43839
  1241
  |> Symtab.dest
blanchet@43839
  1242
  |> sort_wrt fst
blanchet@43839
  1243
  |> rpair []
blanchet@43839
  1244
  |-> fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
blanchet@43839
  1245
                                                     nonmono_Ts type_sys)
blanchet@43414
  1246
blanchet@43708
  1247
fun should_add_ti_ti_helper (Tags (Polymorphic, level, Heavy)) =
blanchet@43702
  1248
    level = Nonmonotonic_Types orelse level = Finite_Types
blanchet@43702
  1249
  | should_add_ti_ti_helper _ = false
blanchet@43702
  1250
blanchet@43780
  1251
fun offset_of_heading_in_problem _ [] j = j
blanchet@43780
  1252
  | offset_of_heading_in_problem needle ((heading, lines) :: problem) j =
blanchet@43780
  1253
    if heading = needle then j
blanchet@43780
  1254
    else offset_of_heading_in_problem needle problem (j + length lines)
blanchet@43780
  1255
blanchet@43839
  1256
val implicit_declsN = "Should-be-implicit typings"
blanchet@43839
  1257
val explicit_declsN = "Explicit typings"
blanchet@41405
  1258
val factsN = "Relevant facts"
blanchet@41405
  1259
val class_relsN = "Class relationships"
blanchet@43414
  1260
val aritiesN = "Arities"
blanchet@41405
  1261
val helpersN = "Helper facts"
blanchet@41405
  1262
val conjsN = "Conjectures"
blanchet@41561
  1263
val free_typesN = "Type variables"
blanchet@41405
  1264
blanchet@43780
  1265
fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_sys
blanchet@43780
  1266
                        explicit_apply hyp_ts concl_t facts =
blanchet@38506
  1267
  let
blanchet@41561
  1268
    val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) =
blanchet@43803
  1269
      translate_formulas ctxt format prem_kind type_sys hyp_ts concl_t facts
blanchet@43905
  1270
    val sym_tab = conjs @ facts |> sym_table_for_facts ctxt explicit_apply
blanchet@43755
  1271
    val nonmono_Ts = conjs @ facts |> nonmonotonic_types_for_facts ctxt type_sys
blanchet@43803
  1272
    val repair = repair_fact ctxt format nonmono_Ts type_sys sym_tab
blanchet@43552
  1273
    val (conjs, facts) = (conjs, facts) |> pairself (map repair)
blanchet@43905
  1274
    val repaired_sym_tab =
blanchet@43905
  1275
      conjs @ facts |> sym_table_for_facts ctxt (SOME false)
blanchet@43444
  1276
    val helpers =
blanchet@43803
  1277
      repaired_sym_tab |> helper_facts_for_sym_table ctxt format type_sys
blanchet@43803
  1278
                       |> map repair
blanchet@43765
  1279
    val lavish_nonmono_Ts =
blanchet@43765
  1280
      if null nonmono_Ts orelse
blanchet@43765
  1281
         polymorphism_of_type_sys type_sys <> Polymorphic then
blanchet@43765
  1282
        nonmono_Ts
blanchet@43765
  1283
      else
blanchet@43765
  1284
        [TVar (("'a", 0), HOLogic.typeS)]
blanchet@43550
  1285
    val sym_decl_lines =
blanchet@43596
  1286
      (conjs, helpers @ facts)
blanchet@43755
  1287
      |> sym_decl_table_for_facts ctxt type_sys repaired_sym_tab
blanchet@43797
  1288
      |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind
blanchet@43797
  1289
                                          lavish_nonmono_Ts type_sys
blanchet@43750
  1290
    val helper_lines =
blanchet@43797
  1291
      0 upto length helpers - 1 ~~ helpers
blanchet@43797
  1292
      |> map (formula_line_for_fact ctxt format helper_prefix lavish_nonmono_Ts
blanchet@43797
  1293
                                    type_sys)
blanchet@43797
  1294
      |> (if should_add_ti_ti_helper type_sys then cons (ti_ti_helper_fact ())
blanchet@43797
  1295
          else I)
blanchet@43393
  1296
    (* Reordering these might confuse the proof reconstruction code or the SPASS
blanchet@43880
  1297
       FLOTTER hack. *)
blanchet@38506
  1298
    val problem =
blanchet@43839
  1299
      [(explicit_declsN, sym_decl_lines),
blanchet@43797
  1300
       (factsN,
blanchet@43797
  1301
        map (formula_line_for_fact ctxt format fact_prefix nonmono_Ts type_sys)
blanchet@43797
  1302
            (0 upto length facts - 1 ~~ facts)),
blanchet@43416
  1303
       (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses),
blanchet@43416
  1304
       (aritiesN, map formula_line_for_arity_clause arity_clauses),
blanchet@43750
  1305
       (helpersN, helper_lines),
blanchet@43803
  1306
       (conjsN,
blanchet@43803
  1307
        map (formula_line_for_conjecture ctxt format nonmono_Ts type_sys)
blanchet@43803
  1308
            conjs),
blanchet@43797
  1309
       (free_typesN,
blanchet@43797
  1310
        formula_lines_for_free_types format type_sys (facts @ conjs))]
blanchet@43414
  1311
    val problem =
blanchet@43432
  1312
      problem
blanchet@43839
  1313
      |> (if format = CNF_UEQ then filter_cnf_ueq_problem else I)
blanchet@43839
  1314
      |> (if is_format_typed format then
blanchet@43839
  1315
            declare_undeclared_syms_in_atp_problem type_decl_prefix
blanchet@43839
  1316
                                                   implicit_declsN
blanchet@43839
  1317
          else
blanchet@43839
  1318
            I)
blanchet@43517
  1319
    val (problem, pool) =
blanchet@43517
  1320
      problem |> nice_atp_problem (Config.get ctxt readable_names)
blanchet@43750
  1321
    val helpers_offset = offset_of_heading_in_problem helpersN problem 0
blanchet@43750
  1322
    val typed_helpers =
blanchet@43750
  1323
      map_filter (fn (j, {name, ...}) =>
blanchet@43750
  1324
                     if String.isSuffix typed_helper_suffix name then SOME j
blanchet@43750
  1325
                     else NONE)
blanchet@43750
  1326
                 ((helpers_offset + 1 upto helpers_offset + length helpers)
blanchet@43750
  1327
                  ~~ helpers)
blanchet@43649
  1328
    fun add_sym_arity (s, {min_ary, ...} : sym_info) =
blanchet@43620
  1329
      if min_ary > 0 then
blanchet@43620
  1330
        case strip_prefix_and_unascii const_prefix s of
blanchet@43620
  1331
          SOME s => Symtab.insert (op =) (s, min_ary)
blanchet@43620
  1332
        | NONE => I
blanchet@43620
  1333
      else
blanchet@43620
  1334
        I
blanchet@38506
  1335
  in
blanchet@38506
  1336
    (problem,
blanchet@38506
  1337
     case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
blanchet@43456
  1338
     offset_of_heading_in_problem conjsN problem 0,
blanchet@43412
  1339
     offset_of_heading_in_problem factsN problem 0,
blanchet@43620
  1340
     fact_names |> Vector.fromList,
blanchet@43750
  1341
     typed_helpers,
blanchet@43620
  1342
     Symtab.empty |> Symtab.fold add_sym_arity sym_tab)
blanchet@38506
  1343
  end
blanchet@38506
  1344
blanchet@41561
  1345
(* FUDGE *)
blanchet@41561
  1346
val conj_weight = 0.0
blanchet@42641
  1347
val hyp_weight = 0.1
blanchet@42641
  1348
val fact_min_weight = 0.2
blanchet@41561
  1349
val fact_max_weight = 1.0
blanchet@43479
  1350
val type_info_default_weight = 0.8
blanchet@41561
  1351
blanchet@41561
  1352
fun add_term_weights weight (ATerm (s, tms)) =
blanchet@43839
  1353
  is_tptp_user_symbol s ? Symtab.default (s, weight)
blanchet@41561
  1354
  #> fold (add_term_weights weight) tms
blanchet@43448
  1355
fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
blanchet@43705
  1356
    formula_fold NONE (K (add_term_weights weight)) phi
blanchet@43399
  1357
  | add_problem_line_weights _ _ = I
blanchet@41561
  1358
blanchet@41561
  1359
fun add_conjectures_weights [] = I
blanchet@41561
  1360
  | add_conjectures_weights conjs =
blanchet@41561
  1361
    let val (hyps, conj) = split_last conjs in
blanchet@41561
  1362
      add_problem_line_weights conj_weight conj
blanchet@41561
  1363
      #> fold (add_problem_line_weights hyp_weight) hyps
blanchet@41561
  1364
    end
blanchet@41561
  1365
blanchet@41561
  1366
fun add_facts_weights facts =
blanchet@41561
  1367
  let
blanchet@41561
  1368
    val num_facts = length facts
blanchet@41561
  1369
    fun weight_of j =
blanchet@41561
  1370
      fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
blanchet@41561
  1371
                        / Real.fromInt num_facts
blanchet@41561
  1372
  in
blanchet@41561
  1373
    map weight_of (0 upto num_facts - 1) ~~ facts
blanchet@41561
  1374
    |> fold (uncurry add_problem_line_weights)
blanchet@41561
  1375
  end
blanchet@41561
  1376
blanchet@41561
  1377
(* Weights are from 0.0 (most important) to 1.0 (least important). *)
blanchet@41561
  1378
fun atp_problem_weights problem =
blanchet@43479
  1379
  let val get = these o AList.lookup (op =) problem in
blanchet@43479
  1380
    Symtab.empty
blanchet@43479
  1381
    |> add_conjectures_weights (get free_typesN @ get conjsN)
blanchet@43479
  1382
    |> add_facts_weights (get factsN)
blanchet@43479
  1383
    |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
blanchet@43839
  1384
            [explicit_declsN, class_relsN, aritiesN]
blanchet@43479
  1385
    |> Symtab.dest
blanchet@43479
  1386
    |> sort (prod_ord Real.compare string_ord o pairself swap)
blanchet@43479
  1387
  end
blanchet@41561
  1388
blanchet@38506
  1389
end;