src/HOL/Tools/ATP/atp_translate.ML
author blanchet
Tue, 31 May 2011 16:38:36 +0200
changeset 43934 40e50afbc203
parent 43933 93ec303e1917
child 43937 f181d66046d4
permissions -rw-r--r--
tuning
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@43926
     9
signature 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@43926
    15
blanchet@43926
    16
  type name = string * string
blanchet@43926
    17
blanchet@43926
    18
  datatype type_literal =
blanchet@43926
    19
    TyLitVar of name * name |
blanchet@43926
    20
    TyLitFree of name * name
blanchet@43926
    21
blanchet@43926
    22
  datatype arity_literal =
blanchet@43926
    23
    TConsLit of name * name * name list |
blanchet@43926
    24
    TVarLit of name * name
blanchet@43926
    25
blanchet@43927
    26
  type arity_clause =
blanchet@43927
    27
    {name: string,
blanchet@43927
    28
     prem_lits: arity_literal list,
blanchet@43927
    29
     concl_lits: arity_literal}
blanchet@43926
    30
blanchet@43927
    31
  type class_rel_clause =
blanchet@43927
    32
    {name: string,
blanchet@43927
    33
     subclass: name,
blanchet@43927
    34
     superclass: name}
blanchet@43926
    35
blanchet@43926
    36
  datatype combterm =
blanchet@43926
    37
    CombConst of name * typ * typ list |
blanchet@43926
    38
    CombVar of name * typ |
blanchet@43926
    39
    CombApp of combterm * combterm
blanchet@43926
    40
blanchet@43926
    41
  datatype locality = General | Intro | Elim | Simp | Local | Assum | Chained
blanchet@43484
    42
blanchet@43484
    43
  datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
blanchet@43484
    44
  datatype type_level =
blanchet@43484
    45
    All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types
blanchet@43708
    46
  datatype type_heaviness = Heavy | Light
blanchet@43484
    47
blanchet@43484
    48
  datatype type_system =
blanchet@43587
    49
    Simple_Types of type_level |
blanchet@43708
    50
    Preds of polymorphism * type_level * type_heaviness |
blanchet@43708
    51
    Tags of polymorphism * type_level * type_heaviness
blanchet@43484
    52
blanchet@40358
    53
  type translated_formula
blanchet@38506
    54
blanchet@43926
    55
  val type_tag_name : string
blanchet@43926
    56
  val bound_var_prefix : string
blanchet@43926
    57
  val schematic_var_prefix: string
blanchet@43926
    58
  val fixed_var_prefix: string
blanchet@43926
    59
  val tvar_prefix: string
blanchet@43926
    60
  val tfree_prefix: string
blanchet@43926
    61
  val const_prefix: string
blanchet@43926
    62
  val type_const_prefix: string
blanchet@43926
    63
  val class_prefix: string
blanchet@43926
    64
  val skolem_const_prefix : string
blanchet@43926
    65
  val old_skolem_const_prefix : string
blanchet@43926
    66
  val new_skolem_const_prefix : string
blanchet@40445
    67
  val fact_prefix : string
blanchet@38506
    68
  val conjecture_prefix : string
blanchet@43750
    69
  val helper_prefix : string
blanchet@43750
    70
  val typed_helper_suffix : string
blanchet@43807
    71
  val predicator_name : string
blanchet@43807
    72
  val app_op_name : string
blanchet@43807
    73
  val type_pred_name : string
blanchet@43803
    74
  val simple_type_prefix : string
blanchet@43926
    75
  val ascii_of: string -> string
blanchet@43926
    76
  val unascii_of: string -> string
blanchet@43926
    77
  val strip_prefix_and_unascii : string -> string -> string option
blanchet@43926
    78
  val proxify_const : string -> (int * (string * string)) option
blanchet@43926
    79
  val invert_const: string -> string
blanchet@43926
    80
  val unproxify_const: string -> string
blanchet@43926
    81
  val make_bound_var : string -> string
blanchet@43926
    82
  val make_schematic_var : string * int -> string
blanchet@43926
    83
  val make_fixed_var : string -> string
blanchet@43926
    84
  val make_schematic_type_var : string * int -> string
blanchet@43926
    85
  val make_fixed_type_var : string -> string
blanchet@43926
    86
  val make_fixed_const : string -> string
blanchet@43926
    87
  val make_fixed_type_const : string -> string
blanchet@43926
    88
  val make_type_class : string -> string
blanchet@43934
    89
  val new_skolem_var_name_from_const : string -> string
blanchet@43934
    90
  val num_type_args : theory -> string -> int
blanchet@43926
    91
  val make_arity_clauses :
blanchet@43926
    92
    theory -> string list -> class list -> class list * arity_clause list
blanchet@43926
    93
  val make_class_rel_clauses :
blanchet@43926
    94
    theory -> class list -> class list -> class_rel_clause list
blanchet@43926
    95
  val combtyp_of : combterm -> typ
blanchet@43926
    96
  val strip_combterm_comb : combterm -> combterm * combterm list
blanchet@43926
    97
  val atyps_of : typ -> typ list
blanchet@43926
    98
  val combterm_from_term :
blanchet@43926
    99
    theory -> (string * typ) list -> term -> combterm * typ list
blanchet@43926
   100
  val is_locality_global : locality -> bool
blanchet@43484
   101
  val type_sys_from_string : string -> type_system
blanchet@43484
   102
  val polymorphism_of_type_sys : type_system -> polymorphism
blanchet@43484
   103
  val level_of_type_sys : type_system -> type_level
blanchet@43484
   104
  val is_type_sys_virtually_sound : type_system -> bool
blanchet@43484
   105
  val is_type_sys_fairly_sound : type_system -> bool
blanchet@43926
   106
  val raw_type_literals_for_types : typ list -> type_literal list
blanchet@43413
   107
  val unmangled_const : string -> string * string fo_term list
blanchet@41336
   108
  val translate_atp_fact :
blanchet@43835
   109
    Proof.context -> format -> type_system -> bool -> (string * locality) * thm
blanchet@43511
   110
    -> translated_formula option * ((string * locality) * thm)
blanchet@43926
   111
  val helper_table : (string * (bool * thm list)) list
blanchet@43926
   112
  val tfree_classes_of_terms : term list -> string list
blanchet@43926
   113
  val tvar_classes_of_terms : term list -> string list
blanchet@43926
   114
  val type_consts_of_terms : theory -> term list -> string list
blanchet@40240
   115
  val prepare_atp_problem :
blanchet@43780
   116
    Proof.context -> format -> formula_kind -> formula_kind -> type_system
blanchet@43933
   117
    -> bool option -> bool -> term list -> term
blanchet@41339
   118
    -> (translated_formula option * ((string * 'a) * thm)) list
blanchet@43412
   119
    -> string problem * string Symtab.table * int * int
blanchet@43750
   120
       * (string * 'a) list vector * int list * int Symtab.table
blanchet@41561
   121
  val atp_problem_weights : string problem -> (string * real) list
blanchet@38506
   122
end;
blanchet@38506
   123
blanchet@43926
   124
structure ATP_Translate : ATP_TRANSLATE =
blanchet@38506
   125
struct
blanchet@38506
   126
blanchet@43926
   127
open ATP_Util
blanchet@38506
   128
open ATP_Problem
blanchet@43926
   129
blanchet@43926
   130
type name = string * string
blanchet@43926
   131
blanchet@43926
   132
(* FIXME: avoid *)
blanchet@43926
   133
fun union_all xss = fold (union (op =)) xss []
blanchet@43511
   134
blanchet@43511
   135
(* experimental *)
blanchet@43511
   136
val generate_useful_info = false
blanchet@38506
   137
blanchet@43748
   138
fun useful_isabelle_info s =
blanchet@43748
   139
  if generate_useful_info then
blanchet@43748
   140
    SOME (ATerm ("[]", [ATerm ("isabelle_" ^ s, [])]))
blanchet@43748
   141
  else
blanchet@43748
   142
    NONE
blanchet@43748
   143
blanchet@43748
   144
val intro_info = useful_isabelle_info "intro"
blanchet@43748
   145
val elim_info = useful_isabelle_info "elim"
blanchet@43748
   146
val simp_info = useful_isabelle_info "simp"
blanchet@43748
   147
blanchet@43926
   148
val type_tag_name = "ti"
blanchet@43926
   149
blanchet@43926
   150
val bound_var_prefix = "B_"
blanchet@43926
   151
val schematic_var_prefix = "V_"
blanchet@43926
   152
val fixed_var_prefix = "v_"
blanchet@43926
   153
blanchet@43926
   154
val tvar_prefix = "T_"
blanchet@43926
   155
val tfree_prefix = "t_"
blanchet@43926
   156
blanchet@43926
   157
val const_prefix = "c_"
blanchet@43926
   158
val type_const_prefix = "tc_"
blanchet@43926
   159
val class_prefix = "cl_"
blanchet@43926
   160
blanchet@43926
   161
val skolem_const_prefix = "Sledgehammer" ^ Long_Name.separator ^ "Sko"
blanchet@43926
   162
val old_skolem_const_prefix = skolem_const_prefix ^ "o"
blanchet@43926
   163
val new_skolem_const_prefix = skolem_const_prefix ^ "n"
blanchet@43926
   164
blanchet@43839
   165
val type_decl_prefix = "ty_"
blanchet@43839
   166
val sym_decl_prefix = "sy_"
blanchet@43839
   167
val sym_formula_prefix = "sym_"
blanchet@40445
   168
val fact_prefix = "fact_"
blanchet@38506
   169
val conjecture_prefix = "conj_"
blanchet@38506
   170
val helper_prefix = "help_"
blanchet@43934
   171
val class_rel_clause_prefix = "crel_"
blanchet@38506
   172
val arity_clause_prefix = "arity_"
blanchet@43926
   173
val tfree_clause_prefix = "tfree_"
blanchet@38506
   174
blanchet@43750
   175
val typed_helper_suffix = "_T"
blanchet@43750
   176
val untyped_helper_suffix = "_U"
blanchet@43750
   177
blanchet@43807
   178
val predicator_name = "hBOOL"
blanchet@43807
   179
val app_op_name = "hAPP"
blanchet@43807
   180
val type_pred_name = "is"
blanchet@43803
   181
val simple_type_prefix = "ty_"
blanchet@43402
   182
blanchet@38506
   183
(* Freshness almost guaranteed! *)
blanchet@38506
   184
val sledgehammer_weak_prefix = "Sledgehammer:"
blanchet@38506
   185
blanchet@43926
   186
(*Escaping of special characters.
blanchet@43926
   187
  Alphanumeric characters are left unchanged.
blanchet@43926
   188
  The character _ goes to __
blanchet@43926
   189
  Characters in the range ASCII space to / go to _A to _P, respectively.
blanchet@43926
   190
  Other characters go to _nnn where nnn is the decimal ASCII code.*)
blanchet@43934
   191
val upper_a_minus_space = Char.ord #"A" - Char.ord #" "
blanchet@43926
   192
blanchet@43926
   193
fun stringN_of_int 0 _ = ""
blanchet@43926
   194
  | stringN_of_int k n =
blanchet@43926
   195
    stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
blanchet@43926
   196
blanchet@43926
   197
fun ascii_of_char c =
blanchet@43926
   198
  if Char.isAlphaNum c then
blanchet@43926
   199
    String.str c
blanchet@43926
   200
  else if c = #"_" then
blanchet@43926
   201
    "__"
blanchet@43926
   202
  else if #" " <= c andalso c <= #"/" then
blanchet@43926
   203
    "_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space))
blanchet@43926
   204
  else
blanchet@43926
   205
    (* fixed width, in case more digits follow *)
blanchet@43926
   206
    "_" ^ stringN_of_int 3 (Char.ord c)
blanchet@43926
   207
blanchet@43926
   208
val ascii_of = String.translate ascii_of_char
blanchet@43926
   209
blanchet@43926
   210
(** Remove ASCII armoring from names in proof files **)
blanchet@43926
   211
blanchet@43926
   212
(* We don't raise error exceptions because this code can run inside a worker
blanchet@43926
   213
   thread. Also, the errors are impossible. *)
blanchet@43926
   214
val unascii_of =
blanchet@43926
   215
  let
blanchet@43926
   216
    fun un rcs [] = String.implode(rev rcs)
blanchet@43926
   217
      | un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)
blanchet@43926
   218
        (* Three types of _ escapes: __, _A to _P, _nnn *)
blanchet@43926
   219
      | un rcs (#"_" :: #"_" :: cs) = un (#"_"::rcs) cs
blanchet@43926
   220
      | un rcs (#"_" :: c :: cs) =
blanchet@43926
   221
        if #"A" <= c andalso c<= #"P" then
blanchet@43926
   222
          (* translation of #" " to #"/" *)
blanchet@43926
   223
          un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs
blanchet@43926
   224
        else
blanchet@43926
   225
          let val digits = List.take (c::cs, 3) handle Subscript => [] in
blanchet@43926
   226
            case Int.fromString (String.implode digits) of
blanchet@43926
   227
              SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))
blanchet@43926
   228
            | NONE => un (c:: #"_"::rcs) cs (* ERROR *)
blanchet@43926
   229
          end
blanchet@43926
   230
      | un rcs (c :: cs) = un (c :: rcs) cs
blanchet@43926
   231
  in un [] o String.explode end
blanchet@43926
   232
blanchet@43926
   233
(* If string s has the prefix s1, return the result of deleting it,
blanchet@43926
   234
   un-ASCII'd. *)
blanchet@43926
   235
fun strip_prefix_and_unascii s1 s =
blanchet@43926
   236
  if String.isPrefix s1 s then
blanchet@43926
   237
    SOME (unascii_of (String.extract (s, size s1, NONE)))
blanchet@43926
   238
  else
blanchet@43926
   239
    NONE
blanchet@43926
   240
blanchet@43926
   241
val proxies =
blanchet@43926
   242
  [("c_False",
blanchet@43926
   243
    (@{const_name False}, (0, ("fFalse", @{const_name ATP.fFalse})))),
blanchet@43926
   244
   ("c_True", (@{const_name True}, (0, ("fTrue", @{const_name ATP.fTrue})))),
blanchet@43926
   245
   ("c_Not", (@{const_name Not}, (1, ("fNot", @{const_name ATP.fNot})))),
blanchet@43926
   246
   ("c_conj", (@{const_name conj}, (2, ("fconj", @{const_name ATP.fconj})))),
blanchet@43926
   247
   ("c_disj", (@{const_name disj}, (2, ("fdisj", @{const_name ATP.fdisj})))),
blanchet@43926
   248
   ("c_implies",
blanchet@43926
   249
    (@{const_name implies}, (2, ("fimplies", @{const_name ATP.fimplies})))),
blanchet@43926
   250
   ("equal",
blanchet@43926
   251
    (@{const_name HOL.eq}, (2, ("fequal", @{const_name ATP.fequal}))))]
blanchet@43926
   252
blanchet@43926
   253
val proxify_const = AList.lookup (op =) proxies #> Option.map snd
blanchet@43926
   254
blanchet@43926
   255
(* Readable names for the more common symbolic functions. Do not mess with the
blanchet@43926
   256
   table unless you know what you are doing. *)
blanchet@43926
   257
val const_trans_table =
blanchet@43926
   258
  [(@{type_name Product_Type.prod}, "prod"),
blanchet@43926
   259
   (@{type_name Sum_Type.sum}, "sum"),
blanchet@43926
   260
   (@{const_name False}, "False"),
blanchet@43926
   261
   (@{const_name True}, "True"),
blanchet@43926
   262
   (@{const_name Not}, "Not"),
blanchet@43926
   263
   (@{const_name conj}, "conj"),
blanchet@43926
   264
   (@{const_name disj}, "disj"),
blanchet@43926
   265
   (@{const_name implies}, "implies"),
blanchet@43926
   266
   (@{const_name HOL.eq}, "equal"),
blanchet@43926
   267
   (@{const_name If}, "If"),
blanchet@43926
   268
   (@{const_name Set.member}, "member"),
blanchet@43926
   269
   (@{const_name Meson.COMBI}, "COMBI"),
blanchet@43926
   270
   (@{const_name Meson.COMBK}, "COMBK"),
blanchet@43926
   271
   (@{const_name Meson.COMBB}, "COMBB"),
blanchet@43926
   272
   (@{const_name Meson.COMBC}, "COMBC"),
blanchet@43926
   273
   (@{const_name Meson.COMBS}, "COMBS")]
blanchet@43926
   274
  |> Symtab.make
blanchet@43926
   275
  |> fold (Symtab.update o swap o snd o snd o snd) proxies
blanchet@43926
   276
blanchet@43926
   277
(* Invert the table of translations between Isabelle and ATPs. *)
blanchet@43926
   278
val const_trans_table_inv =
blanchet@43926
   279
  const_trans_table |> Symtab.dest |> map swap |> Symtab.make
blanchet@43926
   280
val const_trans_table_unprox =
blanchet@43926
   281
  Symtab.empty
blanchet@43926
   282
  |> fold (fn (_, (isa, (_, (_, metis)))) => Symtab.update (metis, isa)) proxies
blanchet@43926
   283
blanchet@43926
   284
val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
blanchet@43926
   285
val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox)
blanchet@43926
   286
blanchet@43926
   287
fun lookup_const c =
blanchet@43926
   288
  case Symtab.lookup const_trans_table c of
blanchet@43926
   289
    SOME c' => c'
blanchet@43926
   290
  | NONE => ascii_of c
blanchet@43926
   291
blanchet@43926
   292
(*Remove the initial ' character from a type variable, if it is present*)
blanchet@43926
   293
fun trim_type_var s =
blanchet@43926
   294
  if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
blanchet@43926
   295
  else raise Fail ("trim_type: Malformed type variable encountered: " ^ s)
blanchet@43926
   296
blanchet@43926
   297
fun ascii_of_indexname (v,0) = ascii_of v
blanchet@43926
   298
  | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ string_of_int i
blanchet@43926
   299
blanchet@43926
   300
fun make_bound_var x = bound_var_prefix ^ ascii_of x
blanchet@43926
   301
fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
blanchet@43926
   302
fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
blanchet@43926
   303
blanchet@43926
   304
fun make_schematic_type_var (x,i) =
blanchet@43926
   305
      tvar_prefix ^ (ascii_of_indexname (trim_type_var x, i))
blanchet@43926
   306
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x))
blanchet@43926
   307
blanchet@43926
   308
(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
blanchet@43926
   309
fun make_fixed_const @{const_name HOL.eq} = "equal"
blanchet@43926
   310
  | make_fixed_const c = const_prefix ^ lookup_const c
blanchet@43926
   311
blanchet@43926
   312
fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
blanchet@43926
   313
blanchet@43926
   314
fun make_type_class clas = class_prefix ^ ascii_of clas
blanchet@43926
   315
blanchet@43934
   316
fun new_skolem_var_name_from_const s =
blanchet@43934
   317
  let val ss = s |> space_explode Long_Name.separator in
blanchet@43934
   318
    nth ss (length ss - 2)
blanchet@43934
   319
  end
blanchet@43934
   320
blanchet@43934
   321
(* The number of type arguments of a constant, zero if it's monomorphic. For
blanchet@43934
   322
   (instances of) Skolem pseudoconstants, this information is encoded in the
blanchet@43934
   323
   constant name. *)
blanchet@43934
   324
fun num_type_args thy s =
blanchet@43934
   325
  if String.isPrefix skolem_const_prefix s then
blanchet@43934
   326
    s |> space_explode Long_Name.separator |> List.last |> Int.fromString |> the
blanchet@43934
   327
  else
blanchet@43934
   328
    (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
blanchet@43934
   329
blanchet@43926
   330
(** Definitions and functions for FOL clauses and formulas for TPTP **)
blanchet@43926
   331
blanchet@43926
   332
(* The first component is the type class; the second is a "TVar" or "TFree". *)
blanchet@43926
   333
datatype type_literal =
blanchet@43926
   334
  TyLitVar of name * name |
blanchet@43926
   335
  TyLitFree of name * name
blanchet@43926
   336
blanchet@43926
   337
blanchet@43926
   338
(** Isabelle arities **)
blanchet@43926
   339
blanchet@43926
   340
datatype arity_literal =
blanchet@43926
   341
  TConsLit of name * name * name list |
blanchet@43926
   342
  TVarLit of name * name
blanchet@43926
   343
blanchet@43926
   344
fun gen_TVars 0 = []
blanchet@43934
   345
  | gen_TVars n = ("T_" ^ string_of_int n) :: gen_TVars (n-1)
blanchet@43926
   346
blanchet@43926
   347
fun pack_sort (_,[])  = []
blanchet@43926
   348
  | pack_sort (tvar, "HOL.type" :: srt) =
blanchet@43926
   349
    pack_sort (tvar, srt) (* IGNORE sort "type" *)
blanchet@43926
   350
  | pack_sort (tvar, cls :: srt) =
blanchet@43926
   351
    (`make_type_class cls, `I tvar) :: pack_sort (tvar, srt)
blanchet@43926
   352
blanchet@43927
   353
type arity_clause =
blanchet@43927
   354
  {name: string,
blanchet@43927
   355
   prem_lits: arity_literal list,
blanchet@43927
   356
   concl_lits: arity_literal}
blanchet@43926
   357
blanchet@43926
   358
(* Arity of type constructor "tcon :: (arg1, ..., argN) res" *)
blanchet@43926
   359
fun make_axiom_arity_clause (tcons, name, (cls, args)) =
blanchet@43926
   360
  let
blanchet@43926
   361
    val tvars = gen_TVars (length args)
blanchet@43926
   362
    val tvars_srts = ListPair.zip (tvars, args)
blanchet@43926
   363
  in
blanchet@43927
   364
    {name = name,
blanchet@43927
   365
     prem_lits = map TVarLit (union_all (map pack_sort tvars_srts)),
blanchet@43927
   366
     concl_lits = TConsLit (`make_type_class cls,
blanchet@43927
   367
                            `make_fixed_type_const tcons,
blanchet@43927
   368
                            tvars ~~ tvars)}
blanchet@43926
   369
  end
blanchet@43926
   370
blanchet@43926
   371
fun arity_clause _ _ (_, []) = []
blanchet@43926
   372
  | arity_clause seen n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
blanchet@43926
   373
      arity_clause seen n (tcons,ars)
blanchet@43926
   374
  | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
blanchet@43926
   375
      if member (op =) seen class then (*multiple arities for the same tycon, class pair*)
blanchet@43926
   376
          make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class ^ "_" ^ string_of_int n, ar) ::
blanchet@43926
   377
          arity_clause seen (n+1) (tcons,ars)
blanchet@43926
   378
      else
blanchet@43926
   379
          make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class, ar) ::
blanchet@43926
   380
          arity_clause (class::seen) n (tcons,ars)
blanchet@43926
   381
blanchet@43926
   382
fun multi_arity_clause [] = []
blanchet@43926
   383
  | multi_arity_clause ((tcons, ars) :: tc_arlists) =
blanchet@43926
   384
      arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
blanchet@43926
   385
blanchet@43926
   386
(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
blanchet@43926
   387
  provided its arguments have the corresponding sorts.*)
blanchet@43926
   388
fun type_class_pairs thy tycons classes =
blanchet@43934
   389
  let
blanchet@43934
   390
    val alg = Sign.classes_of thy
blanchet@43934
   391
    fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
blanchet@43934
   392
    fun add_class tycon class =
blanchet@43934
   393
      cons (class, domain_sorts tycon class)
blanchet@43934
   394
      handle Sorts.CLASS_ERROR _ => I
blanchet@43934
   395
    fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
blanchet@43934
   396
  in map try_classes tycons end
blanchet@43926
   397
blanchet@43926
   398
(*Proving one (tycon, class) membership may require proving others, so iterate.*)
blanchet@43926
   399
fun iter_type_class_pairs _ _ [] = ([], [])
blanchet@43926
   400
  | iter_type_class_pairs thy tycons classes =
blanchet@43926
   401
      let val cpairs = type_class_pairs thy tycons classes
blanchet@43926
   402
          val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
blanchet@43926
   403
            |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
blanchet@43926
   404
          val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
blanchet@43934
   405
      in (union (op =) classes' classes, union (op =) cpairs' cpairs) end
blanchet@43926
   406
blanchet@43926
   407
fun make_arity_clauses thy tycons =
blanchet@43926
   408
  iter_type_class_pairs thy tycons ##> multi_arity_clause
blanchet@43926
   409
blanchet@43926
   410
blanchet@43926
   411
(** Isabelle class relations **)
blanchet@43926
   412
blanchet@43927
   413
type class_rel_clause =
blanchet@43927
   414
  {name: string,
blanchet@43927
   415
   subclass: name,
blanchet@43927
   416
   superclass: name}
blanchet@43926
   417
blanchet@43926
   418
(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
blanchet@43926
   419
fun class_pairs _ [] _ = []
blanchet@43926
   420
  | class_pairs thy subs supers =
blanchet@43926
   421
      let
blanchet@43926
   422
        val class_less = Sorts.class_less (Sign.classes_of thy)
blanchet@43926
   423
        fun add_super sub super = class_less (sub, super) ? cons (sub, super)
blanchet@43926
   424
        fun add_supers sub = fold (add_super sub) supers
blanchet@43926
   425
      in fold add_supers subs [] end
blanchet@43926
   426
blanchet@43926
   427
fun make_class_rel_clause (sub,super) =
blanchet@43927
   428
  {name = sub ^ "_" ^ super,
blanchet@43927
   429
   subclass = `make_type_class sub,
blanchet@43927
   430
   superclass = `make_type_class super}
blanchet@43926
   431
blanchet@43926
   432
fun make_class_rel_clauses thy subs supers =
blanchet@43934
   433
  map make_class_rel_clause (class_pairs thy subs supers)
blanchet@43926
   434
blanchet@43926
   435
datatype combterm =
blanchet@43926
   436
  CombConst of name * typ * typ list |
blanchet@43926
   437
  CombVar of name * typ |
blanchet@43926
   438
  CombApp of combterm * combterm
blanchet@43926
   439
blanchet@43926
   440
fun combtyp_of (CombConst (_, T, _)) = T
blanchet@43926
   441
  | combtyp_of (CombVar (_, T)) = T
blanchet@43926
   442
  | combtyp_of (CombApp (t1, _)) = snd (dest_funT (combtyp_of t1))
blanchet@43926
   443
blanchet@43926
   444
(*gets the head of a combinator application, along with the list of arguments*)
blanchet@43926
   445
fun strip_combterm_comb u =
blanchet@43926
   446
    let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
blanchet@43926
   447
        |   stripc  x =  x
blanchet@43926
   448
    in stripc(u,[]) end
blanchet@43926
   449
blanchet@43926
   450
fun atyps_of T = fold_atyps (insert (op =)) T []
blanchet@43926
   451
blanchet@43926
   452
fun new_skolem_const_name s num_T_args =
blanchet@43926
   453
  [new_skolem_const_prefix, s, string_of_int num_T_args]
blanchet@43926
   454
  |> space_implode Long_Name.separator
blanchet@43926
   455
blanchet@43926
   456
(* Converts a term (with combinators) into a combterm. Also accumulates sort
blanchet@43926
   457
   infomation. *)
blanchet@43926
   458
fun combterm_from_term thy bs (P $ Q) =
blanchet@43926
   459
    let
blanchet@43926
   460
      val (P', P_atomics_Ts) = combterm_from_term thy bs P
blanchet@43926
   461
      val (Q', Q_atomics_Ts) = combterm_from_term thy bs Q
blanchet@43926
   462
    in (CombApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
blanchet@43926
   463
  | combterm_from_term thy _ (Const (c, T)) =
blanchet@43926
   464
    let
blanchet@43926
   465
      val tvar_list =
blanchet@43926
   466
        (if String.isPrefix old_skolem_const_prefix c then
blanchet@43926
   467
           [] |> Term.add_tvarsT T |> map TVar
blanchet@43926
   468
         else
blanchet@43926
   469
           (c, T) |> Sign.const_typargs thy)
blanchet@43926
   470
      val c' = CombConst (`make_fixed_const c, T, tvar_list)
blanchet@43926
   471
    in (c', atyps_of T) end
blanchet@43926
   472
  | combterm_from_term _ _ (Free (v, T)) =
blanchet@43926
   473
    (CombConst (`make_fixed_var v, T, []), atyps_of T)
blanchet@43926
   474
  | combterm_from_term _ _ (Var (v as (s, _), T)) =
blanchet@43926
   475
    (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
blanchet@43926
   476
       let
blanchet@43926
   477
         val Ts = T |> strip_type |> swap |> op ::
blanchet@43926
   478
         val s' = new_skolem_const_name s (length Ts)
blanchet@43926
   479
       in CombConst (`make_fixed_const s', T, Ts) end
blanchet@43926
   480
     else
blanchet@43926
   481
       CombVar ((make_schematic_var v, s), T), atyps_of T)
blanchet@43926
   482
  | combterm_from_term _ bs (Bound j) =
blanchet@43926
   483
    nth bs j
blanchet@43926
   484
    |> (fn (s, T) => (CombConst (`make_bound_var s, T, []), atyps_of T))
blanchet@43926
   485
  | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
blanchet@43926
   486
blanchet@43926
   487
datatype locality = General | Intro | Elim | Simp | Local | Assum | Chained
blanchet@43926
   488
blanchet@43926
   489
(* (quasi-)underapproximation of the truth *)
blanchet@43926
   490
fun is_locality_global Local = false
blanchet@43926
   491
  | is_locality_global Assum = false
blanchet@43926
   492
  | is_locality_global Chained = false
blanchet@43926
   493
  | is_locality_global _ = true
blanchet@43926
   494
blanchet@43484
   495
datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic
blanchet@43484
   496
datatype type_level =
blanchet@43484
   497
  All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types
blanchet@43708
   498
datatype type_heaviness = Heavy | Light
blanchet@43484
   499
blanchet@43484
   500
datatype type_system =
blanchet@43587
   501
  Simple_Types of type_level |
blanchet@43708
   502
  Preds of polymorphism * type_level * type_heaviness |
blanchet@43708
   503
  Tags of polymorphism * type_level * type_heaviness
blanchet@43484
   504
blanchet@43559
   505
fun try_unsuffixes ss s =
blanchet@43559
   506
  fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
blanchet@43559
   507
blanchet@43484
   508
fun type_sys_from_string s =
blanchet@43587
   509
  (case try (unprefix "poly_") s of
blanchet@43587
   510
     SOME s => (SOME Polymorphic, s)
blanchet@43484
   511
   | NONE =>
blanchet@43484
   512
     case try (unprefix "mono_") s of
blanchet@43587
   513
       SOME s => (SOME Monomorphic, s)
blanchet@43587
   514
     | NONE =>
blanchet@43587
   515
       case try (unprefix "mangled_") s of
blanchet@43587
   516
         SOME s => (SOME Mangled_Monomorphic, s)
blanchet@43587
   517
       | NONE => (NONE, s))
blanchet@43484
   518
  ||> (fn s =>
blanchet@43559
   519
          (* "_query" and "_bang" are for the ASCII-challenged Mirabelle. *)
blanchet@43559
   520
          case try_unsuffixes ["?", "_query"] s of
blanchet@43484
   521
            SOME s => (Nonmonotonic_Types, s)
blanchet@43484
   522
          | NONE =>
blanchet@43559
   523
            case try_unsuffixes ["!", "_bang"] s of
blanchet@43484
   524
              SOME s => (Finite_Types, s)
blanchet@43484
   525
            | NONE => (All_Types, s))
blanchet@43699
   526
  ||> apsnd (fn s =>
blanchet@43708
   527
                case try (unsuffix "_heavy") s of
blanchet@43723
   528
                  SOME s => (Heavy, s)
blanchet@43723
   529
                | NONE => (Light, s))
blanchet@43708
   530
  |> (fn (poly, (level, (heaviness, core))) =>
blanchet@43708
   531
         case (core, (poly, level, heaviness)) of
blanchet@43724
   532
           ("simple", (NONE, _, Light)) => Simple_Types level
blanchet@43723
   533
         | ("preds", (SOME poly, _, _)) => Preds (poly, level, heaviness)
blanchet@43720
   534
         | ("tags", (SOME Polymorphic, All_Types, _)) =>
blanchet@43723
   535
           Tags (Polymorphic, All_Types, heaviness)
blanchet@43755
   536
         | ("tags", (SOME Polymorphic, _, _)) =>
blanchet@43755
   537
           (* The actual light encoding is very unsound. *)
blanchet@43755
   538
           Tags (Polymorphic, level, Heavy)
blanchet@43723
   539
         | ("tags", (SOME poly, _, _)) => Tags (poly, level, heaviness)
blanchet@43723
   540
         | ("args", (SOME poly, All_Types (* naja *), Light)) =>
blanchet@43708
   541
           Preds (poly, Const_Arg_Types, Light)
blanchet@43723
   542
         | ("erased", (NONE, All_Types (* naja *), Light)) =>
blanchet@43708
   543
           Preds (Polymorphic, No_Types, Light)
blanchet@43618
   544
         | _ => raise Same.SAME)
blanchet@43618
   545
  handle Same.SAME => error ("Unknown type system: " ^ quote s ^ ".")
blanchet@43484
   546
blanchet@43587
   547
fun polymorphism_of_type_sys (Simple_Types _) = Mangled_Monomorphic
blanchet@43699
   548
  | polymorphism_of_type_sys (Preds (poly, _, _)) = poly
blanchet@43699
   549
  | polymorphism_of_type_sys (Tags (poly, _, _)) = poly
blanchet@43484
   550
blanchet@43587
   551
fun level_of_type_sys (Simple_Types level) = level
blanchet@43699
   552
  | level_of_type_sys (Preds (_, level, _)) = level
blanchet@43699
   553
  | level_of_type_sys (Tags (_, level, _)) = level
blanchet@43699
   554
blanchet@43708
   555
fun heaviness_of_type_sys (Simple_Types _) = Heavy
blanchet@43708
   556
  | heaviness_of_type_sys (Preds (_, _, heaviness)) = heaviness
blanchet@43708
   557
  | heaviness_of_type_sys (Tags (_, _, heaviness)) = heaviness
blanchet@43702
   558
blanchet@43557
   559
fun is_type_level_virtually_sound level =
blanchet@43557
   560
  level = All_Types orelse level = Nonmonotonic_Types
blanchet@43484
   561
val is_type_sys_virtually_sound =
blanchet@43484
   562
  is_type_level_virtually_sound o level_of_type_sys
blanchet@43484
   563
blanchet@43484
   564
fun is_type_level_fairly_sound level =
blanchet@43484
   565
  is_type_level_virtually_sound level orelse level = Finite_Types
blanchet@43484
   566
val is_type_sys_fairly_sound = is_type_level_fairly_sound o level_of_type_sys
blanchet@43484
   567
blanchet@43835
   568
fun is_setting_higher_order THF (Simple_Types _) = true
blanchet@43835
   569
  | is_setting_higher_order _ _ = false
blanchet@43835
   570
blanchet@40358
   571
type translated_formula =
blanchet@38991
   572
  {name: string,
blanchet@43511
   573
   locality: locality,
blanchet@43396
   574
   kind: formula_kind,
blanchet@43433
   575
   combformula: (name, typ, combterm) formula,
blanchet@43433
   576
   atomic_types: typ list}
blanchet@38506
   577
blanchet@43511
   578
fun update_combformula f ({name, locality, kind, combformula, atomic_types}
blanchet@43511
   579
                          : translated_formula) =
blanchet@43511
   580
  {name = name, locality = locality, kind = kind, combformula = f combformula,
blanchet@43433
   581
   atomic_types = atomic_types} : translated_formula
blanchet@43413
   582
blanchet@43429
   583
fun fact_lift f ({combformula, ...} : translated_formula) = f combformula
blanchet@43429
   584
blanchet@43905
   585
val type_instance = Sign.typ_instance o Proof_Context.theory_of
blanchet@43905
   586
blanchet@43905
   587
fun insert_type ctxt get_T x xs =
blanchet@43905
   588
  let val T = get_T x in
blanchet@43905
   589
    if exists (curry (type_instance ctxt) T o get_T) xs then xs
blanchet@43905
   590
    else x :: filter_out (curry (type_instance ctxt o swap) T o get_T) xs
blanchet@43905
   591
  end
blanchet@43547
   592
blanchet@43618
   593
(* The Booleans indicate whether all type arguments should be kept. *)
blanchet@43618
   594
datatype type_arg_policy =
blanchet@43618
   595
  Explicit_Type_Args of bool |
blanchet@43618
   596
  Mangled_Type_Args of bool |
blanchet@43618
   597
  No_Type_Args
blanchet@41384
   598
blanchet@43707
   599
fun should_drop_arg_type_args (Simple_Types _) =
blanchet@43707
   600
    false (* since TFF doesn't support overloading *)
blanchet@43707
   601
  | should_drop_arg_type_args type_sys =
blanchet@43707
   602
    level_of_type_sys type_sys = All_Types andalso
blanchet@43708
   603
    heaviness_of_type_sys type_sys = Heavy
blanchet@43702
   604
blanchet@43460
   605
fun general_type_arg_policy type_sys =
blanchet@43460
   606
  if level_of_type_sys type_sys = No_Types then
blanchet@43460
   607
    No_Type_Args
blanchet@43460
   608
  else if polymorphism_of_type_sys type_sys = Mangled_Monomorphic then
blanchet@43702
   609
    Mangled_Type_Args (should_drop_arg_type_args type_sys)
blanchet@43460
   610
  else
blanchet@43702
   611
    Explicit_Type_Args (should_drop_arg_type_args type_sys)
blanchet@43434
   612
blanchet@43792
   613
fun type_arg_policy type_sys s =
blanchet@43792
   614
  if s = @{const_name HOL.eq} orelse
blanchet@43807
   615
     (s = app_op_name andalso level_of_type_sys type_sys = Const_Arg_Types) then
blanchet@43792
   616
    No_Type_Args
blanchet@43792
   617
  else
blanchet@43792
   618
    general_type_arg_policy type_sys
blanchet@43088
   619
blanchet@43926
   620
(*Make literals for sorted type variables*)
blanchet@43934
   621
fun sorts_on_typs_aux (_, []) = []
blanchet@43934
   622
  | sorts_on_typs_aux ((x, i), s :: ss) =
blanchet@43934
   623
    sorts_on_typs_aux ((x, i), ss)
blanchet@43934
   624
    |> (if s = the_single @{sort HOL.type} then
blanchet@43934
   625
          I
blanchet@43934
   626
        else if i = ~1 then
blanchet@43934
   627
          cons (TyLitFree (`make_type_class s, `make_fixed_type_var x))
blanchet@43934
   628
        else
blanchet@43934
   629
          cons (TyLitVar (`make_type_class s,
blanchet@43934
   630
                          (make_schematic_type_var (x, i), x))))
blanchet@43926
   631
blanchet@43926
   632
fun sorts_on_typs (TFree (a, s)) = sorts_on_typs_aux ((a, ~1), s)
blanchet@43926
   633
  | sorts_on_typs (TVar (v, s)) = sorts_on_typs_aux (v, s)
blanchet@43926
   634
  | sorts_on_typs _ = raise Fail "expected \"TVar\" or \"TFree\""
blanchet@43926
   635
blanchet@43926
   636
(*Given a list of sorted type variables, return a list of type literals.*)
blanchet@43926
   637
val raw_type_literals_for_types = union_all o map sorts_on_typs
blanchet@43926
   638
blanchet@43926
   639
fun type_literals_for_types format type_sys kind Ts =
blanchet@43797
   640
  if level_of_type_sys type_sys = No_Types orelse format = CNF_UEQ then
blanchet@43224
   641
    []
blanchet@43224
   642
  else
blanchet@43926
   643
    Ts |> raw_type_literals_for_types
blanchet@43224
   644
       |> filter (fn TyLitVar _ => kind <> Conjecture
blanchet@43224
   645
                   | TyLitFree _ => kind = Conjecture)
blanchet@41385
   646
blanchet@43405
   647
fun mk_aconns c phis =
blanchet@43405
   648
  let val (phis', phi') = split_last phis in
blanchet@43405
   649
    fold_rev (mk_aconn c) phis' phi'
blanchet@43405
   650
  end
blanchet@38506
   651
fun mk_ahorn [] phi = phi
blanchet@43405
   652
  | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
blanchet@43393
   653
fun mk_aquant _ [] phi = phi
blanchet@43393
   654
  | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
blanchet@43393
   655
    if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
blanchet@43393
   656
  | mk_aquant q xs phi = AQuant (q, xs, phi)
blanchet@38506
   657
blanchet@43393
   658
fun close_universally atom_vars phi =
blanchet@41393
   659
  let
blanchet@41393
   660
    fun formula_vars bounds (AQuant (_, xs, phi)) =
blanchet@43397
   661
        formula_vars (map fst xs @ bounds) phi
blanchet@41393
   662
      | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
blanchet@43393
   663
      | formula_vars bounds (AAtom tm) =
blanchet@43397
   664
        union (op =) (atom_vars tm []
blanchet@43397
   665
                      |> filter_out (member (op =) bounds o fst))
blanchet@43393
   666
  in mk_aquant AForall (formula_vars [] phi []) phi end
blanchet@43393
   667
blanchet@43402
   668
fun combterm_vars (CombApp (tm1, tm2)) = fold combterm_vars [tm1, tm2]
blanchet@43393
   669
  | combterm_vars (CombConst _) = I
blanchet@43445
   670
  | combterm_vars (CombVar (name, T)) = insert (op =) (name, SOME T)
blanchet@43545
   671
fun close_combformula_universally phi = close_universally combterm_vars phi
blanchet@43393
   672
blanchet@43393
   673
fun term_vars (ATerm (name as (s, _), tms)) =
blanchet@43839
   674
  is_tptp_variable s ? insert (op =) (name, NONE) #> fold term_vars tms
blanchet@43545
   675
fun close_formula_universally phi = close_universally term_vars phi
blanchet@41393
   676
blanchet@43835
   677
val homo_infinite_type_name = @{type_name ind} (* any infinite type *)
blanchet@43835
   678
val homo_infinite_type = Type (homo_infinite_type_name, [])
blanchet@43835
   679
blanchet@43835
   680
fun fo_term_from_typ higher_order =
blanchet@43835
   681
  let
blanchet@43835
   682
    fun term (Type (s, Ts)) =
blanchet@43835
   683
      ATerm (case (higher_order, s) of
blanchet@43835
   684
               (true, @{type_name bool}) => `I tptp_bool_type
blanchet@43835
   685
             | (true, @{type_name fun}) => `I tptp_fun_type
blanchet@43835
   686
             | _ => if s = homo_infinite_type_name then `I tptp_individual_type
blanchet@43835
   687
                    else `make_fixed_type_const s,
blanchet@43835
   688
             map term Ts)
blanchet@43835
   689
    | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
blanchet@43835
   690
    | term (TVar ((x as (s, _)), _)) =
blanchet@43835
   691
      ATerm ((make_schematic_type_var x, s), [])
blanchet@43835
   692
  in term end
blanchet@43433
   693
blanchet@43433
   694
(* This shouldn't clash with anything else. *)
blanchet@43413
   695
val mangled_type_sep = "\000"
blanchet@43413
   696
blanchet@43433
   697
fun generic_mangled_type_name f (ATerm (name, [])) = f name
blanchet@43433
   698
  | generic_mangled_type_name f (ATerm (name, tys)) =
blanchet@43626
   699
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
blanchet@43626
   700
    ^ ")"
blanchet@43413
   701
blanchet@43839
   702
val bool_atype = AType (`I tptp_bool_type)
blanchet@43839
   703
blanchet@43926
   704
fun make_simple_type s =
blanchet@43926
   705
  if s = tptp_bool_type orelse s = tptp_fun_type orelse
blanchet@43926
   706
     s = tptp_individual_type then
blanchet@43926
   707
    s
blanchet@43926
   708
  else
blanchet@43926
   709
    simple_type_prefix ^ ascii_of s
blanchet@43926
   710
blanchet@43835
   711
fun ho_type_from_fo_term higher_order pred_sym ary =
blanchet@43804
   712
  let
blanchet@43804
   713
    fun to_atype ty =
blanchet@43804
   714
      AType ((make_simple_type (generic_mangled_type_name fst ty),
blanchet@43804
   715
              generic_mangled_type_name snd ty))
blanchet@43804
   716
    fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
blanchet@43839
   717
    fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
blanchet@43835
   718
      | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
blanchet@43835
   719
    fun to_ho (ty as ATerm ((s, _), tys)) =
blanchet@43835
   720
      if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
blanchet@43835
   721
  in if higher_order then to_ho else to_fo ary end
blanchet@43804
   722
blanchet@43835
   723
fun mangled_type higher_order pred_sym ary =
blanchet@43835
   724
  ho_type_from_fo_term higher_order pred_sym ary o fo_term_from_typ higher_order
blanchet@43804
   725
blanchet@43835
   726
fun mangled_const_name T_args (s, s') =
blanchet@43804
   727
  let
blanchet@43835
   728
    val ty_args = map (fo_term_from_typ false) T_args
blanchet@43804
   729
    fun type_suffix f g =
blanchet@43804
   730
      fold_rev (curry (op ^) o g o prefix mangled_type_sep
blanchet@43804
   731
                o generic_mangled_type_name f) ty_args ""
blanchet@43804
   732
  in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
blanchet@43413
   733
blanchet@43413
   734
val parse_mangled_ident =
blanchet@43413
   735
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
blanchet@43413
   736
blanchet@43413
   737
fun parse_mangled_type x =
blanchet@43413
   738
  (parse_mangled_ident
blanchet@43413
   739
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
blanchet@43413
   740
                    [] >> ATerm) x
blanchet@43413
   741
and parse_mangled_types x =
blanchet@43413
   742
  (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
blanchet@43413
   743
blanchet@43413
   744
fun unmangled_type s =
blanchet@43413
   745
  s |> suffix ")" |> raw_explode
blanchet@43413
   746
    |> Scan.finite Symbol.stopper
blanchet@43413
   747
           (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
blanchet@43413
   748
                                                quote s)) parse_mangled_type))
blanchet@43413
   749
    |> fst
blanchet@43413
   750
blanchet@43432
   751
val unmangled_const_name = space_explode mangled_type_sep #> hd
blanchet@43413
   752
fun unmangled_const s =
blanchet@43413
   753
  let val ss = space_explode mangled_type_sep s in
blanchet@43413
   754
    (hd ss, map unmangled_type (tl ss))
blanchet@43413
   755
  end
blanchet@43413
   756
blanchet@43858
   757
fun introduce_proxies format type_sys =
blanchet@43439
   758
  let
blanchet@43858
   759
    fun intro top_level (CombApp (tm1, tm2)) =
blanchet@43858
   760
        CombApp (intro top_level tm1, intro false tm2)
blanchet@43858
   761
      | intro top_level (CombConst (name as (s, _), T, T_args)) =
blanchet@43441
   762
        (case proxify_const s of
blanchet@43858
   763
           SOME (_, proxy_base) =>
blanchet@43841
   764
           if top_level orelse is_setting_higher_order format type_sys then
blanchet@43841
   765
             case (top_level, s) of
blanchet@43841
   766
               (_, "c_False") => (`I tptp_false, [])
blanchet@43841
   767
             | (_, "c_True") => (`I tptp_true, [])
blanchet@43841
   768
             | (false, "c_Not") => (`I tptp_not, [])
blanchet@43841
   769
             | (false, "c_conj") => (`I tptp_and, [])
blanchet@43841
   770
             | (false, "c_disj") => (`I tptp_or, [])
blanchet@43841
   771
             | (false, "c_implies") => (`I tptp_implies, [])
blanchet@43841
   772
             | (false, s) =>
blanchet@43858
   773
               if is_tptp_equal s then (`I tptp_equal, [])
blanchet@43858
   774
               else (proxy_base |>> prefix const_prefix, T_args)
blanchet@43841
   775
             | _ => (name, [])
blanchet@43440
   776
           else
blanchet@43445
   777
             (proxy_base |>> prefix const_prefix, T_args)
blanchet@43445
   778
          | NONE => (name, T_args))
blanchet@43445
   779
        |> (fn (name, T_args) => CombConst (name, T, T_args))
blanchet@43858
   780
      | intro _ tm = tm
blanchet@43858
   781
  in intro true end
blanchet@43439
   782
blanchet@43835
   783
fun combformula_from_prop thy format type_sys eq_as_iff =
blanchet@38506
   784
  let
blanchet@43439
   785
    fun do_term bs t atomic_types =
blanchet@41388
   786
      combterm_from_term thy bs (Envir.eta_contract t)
blanchet@43835
   787
      |>> (introduce_proxies format type_sys #> AAtom)
blanchet@43439
   788
      ||> union (op =) atomic_types
blanchet@38506
   789
    fun do_quant bs q s T t' =
blanchet@38743
   790
      let val s = Name.variant (map fst bs) s in
blanchet@38743
   791
        do_formula ((s, T) :: bs) t'
blanchet@43433
   792
        #>> mk_aquant q [(`make_bound_var s, SOME T)]
blanchet@38743
   793
      end
blanchet@38506
   794
    and do_conn bs c t1 t2 =
blanchet@38506
   795
      do_formula bs t1 ##>> do_formula bs t2
blanchet@43402
   796
      #>> uncurry (mk_aconn c)
blanchet@38506
   797
    and do_formula bs t =
blanchet@38506
   798
      case t of
blanchet@43402
   799
        @{const Not} $ t1 => do_formula bs t1 #>> mk_anot
blanchet@38506
   800
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@38506
   801
        do_quant bs AForall s T t'
blanchet@38506
   802
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@38506
   803
        do_quant bs AExists s T t'
haftmann@39028
   804
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
haftmann@39028
   805
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
haftmann@39019
   806
      | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
haftmann@39093
   807
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@41388
   808
        if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t
blanchet@41388
   809
      | _ => do_term bs t
blanchet@38506
   810
  in do_formula [] end
blanchet@38506
   811
blanchet@43615
   812
fun presimplify_term ctxt =
blanchet@43615
   813
  Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
blanchet@43615
   814
  #> Meson.presimplify ctxt
blanchet@43615
   815
  #> prop_of
blanchet@38506
   816
wenzelm@41739
   817
fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j
blanchet@38506
   818
fun conceal_bounds Ts t =
blanchet@38506
   819
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@38506
   820
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@38506
   821
fun reveal_bounds Ts =
blanchet@38506
   822
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@38506
   823
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@38506
   824
blanchet@43612
   825
fun extensionalize_term ctxt t =
blanchet@43612
   826
  let val thy = Proof_Context.theory_of ctxt in
blanchet@43612
   827
    t |> cterm_of thy |> Meson.extensionalize_conv ctxt
blanchet@43612
   828
      |> prop_of |> Logic.dest_equals |> snd
blanchet@43612
   829
  end
blanchet@38831
   830
blanchet@38506
   831
fun introduce_combinators_in_term ctxt kind t =
wenzelm@43232
   832
  let val thy = Proof_Context.theory_of ctxt in
blanchet@38716
   833
    if Meson.is_fol_term thy t then
blanchet@38716
   834
      t
blanchet@38716
   835
    else
blanchet@38716
   836
      let
blanchet@38716
   837
        fun aux Ts t =
blanchet@38716
   838
          case t of
blanchet@38716
   839
            @{const Not} $ t1 => @{const Not} $ aux Ts t1
blanchet@38716
   840
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@38716
   841
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   842
          | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@38890
   843
            aux Ts (t0 $ eta_expand Ts t1 1)
blanchet@38716
   844
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@38716
   845
            t0 $ Abs (s, T, aux (T :: Ts) t')
blanchet@38890
   846
          | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@38890
   847
            aux Ts (t0 $ eta_expand Ts t1 1)
haftmann@39028
   848
          | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39028
   849
          | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39019
   850
          | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
haftmann@39093
   851
          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
blanchet@38716
   852
              $ t1 $ t2 =>
blanchet@38716
   853
            t0 $ aux Ts t1 $ aux Ts t2
blanchet@38716
   854
          | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
blanchet@38716
   855
                   t
blanchet@38716
   856
                 else
blanchet@38716
   857
                   t |> conceal_bounds Ts
blanchet@38716
   858
                     |> Envir.eta_contract
blanchet@38716
   859
                     |> cterm_of thy
blanchet@40071
   860
                     |> Meson_Clausify.introduce_combinators_in_cterm
blanchet@38716
   861
                     |> prop_of |> Logic.dest_equals |> snd
blanchet@38716
   862
                     |> reveal_bounds Ts
blanchet@39616
   863
        val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
blanchet@38716
   864
      in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@38716
   865
      handle THM _ =>
blanchet@38716
   866
             (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@38836
   867
             if kind = Conjecture then HOLogic.false_const
blanchet@38836
   868
             else HOLogic.true_const
blanchet@38716
   869
  end
blanchet@38506
   870
blanchet@38506
   871
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@43224
   872
   same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
blanchet@38506
   873
fun freeze_term t =
blanchet@38506
   874
  let
blanchet@38506
   875
    fun aux (t $ u) = aux t $ aux u
blanchet@38506
   876
      | aux (Abs (s, T, t)) = Abs (s, T, aux t)
blanchet@38506
   877
      | aux (Var ((s, i), T)) =
blanchet@38506
   878
        Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@38506
   879
      | aux t = t
blanchet@38506
   880
  in t |> exists_subterm is_Var t ? aux end
blanchet@38506
   881
blanchet@40445
   882
(* making fact and conjecture formulas *)
blanchet@43835
   883
fun make_formula ctxt format type_sys eq_as_iff presimp name loc kind t =
blanchet@38506
   884
  let
wenzelm@43232
   885
    val thy = Proof_Context.theory_of ctxt
blanchet@38831
   886
    val t = t |> Envir.beta_eta_contract
blanchet@43785
   887
              |> transform_elim_prop
blanchet@41459
   888
              |> Object_Logic.atomize_term thy
blanchet@43434
   889
    val need_trueprop = (fastype_of t = @{typ bool})
blanchet@38890
   890
    val t = t |> need_trueprop ? HOLogic.mk_Trueprop
blanchet@43607
   891
              |> Raw_Simplifier.rewrite_term thy
blanchet@43607
   892
                     (Meson.unfold_set_const_simps ctxt) []
blanchet@43612
   893
              |> extensionalize_term ctxt
blanchet@43615
   894
              |> presimp ? presimplify_term ctxt
blanchet@38506
   895
              |> perhaps (try (HOLogic.dest_Trueprop))
blanchet@38506
   896
              |> introduce_combinators_in_term ctxt kind
blanchet@38836
   897
              |> kind <> Axiom ? freeze_term
blanchet@43803
   898
    val (combformula, atomic_types) =
blanchet@43835
   899
      combformula_from_prop thy format type_sys eq_as_iff t []
blanchet@38506
   900
  in
blanchet@43511
   901
    {name = name, locality = loc, kind = kind, combformula = combformula,
blanchet@43433
   902
     atomic_types = atomic_types}
blanchet@38506
   903
  end
blanchet@38506
   904
blanchet@43835
   905
fun make_fact ctxt format type_sys keep_trivial eq_as_iff presimp
blanchet@43835
   906
              ((name, loc), t) =
blanchet@43803
   907
  case (keep_trivial,
blanchet@43835
   908
        make_formula ctxt format type_sys eq_as_iff presimp name loc Axiom t) of
blanchet@43842
   909
    (false, formula as {combformula = AAtom (CombConst ((s, _), _, _)), ...}) =>
blanchet@43842
   910
    if s = tptp_true then NONE else SOME formula
blanchet@42861
   911
  | (_, formula) => SOME formula
blanchet@43432
   912
blanchet@43835
   913
fun make_conjecture ctxt format prem_kind type_sys ts =
blanchet@38836
   914
  let val last = length ts - 1 in
blanchet@43580
   915
    map2 (fn j => fn t =>
blanchet@43580
   916
             let
blanchet@43580
   917
               val (kind, maybe_negate) =
blanchet@43580
   918
                 if j = last then
blanchet@43580
   919
                   (Conjecture, I)
blanchet@43580
   920
                 else
blanchet@43580
   921
                   (prem_kind,
blanchet@43580
   922
                    if prem_kind = Conjecture then update_combformula mk_anot
blanchet@43580
   923
                    else I)
blanchet@43580
   924
              in
blanchet@43835
   925
                t |> make_formula ctxt format type_sys true true
blanchet@43835
   926
                                  (string_of_int j) General kind
blanchet@43803
   927
                  |> maybe_negate
blanchet@43580
   928
              end)
blanchet@38836
   929
         (0 upto last) ts
blanchet@38836
   930
  end
blanchet@38506
   931
blanchet@43552
   932
(** Finite and infinite type inference **)
blanchet@43552
   933
blanchet@43755
   934
fun deep_freeze_atyp (TVar (_, S)) = TFree ("v", S)
blanchet@43755
   935
  | deep_freeze_atyp T = T
blanchet@43755
   936
val deep_freeze_type = map_atyps deep_freeze_atyp
blanchet@43755
   937
blanchet@43552
   938
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
blanchet@43552
   939
   dangerous because their "exhaust" properties can easily lead to unsound ATP
blanchet@43552
   940
   proofs. On the other hand, all HOL infinite types can be given the same
blanchet@43552
   941
   models in first-order logic (via Löwenheim-Skolem). *)
blanchet@43552
   942
blanchet@43755
   943
fun should_encode_type ctxt (nonmono_Ts as _ :: _) _ T =
blanchet@43755
   944
    exists (curry (type_instance ctxt) (deep_freeze_type T)) nonmono_Ts
blanchet@43707
   945
  | should_encode_type _ _ All_Types _ = true
blanchet@43552
   946
  | should_encode_type ctxt _ Finite_Types T = is_type_surely_finite ctxt T
blanchet@43552
   947
  | should_encode_type _ _ _ _ = false
blanchet@43552
   948
blanchet@43708
   949
fun should_predicate_on_type ctxt nonmono_Ts (Preds (_, level, heaviness))
blanchet@43705
   950
                             should_predicate_on_var T =
blanchet@43747
   951
    (heaviness = Heavy orelse should_predicate_on_var ()) andalso
blanchet@43747
   952
    should_encode_type ctxt nonmono_Ts level T
blanchet@43705
   953
  | should_predicate_on_type _ _ _ _ _ = false
blanchet@43552
   954
blanchet@43707
   955
fun is_var_or_bound_var (CombConst ((s, _), _, _)) =
blanchet@43707
   956
    String.isPrefix bound_var_prefix s
blanchet@43707
   957
  | is_var_or_bound_var (CombVar _) = true
blanchet@43707
   958
  | is_var_or_bound_var _ = false
blanchet@43707
   959
blanchet@43700
   960
datatype tag_site = Top_Level | Eq_Arg | Elsewhere
blanchet@43700
   961
blanchet@43700
   962
fun should_tag_with_type _ _ _ Top_Level _ _ = false
blanchet@43708
   963
  | should_tag_with_type ctxt nonmono_Ts (Tags (_, level, heaviness)) site u T =
blanchet@43708
   964
    (case heaviness of
blanchet@43708
   965
       Heavy => should_encode_type ctxt nonmono_Ts level T
blanchet@43708
   966
     | Light =>
blanchet@43707
   967
       case (site, is_var_or_bound_var u) of
blanchet@43707
   968
         (Eq_Arg, true) => should_encode_type ctxt nonmono_Ts level T
blanchet@43700
   969
       | _ => false)
blanchet@43700
   970
  | should_tag_with_type _ _ _ _ _ _ = false
blanchet@43552
   971
blanchet@43835
   972
fun homogenized_type ctxt nonmono_Ts level =
blanchet@43835
   973
  let
blanchet@43835
   974
    val should_encode = should_encode_type ctxt nonmono_Ts level
blanchet@43835
   975
    fun homo 0 T = if should_encode T then T else homo_infinite_type
blanchet@43835
   976
      | homo ary (Type (@{type_name fun}, [T1, T2])) =
blanchet@43835
   977
        homo 0 T1 --> homo (ary - 1) T2
blanchet@43835
   978
      | homo _ _ = raise Fail "expected function type"
blanchet@43835
   979
  in homo end
blanchet@43552
   980
blanchet@43444
   981
(** "hBOOL" and "hAPP" **)
blanchet@41561
   982
blanchet@43445
   983
type sym_info =
blanchet@43905
   984
  {pred_sym : bool, min_ary : int, max_ary : int, types : typ list}
blanchet@43434
   985
blanchet@43905
   986
fun add_combterm_syms_to_table ctxt explicit_apply =
blanchet@43429
   987
  let
blanchet@43905
   988
    fun consider_var_arity const_T var_T max_ary =
blanchet@43905
   989
      let
blanchet@43905
   990
        fun iter ary T =
blanchet@43905
   991
          if ary = max_ary orelse type_instance ctxt (var_T, T) then ary
blanchet@43905
   992
          else iter (ary + 1) (range_type T)
blanchet@43905
   993
      in iter 0 const_T end
blanchet@43905
   994
    fun add top_level tm (accum as (ho_var_Ts, sym_tab)) =
blanchet@43429
   995
      let val (head, args) = strip_combterm_comb tm in
blanchet@43429
   996
        (case head of
blanchet@43434
   997
           CombConst ((s, _), T, _) =>
blanchet@43429
   998
           if String.isPrefix bound_var_prefix s then
blanchet@43905
   999
             if explicit_apply = NONE andalso can dest_funT T then
blanchet@43905
  1000
               let
blanchet@43905
  1001
                 fun repair_min_arity {pred_sym, min_ary, max_ary, types} =
blanchet@43905
  1002
                   {pred_sym = pred_sym,
blanchet@43905
  1003
                    min_ary =
blanchet@43905
  1004
                      fold (fn T' => consider_var_arity T' T) types min_ary,
blanchet@43905
  1005
                    max_ary = max_ary, types = types}
blanchet@43905
  1006
                 val ho_var_Ts' = ho_var_Ts |> insert_type ctxt I T
blanchet@43905
  1007
               in
blanchet@43905
  1008
                 if pointer_eq (ho_var_Ts', ho_var_Ts) then accum
blanchet@43905
  1009
                 else (ho_var_Ts', Symtab.map (K repair_min_arity) sym_tab)
blanchet@43905
  1010
               end
blanchet@43905
  1011
             else
blanchet@43905
  1012
               accum
blanchet@43429
  1013
           else
blanchet@43905
  1014
             let
blanchet@43905
  1015
               val ary = length args
blanchet@43905
  1016
             in
blanchet@43905
  1017
               (ho_var_Ts,
blanchet@43905
  1018
                case Symtab.lookup sym_tab s of
blanchet@43905
  1019
                  SOME {pred_sym, min_ary, max_ary, types} =>
blanchet@43905
  1020
                  let
blanchet@43905
  1021
                    val types' = types |> insert_type ctxt I T
blanchet@43905
  1022
                    val min_ary =
blanchet@43905
  1023
                      if is_some explicit_apply orelse
blanchet@43905
  1024
                         pointer_eq (types', types) then
blanchet@43905
  1025
                        min_ary
blanchet@43905
  1026
                      else
blanchet@43905
  1027
                        fold (consider_var_arity T) ho_var_Ts min_ary
blanchet@43905
  1028
                  in
blanchet@43905
  1029
                    Symtab.update (s, {pred_sym = pred_sym andalso top_level,
blanchet@43905
  1030
                                       min_ary = Int.min (ary, min_ary),
blanchet@43905
  1031
                                       max_ary = Int.max (ary, max_ary),
blanchet@43905
  1032
                                       types = types'})
blanchet@43905
  1033
                                  sym_tab
blanchet@43905
  1034
                  end
blanchet@43905
  1035
                | NONE =>
blanchet@43905
  1036
                  let
blanchet@43905
  1037
                    val min_ary =
blanchet@43905
  1038
                      case explicit_apply of
blanchet@43905
  1039
                        SOME true => 0
blanchet@43905
  1040
                      | SOME false => ary
blanchet@43905
  1041
                      | NONE => fold (consider_var_arity T) ho_var_Ts ary
blanchet@43905
  1042
                  in
blanchet@43905
  1043
                    Symtab.update_new (s, {pred_sym = top_level,
blanchet@43905
  1044
                                           min_ary = min_ary, max_ary = ary,
blanchet@43905
  1045
                                           types = [T]})
blanchet@43905
  1046
                                      sym_tab
blanchet@43905
  1047
                  end)
blanchet@43905
  1048
             end
blanchet@43905
  1049
         | _ => accum)
blanchet@43905
  1050
        |> fold (add false) args
blanchet@43429
  1051
      end
blanchet@43905
  1052
  in add true end
blanchet@43905
  1053
fun add_fact_syms_to_table ctxt explicit_apply =
blanchet@43905
  1054
  fact_lift (formula_fold NONE
blanchet@43905
  1055
                          (K (add_combterm_syms_to_table ctxt explicit_apply)))
blanchet@38506
  1056
blanchet@43546
  1057
val default_sym_table_entries : (string * sym_info) list =
blanchet@43905
  1058
  [(tptp_equal, {pred_sym = true, min_ary = 2, max_ary = 2, types = []}),
blanchet@43905
  1059
   (tptp_old_equal, {pred_sym = true, min_ary = 2, max_ary = 2, types = []}),
blanchet@43807
  1060
   (make_fixed_const predicator_name,
blanchet@43905
  1061
    {pred_sym = true, min_ary = 1, max_ary = 1, types = []})] @
blanchet@43439
  1062
  ([tptp_false, tptp_true]
blanchet@43905
  1063
   |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = []}))
blanchet@41388
  1064
blanchet@43905
  1065
fun sym_table_for_facts ctxt explicit_apply facts =
blanchet@43905
  1066
  Symtab.empty
blanchet@43905
  1067
  |> fold Symtab.default default_sym_table_entries
blanchet@43905
  1068
  |> pair [] |> fold (add_fact_syms_to_table ctxt explicit_apply) facts |> snd
blanchet@38506
  1069
blanchet@43429
  1070
fun min_arity_of sym_tab s =
blanchet@43429
  1071
  case Symtab.lookup sym_tab s of
blanchet@43445
  1072
    SOME ({min_ary, ...} : sym_info) => min_ary
blanchet@43429
  1073
  | NONE =>
blanchet@43429
  1074
    case strip_prefix_and_unascii const_prefix s of
blanchet@43418
  1075
      SOME s =>
blanchet@43441
  1076
      let val s = s |> unmangled_const_name |> invert_const in
blanchet@43807
  1077
        if s = predicator_name then 1
blanchet@43807
  1078
        else if s = app_op_name then 2
blanchet@43807
  1079
        else if s = type_pred_name then 1
blanchet@43428
  1080
        else 0
blanchet@43418
  1081
      end
blanchet@38506
  1082
    | NONE => 0
blanchet@38506
  1083
blanchet@38506
  1084
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
  1085
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
  1086
   type instantiations). If false, the constant always receives all of its
blanchet@38506
  1087
   arguments and is used as a predicate. *)
blanchet@43429
  1088
fun is_pred_sym sym_tab s =
blanchet@43429
  1089
  case Symtab.lookup sym_tab s of
blanchet@43445
  1090
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
blanchet@43445
  1091
    pred_sym andalso min_ary = max_ary
blanchet@43429
  1092
  | NONE => false
blanchet@38506
  1093
blanchet@43439
  1094
val predicator_combconst =
blanchet@43807
  1095
  CombConst (`make_fixed_const predicator_name, @{typ "bool => bool"}, [])
blanchet@43439
  1096
fun predicator tm = CombApp (predicator_combconst, tm)
blanchet@38506
  1097
blanchet@43439
  1098
fun introduce_predicators_in_combterm sym_tab tm =
blanchet@43413
  1099
  case strip_combterm_comb tm of
blanchet@43413
  1100
    (CombConst ((s, _), _, _), _) =>
blanchet@43439
  1101
    if is_pred_sym sym_tab s then tm else predicator tm
blanchet@43439
  1102
  | _ => predicator tm
blanchet@38506
  1103
blanchet@43415
  1104
fun list_app head args = fold (curry (CombApp o swap)) args head
blanchet@38506
  1105
blanchet@43415
  1106
fun explicit_app arg head =
blanchet@43415
  1107
  let
blanchet@43433
  1108
    val head_T = combtyp_of head
blanchet@43563
  1109
    val (arg_T, res_T) = dest_funT head_T
blanchet@43415
  1110
    val explicit_app =
blanchet@43807
  1111
      CombConst (`make_fixed_const app_op_name, head_T --> head_T,
blanchet@43563
  1112
                 [arg_T, res_T])
blanchet@43415
  1113
  in list_app explicit_app [head, arg] end
blanchet@43415
  1114
fun list_explicit_app head args = fold explicit_app args head
blanchet@43415
  1115
blanchet@43436
  1116
fun introduce_explicit_apps_in_combterm sym_tab =
blanchet@43415
  1117
  let
blanchet@43415
  1118
    fun aux tm =
blanchet@43415
  1119
      case strip_combterm_comb tm of
blanchet@43415
  1120
        (head as CombConst ((s, _), _, _), args) =>
blanchet@43415
  1121
        args |> map aux
blanchet@43428
  1122
             |> chop (min_arity_of sym_tab s)
blanchet@43415
  1123
             |>> list_app head
blanchet@43415
  1124
             |-> list_explicit_app
blanchet@43415
  1125
      | (head, args) => list_explicit_app head (map aux args)
blanchet@43415
  1126
  in aux end
blanchet@43415
  1127
blanchet@43618
  1128
fun chop_fun 0 T = ([], T)
blanchet@43618
  1129
  | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
blanchet@43618
  1130
    chop_fun (n - 1) ran_T |>> cons dom_T
blanchet@43618
  1131
  | chop_fun _ _ = raise Fail "unexpected non-function"
blanchet@43618
  1132
blanchet@43651
  1133
fun filter_type_args _ _ _ [] = []
blanchet@43651
  1134
  | filter_type_args thy s arity T_args =
blanchet@43705
  1135
    let
blanchet@43705
  1136
      (* will throw "TYPE" for pseudo-constants *)
blanchet@43807
  1137
      val U = if s = app_op_name then
blanchet@43705
  1138
                @{typ "('a => 'b) => 'a => 'b"} |> Logic.varifyT_global
blanchet@43705
  1139
              else
blanchet@43705
  1140
                s |> Sign.the_const_type thy
blanchet@43705
  1141
    in
blanchet@43652
  1142
      case Term.add_tvarsT (U |> chop_fun arity |> snd) [] of
blanchet@43652
  1143
        [] => []
blanchet@43652
  1144
      | res_U_vars =>
blanchet@43652
  1145
        let val U_args = (s, U) |> Sign.const_typargs thy in
blanchet@43652
  1146
          U_args ~~ T_args
blanchet@43652
  1147
          |> map_filter (fn (U, T) =>
blanchet@43652
  1148
                            if member (op =) res_U_vars (dest_TVar U) then
blanchet@43652
  1149
                              SOME T
blanchet@43652
  1150
                            else
blanchet@43652
  1151
                              NONE)
blanchet@43652
  1152
        end
blanchet@43651
  1153
    end
blanchet@43651
  1154
    handle TYPE _ => T_args
blanchet@43618
  1155
blanchet@43835
  1156
fun enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys =
blanchet@43618
  1157
  let
blanchet@43618
  1158
    val thy = Proof_Context.theory_of ctxt
blanchet@43618
  1159
    fun aux arity (CombApp (tm1, tm2)) =
blanchet@43618
  1160
        CombApp (aux (arity + 1) tm1, aux 0 tm2)
blanchet@43618
  1161
      | aux arity (CombConst (name as (s, _), T, T_args)) =
blanchet@43571
  1162
        let
blanchet@43571
  1163
          val level = level_of_type_sys type_sys
blanchet@43571
  1164
          val (T, T_args) =
blanchet@43571
  1165
            (* Aggressively merge most "hAPPs" if the type system is unsound
blanchet@43571
  1166
               anyway, by distinguishing overloads only on the homogenized
blanchet@43708
  1167
               result type. Don't do it for lightweight type systems, though,
blanchet@43708
  1168
               since it leads to too many unsound proofs. *)
blanchet@43807
  1169
            if s = const_prefix ^ app_op_name andalso
blanchet@43591
  1170
               length T_args = 2 andalso
blanchet@43707
  1171
               not (is_type_sys_virtually_sound type_sys) andalso
blanchet@43708
  1172
               heaviness_of_type_sys type_sys = Heavy then
blanchet@43835
  1173
              T_args |> map (homogenized_type ctxt nonmono_Ts level 0)
blanchet@43571
  1174
                     |> (fn Ts => let val T = hd Ts --> nth Ts 1 in
blanchet@43702
  1175
                                    (T --> T, tl Ts)
blanchet@43571
  1176
                                  end)
blanchet@43571
  1177
            else
blanchet@43571
  1178
              (T, T_args)
blanchet@43571
  1179
        in
blanchet@43571
  1180
          (case strip_prefix_and_unascii const_prefix s of
blanchet@43571
  1181
             NONE => (name, T_args)
blanchet@43571
  1182
           | SOME s'' =>
blanchet@43618
  1183
             let
blanchet@43618
  1184
               val s'' = invert_const s''
blanchet@43702
  1185
               fun filtered_T_args false = T_args
blanchet@43702
  1186
                 | filtered_T_args true = filter_type_args thy s'' arity T_args
blanchet@43618
  1187
             in
blanchet@43571
  1188
               case type_arg_policy type_sys s'' of
blanchet@43702
  1189
                 Explicit_Type_Args drop_args =>
blanchet@43702
  1190
                 (name, filtered_T_args drop_args)
blanchet@43702
  1191
               | Mangled_Type_Args drop_args =>
blanchet@43835
  1192
                 (mangled_const_name (filtered_T_args drop_args) name, [])
blanchet@43618
  1193
               | No_Type_Args => (name, [])
blanchet@43571
  1194
             end)
blanchet@43571
  1195
          |> (fn (name, T_args) => CombConst (name, T, T_args))
blanchet@43571
  1196
        end
blanchet@43618
  1197
      | aux _ tm = tm
blanchet@43618
  1198
  in aux 0 end
blanchet@43444
  1199
blanchet@43803
  1200
fun repair_combterm ctxt format nonmono_Ts type_sys sym_tab =
blanchet@43835
  1201
  not (is_setting_higher_order format type_sys)
blanchet@43835
  1202
  ? (introduce_explicit_apps_in_combterm sym_tab
blanchet@43835
  1203
     #> introduce_predicators_in_combterm sym_tab)
blanchet@43835
  1204
  #> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43803
  1205
fun repair_fact ctxt format nonmono_Ts type_sys sym_tab =
blanchet@43571
  1206
  update_combformula (formula_map
blanchet@43803
  1207
      (repair_combterm ctxt format nonmono_Ts type_sys sym_tab))
blanchet@43444
  1208
blanchet@43444
  1209
(** Helper facts **)
blanchet@43444
  1210
blanchet@43926
  1211
(* The Boolean indicates that a fairly sound type encoding is needed. *)
blanchet@43926
  1212
val helper_table =
blanchet@43926
  1213
  [("COMBI", (false, @{thms Meson.COMBI_def})),
blanchet@43926
  1214
   ("COMBK", (false, @{thms Meson.COMBK_def})),
blanchet@43926
  1215
   ("COMBB", (false, @{thms Meson.COMBB_def})),
blanchet@43926
  1216
   ("COMBC", (false, @{thms Meson.COMBC_def})),
blanchet@43926
  1217
   ("COMBS", (false, @{thms Meson.COMBS_def})),
blanchet@43926
  1218
   ("fequal",
blanchet@43926
  1219
    (* This is a lie: Higher-order equality doesn't need a sound type encoding.
blanchet@43926
  1220
       However, this is done so for backward compatibility: Including the
blanchet@43926
  1221
       equality helpers by default in Metis breaks a few existing proofs. *)
blanchet@43926
  1222
    (true, @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
blanchet@43926
  1223
                  fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]})),
blanchet@43926
  1224
   ("fFalse", (true, @{thms True_or_False})),
blanchet@43926
  1225
   ("fFalse", (false, [@{lemma "~ fFalse" by (unfold fFalse_def) fast}])),
blanchet@43926
  1226
   ("fTrue", (true, @{thms True_or_False})),
blanchet@43926
  1227
   ("fTrue", (false, [@{lemma "fTrue" by (unfold fTrue_def) fast}])),
blanchet@43926
  1228
   ("fNot",
blanchet@43926
  1229
    (false, @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
blanchet@43926
  1230
                   fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]})),
blanchet@43926
  1231
   ("fconj",
blanchet@43926
  1232
    (false,
blanchet@43926
  1233
     @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
blanchet@43926
  1234
         by (unfold fconj_def) fast+})),
blanchet@43926
  1235
   ("fdisj",
blanchet@43926
  1236
    (false,
blanchet@43926
  1237
     @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
blanchet@43926
  1238
         by (unfold fdisj_def) fast+})),
blanchet@43926
  1239
   ("fimplies",
blanchet@43926
  1240
    (false, @{lemma "P | fimplies P Q" "~ Q | fimplies P Q"
blanchet@43926
  1241
                    "~ fimplies P Q | ~ P | Q"
blanchet@43926
  1242
                by (unfold fimplies_def) fast+})),
blanchet@43926
  1243
   ("If", (true, @{thms if_True if_False True_or_False}))]
blanchet@43926
  1244
blanchet@43444
  1245
fun ti_ti_helper_fact () =
blanchet@43444
  1246
  let
blanchet@43444
  1247
    fun var s = ATerm (`I s, [])
blanchet@43460
  1248
    fun tag tm = ATerm (`make_fixed_const type_tag_name, [var "X", tm])
blanchet@43444
  1249
  in
blanchet@43483
  1250
    Formula (helper_prefix ^ "ti_ti", Axiom,
blanchet@43841
  1251
             AAtom (ATerm (`I tptp_equal, [tag (tag (var "Y")), tag (var "Y")]))
blanchet@43748
  1252
             |> close_formula_universally, simp_info, NONE)
blanchet@43444
  1253
  end
blanchet@43444
  1254
blanchet@43905
  1255
fun helper_facts_for_sym ctxt format type_sys (s, {types, ...} : sym_info) =
blanchet@43444
  1256
  case strip_prefix_and_unascii const_prefix s of
blanchet@43444
  1257
    SOME mangled_s =>
blanchet@43444
  1258
    let
blanchet@43444
  1259
      val thy = Proof_Context.theory_of ctxt
blanchet@43444
  1260
      val unmangled_s = mangled_s |> unmangled_const_name
blanchet@43764
  1261
      fun dub_and_inst c needs_fairly_sound (th, j) =
blanchet@43750
  1262
        ((c ^ "_" ^ string_of_int j ^
blanchet@43764
  1263
          (if needs_fairly_sound then typed_helper_suffix
blanchet@43750
  1264
           else untyped_helper_suffix),
blanchet@43750
  1265
          General),
blanchet@43444
  1266
         let val t = th |> prop_of in
blanchet@43618
  1267
           t |> ((case general_type_arg_policy type_sys of
blanchet@43618
  1268
                    Mangled_Type_Args _ => true
blanchet@43618
  1269
                  | _ => false) andalso
blanchet@43444
  1270
                 not (null (Term.hidden_polymorphism t)))
blanchet@43905
  1271
                ? (case types of
blanchet@43905
  1272
                     [T] => specialize_type thy (invert_const unmangled_s, T)
blanchet@43905
  1273
                   | _ => I)
blanchet@43444
  1274
         end)
blanchet@43444
  1275
      fun make_facts eq_as_iff =
blanchet@43835
  1276
        map_filter (make_fact ctxt format type_sys false eq_as_iff false)
blanchet@43764
  1277
      val fairly_sound = is_type_sys_fairly_sound type_sys
blanchet@43444
  1278
    in
blanchet@43926
  1279
      helper_table
blanchet@43765
  1280
      |> maps (fn (metis_s, (needs_fairly_sound, ths)) =>
blanchet@43444
  1281
                  if metis_s <> unmangled_s orelse
blanchet@43765
  1282
                     (needs_fairly_sound andalso not fairly_sound) then
blanchet@43444
  1283
                    []
blanchet@43444
  1284
                  else
blanchet@43444
  1285
                    ths ~~ (1 upto length ths)
blanchet@43764
  1286
                    |> map (dub_and_inst mangled_s needs_fairly_sound)
blanchet@43764
  1287
                    |> make_facts (not needs_fairly_sound))
blanchet@43444
  1288
    end
blanchet@43444
  1289
  | NONE => []
blanchet@43803
  1290
fun helper_facts_for_sym_table ctxt format type_sys sym_tab =
blanchet@43803
  1291
  Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_sys) sym_tab
blanchet@43803
  1292
                  []
blanchet@43444
  1293
blanchet@43835
  1294
fun translate_atp_fact ctxt format type_sys keep_trivial =
blanchet@43835
  1295
  `(make_fact ctxt format type_sys keep_trivial true true o apsnd prop_of)
blanchet@43444
  1296
blanchet@43926
  1297
(***************************************************************)
blanchet@43926
  1298
(* Type Classes Present in the Axiom or Conjecture Clauses     *)
blanchet@43926
  1299
(***************************************************************)
blanchet@43926
  1300
blanchet@43926
  1301
fun set_insert (x, s) = Symtab.update (x, ()) s
blanchet@43926
  1302
blanchet@43926
  1303
fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
blanchet@43926
  1304
blanchet@43926
  1305
(* Remove this trivial type class (FIXME: similar code elsewhere) *)
blanchet@43926
  1306
fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
blanchet@43926
  1307
blanchet@43934
  1308
fun classes_of_terms get_Ts =
blanchet@43934
  1309
  map (map #2 o get_Ts)
blanchet@43934
  1310
  #> List.foldl add_classes Symtab.empty
blanchet@43934
  1311
  #> delete_type #> Symtab.keys
blanchet@43926
  1312
blanchet@43934
  1313
val tfree_classes_of_terms = classes_of_terms OldTerm.term_tfrees
blanchet@43934
  1314
val tvar_classes_of_terms = classes_of_terms OldTerm.term_tvars
blanchet@43926
  1315
blanchet@43926
  1316
(*fold type constructors*)
blanchet@43926
  1317
fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
blanchet@43934
  1318
  | fold_type_consts _ _ x = x
blanchet@43926
  1319
blanchet@43926
  1320
(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
blanchet@43926
  1321
fun add_type_consts_in_term thy =
blanchet@43926
  1322
  let
blanchet@43926
  1323
    fun aux (Const (@{const_name Meson.skolem}, _) $ _) = I
blanchet@43926
  1324
      | aux (t $ u) = aux t #> aux u
blanchet@43926
  1325
      | aux (Const x) =
blanchet@43926
  1326
        fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
blanchet@43926
  1327
      | aux (Abs (_, _, u)) = aux u
blanchet@43926
  1328
      | aux _ = I
blanchet@43926
  1329
  in aux end
blanchet@43926
  1330
blanchet@43926
  1331
fun type_consts_of_terms thy ts =
blanchet@43934
  1332
  Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty)
blanchet@43926
  1333
blanchet@43926
  1334
blanchet@43803
  1335
fun translate_formulas ctxt format prem_kind type_sys hyp_ts concl_t
blanchet@43803
  1336
                       rich_facts =
blanchet@43444
  1337
  let
blanchet@43444
  1338
    val thy = Proof_Context.theory_of ctxt
blanchet@43444
  1339
    val fact_ts = map (prop_of o snd o snd) rich_facts
blanchet@43444
  1340
    val (facts, fact_names) =
blanchet@43444
  1341
      rich_facts
blanchet@43444
  1342
      |> map_filter (fn (NONE, _) => NONE
blanchet@43444
  1343
                      | (SOME fact, (name, _)) => SOME (fact, name))
blanchet@43444
  1344
      |> ListPair.unzip
blanchet@43444
  1345
    (* Remove existing facts from the conjecture, as this can dramatically
blanchet@43444
  1346
       boost an ATP's performance (for some reason). *)
blanchet@43444
  1347
    val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts)
blanchet@43444
  1348
    val goal_t = Logic.list_implies (hyp_ts, concl_t)
blanchet@43444
  1349
    val all_ts = goal_t :: fact_ts
blanchet@43444
  1350
    val subs = tfree_classes_of_terms all_ts
blanchet@43444
  1351
    val supers = tvar_classes_of_terms all_ts
blanchet@43444
  1352
    val tycons = type_consts_of_terms thy all_ts
blanchet@43835
  1353
    val conjs =
blanchet@43835
  1354
      hyp_ts @ [concl_t] |> make_conjecture ctxt format prem_kind type_sys
blanchet@43444
  1355
    val (supers', arity_clauses) =
blanchet@43460
  1356
      if level_of_type_sys type_sys = No_Types then ([], [])
blanchet@43444
  1357
      else make_arity_clauses thy tycons supers
blanchet@43444
  1358
    val class_rel_clauses = make_class_rel_clauses thy subs supers'
blanchet@43444
  1359
  in
blanchet@43444
  1360
    (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses))
blanchet@43444
  1361
  end
blanchet@43444
  1362
blanchet@43444
  1363
fun fo_literal_from_type_literal (TyLitVar (class, name)) =
blanchet@43444
  1364
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@43444
  1365
  | fo_literal_from_type_literal (TyLitFree (class, name)) =
blanchet@43444
  1366
    (true, ATerm (class, [ATerm (name, [])]))
blanchet@43444
  1367
blanchet@43444
  1368
fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
blanchet@43444
  1369
blanchet@43835
  1370
fun type_pred_combterm ctxt nonmono_Ts type_sys T tm =
blanchet@43807
  1371
  CombApp (CombConst (`make_fixed_const type_pred_name, T --> @{typ bool}, [T])
blanchet@43835
  1372
           |> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys,
blanchet@43444
  1373
           tm)
blanchet@43444
  1374
blanchet@43747
  1375
fun var_occurs_positively_naked_in_term _ (SOME false) _ accum = accum
blanchet@43747
  1376
  | var_occurs_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
blanchet@43841
  1377
    accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
blanchet@43747
  1378
fun is_var_nonmonotonic_in_formula _ _ (SOME false) _ = false
blanchet@43747
  1379
  | is_var_nonmonotonic_in_formula pos phi _ name =
blanchet@43747
  1380
    formula_fold pos (var_occurs_positively_naked_in_term name) phi false
blanchet@43705
  1381
blanchet@43835
  1382
fun mk_const_aterm x T_args args =
blanchet@43835
  1383
  ATerm (x, map (fo_term_from_typ false) T_args @ args)
blanchet@43835
  1384
blanchet@43803
  1385
fun tag_with_type ctxt format nonmono_Ts type_sys T tm =
blanchet@43700
  1386
  CombConst (`make_fixed_const type_tag_name, T --> T, [T])
blanchet@43835
  1387
  |> enforce_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys
blanchet@43803
  1388
  |> term_from_combterm ctxt format nonmono_Ts type_sys Top_Level
blanchet@43700
  1389
  |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm]))
blanchet@43803
  1390
and term_from_combterm ctxt format nonmono_Ts type_sys =
blanchet@43444
  1391
  let
blanchet@43803
  1392
    fun aux site u =
blanchet@43803
  1393
      let
blanchet@43803
  1394
        val (head, args) = strip_combterm_comb u
blanchet@43803
  1395
        val (x as (s, _), T_args) =
blanchet@43803
  1396
          case head of
blanchet@43803
  1397
            CombConst (name, _, T_args) => (name, T_args)
blanchet@43803
  1398
          | CombVar (name, _) => (name, [])
blanchet@43803
  1399
          | CombApp _ => raise Fail "impossible \"CombApp\""
blanchet@43841
  1400
        val arg_site = if site = Top_Level andalso is_tptp_equal s then Eq_Arg
blanchet@43803
  1401
                       else Elsewhere
blanchet@43835
  1402
        val t = mk_const_aterm x T_args (map (aux arg_site) args)
blanchet@43803
  1403
        val T = combtyp_of u
blanchet@43803
  1404
      in
blanchet@43803
  1405
        t |> (if should_tag_with_type ctxt nonmono_Ts type_sys site u T then
blanchet@43803
  1406
                tag_with_type ctxt format nonmono_Ts type_sys T
blanchet@43803
  1407
              else
blanchet@43803
  1408
                I)
blanchet@43803
  1409
      end
blanchet@43803
  1410
  in aux end
blanchet@43803
  1411
and formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43803
  1412
                             should_predicate_on_var =
blanchet@43700
  1413
  let
blanchet@43835
  1414
    val higher_order = is_setting_higher_order format type_sys
blanchet@43803
  1415
    val do_term = term_from_combterm ctxt format nonmono_Ts type_sys Top_Level
blanchet@43444
  1416
    val do_bound_type =
blanchet@43552
  1417
      case type_sys of
blanchet@43587
  1418
        Simple_Types level =>
blanchet@43835
  1419
        homogenized_type ctxt nonmono_Ts level 0
blanchet@43835
  1420
        #> mangled_type higher_order false 0 #> SOME
blanchet@43552
  1421
      | _ => K NONE
blanchet@43747
  1422
    fun do_out_of_bound_type pos phi universal (name, T) =
blanchet@43705
  1423
      if should_predicate_on_type ctxt nonmono_Ts type_sys
blanchet@43747
  1424
             (fn () => should_predicate_on_var pos phi universal name) T then
blanchet@43705
  1425
        CombVar (name, T)
blanchet@43835
  1426
        |> type_pred_combterm ctxt nonmono_Ts type_sys T
blanchet@43747
  1427
        |> do_term |> AAtom |> SOME
blanchet@43444
  1428
      else
blanchet@43444
  1429
        NONE
blanchet@43747
  1430
    fun do_formula pos (AQuant (q, xs, phi)) =
blanchet@43747
  1431
        let
blanchet@43747
  1432
          val phi = phi |> do_formula pos
blanchet@43747
  1433
          val universal = Option.map (q = AExists ? not) pos
blanchet@43747
  1434
        in
blanchet@43705
  1435
          AQuant (q, xs |> map (apsnd (fn NONE => NONE
blanchet@43705
  1436
                                        | SOME T => do_bound_type T)),
blanchet@43705
  1437
                  (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
blanchet@43705
  1438
                      (map_filter
blanchet@43705
  1439
                           (fn (_, NONE) => NONE
blanchet@43705
  1440
                             | (s, SOME T) =>
blanchet@43747
  1441
                               do_out_of_bound_type pos phi universal (s, T))
blanchet@43747
  1442
                           xs)
blanchet@43705
  1443
                      phi)
blanchet@43705
  1444
        end
blanchet@43747
  1445
      | do_formula pos (AConn conn) = aconn_map pos do_formula conn
blanchet@43747
  1446
      | do_formula _ (AAtom tm) = AAtom (do_term tm)
blanchet@43747
  1447
  in do_formula o SOME end
blanchet@43444
  1448
blanchet@43797
  1449
fun bound_atomic_types format type_sys Ts =
blanchet@43592
  1450
  mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal)
blanchet@43926
  1451
                (type_literals_for_types format type_sys Axiom Ts))
blanchet@43592
  1452
blanchet@43797
  1453
fun formula_for_fact ctxt format nonmono_Ts type_sys
blanchet@43444
  1454
                     ({combformula, atomic_types, ...} : translated_formula) =
blanchet@43592
  1455
  combformula
blanchet@43592
  1456
  |> close_combformula_universally
blanchet@43803
  1457
  |> formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43747
  1458
                              is_var_nonmonotonic_in_formula true
blanchet@43797
  1459
  |> bound_atomic_types format type_sys atomic_types
blanchet@43444
  1460
  |> close_formula_universally
blanchet@43444
  1461
blanchet@43444
  1462
(* Each fact is given a unique fact number to avoid name clashes (e.g., because
blanchet@43444
  1463
   of monomorphization). The TPTP explicitly forbids name clashes, and some of
blanchet@43444
  1464
   the remote provers might care. *)
blanchet@43797
  1465
fun formula_line_for_fact ctxt format prefix nonmono_Ts type_sys
blanchet@43511
  1466
                          (j, formula as {name, locality, kind, ...}) =
blanchet@43550
  1467
  Formula (prefix ^ (if polymorphism_of_type_sys type_sys = Polymorphic then ""
blanchet@43550
  1468
                     else string_of_int j ^ "_") ^
blanchet@43518
  1469
           ascii_of name,
blanchet@43797
  1470
           kind, formula_for_fact ctxt format nonmono_Ts type_sys formula, NONE,
blanchet@43748
  1471
           case locality of
blanchet@43748
  1472
             Intro => intro_info
blanchet@43748
  1473
           | Elim => elim_info
blanchet@43748
  1474
           | Simp => simp_info
blanchet@43748
  1475
           | _ => NONE)
blanchet@43444
  1476
blanchet@43927
  1477
fun formula_line_for_class_rel_clause ({name, subclass, superclass, ...}
blanchet@43927
  1478
                                       : class_rel_clause) =
blanchet@43444
  1479
  let val ty_arg = ATerm (`I "T", []) in
blanchet@43448
  1480
    Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
blanchet@43444
  1481
             AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
blanchet@43444
  1482
                               AAtom (ATerm (superclass, [ty_arg]))])
blanchet@43748
  1483
             |> close_formula_universally, intro_info, NONE)
blanchet@43444
  1484
  end
blanchet@43444
  1485
blanchet@43444
  1486
fun fo_literal_from_arity_literal (TConsLit (c, t, args)) =
blanchet@43444
  1487
    (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
blanchet@43444
  1488
  | fo_literal_from_arity_literal (TVarLit (c, sort)) =
blanchet@43444
  1489
    (false, ATerm (c, [ATerm (sort, [])]))
blanchet@43444
  1490
blanchet@43927
  1491
fun formula_line_for_arity_clause ({name, prem_lits, concl_lits, ...}
blanchet@43927
  1492
                                   : arity_clause) =
blanchet@43448
  1493
  Formula (arity_clause_prefix ^ ascii_of name, Axiom,
blanchet@43444
  1494
           mk_ahorn (map (formula_from_fo_literal o apfst not
blanchet@43766
  1495
                          o fo_literal_from_arity_literal) prem_lits)
blanchet@43444
  1496
                    (formula_from_fo_literal
blanchet@43766
  1497
                         (fo_literal_from_arity_literal concl_lits))
blanchet@43748
  1498
           |> close_formula_universally, intro_info, NONE)
blanchet@43444
  1499
blanchet@43803
  1500
fun formula_line_for_conjecture ctxt format nonmono_Ts type_sys
blanchet@43444
  1501
        ({name, kind, combformula, ...} : translated_formula) =
blanchet@43448
  1502
  Formula (conjecture_prefix ^ name, kind,
blanchet@43803
  1503
           formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43780
  1504
               is_var_nonmonotonic_in_formula false
blanchet@43780
  1505
               (close_combformula_universally combformula)
blanchet@43444
  1506
           |> close_formula_universally, NONE, NONE)
blanchet@43444
  1507
blanchet@43797
  1508
fun free_type_literals format type_sys
blanchet@43797
  1509
                       ({atomic_types, ...} : translated_formula) =
blanchet@43926
  1510
  atomic_types |> type_literals_for_types format type_sys Conjecture
blanchet@43444
  1511
               |> map fo_literal_from_type_literal
blanchet@43444
  1512
blanchet@43444
  1513
fun formula_line_for_free_type j lit =
blanchet@43926
  1514
  Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis,
blanchet@43444
  1515
           formula_from_fo_literal lit, NONE, NONE)
blanchet@43797
  1516
fun formula_lines_for_free_types format type_sys facts =
blanchet@43444
  1517
  let
blanchet@43797
  1518
    val litss = map (free_type_literals format type_sys) facts
blanchet@43444
  1519
    val lits = fold (union (op =)) litss []
blanchet@43444
  1520
  in map2 formula_line_for_free_type (0 upto length lits - 1) lits end
blanchet@43444
  1521
blanchet@43444
  1522
(** Symbol declarations **)
blanchet@43415
  1523
blanchet@43445
  1524
fun should_declare_sym type_sys pred_sym s =
blanchet@43839
  1525
  is_tptp_user_symbol s andalso not (String.isPrefix bound_var_prefix s) andalso
blanchet@43765
  1526
  (case type_sys of
blanchet@43765
  1527
     Simple_Types _ => true
blanchet@43765
  1528
   | Tags (_, _, Light) => true
blanchet@43765
  1529
   | _ => not pred_sym)
blanchet@43413
  1530
blanchet@43755
  1531
fun sym_decl_table_for_facts ctxt type_sys repaired_sym_tab (conjs, facts) =
blanchet@43445
  1532
  let
blanchet@43568
  1533
    fun add_combterm in_conj tm =
blanchet@43445
  1534
      let val (head, args) = strip_combterm_comb tm in
blanchet@43445
  1535
        (case head of
blanchet@43445
  1536
           CombConst ((s, s'), T, T_args) =>
blanchet@43445
  1537
           let val pred_sym = is_pred_sym repaired_sym_tab s in
blanchet@43445
  1538
             if should_declare_sym type_sys pred_sym s then
blanchet@43447
  1539
               Symtab.map_default (s, [])
blanchet@43755
  1540
                   (insert_type ctxt #3 (s', T_args, T, pred_sym, length args,
blanchet@43755
  1541
                                         in_conj))
blanchet@43445
  1542
             else
blanchet@43445
  1543
               I
blanchet@43445
  1544
           end
blanchet@43445
  1545
         | _ => I)
blanchet@43568
  1546
        #> fold (add_combterm in_conj) args
blanchet@43445
  1547
      end
blanchet@43568
  1548
    fun add_fact in_conj =
blanchet@43705
  1549
      fact_lift (formula_fold NONE (K (add_combterm in_conj)))
blanchet@43568
  1550
  in
blanchet@43568
  1551
    Symtab.empty
blanchet@43568
  1552
    |> is_type_sys_fairly_sound type_sys
blanchet@43568
  1553
       ? (fold (add_fact true) conjs #> fold (add_fact false) facts)
blanchet@43568
  1554
  end
blanchet@43445
  1555
blanchet@43755
  1556
(* These types witness that the type classes they belong to allow infinite
blanchet@43755
  1557
   models and hence that any types with these type classes is monotonic. *)
blanchet@43926
  1558
val known_infinite_types =
blanchet@43926
  1559
  [@{typ nat}, Type ("Int.int", []), @{typ "nat => bool"}]
blanchet@43755
  1560
blanchet@43555
  1561
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
blanchet@43555
  1562
   out with monotonicity" paper presented at CADE 2011. *)
blanchet@43755
  1563
fun add_combterm_nonmonotonic_types _ _ (SOME false) _ = I
blanchet@43700
  1564
  | add_combterm_nonmonotonic_types ctxt level _
blanchet@43841
  1565
        (CombApp (CombApp (CombConst ((s, _), Type (_, [T, _]), _), tm1), tm2)) =
blanchet@43841
  1566
    (is_tptp_equal s andalso exists is_var_or_bound_var [tm1, tm2] andalso
blanchet@43700
  1567
     (case level of
blanchet@43755
  1568
        Nonmonotonic_Types =>
blanchet@43755
  1569
        not (is_type_surely_infinite ctxt known_infinite_types T)
blanchet@43700
  1570
      | Finite_Types => is_type_surely_finite ctxt T
blanchet@43755
  1571
      | _ => true)) ? insert_type ctxt I (deep_freeze_type T)
blanchet@43700
  1572
  | add_combterm_nonmonotonic_types _ _ _ _ = I
blanchet@43700
  1573
fun add_fact_nonmonotonic_types ctxt level ({kind, combformula, ...}
blanchet@43700
  1574
                                            : translated_formula) =
blanchet@43705
  1575
  formula_fold (SOME (kind <> Conjecture))
blanchet@43700
  1576
               (add_combterm_nonmonotonic_types ctxt level) combformula
blanchet@43755
  1577
fun nonmonotonic_types_for_facts ctxt type_sys facts =
blanchet@43700
  1578
  let val level = level_of_type_sys type_sys in
blanchet@43755
  1579
    if level = Nonmonotonic_Types orelse level = Finite_Types then
blanchet@43755
  1580
      [] |> fold (add_fact_nonmonotonic_types ctxt level) facts
blanchet@43755
  1581
         (* We must add "bool" in case the helper "True_or_False" is added
blanchet@43755
  1582
            later. In addition, several places in the code rely on the list of
blanchet@43755
  1583
            nonmonotonic types not being empty. *)
blanchet@43755
  1584
         |> insert_type ctxt I @{typ bool}
blanchet@43755
  1585
    else
blanchet@43755
  1586
      []
blanchet@43700
  1587
  end
blanchet@43547
  1588
blanchet@43835
  1589
fun decl_line_for_sym ctxt format nonmono_Ts type_sys s
blanchet@43835
  1590
                      (s', T_args, T, pred_sym, ary, _) =
blanchet@43835
  1591
  let
blanchet@43835
  1592
    val (higher_order, T_arg_Ts, level) =
blanchet@43835
  1593
      case type_sys of
blanchet@43835
  1594
        Simple_Types level => (format = THF, [], level)
blanchet@43835
  1595
      | _ => (false, replicate (length T_args) homo_infinite_type, No_Types)
blanchet@43835
  1596
  in
blanchet@43839
  1597
    Decl (sym_decl_prefix ^ s, (s, s'),
blanchet@43835
  1598
          (T_arg_Ts ---> (T |> homogenized_type ctxt nonmono_Ts level ary))
blanchet@43835
  1599
          |> mangled_type higher_order pred_sym (length T_arg_Ts + ary))
blanchet@43835
  1600
  end
blanchet@43450
  1601
blanchet@43463
  1602
fun is_polymorphic_type T = fold_atyps (fn TVar _ => K true | _ => I) T false
blanchet@43463
  1603
blanchet@43797
  1604
fun formula_line_for_pred_sym_decl ctxt format conj_sym_kind nonmono_Ts type_sys
blanchet@43797
  1605
                                   n s j (s', T_args, T, _, ary, in_conj) =
blanchet@43450
  1606
  let
blanchet@43580
  1607
    val (kind, maybe_negate) =
blanchet@43580
  1608
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43580
  1609
      else (Axiom, I)
blanchet@43618
  1610
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@43450
  1611
    val bound_names =
blanchet@43450
  1612
      1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  1613
    val bounds =
blanchet@43450
  1614
      bound_names ~~ arg_Ts |> map (fn (name, T) => CombConst (name, T, []))
blanchet@43450
  1615
    val bound_Ts =
blanchet@43463
  1616
      arg_Ts |> map (fn T => if n > 1 orelse is_polymorphic_type T then SOME T
blanchet@43463
  1617
                             else NONE)
blanchet@43450
  1618
  in
blanchet@43839
  1619
    Formula (sym_formula_prefix ^ s ^
blanchet@43580
  1620
             (if n > 1 then "_" ^ string_of_int j else ""), kind,
blanchet@43450
  1621
             CombConst ((s, s'), T, T_args)
blanchet@43700
  1622
             |> fold (curry (CombApp o swap)) bounds
blanchet@43835
  1623
             |> type_pred_combterm ctxt nonmono_Ts type_sys res_T
blanchet@43804
  1624
             |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
blanchet@43803
  1625
             |> formula_from_combformula ctxt format nonmono_Ts type_sys
blanchet@43747
  1626
                                         (K (K (K (K true)))) true
blanchet@43797
  1627
             |> n > 1 ? bound_atomic_types format type_sys (atyps_of T)
blanchet@43580
  1628
             |> close_formula_universally
blanchet@43580
  1629
             |> maybe_negate,
blanchet@43748
  1630
             intro_info, NONE)
blanchet@43450
  1631
  end
blanchet@43450
  1632
blanchet@43797
  1633
fun formula_lines_for_tag_sym_decl ctxt format conj_sym_kind nonmono_Ts type_sys
blanchet@43797
  1634
        n s (j, (s', T_args, T, pred_sym, ary, in_conj)) =
blanchet@43700
  1635
  let
blanchet@43700
  1636
    val ident_base =
blanchet@43839
  1637
      sym_formula_prefix ^ s ^ (if n > 1 then "_" ^ string_of_int j else "")
blanchet@43721
  1638
    val (kind, maybe_negate) =
blanchet@43721
  1639
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43721
  1640
      else (Axiom, I)
blanchet@43700
  1641
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@43700
  1642
    val bound_names =
blanchet@43700
  1643
      1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  1644
    val bounds = bound_names |> map (fn name => ATerm (name, []))
blanchet@43835
  1645
    val cst = mk_const_aterm (s, s') T_args
blanchet@43701
  1646
    val atomic_Ts = atyps_of T
blanchet@43705
  1647
    fun eq tms =
blanchet@43705
  1648
      (if pred_sym then AConn (AIff, map AAtom tms)
blanchet@43841
  1649
       else AAtom (ATerm (`I tptp_equal, tms)))
blanchet@43797
  1650
      |> bound_atomic_types format type_sys atomic_Ts
blanchet@43701
  1651
      |> close_formula_universally
blanchet@43721
  1652
      |> maybe_negate
blanchet@43707
  1653
    val should_encode = should_encode_type ctxt nonmono_Ts All_Types
blanchet@43803
  1654
    val tag_with = tag_with_type ctxt format nonmono_Ts type_sys
blanchet@43700
  1655
    val add_formula_for_res =
blanchet@43700
  1656
      if should_encode res_T then
blanchet@43721
  1657
        cons (Formula (ident_base ^ "_res", kind,
blanchet@43835
  1658
                       eq [tag_with res_T (cst bounds), cst bounds],
blanchet@43748
  1659
                       simp_info, NONE))
blanchet@43700
  1660
      else
blanchet@43700
  1661
        I
blanchet@43700
  1662
    fun add_formula_for_arg k =
blanchet@43700
  1663
      let val arg_T = nth arg_Ts k in
blanchet@43700
  1664
        if should_encode arg_T then
blanchet@43700
  1665
          case chop k bounds of
blanchet@43700
  1666
            (bounds1, bound :: bounds2) =>
blanchet@43721
  1667
            cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,
blanchet@43835
  1668
                           eq [cst (bounds1 @ tag_with arg_T bound :: bounds2),
blanchet@43835
  1669
                               cst bounds],
blanchet@43748
  1670
                           simp_info, NONE))
blanchet@43700
  1671
          | _ => raise Fail "expected nonempty tail"
blanchet@43700
  1672
        else
blanchet@43700
  1673
          I
blanchet@43700
  1674
      end
blanchet@43700
  1675
  in
blanchet@43705
  1676
    [] |> not pred_sym ? add_formula_for_res
blanchet@43700
  1677
       |> fold add_formula_for_arg (ary - 1 downto 0)
blanchet@43700
  1678
  end
blanchet@43700
  1679
blanchet@43707
  1680
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
blanchet@43707
  1681
blanchet@43797
  1682
fun problem_lines_for_sym_decls ctxt format conj_sym_kind nonmono_Ts type_sys
blanchet@43580
  1683
                                (s, decls) =
blanchet@43839
  1684
  case type_sys of
blanchet@43839
  1685
    Simple_Types _ =>
blanchet@43839
  1686
    decls |> map (decl_line_for_sym ctxt format nonmono_Ts type_sys s)
blanchet@43839
  1687
  | Preds _ =>
blanchet@43839
  1688
    let
blanchet@43839
  1689
      val decls =
blanchet@43839
  1690
        case decls of
blanchet@43839
  1691
          decl :: (decls' as _ :: _) =>
blanchet@43839
  1692
          let val T = result_type_of_decl decl in
blanchet@43839
  1693
            if forall (curry (type_instance ctxt o swap) T
blanchet@43839
  1694
                       o result_type_of_decl) decls' then
blanchet@43839
  1695
              [decl]
blanchet@43839
  1696
            else
blanchet@43839
  1697
              decls
blanchet@43839
  1698
          end
blanchet@43839
  1699
        | _ => decls
blanchet@43839
  1700
      val n = length decls
blanchet@43839
  1701
      val decls =
blanchet@43839
  1702
        decls
blanchet@43839
  1703
        |> filter (should_predicate_on_type ctxt nonmono_Ts type_sys (K true)
blanchet@43839
  1704
                   o result_type_of_decl)
blanchet@43839
  1705
    in
blanchet@43839
  1706
      (0 upto length decls - 1, decls)
blanchet@43839
  1707
      |-> map2 (formula_line_for_pred_sym_decl ctxt format conj_sym_kind
blanchet@43839
  1708
                                               nonmono_Ts type_sys n s)
blanchet@43839
  1709
    end
blanchet@43839
  1710
  | Tags (_, _, heaviness) =>
blanchet@43839
  1711
    (case heaviness of
blanchet@43839
  1712
       Heavy => []
blanchet@43839
  1713
     | Light =>
blanchet@43839
  1714
       let val n = length decls in
blanchet@43839
  1715
         (0 upto n - 1 ~~ decls)
blanchet@43839
  1716
         |> maps (formula_lines_for_tag_sym_decl ctxt format conj_sym_kind
blanchet@43839
  1717
                                                 nonmono_Ts type_sys n s)
blanchet@43839
  1718
       end)
blanchet@43450
  1719
blanchet@43797
  1720
fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind nonmono_Ts
blanchet@43797
  1721
                                     type_sys sym_decl_tab =
blanchet@43839
  1722
  sym_decl_tab
blanchet@43839
  1723
  |> Symtab.dest
blanchet@43839
  1724
  |> sort_wrt fst
blanchet@43839
  1725
  |> rpair []
blanchet@43839
  1726
  |-> fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
blanchet@43839
  1727
                                                     nonmono_Ts type_sys)
blanchet@43414
  1728
blanchet@43708
  1729
fun should_add_ti_ti_helper (Tags (Polymorphic, level, Heavy)) =
blanchet@43702
  1730
    level = Nonmonotonic_Types orelse level = Finite_Types
blanchet@43702
  1731
  | should_add_ti_ti_helper _ = false
blanchet@43702
  1732
blanchet@43780
  1733
fun offset_of_heading_in_problem _ [] j = j
blanchet@43780
  1734
  | offset_of_heading_in_problem needle ((heading, lines) :: problem) j =
blanchet@43780
  1735
    if heading = needle then j
blanchet@43780
  1736
    else offset_of_heading_in_problem needle problem (j + length lines)
blanchet@43780
  1737
blanchet@43839
  1738
val implicit_declsN = "Should-be-implicit typings"
blanchet@43839
  1739
val explicit_declsN = "Explicit typings"
blanchet@41405
  1740
val factsN = "Relevant facts"
blanchet@41405
  1741
val class_relsN = "Class relationships"
blanchet@43414
  1742
val aritiesN = "Arities"
blanchet@41405
  1743
val helpersN = "Helper facts"
blanchet@41405
  1744
val conjsN = "Conjectures"
blanchet@41561
  1745
val free_typesN = "Type variables"
blanchet@41405
  1746
blanchet@43780
  1747
fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_sys
blanchet@43933
  1748
                        explicit_apply readable_names hyp_ts concl_t facts =
blanchet@38506
  1749
  let
blanchet@41561
  1750
    val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) =
blanchet@43803
  1751
      translate_formulas ctxt format prem_kind type_sys hyp_ts concl_t facts
blanchet@43905
  1752
    val sym_tab = conjs @ facts |> sym_table_for_facts ctxt explicit_apply
blanchet@43755
  1753
    val nonmono_Ts = conjs @ facts |> nonmonotonic_types_for_facts ctxt type_sys
blanchet@43803
  1754
    val repair = repair_fact ctxt format nonmono_Ts type_sys sym_tab
blanchet@43552
  1755
    val (conjs, facts) = (conjs, facts) |> pairself (map repair)
blanchet@43905
  1756
    val repaired_sym_tab =
blanchet@43905
  1757
      conjs @ facts |> sym_table_for_facts ctxt (SOME false)
blanchet@43444
  1758
    val helpers =
blanchet@43803
  1759
      repaired_sym_tab |> helper_facts_for_sym_table ctxt format type_sys
blanchet@43803
  1760
                       |> map repair
blanchet@43765
  1761
    val lavish_nonmono_Ts =
blanchet@43765
  1762
      if null nonmono_Ts orelse
blanchet@43765
  1763
         polymorphism_of_type_sys type_sys <> Polymorphic then
blanchet@43765
  1764
        nonmono_Ts
blanchet@43765
  1765
      else
blanchet@43765
  1766
        [TVar (("'a", 0), HOLogic.typeS)]
blanchet@43550
  1767
    val sym_decl_lines =
blanchet@43596
  1768
      (conjs, helpers @ facts)
blanchet@43755
  1769
      |> sym_decl_table_for_facts ctxt type_sys repaired_sym_tab
blanchet@43797
  1770
      |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind
blanchet@43797
  1771
                                          lavish_nonmono_Ts type_sys
blanchet@43750
  1772
    val helper_lines =
blanchet@43797
  1773
      0 upto length helpers - 1 ~~ helpers
blanchet@43797
  1774
      |> map (formula_line_for_fact ctxt format helper_prefix lavish_nonmono_Ts
blanchet@43797
  1775
                                    type_sys)
blanchet@43797
  1776
      |> (if should_add_ti_ti_helper type_sys then cons (ti_ti_helper_fact ())
blanchet@43797
  1777
          else I)
blanchet@43393
  1778
    (* Reordering these might confuse the proof reconstruction code or the SPASS
blanchet@43880
  1779
       FLOTTER hack. *)
blanchet@38506
  1780
    val problem =
blanchet@43839
  1781
      [(explicit_declsN, sym_decl_lines),
blanchet@43797
  1782
       (factsN,
blanchet@43797
  1783
        map (formula_line_for_fact ctxt format fact_prefix nonmono_Ts type_sys)
blanchet@43797
  1784
            (0 upto length facts - 1 ~~ facts)),
blanchet@43416
  1785
       (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses),
blanchet@43416
  1786
       (aritiesN, map formula_line_for_arity_clause arity_clauses),
blanchet@43750
  1787
       (helpersN, helper_lines),
blanchet@43803
  1788
       (conjsN,
blanchet@43803
  1789
        map (formula_line_for_conjecture ctxt format nonmono_Ts type_sys)
blanchet@43803
  1790
            conjs),
blanchet@43797
  1791
       (free_typesN,
blanchet@43797
  1792
        formula_lines_for_free_types format type_sys (facts @ conjs))]
blanchet@43414
  1793
    val problem =
blanchet@43432
  1794
      problem
blanchet@43933
  1795
      |> (case format of
blanchet@43933
  1796
            CNF => ensure_cnf_problem
blanchet@43933
  1797
          | CNF_UEQ => filter_cnf_ueq_problem
blanchet@43933
  1798
          | _ => I)
blanchet@43839
  1799
      |> (if is_format_typed format then
blanchet@43839
  1800
            declare_undeclared_syms_in_atp_problem type_decl_prefix
blanchet@43839
  1801
                                                   implicit_declsN
blanchet@43839
  1802
          else
blanchet@43839
  1803
            I)
blanchet@43933
  1804
    val (problem, pool) = problem |> nice_atp_problem readable_names
blanchet@43750
  1805
    val helpers_offset = offset_of_heading_in_problem helpersN problem 0
blanchet@43750
  1806
    val typed_helpers =
blanchet@43750
  1807
      map_filter (fn (j, {name, ...}) =>
blanchet@43750
  1808
                     if String.isSuffix typed_helper_suffix name then SOME j
blanchet@43750
  1809
                     else NONE)
blanchet@43750
  1810
                 ((helpers_offset + 1 upto helpers_offset + length helpers)
blanchet@43750
  1811
                  ~~ helpers)
blanchet@43649
  1812
    fun add_sym_arity (s, {min_ary, ...} : sym_info) =
blanchet@43620
  1813
      if min_ary > 0 then
blanchet@43620
  1814
        case strip_prefix_and_unascii const_prefix s of
blanchet@43620
  1815
          SOME s => Symtab.insert (op =) (s, min_ary)
blanchet@43620
  1816
        | NONE => I
blanchet@43620
  1817
      else
blanchet@43620
  1818
        I
blanchet@38506
  1819
  in
blanchet@38506
  1820
    (problem,
blanchet@38506
  1821
     case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
blanchet@43456
  1822
     offset_of_heading_in_problem conjsN problem 0,
blanchet@43412
  1823
     offset_of_heading_in_problem factsN problem 0,
blanchet@43620
  1824
     fact_names |> Vector.fromList,
blanchet@43750
  1825
     typed_helpers,
blanchet@43620
  1826
     Symtab.empty |> Symtab.fold add_sym_arity sym_tab)
blanchet@38506
  1827
  end
blanchet@38506
  1828
blanchet@41561
  1829
(* FUDGE *)
blanchet@41561
  1830
val conj_weight = 0.0
blanchet@42641
  1831
val hyp_weight = 0.1
blanchet@42641
  1832
val fact_min_weight = 0.2
blanchet@41561
  1833
val fact_max_weight = 1.0
blanchet@43479
  1834
val type_info_default_weight = 0.8
blanchet@41561
  1835
blanchet@41561
  1836
fun add_term_weights weight (ATerm (s, tms)) =
blanchet@43839
  1837
  is_tptp_user_symbol s ? Symtab.default (s, weight)
blanchet@41561
  1838
  #> fold (add_term_weights weight) tms
blanchet@43448
  1839
fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
blanchet@43705
  1840
    formula_fold NONE (K (add_term_weights weight)) phi
blanchet@43399
  1841
  | add_problem_line_weights _ _ = I
blanchet@41561
  1842
blanchet@41561
  1843
fun add_conjectures_weights [] = I
blanchet@41561
  1844
  | add_conjectures_weights conjs =
blanchet@41561
  1845
    let val (hyps, conj) = split_last conjs in
blanchet@41561
  1846
      add_problem_line_weights conj_weight conj
blanchet@41561
  1847
      #> fold (add_problem_line_weights hyp_weight) hyps
blanchet@41561
  1848
    end
blanchet@41561
  1849
blanchet@41561
  1850
fun add_facts_weights facts =
blanchet@41561
  1851
  let
blanchet@41561
  1852
    val num_facts = length facts
blanchet@41561
  1853
    fun weight_of j =
blanchet@41561
  1854
      fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
blanchet@41561
  1855
                        / Real.fromInt num_facts
blanchet@41561
  1856
  in
blanchet@41561
  1857
    map weight_of (0 upto num_facts - 1) ~~ facts
blanchet@41561
  1858
    |> fold (uncurry add_problem_line_weights)
blanchet@41561
  1859
  end
blanchet@41561
  1860
blanchet@41561
  1861
(* Weights are from 0.0 (most important) to 1.0 (least important). *)
blanchet@41561
  1862
fun atp_problem_weights problem =
blanchet@43479
  1863
  let val get = these o AList.lookup (op =) problem in
blanchet@43479
  1864
    Symtab.empty
blanchet@43479
  1865
    |> add_conjectures_weights (get free_typesN @ get conjsN)
blanchet@43479
  1866
    |> add_facts_weights (get factsN)
blanchet@43479
  1867
    |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
blanchet@43839
  1868
            [explicit_declsN, class_relsN, aritiesN]
blanchet@43479
  1869
    |> Symtab.dest
blanchet@43479
  1870
    |> sort (prod_ord Real.compare string_ord o pairself swap)
blanchet@43479
  1871
  end
blanchet@41561
  1872
blanchet@38506
  1873
end;