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