src/HOL/Tools/ATP/atp_problem_generate.ML
author blanchet
Wed, 01 Feb 2012 14:53:46 +0100
changeset 47217 a7538ad74997
parent 47213 0ccf458a3633
child 47220 676a4b4b6e73
permissions -rw-r--r--
tuning
blanchet@47148
     1
(*  Title:      HOL/Tools/ATP/atp_problem_generate.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@47148
     9
signature ATP_PROBLEM_GENERATE =
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@46172
    14
  type atp_format = ATP_Problem.atp_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@47168
    18
  datatype scope = Global | Local | Assum | Chained
blanchet@47169
    19
  datatype status = General | Induct | Intro | Elim | Simp | Eq
blanchet@47168
    20
  type stature = scope * status
blanchet@43484
    21
blanchet@45349
    22
  datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic
blanchet@47129
    23
  datatype strictness = Strict | Non_Strict
blanchet@45673
    24
  datatype granularity = All_Vars | Positively_Naked_Vars | Ghost_Type_Arg_Vars
blanchet@43484
    25
  datatype type_level =
blanchet@45256
    26
    All_Types |
blanchet@47129
    27
    Noninf_Nonmono_Types of strictness * granularity |
blanchet@45673
    28
    Fin_Nonmono_Types of granularity |
blanchet@45256
    29
    Const_Arg_Types |
blanchet@44233
    30
    No_Types
blanchet@45650
    31
  type type_enc
blanchet@43484
    32
blanchet@45351
    33
  val type_tag_idempotence : bool Config.T
blanchet@45351
    34
  val type_tag_arguments : bool Config.T
blanchet@46384
    35
  val no_lamsN : string
blanchet@46384
    36
  val hide_lamsN : string
blanchet@47193
    37
  val liftingN : string
blanchet@47193
    38
  val combsN : string
blanchet@47193
    39
  val combs_and_liftingN : string
blanchet@47193
    40
  val combs_or_liftingN : string
blanchet@46384
    41
  val lam_liftingN : string
blanchet@46384
    42
  val keep_lamsN : string
blanchet@44367
    43
  val schematic_var_prefix : string
blanchet@44367
    44
  val fixed_var_prefix : string
blanchet@44367
    45
  val tvar_prefix : string
blanchet@44367
    46
  val tfree_prefix : string
blanchet@44367
    47
  val const_prefix : string
blanchet@44367
    48
  val type_const_prefix : string
blanchet@44367
    49
  val class_prefix : string
blanchet@46425
    50
  val lam_lifted_prefix : string
blanchet@46425
    51
  val lam_lifted_mono_prefix : string
blanchet@46425
    52
  val lam_lifted_poly_prefix : string
blanchet@43926
    53
  val skolem_const_prefix : string
blanchet@43926
    54
  val old_skolem_const_prefix : string
blanchet@43926
    55
  val new_skolem_const_prefix : string
blanchet@46425
    56
  val combinator_prefix : string
blanchet@43966
    57
  val type_decl_prefix : string
blanchet@43966
    58
  val sym_decl_prefix : string
blanchet@44860
    59
  val guards_sym_formula_prefix : string
blanchet@45255
    60
  val tags_sym_formula_prefix : string
blanchet@40445
    61
  val fact_prefix : string
blanchet@38506
    62
  val conjecture_prefix : string
blanchet@43750
    63
  val helper_prefix : string
blanchet@43966
    64
  val class_rel_clause_prefix : string
blanchet@43966
    65
  val arity_clause_prefix : string
blanchet@43966
    66
  val tfree_clause_prefix : string
blanchet@46425
    67
  val lam_fact_prefix : string
blanchet@43750
    68
  val typed_helper_suffix : string
blanchet@43966
    69
  val untyped_helper_suffix : string
blanchet@44000
    70
  val type_tag_idempotence_helper_name : string
blanchet@43807
    71
  val predicator_name : string
blanchet@43807
    72
  val app_op_name : string
blanchet@45255
    73
  val type_guard_name : string
blanchet@43945
    74
  val type_tag_name : string
blanchet@43803
    75
  val simple_type_prefix : string
blanchet@44015
    76
  val prefixed_predicator_name : string
blanchet@43971
    77
  val prefixed_app_op_name : string
blanchet@43971
    78
  val prefixed_type_tag_name : string
blanchet@44367
    79
  val ascii_of : string -> string
blanchet@44367
    80
  val unascii_of : string -> string
blanchet@46382
    81
  val unprefix_and_unascii : string -> string -> string option
blanchet@44000
    82
  val proxy_table : (string * (string * (thm * (string * string)))) list
blanchet@44000
    83
  val proxify_const : string -> (string * string) option
blanchet@44367
    84
  val invert_const : string -> string
blanchet@44367
    85
  val unproxify_const : string -> string
blanchet@43934
    86
  val new_skolem_var_name_from_const : string -> string
blanchet@44089
    87
  val atp_irrelevant_consts : string list
blanchet@44089
    88
  val atp_schematic_consts_of : term -> typ list Symtab.table
blanchet@44692
    89
  val is_type_enc_higher_order : type_enc -> bool
blanchet@44493
    90
  val polymorphism_of_type_enc : type_enc -> polymorphism
blanchet@44493
    91
  val level_of_type_enc : type_enc -> type_level
blanchet@45256
    92
  val is_type_enc_quasi_sound : type_enc -> bool
blanchet@44493
    93
  val is_type_enc_fairly_sound : type_enc -> bool
blanchet@47129
    94
  val type_enc_from_string : strictness -> string -> type_enc
blanchet@46172
    95
  val adjust_type_enc : atp_format -> type_enc -> type_enc
blanchet@43977
    96
  val mk_aconns :
blanchet@43977
    97
    connective -> ('a, 'b, 'c) formula list -> ('a, 'b, 'c) formula
nik@44535
    98
  val unmangled_const : string -> string * (string, 'b) ho_term list
blanchet@43971
    99
  val unmangled_const_name : string -> string
blanchet@44035
   100
  val helper_table : ((string * bool) * thm list) list
blanchet@46385
   101
  val trans_lams_from_string :
blanchet@46385
   102
    Proof.context -> type_enc -> string -> term list -> term list * term list
blanchet@44372
   103
  val factsN : string
blanchet@40240
   104
  val prepare_atp_problem :
blanchet@46172
   105
    Proof.context -> atp_format -> formula_kind -> formula_kind -> type_enc
blanchet@45256
   106
    -> bool -> string -> bool -> bool -> term list -> term
blanchet@47168
   107
    -> ((string * stature) * term) list
blanchet@47168
   108
    -> string problem * string Symtab.table * (string * stature) list vector
blanchet@46422
   109
       * (string * term) list * int Symtab.table
blanchet@41561
   110
  val atp_problem_weights : string problem -> (string * real) list
blanchet@38506
   111
end;
blanchet@38506
   112
blanchet@47148
   113
structure ATP_Problem_Generate : ATP_PROBLEM_GENERATE =
blanchet@38506
   114
struct
blanchet@38506
   115
blanchet@43926
   116
open ATP_Util
blanchet@38506
   117
open ATP_Problem
blanchet@43926
   118
blanchet@43926
   119
type name = string * string
blanchet@43926
   120
blanchet@45351
   121
val type_tag_idempotence =
blanchet@45411
   122
  Attrib.setup_config_bool @{binding atp_type_tag_idempotence} (K false)
blanchet@45351
   123
val type_tag_arguments =
blanchet@45411
   124
  Attrib.setup_config_bool @{binding atp_type_tag_arguments} (K false)
blanchet@45351
   125
blanchet@46387
   126
val no_lamsN = "no_lams" (* used internally; undocumented *)
blanchet@46384
   127
val hide_lamsN = "hide_lams"
blanchet@47193
   128
val liftingN = "lifting"
blanchet@47193
   129
val combsN = "combs"
blanchet@47193
   130
val combs_and_liftingN = "combs_and_lifting"
blanchet@47193
   131
val combs_or_liftingN = "combs_or_lifting"
blanchet@46384
   132
val keep_lamsN = "keep_lams"
blanchet@47193
   133
val lam_liftingN = "lam_lifting" (* legacy *)
blanchet@44959
   134
blanchet@46650
   135
(* It's still unclear whether all TFF1 implementations will support type
blanchet@46650
   136
   signatures such as "!>[A : $tType] : $o", with ghost type variables. *)
blanchet@46650
   137
val avoid_first_order_ghost_type_vars = false
blanchet@45480
   138
blanchet@43926
   139
val bound_var_prefix = "B_"
blanchet@45262
   140
val all_bound_var_prefix = "BA_"
blanchet@45262
   141
val exist_bound_var_prefix = "BE_"
blanchet@43926
   142
val schematic_var_prefix = "V_"
blanchet@43926
   143
val fixed_var_prefix = "v_"
blanchet@43926
   144
val tvar_prefix = "T_"
blanchet@43926
   145
val tfree_prefix = "t_"
blanchet@43926
   146
val const_prefix = "c_"
blanchet@43926
   147
val type_const_prefix = "tc_"
blanchet@45346
   148
val simple_type_prefix = "s_"
blanchet@43926
   149
val class_prefix = "cl_"
blanchet@43926
   150
blanchet@46380
   151
(* Freshness almost guaranteed! *)
blanchet@46380
   152
val atp_weak_prefix = "ATP:"
blanchet@46380
   153
blanchet@46425
   154
val lam_lifted_prefix = atp_weak_prefix ^ "Lam"
blanchet@46425
   155
val lam_lifted_mono_prefix = lam_lifted_prefix ^ "m"
blanchet@46425
   156
val lam_lifted_poly_prefix = lam_lifted_prefix ^ "p"
blanchet@44778
   157
blanchet@44734
   158
val skolem_const_prefix = "ATP" ^ Long_Name.separator ^ "Sko"
blanchet@43926
   159
val old_skolem_const_prefix = skolem_const_prefix ^ "o"
blanchet@43926
   160
val new_skolem_const_prefix = skolem_const_prefix ^ "n"
blanchet@43926
   161
blanchet@46425
   162
val combinator_prefix = "COMB"
blanchet@46425
   163
blanchet@43839
   164
val type_decl_prefix = "ty_"
blanchet@43839
   165
val sym_decl_prefix = "sy_"
blanchet@44860
   166
val guards_sym_formula_prefix = "gsy_"
blanchet@45255
   167
val tags_sym_formula_prefix = "tsy_"
blanchet@40445
   168
val fact_prefix = "fact_"
blanchet@38506
   169
val conjecture_prefix = "conj_"
blanchet@38506
   170
val helper_prefix = "help_"
blanchet@44000
   171
val class_rel_clause_prefix = "clar_"
blanchet@38506
   172
val arity_clause_prefix = "arity_"
blanchet@43926
   173
val tfree_clause_prefix = "tfree_"
blanchet@38506
   174
blanchet@46425
   175
val lam_fact_prefix = "ATP.lambda_"
blanchet@43750
   176
val typed_helper_suffix = "_T"
blanchet@43750
   177
val untyped_helper_suffix = "_U"
blanchet@44000
   178
val type_tag_idempotence_helper_name = helper_prefix ^ "ti_idem"
blanchet@43750
   179
blanchet@45346
   180
val predicator_name = "pp"
blanchet@45346
   181
val app_op_name = "aa"
blanchet@45346
   182
val type_guard_name = "gg"
blanchet@45346
   183
val type_tag_name = "tt"
blanchet@43402
   184
blanchet@44015
   185
val prefixed_predicator_name = const_prefix ^ predicator_name
blanchet@43971
   186
val prefixed_app_op_name = const_prefix ^ app_op_name
blanchet@43971
   187
val prefixed_type_tag_name = const_prefix ^ type_tag_name
blanchet@43971
   188
blanchet@43926
   189
(*Escaping of special characters.
blanchet@43926
   190
  Alphanumeric characters are left unchanged.
blanchet@43926
   191
  The character _ goes to __
blanchet@43926
   192
  Characters in the range ASCII space to / go to _A to _P, respectively.
blanchet@43926
   193
  Other characters go to _nnn where nnn is the decimal ASCII code.*)
blanchet@43934
   194
val upper_a_minus_space = Char.ord #"A" - Char.ord #" "
blanchet@43926
   195
blanchet@43926
   196
fun stringN_of_int 0 _ = ""
blanchet@43926
   197
  | stringN_of_int k n =
blanchet@43926
   198
    stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
blanchet@43926
   199
blanchet@43926
   200
fun ascii_of_char c =
blanchet@43926
   201
  if Char.isAlphaNum c then
blanchet@43926
   202
    String.str c
blanchet@43926
   203
  else if c = #"_" then
blanchet@43926
   204
    "__"
blanchet@43926
   205
  else if #" " <= c andalso c <= #"/" then
blanchet@43926
   206
    "_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space))
blanchet@43926
   207
  else
blanchet@43926
   208
    (* fixed width, in case more digits follow *)
blanchet@43926
   209
    "_" ^ stringN_of_int 3 (Char.ord c)
blanchet@43926
   210
blanchet@43926
   211
val ascii_of = String.translate ascii_of_char
blanchet@43926
   212
blanchet@43926
   213
(** Remove ASCII armoring from names in proof files **)
blanchet@43926
   214
blanchet@43926
   215
(* We don't raise error exceptions because this code can run inside a worker
blanchet@43926
   216
   thread. Also, the errors are impossible. *)
blanchet@43926
   217
val unascii_of =
blanchet@43926
   218
  let
blanchet@43926
   219
    fun un rcs [] = String.implode(rev rcs)
blanchet@43926
   220
      | un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)
blanchet@43926
   221
        (* Three types of _ escapes: __, _A to _P, _nnn *)
blanchet@44367
   222
      | un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs
blanchet@43926
   223
      | un rcs (#"_" :: c :: cs) =
blanchet@43926
   224
        if #"A" <= c andalso c<= #"P" then
blanchet@43926
   225
          (* translation of #" " to #"/" *)
blanchet@43926
   226
          un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs
blanchet@43926
   227
        else
blanchet@44367
   228
          let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in
blanchet@43926
   229
            case Int.fromString (String.implode digits) of
blanchet@43926
   230
              SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))
blanchet@44367
   231
            | NONE => un (c :: #"_" :: rcs) cs (* ERROR *)
blanchet@43926
   232
          end
blanchet@43926
   233
      | un rcs (c :: cs) = un (c :: rcs) cs
blanchet@43926
   234
  in un [] o String.explode end
blanchet@43926
   235
blanchet@43926
   236
(* If string s has the prefix s1, return the result of deleting it,
blanchet@43926
   237
   un-ASCII'd. *)
blanchet@46382
   238
fun unprefix_and_unascii s1 s =
blanchet@43926
   239
  if String.isPrefix s1 s then
blanchet@43926
   240
    SOME (unascii_of (String.extract (s, size s1, NONE)))
blanchet@43926
   241
  else
blanchet@43926
   242
    NONE
blanchet@43926
   243
blanchet@44000
   244
val proxy_table =
blanchet@44000
   245
  [("c_False", (@{const_name False}, (@{thm fFalse_def},
blanchet@44000
   246
       ("fFalse", @{const_name ATP.fFalse})))),
blanchet@44000
   247
   ("c_True", (@{const_name True}, (@{thm fTrue_def},
blanchet@44000
   248
       ("fTrue", @{const_name ATP.fTrue})))),
blanchet@44000
   249
   ("c_Not", (@{const_name Not}, (@{thm fNot_def},
blanchet@44000
   250
       ("fNot", @{const_name ATP.fNot})))),
blanchet@44000
   251
   ("c_conj", (@{const_name conj}, (@{thm fconj_def},
blanchet@44000
   252
       ("fconj", @{const_name ATP.fconj})))),
blanchet@44000
   253
   ("c_disj", (@{const_name disj}, (@{thm fdisj_def},
blanchet@44000
   254
       ("fdisj", @{const_name ATP.fdisj})))),
blanchet@44000
   255
   ("c_implies", (@{const_name implies}, (@{thm fimplies_def},
blanchet@44000
   256
       ("fimplies", @{const_name ATP.fimplies})))),
blanchet@44000
   257
   ("equal", (@{const_name HOL.eq}, (@{thm fequal_def},
nik@44537
   258
       ("fequal", @{const_name ATP.fequal})))),
nik@44537
   259
   ("c_All", (@{const_name All}, (@{thm fAll_def},
nik@44537
   260
       ("fAll", @{const_name ATP.fAll})))),
nik@44537
   261
   ("c_Ex", (@{const_name Ex}, (@{thm fEx_def},
nik@44537
   262
       ("fEx", @{const_name ATP.fEx}))))]
blanchet@43926
   263
blanchet@44000
   264
val proxify_const = AList.lookup (op =) proxy_table #> Option.map (snd o snd)
blanchet@43926
   265
blanchet@43926
   266
(* Readable names for the more common symbolic functions. Do not mess with the
blanchet@43926
   267
   table unless you know what you are doing. *)
blanchet@43926
   268
val const_trans_table =
blanchet@43926
   269
  [(@{type_name Product_Type.prod}, "prod"),
blanchet@43926
   270
   (@{type_name Sum_Type.sum}, "sum"),
blanchet@43926
   271
   (@{const_name False}, "False"),
blanchet@43926
   272
   (@{const_name True}, "True"),
blanchet@43926
   273
   (@{const_name Not}, "Not"),
blanchet@43926
   274
   (@{const_name conj}, "conj"),
blanchet@43926
   275
   (@{const_name disj}, "disj"),
blanchet@43926
   276
   (@{const_name implies}, "implies"),
blanchet@43926
   277
   (@{const_name HOL.eq}, "equal"),
nik@44537
   278
   (@{const_name All}, "All"),
nik@44537
   279
   (@{const_name Ex}, "Ex"),
blanchet@43926
   280
   (@{const_name If}, "If"),
blanchet@43926
   281
   (@{const_name Set.member}, "member"),
blanchet@46425
   282
   (@{const_name Meson.COMBI}, combinator_prefix ^ "I"),
blanchet@46425
   283
   (@{const_name Meson.COMBK}, combinator_prefix ^ "K"),
blanchet@46425
   284
   (@{const_name Meson.COMBB}, combinator_prefix ^ "B"),
blanchet@46425
   285
   (@{const_name Meson.COMBC}, combinator_prefix ^ "C"),
blanchet@46425
   286
   (@{const_name Meson.COMBS}, combinator_prefix ^ "S")]
blanchet@43926
   287
  |> Symtab.make
blanchet@44000
   288
  |> fold (Symtab.update o swap o snd o snd o snd) proxy_table
blanchet@43926
   289
blanchet@43926
   290
(* Invert the table of translations between Isabelle and ATPs. *)
blanchet@43926
   291
val const_trans_table_inv =
blanchet@43926
   292
  const_trans_table |> Symtab.dest |> map swap |> Symtab.make
blanchet@43926
   293
val const_trans_table_unprox =
blanchet@43926
   294
  Symtab.empty
blanchet@44000
   295
  |> fold (fn (_, (isa, (_, (_, atp)))) => Symtab.update (atp, isa)) proxy_table
blanchet@43926
   296
blanchet@43926
   297
val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
blanchet@43926
   298
val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox)
blanchet@43926
   299
blanchet@43926
   300
fun lookup_const c =
blanchet@43926
   301
  case Symtab.lookup const_trans_table c of
blanchet@43926
   302
    SOME c' => c'
blanchet@43926
   303
  | NONE => ascii_of c
blanchet@43926
   304
blanchet@44489
   305
fun ascii_of_indexname (v, 0) = ascii_of v
blanchet@44489
   306
  | ascii_of_indexname (v, i) = ascii_of v ^ "_" ^ string_of_int i
blanchet@43926
   307
blanchet@43926
   308
fun make_bound_var x = bound_var_prefix ^ ascii_of x
blanchet@45262
   309
fun make_all_bound_var x = all_bound_var_prefix ^ ascii_of x
blanchet@45262
   310
fun make_exist_bound_var x = exist_bound_var_prefix ^ ascii_of x
blanchet@43926
   311
fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
blanchet@43926
   312
fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
blanchet@43926
   313
blanchet@44489
   314
fun make_schematic_type_var (x, i) =
blanchet@45459
   315
  tvar_prefix ^ (ascii_of_indexname (unprefix "'" x, i))
blanchet@44489
   316
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (unprefix "'" x))
blanchet@43926
   317
blanchet@46172
   318
(* "HOL.eq" and choice are mapped to the ATP's equivalents *)
nik@45451
   319
local
nik@45451
   320
  val choice_const = (fst o dest_Const o HOLogic.choice_const) Term.dummyT
nik@45451
   321
  fun default c = const_prefix ^ lookup_const c
nik@45451
   322
in
nik@45451
   323
  fun make_fixed_const _ @{const_name HOL.eq} = tptp_old_equal
blanchet@45618
   324
    | make_fixed_const (SOME (THF (_, _, THF_With_Choice))) c =
blanchet@45618
   325
      if c = choice_const then tptp_choice else default c
nik@45451
   326
    | make_fixed_const _ c = default c
nik@45451
   327
end
blanchet@43926
   328
blanchet@43926
   329
fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
blanchet@43926
   330
blanchet@43926
   331
fun make_type_class clas = class_prefix ^ ascii_of clas
blanchet@43926
   332
blanchet@43934
   333
fun new_skolem_var_name_from_const s =
blanchet@43934
   334
  let val ss = s |> space_explode Long_Name.separator in
blanchet@43934
   335
    nth ss (length ss - 2)
blanchet@43934
   336
  end
blanchet@43934
   337
blanchet@44089
   338
(* These are either simplified away by "Meson.presimplify" (most of the time) or
blanchet@44089
   339
   handled specially via "fFalse", "fTrue", ..., "fequal". *)
blanchet@44089
   340
val atp_irrelevant_consts =
blanchet@44089
   341
  [@{const_name False}, @{const_name True}, @{const_name Not},
blanchet@44089
   342
   @{const_name conj}, @{const_name disj}, @{const_name implies},
blanchet@44089
   343
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
blanchet@44089
   344
blanchet@44089
   345
val atp_monomorph_bad_consts =
blanchet@44089
   346
  atp_irrelevant_consts @
blanchet@44089
   347
  (* These are ignored anyway by the relevance filter (unless they appear in
blanchet@44089
   348
     higher-order places) but not by the monomorphizer. *)
blanchet@44089
   349
  [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
blanchet@44089
   350
   @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
blanchet@44089
   351
   @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
blanchet@44089
   352
blanchet@44099
   353
fun add_schematic_const (x as (_, T)) =
blanchet@44099
   354
  Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x
blanchet@44099
   355
val add_schematic_consts_of =
blanchet@44099
   356
  Term.fold_aterms (fn Const (x as (s, _)) =>
blanchet@44099
   357
                       not (member (op =) atp_monomorph_bad_consts s)
blanchet@44099
   358
                       ? add_schematic_const x
blanchet@44099
   359
                      | _ => I)
blanchet@44099
   360
fun atp_schematic_consts_of t = add_schematic_consts_of t Symtab.empty
blanchet@44089
   361
blanchet@43926
   362
(** Definitions and functions for FOL clauses and formulas for TPTP **)
blanchet@43926
   363
blanchet@43926
   364
(** Isabelle arities **)
blanchet@43926
   365
blanchet@45483
   366
type arity_atom = name * name * name list
blanchet@43926
   367
blanchet@44104
   368
val type_class = the_single @{sort type}
blanchet@44104
   369
blanchet@43927
   370
type arity_clause =
blanchet@44367
   371
  {name : string,
blanchet@45483
   372
   prem_atoms : arity_atom list,
blanchet@45483
   373
   concl_atom : arity_atom}
blanchet@45483
   374
blanchet@45483
   375
fun add_prem_atom tvar =
blanchet@45483
   376
  fold (fn s => s <> type_class ? cons (`make_type_class s, `I tvar, []))
blanchet@43926
   377
blanchet@43926
   378
(* Arity of type constructor "tcon :: (arg1, ..., argN) res" *)
blanchet@43926
   379
fun make_axiom_arity_clause (tcons, name, (cls, args)) =
blanchet@43926
   380
  let
blanchet@45459
   381
    val tvars = map (prefix tvar_prefix o string_of_int) (1 upto length args)
blanchet@43926
   382
    val tvars_srts = ListPair.zip (tvars, args)
blanchet@43926
   383
  in
blanchet@43927
   384
    {name = name,
blanchet@45483
   385
     prem_atoms = [] |> fold (uncurry add_prem_atom) tvars_srts,
blanchet@45483
   386
     concl_atom = (`make_type_class cls, `make_fixed_type_const tcons,
blanchet@45483
   387
                   tvars ~~ tvars)}
blanchet@43926
   388
  end
blanchet@43926
   389
blanchet@43926
   390
fun arity_clause _ _ (_, []) = []
blanchet@44366
   391
  | arity_clause seen n (tcons, ("HOL.type", _) :: ars) =  (* ignore *)
blanchet@44366
   392
    arity_clause seen n (tcons, ars)
blanchet@44366
   393
  | arity_clause seen n (tcons, (ar as (class, _)) :: ars) =
blanchet@44366
   394
    if member (op =) seen class then
blanchet@44366
   395
      (* multiple arities for the same (tycon, class) pair *)
blanchet@44366
   396
      make_axiom_arity_clause (tcons,
blanchet@44366
   397
          lookup_const tcons ^ "___" ^ ascii_of class ^ "_" ^ string_of_int n,
blanchet@44366
   398
          ar) ::
blanchet@44366
   399
      arity_clause seen (n + 1) (tcons, ars)
blanchet@44366
   400
    else
blanchet@44366
   401
      make_axiom_arity_clause (tcons, lookup_const tcons ^ "___" ^
blanchet@44366
   402
                               ascii_of class, ar) ::
blanchet@44366
   403
      arity_clause (class :: seen) n (tcons, ars)
blanchet@43926
   404
blanchet@43926
   405
fun multi_arity_clause [] = []
blanchet@43926
   406
  | multi_arity_clause ((tcons, ars) :: tc_arlists) =
blanchet@45643
   407
    arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
blanchet@43926
   408
blanchet@44489
   409
(* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in
blanchet@44489
   410
   theory thy provided its arguments have the corresponding sorts. *)
blanchet@43926
   411
fun type_class_pairs thy tycons classes =
blanchet@43934
   412
  let
blanchet@43934
   413
    val alg = Sign.classes_of thy
blanchet@43934
   414
    fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
blanchet@43934
   415
    fun add_class tycon class =
blanchet@43934
   416
      cons (class, domain_sorts tycon class)
blanchet@43934
   417
      handle Sorts.CLASS_ERROR _ => I
blanchet@43934
   418
    fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
blanchet@43934
   419
  in map try_classes tycons end
blanchet@43926
   420
blanchet@43926
   421
(*Proving one (tycon, class) membership may require proving others, so iterate.*)
blanchet@43926
   422
fun iter_type_class_pairs _ _ [] = ([], [])
blanchet@43926
   423
  | iter_type_class_pairs thy tycons classes =
blanchet@44104
   424
      let
blanchet@44104
   425
        fun maybe_insert_class s =
blanchet@44104
   426
          (s <> type_class andalso not (member (op =) classes s))
blanchet@44104
   427
          ? insert (op =) s
blanchet@44104
   428
        val cpairs = type_class_pairs thy tycons classes
blanchet@44104
   429
        val newclasses =
blanchet@44104
   430
          [] |> fold (fold (fold (fold maybe_insert_class) o snd) o snd) cpairs
blanchet@44104
   431
        val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
blanchet@44107
   432
      in (classes' @ classes, union (op =) cpairs' cpairs) end
blanchet@43926
   433
blanchet@43926
   434
fun make_arity_clauses thy tycons =
blanchet@43926
   435
  iter_type_class_pairs thy tycons ##> multi_arity_clause
blanchet@43926
   436
blanchet@43926
   437
blanchet@43926
   438
(** Isabelle class relations **)
blanchet@43926
   439
blanchet@43927
   440
type class_rel_clause =
blanchet@44367
   441
  {name : string,
blanchet@44367
   442
   subclass : name,
blanchet@44367
   443
   superclass : name}
blanchet@43926
   444
blanchet@44489
   445
(* Generate all pairs (sub, super) such that sub is a proper subclass of super
blanchet@44489
   446
   in theory "thy". *)
blanchet@43926
   447
fun class_pairs _ [] _ = []
blanchet@43926
   448
  | class_pairs thy subs supers =
blanchet@43926
   449
      let
blanchet@43926
   450
        val class_less = Sorts.class_less (Sign.classes_of thy)
blanchet@43926
   451
        fun add_super sub super = class_less (sub, super) ? cons (sub, super)
blanchet@43926
   452
        fun add_supers sub = fold (add_super sub) supers
blanchet@43926
   453
      in fold add_supers subs [] end
blanchet@43926
   454
blanchet@44489
   455
fun make_class_rel_clause (sub, super) =
blanchet@44489
   456
  {name = sub ^ "_" ^ super, subclass = `make_type_class sub,
blanchet@43927
   457
   superclass = `make_type_class super}
blanchet@43926
   458
blanchet@43926
   459
fun make_class_rel_clauses thy subs supers =
blanchet@43934
   460
  map make_class_rel_clause (class_pairs thy subs supers)
blanchet@43926
   461
blanchet@44730
   462
(* intermediate terms *)
blanchet@44730
   463
datatype iterm =
blanchet@44730
   464
  IConst of name * typ * typ list |
blanchet@44730
   465
  IVar of name * typ |
blanchet@44730
   466
  IApp of iterm * iterm |
blanchet@44730
   467
  IAbs of (name * typ) * iterm
blanchet@43926
   468
blanchet@44730
   469
fun ityp_of (IConst (_, T, _)) = T
blanchet@44730
   470
  | ityp_of (IVar (_, T)) = T
blanchet@44730
   471
  | ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1))
blanchet@44730
   472
  | ityp_of (IAbs ((_, T), tm)) = T --> ityp_of tm
blanchet@43926
   473
blanchet@43926
   474
(*gets the head of a combinator application, along with the list of arguments*)
blanchet@44730
   475
fun strip_iterm_comb u =
blanchet@44367
   476
  let
blanchet@44730
   477
    fun stripc (IApp (t, u), ts) = stripc (t, u :: ts)
blanchet@44367
   478
      | stripc x = x
blanchet@44367
   479
  in stripc (u, []) end
blanchet@43926
   480
blanchet@46187
   481
fun atomic_types_of T = fold_atyps (insert (op =)) T []
blanchet@43926
   482
blanchet@46380
   483
val tvar_a_str = "'a"
blanchet@46380
   484
val tvar_a = TVar ((tvar_a_str, 0), HOLogic.typeS)
blanchet@46380
   485
val tvar_a_name = (make_schematic_type_var (tvar_a_str, 0), tvar_a_str)
blanchet@46380
   486
val itself_name = `make_fixed_type_const @{type_name itself}
blanchet@46380
   487
val TYPE_name = `(make_fixed_const NONE) @{const_name TYPE}
blanchet@46380
   488
val tvar_a_atype = AType (tvar_a_name, [])
blanchet@46380
   489
val a_itself_atype = AType (itself_name, [tvar_a_atype])
blanchet@46380
   490
blanchet@43926
   491
fun new_skolem_const_name s num_T_args =
blanchet@43926
   492
  [new_skolem_const_prefix, s, string_of_int num_T_args]
blanchet@43926
   493
  |> space_implode Long_Name.separator
blanchet@43926
   494
blanchet@45458
   495
fun robust_const_type thy s =
blanchet@46380
   496
  if s = app_op_name then
blanchet@46380
   497
    Logic.varifyT_global @{typ "('a => 'b) => 'a => 'b"}
blanchet@46425
   498
  else if String.isPrefix lam_lifted_prefix s then
blanchet@46380
   499
    Logic.varifyT_global @{typ "'a => 'b"}
blanchet@46380
   500
  else
blanchet@46380
   501
    (* Old Skolems throw a "TYPE" exception here, which will be caught. *)
blanchet@46380
   502
    s |> Sign.the_const_type thy
blanchet@45458
   503
blanchet@45458
   504
(* This function only makes sense if "T" is as general as possible. *)
blanchet@45458
   505
fun robust_const_typargs thy (s, T) =
blanchet@46380
   506
  if s = app_op_name then
blanchet@46380
   507
    let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end
blanchet@46380
   508
  else if String.isPrefix old_skolem_const_prefix s then
blanchet@46380
   509
    [] |> Term.add_tvarsT T |> rev |> map TVar
blanchet@46425
   510
  else if String.isPrefix lam_lifted_prefix s then
blanchet@46425
   511
    if String.isPrefix lam_lifted_poly_prefix s then
blanchet@46382
   512
      let val (T1, T2) = T |> dest_funT in [T1, T2] end
blanchet@46382
   513
    else
blanchet@46382
   514
      []
blanchet@46380
   515
  else
blanchet@46380
   516
    (s, T) |> Sign.const_typargs thy
blanchet@45458
   517
blanchet@44730
   518
(* Converts an Isabelle/HOL term (with combinators) into an intermediate term.
blanchet@44730
   519
   Also accumulates sort infomation. *)
nik@45350
   520
fun iterm_from_term thy format bs (P $ Q) =
blanchet@43926
   521
    let
nik@45350
   522
      val (P', P_atomics_Ts) = iterm_from_term thy format bs P
nik@45350
   523
      val (Q', Q_atomics_Ts) = iterm_from_term thy format bs Q
blanchet@44730
   524
    in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
nik@45350
   525
  | iterm_from_term thy format _ (Const (c, T)) =
nik@45350
   526
    (IConst (`(make_fixed_const (SOME format)) c, T,
blanchet@45458
   527
             robust_const_typargs thy (c, T)),
blanchet@46187
   528
     atomic_types_of T)
nik@45350
   529
  | iterm_from_term _ _ _ (Free (s, T)) =
blanchet@46380
   530
    (IConst (`make_fixed_var s, T, []), atomic_types_of T)
nik@45350
   531
  | iterm_from_term _ format _ (Var (v as (s, _), T)) =
blanchet@43926
   532
    (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
blanchet@43926
   533
       let
blanchet@43926
   534
         val Ts = T |> strip_type |> swap |> op ::
blanchet@43926
   535
         val s' = new_skolem_const_name s (length Ts)
nik@45350
   536
       in IConst (`(make_fixed_const (SOME format)) s', T, Ts) end
blanchet@43926
   537
     else
blanchet@46187
   538
       IVar ((make_schematic_var v, s), T), atomic_types_of T)
nik@45350
   539
  | iterm_from_term _ _ bs (Bound j) =
blanchet@46187
   540
    nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T))
nik@45350
   541
  | iterm_from_term thy format bs (Abs (s, T, t)) =
nik@44537
   542
    let
nik@44537
   543
      fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string
nik@44537
   544
      val s = vary s
blanchet@45262
   545
      val name = `make_bound_var s
nik@45350
   546
      val (tm, atomic_Ts) = iterm_from_term thy format ((s, (name, T)) :: bs) t
blanchet@46187
   547
    in (IAbs ((name, T), tm), union (op =) atomic_Ts (atomic_types_of T)) end
blanchet@43926
   548
blanchet@47168
   549
datatype scope = Global | Local | Assum | Chained
blanchet@47169
   550
datatype status = General | Induct | Intro | Elim | Simp | Eq
blanchet@47168
   551
type stature = scope * status
blanchet@43926
   552
blanchet@44491
   553
datatype order = First_Order | Higher_Order
blanchet@45349
   554
datatype polymorphism = Polymorphic | Raw_Monomorphic | Mangled_Monomorphic
blanchet@47129
   555
datatype strictness = Strict | Non_Strict
blanchet@45673
   556
datatype granularity = All_Vars | Positively_Naked_Vars | Ghost_Type_Arg_Vars
blanchet@43484
   557
datatype type_level =
blanchet@45256
   558
  All_Types |
blanchet@47129
   559
  Noninf_Nonmono_Types of strictness * granularity |
blanchet@45673
   560
  Fin_Nonmono_Types of granularity |
blanchet@45256
   561
  Const_Arg_Types |
blanchet@44233
   562
  No_Types
blanchet@43484
   563
blanchet@44493
   564
datatype type_enc =
blanchet@45455
   565
  Simple_Types of order * polymorphism * type_level |
blanchet@45639
   566
  Guards of polymorphism * type_level |
blanchet@45639
   567
  Tags of polymorphism * type_level
blanchet@45639
   568
blanchet@45639
   569
fun is_type_enc_higher_order (Simple_Types (Higher_Order, _, _)) = true
blanchet@45639
   570
  | is_type_enc_higher_order _ = false
blanchet@45639
   571
blanchet@45639
   572
fun polymorphism_of_type_enc (Simple_Types (_, poly, _)) = poly
blanchet@45639
   573
  | polymorphism_of_type_enc (Guards (poly, _)) = poly
blanchet@45639
   574
  | polymorphism_of_type_enc (Tags (poly, _)) = poly
blanchet@45639
   575
blanchet@45639
   576
fun level_of_type_enc (Simple_Types (_, _, level)) = level
blanchet@45639
   577
  | level_of_type_enc (Guards (_, level)) = level
blanchet@45639
   578
  | level_of_type_enc (Tags (_, level)) = level
blanchet@45639
   579
blanchet@45673
   580
fun granularity_of_type_level (Noninf_Nonmono_Types (_, grain)) = grain
blanchet@45673
   581
  | granularity_of_type_level (Fin_Nonmono_Types grain) = grain
blanchet@45673
   582
  | granularity_of_type_level _ = All_Vars
blanchet@45639
   583
blanchet@45639
   584
fun is_type_level_quasi_sound All_Types = true
blanchet@45639
   585
  | is_type_level_quasi_sound (Noninf_Nonmono_Types _) = true
blanchet@45639
   586
  | is_type_level_quasi_sound _ = false
blanchet@45639
   587
val is_type_enc_quasi_sound = is_type_level_quasi_sound o level_of_type_enc
blanchet@45639
   588
blanchet@45639
   589
fun is_type_level_fairly_sound (Fin_Nonmono_Types _) = true
blanchet@45639
   590
  | is_type_level_fairly_sound level = is_type_level_quasi_sound level
blanchet@45639
   591
val is_type_enc_fairly_sound = is_type_level_fairly_sound o level_of_type_enc
blanchet@45639
   592
blanchet@45639
   593
fun is_type_level_monotonicity_based (Noninf_Nonmono_Types _) = true
blanchet@45639
   594
  | is_type_level_monotonicity_based (Fin_Nonmono_Types _) = true
blanchet@45639
   595
  | is_type_level_monotonicity_based _ = false
blanchet@43484
   596
blanchet@45653
   597
(* "_query", "_bang", and "_at" are for the ASCII-challenged Metis and
blanchet@45653
   598
   Mirabelle. *)
blanchet@45653
   599
val queries = ["?", "_query"]
blanchet@45653
   600
val bangs = ["!", "_bang"]
blanchet@45653
   601
val ats = ["@", "_at"]
blanchet@45653
   602
blanchet@43559
   603
fun try_unsuffixes ss s =
blanchet@43559
   604
  fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
blanchet@43559
   605
blanchet@45653
   606
fun try_nonmono constr suffixes fallback s =
blanchet@45653
   607
  case try_unsuffixes suffixes s of
blanchet@45653
   608
    SOME s =>
blanchet@45653
   609
    (case try_unsuffixes suffixes s of
blanchet@45673
   610
       SOME s => (constr Positively_Naked_Vars, s)
blanchet@45653
   611
     | NONE =>
blanchet@45653
   612
       case try_unsuffixes ats s of
blanchet@45673
   613
         SOME s => (constr Ghost_Type_Arg_Vars, s)
blanchet@45673
   614
       | NONE => (constr All_Vars, s))
blanchet@45653
   615
  | NONE => fallback s
blanchet@45639
   616
blanchet@47129
   617
fun type_enc_from_string strictness s =
blanchet@43587
   618
  (case try (unprefix "poly_") s of
blanchet@43587
   619
     SOME s => (SOME Polymorphic, s)
blanchet@43484
   620
   | NONE =>
blanchet@45349
   621
     case try (unprefix "raw_mono_") s of
blanchet@45349
   622
       SOME s => (SOME Raw_Monomorphic, s)
blanchet@43587
   623
     | NONE =>
blanchet@45349
   624
       case try (unprefix "mono_") s of
blanchet@43587
   625
         SOME s => (SOME Mangled_Monomorphic, s)
blanchet@43587
   626
       | NONE => (NONE, s))
blanchet@45653
   627
  ||> (pair All_Types
blanchet@45653
   628
       |> try_nonmono Fin_Nonmono_Types bangs
blanchet@47129
   629
       |> try_nonmono (curry Noninf_Nonmono_Types strictness) queries)
blanchet@45639
   630
  |> (fn (poly, (level, core)) =>
blanchet@45639
   631
         case (core, (poly, level)) of
blanchet@45639
   632
           ("simple", (SOME poly, _)) =>
blanchet@45606
   633
           (case (poly, level) of
blanchet@45606
   634
              (Polymorphic, All_Types) =>
blanchet@45606
   635
              Simple_Types (First_Order, Polymorphic, All_Types)
blanchet@45606
   636
            | (Mangled_Monomorphic, _) =>
blanchet@45673
   637
              if granularity_of_type_level level = All_Vars then
blanchet@45639
   638
                Simple_Types (First_Order, Mangled_Monomorphic, level)
blanchet@45639
   639
              else
blanchet@45639
   640
                raise Same.SAME
blanchet@45606
   641
            | _ => raise Same.SAME)
blanchet@45639
   642
         | ("simple_higher", (SOME poly, _)) =>
blanchet@45455
   643
           (case (poly, level) of
blanchet@45618
   644
              (Polymorphic, All_Types) =>
blanchet@45618
   645
              Simple_Types (Higher_Order, Polymorphic, All_Types)
blanchet@45618
   646
            | (_, Noninf_Nonmono_Types _) => raise Same.SAME
blanchet@45606
   647
            | (Mangled_Monomorphic, _) =>
blanchet@45673
   648
              if granularity_of_type_level level = All_Vars then
blanchet@45639
   649
                Simple_Types (Higher_Order, Mangled_Monomorphic, level)
blanchet@45639
   650
              else
blanchet@45639
   651
                raise Same.SAME
blanchet@45606
   652
            | _ => raise Same.SAME)
blanchet@45672
   653
         | ("guards", (SOME poly, _)) =>
blanchet@46820
   654
           if poly = Mangled_Monomorphic andalso
blanchet@46820
   655
              granularity_of_type_level level = Ghost_Type_Arg_Vars then
blanchet@46820
   656
             raise Same.SAME
blanchet@46820
   657
           else
blanchet@46820
   658
             Guards (poly, level)
blanchet@45672
   659
         | ("tags", (SOME poly, _)) =>
blanchet@46820
   660
           if granularity_of_type_level level = Ghost_Type_Arg_Vars then
blanchet@46820
   661
             raise Same.SAME
blanchet@46820
   662
           else
blanchet@46820
   663
             Tags (poly, level)
blanchet@45639
   664
         | ("args", (SOME poly, All_Types (* naja *))) =>
blanchet@45639
   665
           Guards (poly, Const_Arg_Types)
blanchet@45639
   666
         | ("erased", (NONE, All_Types (* naja *))) =>
blanchet@45639
   667
           Guards (Polymorphic, No_Types)
blanchet@43618
   668
         | _ => raise Same.SAME)
blanchet@45653
   669
  handle Same.SAME => error ("Unknown type encoding: " ^ quote s ^ ".")
blanchet@43484
   670
blanchet@45618
   671
fun adjust_type_enc (THF (TPTP_Monomorphic, _, _))
blanchet@45618
   672
                    (Simple_Types (order, _, level)) =
blanchet@45455
   673
    Simple_Types (order, Mangled_Monomorphic, level)
blanchet@45618
   674
  | adjust_type_enc (THF _) type_enc = type_enc
blanchet@45618
   675
  | adjust_type_enc (TFF (TPTP_Monomorphic, _)) (Simple_Types (_, _, level)) =
blanchet@45455
   676
    Simple_Types (First_Order, Mangled_Monomorphic, level)
blanchet@46174
   677
  | adjust_type_enc (DFG DFG_Sorted) (Simple_Types (_, _, level)) =
blanchet@46172
   678
    Simple_Types (First_Order, Mangled_Monomorphic, level)
blanchet@45618
   679
  | adjust_type_enc (TFF _) (Simple_Types (_, poly, level)) =
blanchet@45455
   680
    Simple_Types (First_Order, poly, level)
blanchet@45455
   681
  | adjust_type_enc format (Simple_Types (_, poly, level)) =
blanchet@45639
   682
    adjust_type_enc format (Guards (poly, level))
blanchet@45275
   683
  | adjust_type_enc CNF_UEQ (type_enc as Guards stuff) =
blanchet@45275
   684
    (if is_type_enc_fairly_sound type_enc then Tags else Guards) stuff
blanchet@45275
   685
  | adjust_type_enc _ type_enc = type_enc
blanchet@43942
   686
blanchet@46380
   687
fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u
blanchet@46380
   688
  | constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t)
blanchet@46380
   689
  | constify_lifted (Free (x as (s, _))) =
blanchet@46425
   690
    (if String.isPrefix lam_lifted_prefix s then Const else Free) x
blanchet@46380
   691
  | constify_lifted t = t
blanchet@46380
   692
blanchet@47203
   693
(* Requires bound variables not to clash with any schematic variables (as should
blanchet@47203
   694
   be the case right after lambda-lifting). *)
blanchet@47213
   695
fun open_form unprefix (t as Const (@{const_name All}, _) $ Abs (s, T, t')) =
blanchet@47213
   696
    (case try unprefix s of
blanchet@47213
   697
       SOME s =>
blanchet@47213
   698
       let
blanchet@47213
   699
         val names = Name.make_context (map fst (Term.add_var_names t' []))
blanchet@47213
   700
         val (s, _) = Name.variant s names
blanchet@47213
   701
       in open_form unprefix (subst_bound (Var ((s, 0), T), t')) end
blanchet@47213
   702
     | NONE => t)
blanchet@47213
   703
  | open_form _ t = t
blanchet@47203
   704
blanchet@46425
   705
fun lift_lams_part_1 ctxt type_enc =
blanchet@46439
   706
  map close_form #> rpair ctxt
blanchet@44959
   707
  #-> Lambda_Lifting.lift_lambdas
blanchet@46435
   708
          (SOME ((if polymorphism_of_type_enc type_enc = Polymorphic then
blanchet@46435
   709
                    lam_lifted_poly_prefix
blanchet@46435
   710
                  else
blanchet@46435
   711
                    lam_lifted_mono_prefix) ^ "_a"))
blanchet@44959
   712
          Lambda_Lifting.is_quantifier
blanchet@46425
   713
  #> fst
blanchet@47213
   714
fun lift_lams_part_2 (facts, lifted) =
blanchet@47213
   715
  (map (open_form (unprefix close_form_prefix) o constify_lifted) facts,
blanchet@47213
   716
   map (open_form I o constify_lifted) lifted)
blanchet@46425
   717
val lift_lams = lift_lams_part_2 ooo lift_lams_part_1
blanchet@44959
   718
blanchet@44959
   719
fun intentionalize_def (Const (@{const_name All}, _) $ Abs (_, _, t)) =
blanchet@44959
   720
    intentionalize_def t
blanchet@44959
   721
  | intentionalize_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
blanchet@44959
   722
    let
blanchet@44959
   723
      fun lam T t = Abs (Name.uu, T, t)
blanchet@44959
   724
      val (head, args) = strip_comb t ||> rev
blanchet@44959
   725
      val head_T = fastype_of head
blanchet@44959
   726
      val n = length args
blanchet@44959
   727
      val arg_Ts = head_T |> binder_types |> take n |> rev
blanchet@44959
   728
      val u = u |> subst_atomic (args ~~ map Bound (0 upto n - 1))
blanchet@44959
   729
    in HOLogic.eq_const head_T $ head $ fold lam arg_Ts u end
blanchet@44959
   730
  | intentionalize_def t = t
blanchet@44959
   731
blanchet@40358
   732
type translated_formula =
blanchet@44367
   733
  {name : string,
blanchet@47168
   734
   stature : stature,
blanchet@44367
   735
   kind : formula_kind,
blanchet@44730
   736
   iformula : (name, typ, iterm) formula,
blanchet@44367
   737
   atomic_types : typ list}
blanchet@38506
   738
blanchet@47168
   739
fun update_iformula f ({name, stature, kind, iformula, atomic_types}
blanchet@44730
   740
                       : translated_formula) =
blanchet@47168
   741
  {name = name, stature = stature, kind = kind, iformula = f iformula,
blanchet@43433
   742
   atomic_types = atomic_types} : translated_formula
blanchet@43413
   743
blanchet@44730
   744
fun fact_lift f ({iformula, ...} : translated_formula) = f iformula
blanchet@43429
   745
blanchet@43905
   746
fun insert_type ctxt get_T x xs =
blanchet@43905
   747
  let val T = get_T x in
blanchet@45258
   748
    if exists (type_instance ctxt T o get_T) xs then xs
blanchet@45258
   749
    else x :: filter_out (type_generalization ctxt T o get_T) xs
blanchet@43905
   750
  end
blanchet@43547
   751
blanchet@43618
   752
(* The Booleans indicate whether all type arguments should be kept. *)
blanchet@43618
   753
datatype type_arg_policy =
blanchet@46808
   754
  Explicit_Type_Args of bool (* infer_from_term_args *) |
blanchet@45642
   755
  Mangled_Type_Args |
blanchet@43618
   756
  No_Type_Args
blanchet@41384
   757
blanchet@46816
   758
fun type_arg_policy monom_constrs type_enc s =
blanchet@46186
   759
  let val poly = polymorphism_of_type_enc type_enc in
blanchet@45645
   760
    if s = type_tag_name then
blanchet@46186
   761
      if poly = Mangled_Monomorphic then Mangled_Type_Args
blanchet@46186
   762
      else Explicit_Type_Args false
blanchet@45645
   763
    else case type_enc of
blanchet@46186
   764
      Simple_Types (_, Polymorphic, _) => Explicit_Type_Args false
blanchet@46186
   765
    | Tags (_, All_Types) => No_Type_Args
blanchet@45645
   766
    | _ =>
blanchet@45645
   767
      let val level = level_of_type_enc type_enc in
blanchet@45645
   768
        if level = No_Types orelse s = @{const_name HOL.eq} orelse
blanchet@45645
   769
           (s = app_op_name andalso level = Const_Arg_Types) then
blanchet@45645
   770
          No_Type_Args
blanchet@46186
   771
        else if poly = Mangled_Monomorphic then
blanchet@45645
   772
          Mangled_Type_Args
blanchet@46816
   773
        else if member (op =) monom_constrs s andalso
blanchet@46816
   774
                granularity_of_type_level level = Positively_Naked_Vars then
blanchet@46816
   775
          No_Type_Args
blanchet@45645
   776
        else
blanchet@45673
   777
          Explicit_Type_Args
blanchet@45673
   778
              (level = All_Types orelse
blanchet@45673
   779
               granularity_of_type_level level = Ghost_Type_Arg_Vars)
blanchet@45645
   780
      end
blanchet@45645
   781
  end
blanchet@43088
   782
blanchet@45483
   783
(* Make atoms for sorted type variables. *)
blanchet@44104
   784
fun generic_add_sorts_on_type (_, []) = I
blanchet@44104
   785
  | generic_add_sorts_on_type ((x, i), s :: ss) =
blanchet@44104
   786
    generic_add_sorts_on_type ((x, i), ss)
blanchet@44104
   787
    #> (if s = the_single @{sort HOL.type} then
blanchet@43934
   788
          I
blanchet@43934
   789
        else if i = ~1 then
blanchet@45481
   790
          insert (op =) (`make_type_class s, `make_fixed_type_var x)
blanchet@43934
   791
        else
blanchet@45481
   792
          insert (op =) (`make_type_class s,
blanchet@45481
   793
                         (make_schematic_type_var (x, i), x)))
blanchet@44104
   794
fun add_sorts_on_tfree (TFree (s, S)) = generic_add_sorts_on_type ((s, ~1), S)
blanchet@44104
   795
  | add_sorts_on_tfree _ = I
blanchet@44104
   796
fun add_sorts_on_tvar (TVar z) = generic_add_sorts_on_type z
blanchet@44104
   797
  | add_sorts_on_tvar _ = I
blanchet@43926
   798
blanchet@45483
   799
fun type_class_formula type_enc class arg =
blanchet@45483
   800
  AAtom (ATerm (class, arg ::
blanchet@45483
   801
      (case type_enc of
blanchet@45618
   802
         Simple_Types (First_Order, Polymorphic, _) =>
blanchet@46172
   803
         if avoid_first_order_ghost_type_vars then [ATerm (TYPE_name, [arg])]
blanchet@45618
   804
         else []
blanchet@45483
   805
       | _ => [])))
blanchet@45483
   806
fun formulas_for_types type_enc add_sorts_on_typ Ts =
blanchet@44493
   807
  [] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts
blanchet@45483
   808
     |> map (fn (class, name) =>
blanchet@45483
   809
                type_class_formula type_enc class (ATerm (name, [])))
blanchet@41385
   810
blanchet@43405
   811
fun mk_aconns c phis =
blanchet@43405
   812
  let val (phis', phi') = split_last phis in
blanchet@43405
   813
    fold_rev (mk_aconn c) phis' phi'
blanchet@43405
   814
  end
blanchet@38506
   815
fun mk_ahorn [] phi = phi
blanchet@43405
   816
  | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
blanchet@43393
   817
fun mk_aquant _ [] phi = phi
blanchet@43393
   818
  | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
blanchet@43393
   819
    if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
blanchet@43393
   820
  | mk_aquant q xs phi = AQuant (q, xs, phi)
blanchet@38506
   821
blanchet@46186
   822
fun close_universally add_term_vars phi =
blanchet@41393
   823
  let
blanchet@46186
   824
    fun add_formula_vars bounds (AQuant (_, xs, phi)) =
blanchet@46186
   825
        add_formula_vars (map fst xs @ bounds) phi
blanchet@46186
   826
      | add_formula_vars bounds (AConn (_, phis)) =
blanchet@46186
   827
        fold (add_formula_vars bounds) phis
blanchet@46186
   828
      | add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm
blanchet@46186
   829
  in mk_aquant AForall (add_formula_vars [] phi []) phi end
blanchet@43393
   830
blanchet@46248
   831
fun add_term_vars bounds (ATerm (name as (s, _), tms)) =
blanchet@46248
   832
    (if is_tptp_variable s andalso
blanchet@46248
   833
        not (String.isPrefix tvar_prefix s) andalso
blanchet@46248
   834
        not (member (op =) bounds name) then
blanchet@46248
   835
       insert (op =) (name, NONE)
blanchet@46248
   836
     else
blanchet@46248
   837
       I)
blanchet@46248
   838
    #> fold (add_term_vars bounds) tms
blanchet@46248
   839
  | add_term_vars bounds (AAbs ((name, _), tm)) =
blanchet@46248
   840
    add_term_vars (name :: bounds) tm
blanchet@46272
   841
fun close_formula_universally phi = close_universally add_term_vars phi
blanchet@46186
   842
blanchet@46186
   843
fun add_iterm_vars bounds (IApp (tm1, tm2)) =
blanchet@46186
   844
    fold (add_iterm_vars bounds) [tm1, tm2]
blanchet@46186
   845
  | add_iterm_vars _ (IConst _) = I
blanchet@46186
   846
  | add_iterm_vars bounds (IVar (name, T)) =
blanchet@46186
   847
    not (member (op =) bounds name) ? insert (op =) (name, SOME T)
blanchet@46186
   848
  | add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm
blanchet@46186
   849
fun close_iformula_universally phi = close_universally add_iterm_vars phi
blanchet@41393
   850
blanchet@47166
   851
val fused_infinite_type_name = "ATP.fused_inf" (* shouldn't clash *)
blanchet@45458
   852
val fused_infinite_type = Type (fused_infinite_type_name, [])
blanchet@45458
   853
blanchet@45458
   854
fun tvar_name (x as (s, _)) = (make_schematic_type_var x, s)
blanchet@43835
   855
nik@44535
   856
fun ho_term_from_typ format type_enc =
blanchet@43835
   857
  let
blanchet@43835
   858
    fun term (Type (s, Ts)) =
blanchet@44493
   859
      ATerm (case (is_type_enc_higher_order type_enc, s) of
blanchet@43835
   860
               (true, @{type_name bool}) => `I tptp_bool_type
blanchet@43835
   861
             | (true, @{type_name fun}) => `I tptp_fun_type
blanchet@45458
   862
             | _ => if s = fused_infinite_type_name andalso
blanchet@45106
   863
                       is_format_typed format then
blanchet@44019
   864
                      `I tptp_individual_type
blanchet@44019
   865
                    else
blanchet@44019
   866
                      `make_fixed_type_const s,
blanchet@43835
   867
             map term Ts)
blanchet@43835
   868
    | term (TFree (s, _)) = ATerm (`make_fixed_type_var s, [])
blanchet@45458
   869
    | term (TVar (x, _)) = ATerm (tvar_name x, [])
blanchet@43835
   870
  in term end
blanchet@43433
   871
nik@44535
   872
fun ho_term_for_type_arg format type_enc T =
nik@44535
   873
  if T = dummyT then NONE else SOME (ho_term_from_typ format type_enc T)
blanchet@44264
   874
blanchet@43433
   875
(* This shouldn't clash with anything else. *)
blanchet@43413
   876
val mangled_type_sep = "\000"
blanchet@43413
   877
blanchet@43433
   878
fun generic_mangled_type_name f (ATerm (name, [])) = f name
blanchet@43433
   879
  | generic_mangled_type_name f (ATerm (name, tys)) =
blanchet@43626
   880
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
blanchet@43626
   881
    ^ ")"
blanchet@44558
   882
  | generic_mangled_type_name _ _ = raise Fail "unexpected type abstraction"
blanchet@43413
   883
blanchet@45255
   884
fun mangled_type format type_enc =
blanchet@45255
   885
  generic_mangled_type_name fst o ho_term_from_typ format type_enc
blanchet@45255
   886
blanchet@43926
   887
fun make_simple_type s =
blanchet@43926
   888
  if s = tptp_bool_type orelse s = tptp_fun_type orelse
blanchet@43926
   889
     s = tptp_individual_type then
blanchet@43926
   890
    s
blanchet@43926
   891
  else
blanchet@43926
   892
    simple_type_prefix ^ ascii_of s
blanchet@43926
   893
nik@44535
   894
fun ho_type_from_ho_term type_enc pred_sym ary =
blanchet@43804
   895
  let
blanchet@45457
   896
    fun to_mangled_atype ty =
blanchet@43804
   897
      AType ((make_simple_type (generic_mangled_type_name fst ty),
blanchet@45457
   898
              generic_mangled_type_name snd ty), [])
blanchet@45457
   899
    fun to_poly_atype (ATerm (name, tys)) = AType (name, map to_poly_atype tys)
blanchet@45457
   900
      | to_poly_atype _ = raise Fail "unexpected type abstraction"
blanchet@45457
   901
    val to_atype =
blanchet@45457
   902
      if polymorphism_of_type_enc type_enc = Polymorphic then to_poly_atype
blanchet@45457
   903
      else to_mangled_atype
blanchet@43804
   904
    fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
blanchet@43839
   905
    fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
blanchet@43835
   906
      | to_fo ary (ATerm (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
blanchet@44558
   907
      | to_fo _ _ = raise Fail "unexpected type abstraction"
blanchet@43835
   908
    fun to_ho (ty as ATerm ((s, _), tys)) =
nik@44535
   909
        if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
nik@44535
   910
      | to_ho _ = raise Fail "unexpected type abstraction"
blanchet@44493
   911
  in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end
blanchet@43804
   912
nik@44536
   913
fun ho_type_from_typ format type_enc pred_sym ary =
nik@44535
   914
  ho_type_from_ho_term type_enc pred_sym ary
nik@44535
   915
  o ho_term_from_typ format type_enc
blanchet@43804
   916
blanchet@44493
   917
fun mangled_const_name format type_enc T_args (s, s') =
blanchet@43804
   918
  let
nik@44535
   919
    val ty_args = T_args |> map_filter (ho_term_for_type_arg format type_enc)
blanchet@43804
   920
    fun type_suffix f g =
blanchet@43804
   921
      fold_rev (curry (op ^) o g o prefix mangled_type_sep
blanchet@43804
   922
                o generic_mangled_type_name f) ty_args ""
blanchet@43804
   923
  in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
blanchet@43413
   924
blanchet@43413
   925
val parse_mangled_ident =
blanchet@43413
   926
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
blanchet@43413
   927
blanchet@43413
   928
fun parse_mangled_type x =
blanchet@43413
   929
  (parse_mangled_ident
blanchet@43413
   930
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
blanchet@43413
   931
                    [] >> ATerm) x
blanchet@43413
   932
and parse_mangled_types x =
blanchet@43413
   933
  (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
blanchet@43413
   934
blanchet@43413
   935
fun unmangled_type s =
blanchet@43413
   936
  s |> suffix ")" |> raw_explode
blanchet@43413
   937
    |> Scan.finite Symbol.stopper
blanchet@43413
   938
           (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
blanchet@43413
   939
                                                quote s)) parse_mangled_type))
blanchet@43413
   940
    |> fst
blanchet@43413
   941
blanchet@43432
   942
val unmangled_const_name = space_explode mangled_type_sep #> hd
blanchet@43413
   943
fun unmangled_const s =
blanchet@43413
   944
  let val ss = space_explode mangled_type_sep s in
blanchet@43413
   945
    (hd ss, map unmangled_type (tl ss))
blanchet@43413
   946
  end
blanchet@43413
   947
blanchet@45644
   948
fun introduce_proxies_in_iterm type_enc =
blanchet@43439
   949
  let
blanchet@44858
   950
    fun tweak_ho_quant ho_quant T [IAbs _] = IConst (`I ho_quant, T, [])
blanchet@44858
   951
      | tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _]))
blanchet@44858
   952
                       _ =
blanchet@44858
   953
        (* Eta-expand "!!" and "??", to work around LEO-II 1.2.8 parser
blanchet@44858
   954
           limitation. This works in conjuction with special code in
blanchet@44858
   955
           "ATP_Problem" that uses the syntactic sugar "!" and "?" whenever
blanchet@44858
   956
           possible. *)
blanchet@44858
   957
        IAbs ((`I "P", p_T),
blanchet@44858
   958
              IApp (IConst (`I ho_quant, T, []),
blanchet@44858
   959
                    IAbs ((`I "X", x_T),
blanchet@44858
   960
                          IApp (IConst (`I "P", p_T, []),
blanchet@44858
   961
                                IConst (`I "X", x_T, [])))))
blanchet@44858
   962
      | tweak_ho_quant _ _ _ = raise Fail "unexpected type for quantifier"
blanchet@44858
   963
    fun intro top_level args (IApp (tm1, tm2)) =
blanchet@44858
   964
        IApp (intro top_level (tm2 :: args) tm1, intro false [] tm2)
blanchet@44858
   965
      | intro top_level args (IConst (name as (s, _), T, T_args)) =
blanchet@43441
   966
        (case proxify_const s of
blanchet@44000
   967
           SOME proxy_base =>
blanchet@44493
   968
           if top_level orelse is_type_enc_higher_order type_enc then
blanchet@43841
   969
             case (top_level, s) of
blanchet@44858
   970
               (_, "c_False") => IConst (`I tptp_false, T, [])
blanchet@44858
   971
             | (_, "c_True") => IConst (`I tptp_true, T, [])
blanchet@44858
   972
             | (false, "c_Not") => IConst (`I tptp_not, T, [])
blanchet@44858
   973
             | (false, "c_conj") => IConst (`I tptp_and, T, [])
blanchet@44858
   974
             | (false, "c_disj") => IConst (`I tptp_or, T, [])
blanchet@44858
   975
             | (false, "c_implies") => IConst (`I tptp_implies, T, [])
blanchet@44858
   976
             | (false, "c_All") => tweak_ho_quant tptp_ho_forall T args
blanchet@44858
   977
             | (false, "c_Ex") => tweak_ho_quant tptp_ho_exists T args
blanchet@43841
   978
             | (false, s) =>
blanchet@44968
   979
               if is_tptp_equal s andalso length args = 2 then
blanchet@44968
   980
                 IConst (`I tptp_equal, T, [])
blanchet@44968
   981
               else
blanchet@45453
   982
                 (* Use a proxy even for partially applied THF0 equality,
blanchet@45453
   983
                    because the LEO-II and Satallax parsers complain about not
blanchet@45453
   984
                    being able to infer the type of "=". *)
blanchet@44968
   985
                 IConst (proxy_base |>> prefix const_prefix, T, T_args)
blanchet@44858
   986
             | _ => IConst (name, T, [])
blanchet@43440
   987
           else
blanchet@44858
   988
             IConst (proxy_base |>> prefix const_prefix, T, T_args)
blanchet@46038
   989
          | NONE => if s = tptp_choice then tweak_ho_quant tptp_choice T args
blanchet@46038
   990
                    else IConst (name, T, T_args))
blanchet@44858
   991
      | intro _ _ (IAbs (bound, tm)) = IAbs (bound, intro false [] tm)
blanchet@44858
   992
      | intro _ _ tm = tm
blanchet@44858
   993
  in intro true [] end
blanchet@43439
   994
blanchet@45645
   995
fun mangle_type_args_in_iterm format type_enc =
blanchet@45645
   996
  if polymorphism_of_type_enc type_enc = Mangled_Monomorphic then
blanchet@45645
   997
    let
blanchet@45645
   998
      fun mangle (IApp (tm1, tm2)) = IApp (mangle tm1, mangle tm2)
blanchet@45645
   999
        | mangle (tm as IConst (_, _, [])) = tm
blanchet@45645
  1000
        | mangle (tm as IConst (name as (s, _), T, T_args)) =
blanchet@46382
  1001
          (case unprefix_and_unascii const_prefix s of
blanchet@45645
  1002
             NONE => tm
blanchet@45645
  1003
           | SOME s'' =>
blanchet@46816
  1004
             case type_arg_policy [] type_enc (invert_const s'') of
blanchet@45645
  1005
               Mangled_Type_Args =>
blanchet@45645
  1006
               IConst (mangled_const_name format type_enc T_args name, T, [])
blanchet@45645
  1007
             | _ => tm)
blanchet@45645
  1008
        | mangle (IAbs (bound, tm)) = IAbs (bound, mangle tm)
blanchet@45645
  1009
        | mangle tm = tm
blanchet@45645
  1010
    in mangle end
blanchet@45645
  1011
  else
blanchet@45645
  1012
    I
blanchet@45645
  1013
blanchet@45644
  1014
fun chop_fun 0 T = ([], T)
blanchet@45644
  1015
  | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
blanchet@45644
  1016
    chop_fun (n - 1) ran_T |>> cons dom_T
blanchet@45644
  1017
  | chop_fun _ T = ([], T)
blanchet@45644
  1018
blanchet@45645
  1019
fun filter_const_type_args _ _ _ [] = []
blanchet@45645
  1020
  | filter_const_type_args thy s ary T_args =
blanchet@45644
  1021
    let
blanchet@45644
  1022
      val U = robust_const_type thy s
blanchet@45644
  1023
      val arg_U_vars = fold Term.add_tvarsT (U |> chop_fun ary |> fst) []
blanchet@45644
  1024
      val U_args = (s, U) |> robust_const_typargs thy
blanchet@45644
  1025
    in
blanchet@45644
  1026
      U_args ~~ T_args
blanchet@45644
  1027
      |> map (fn (U, T) =>
blanchet@45644
  1028
                 if member (op =) arg_U_vars (dest_TVar U) then dummyT else T)
blanchet@45644
  1029
    end
blanchet@45644
  1030
    handle TYPE _ => T_args
blanchet@45644
  1031
blanchet@46816
  1032
fun filter_type_args_in_iterm thy monom_constrs type_enc =
blanchet@38506
  1033
  let
blanchet@45645
  1034
    fun filt ary (IApp (tm1, tm2)) = IApp (filt (ary + 1) tm1, filt 0 tm2)
blanchet@45645
  1035
      | filt _ (tm as IConst (_, _, [])) = tm
blanchet@45645
  1036
      | filt ary (IConst (name as (s, _), T, T_args)) =
blanchet@46382
  1037
        (case unprefix_and_unascii const_prefix s of
blanchet@45644
  1038
           NONE =>
blanchet@45645
  1039
           (name,
blanchet@45645
  1040
            if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then
blanchet@45645
  1041
              []
blanchet@45645
  1042
            else
blanchet@45645
  1043
              T_args)
blanchet@45644
  1044
         | SOME s'' =>
blanchet@45644
  1045
           let
blanchet@45644
  1046
             val s'' = invert_const s''
blanchet@45644
  1047
             fun filter_T_args false = T_args
blanchet@45645
  1048
               | filter_T_args true = filter_const_type_args thy s'' ary T_args
blanchet@45644
  1049
           in
blanchet@46816
  1050
             case type_arg_policy monom_constrs type_enc s'' of
blanchet@46808
  1051
               Explicit_Type_Args infer_from_term_args =>
blanchet@46808
  1052
               (name, filter_T_args infer_from_term_args)
blanchet@45644
  1053
             | No_Type_Args => (name, [])
blanchet@45645
  1054
             | Mangled_Type_Args => raise Fail "unexpected (un)mangled symbol"
blanchet@45644
  1055
           end)
blanchet@45644
  1056
        |> (fn (name, T_args) => IConst (name, T, T_args))
blanchet@45645
  1057
      | filt _ (IAbs (bound, tm)) = IAbs (bound, filt 0 tm)
blanchet@45645
  1058
      | filt _ tm = tm
blanchet@45645
  1059
  in filt 0 end
blanchet@45644
  1060
blanchet@45644
  1061
fun iformula_from_prop ctxt format type_enc eq_as_iff =
blanchet@45644
  1062
  let
blanchet@45644
  1063
    val thy = Proof_Context.theory_of ctxt
blanchet@46187
  1064
    fun do_term bs t atomic_Ts =
nik@45350
  1065
      iterm_from_term thy format bs (Envir.eta_contract t)
blanchet@45644
  1066
      |>> (introduce_proxies_in_iterm type_enc
blanchet@45645
  1067
           #> mangle_type_args_in_iterm format type_enc
blanchet@45644
  1068
           #> AAtom)
blanchet@46187
  1069
      ||> union (op =) atomic_Ts
blanchet@45262
  1070
    fun do_quant bs q pos s T t' =
blanchet@45262
  1071
      let
blanchet@45262
  1072
        val s = singleton (Name.variant_list (map fst bs)) s
blanchet@45262
  1073
        val universal = Option.map (q = AExists ? not) pos
blanchet@45262
  1074
        val name =
blanchet@45262
  1075
          s |> `(case universal of
blanchet@45262
  1076
                   SOME true => make_all_bound_var
blanchet@45262
  1077
                 | SOME false => make_exist_bound_var
blanchet@45262
  1078
                 | NONE => make_bound_var)
blanchet@45262
  1079
      in
blanchet@45262
  1080
        do_formula ((s, (name, T)) :: bs) pos t'
blanchet@45262
  1081
        #>> mk_aquant q [(name, SOME T)]
blanchet@46187
  1082
        ##> union (op =) (atomic_types_of T)
blanchet@38743
  1083
      end
blanchet@45262
  1084
    and do_conn bs c pos1 t1 pos2 t2 =
blanchet@45262
  1085
      do_formula bs pos1 t1 ##>> do_formula bs pos2 t2 #>> uncurry (mk_aconn c)
blanchet@45262
  1086
    and do_formula bs pos t =
blanchet@38506
  1087
      case t of
blanchet@45262
  1088
        @{const Trueprop} $ t1 => do_formula bs pos t1
blanchet@45262
  1089
      | @{const Not} $ t1 => do_formula bs (Option.map not pos) t1 #>> mk_anot
blanchet@38506
  1090
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@45262
  1091
        do_quant bs AForall pos s T t'
blanchet@46038
  1092
      | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@46038
  1093
        do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
blanchet@38506
  1094
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@45262
  1095
        do_quant bs AExists pos s T t'
blanchet@46038
  1096
      | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@46038
  1097
        do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
blanchet@45262
  1098
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd pos t1 pos t2
blanchet@45262
  1099
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr pos t1 pos t2
blanchet@45262
  1100
      | @{const HOL.implies} $ t1 $ t2 =>
blanchet@45262
  1101
        do_conn bs AImplies (Option.map not pos) t1 pos t2
haftmann@39093
  1102
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@45262
  1103
        if eq_as_iff then do_conn bs AIff NONE t1 NONE t2 else do_term bs t
blanchet@41388
  1104
      | _ => do_term bs t
blanchet@38506
  1105
  in do_formula [] end
blanchet@38506
  1106
blanchet@46964
  1107
fun presimplify_term ctxt t =
blanchet@46964
  1108
  t |> exists_Const (member (op =) Meson.presimplified_consts o fst) t
blanchet@46964
  1109
       ? (Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
blanchet@46964
  1110
          #> Meson.presimplify
blanchet@46964
  1111
          #> prop_of)
blanchet@38506
  1112
blanchet@44734
  1113
fun concealed_bound_name j = atp_weak_prefix ^ string_of_int j
blanchet@38506
  1114
fun conceal_bounds Ts t =
blanchet@38506
  1115
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@38506
  1116
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@38506
  1117
fun reveal_bounds Ts =
blanchet@38506
  1118
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@38506
  1119
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@38506
  1120
blanchet@44106
  1121
fun is_fun_equality (@{const_name HOL.eq},
blanchet@44106
  1122
                     Type (_, [Type (@{type_name fun}, _), _])) = true
blanchet@44106
  1123
  | is_fun_equality _ = false
blanchet@44106
  1124
blanchet@43612
  1125
fun extensionalize_term ctxt t =
blanchet@44106
  1126
  if exists_Const is_fun_equality t then
blanchet@44106
  1127
    let val thy = Proof_Context.theory_of ctxt in
blanchet@44106
  1128
      t |> cterm_of thy |> Meson.extensionalize_conv ctxt
blanchet@44106
  1129
        |> prop_of |> Logic.dest_equals |> snd
blanchet@44106
  1130
    end
blanchet@44106
  1131
  else
blanchet@44106
  1132
    t
blanchet@38831
  1133
blanchet@44733
  1134
fun simple_translate_lambdas do_lambdas ctxt t =
blanchet@44734
  1135
  let val thy = Proof_Context.theory_of ctxt in
blanchet@44734
  1136
    if Meson.is_fol_term thy t then
blanchet@44734
  1137
      t
blanchet@44734
  1138
    else
blanchet@44734
  1139
      let
blanchet@45676
  1140
        fun trans Ts t =
blanchet@44734
  1141
          case t of
blanchet@45676
  1142
            @{const Not} $ t1 => @{const Not} $ trans Ts t1
blanchet@44734
  1143
          | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@45676
  1144
            t0 $ Abs (s, T, trans (T :: Ts) t')
blanchet@44734
  1145
          | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@45676
  1146
            trans Ts (t0 $ eta_expand Ts t1 1)
blanchet@44734
  1147
          | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@45676
  1148
            t0 $ Abs (s, T, trans (T :: Ts) t')
blanchet@44734
  1149
          | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@45676
  1150
            trans Ts (t0 $ eta_expand Ts t1 1)
blanchet@45676
  1151
          | (t0 as @{const HOL.conj}) $ t1 $ t2 =>
blanchet@45676
  1152
            t0 $ trans Ts t1 $ trans Ts t2
blanchet@45676
  1153
          | (t0 as @{const HOL.disj}) $ t1 $ t2 =>
blanchet@45676
  1154
            t0 $ trans Ts t1 $ trans Ts t2
blanchet@45676
  1155
          | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
blanchet@45676
  1156
            t0 $ trans Ts t1 $ trans Ts t2
blanchet@44734
  1157
          | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
blanchet@44734
  1158
              $ t1 $ t2 =>
blanchet@45676
  1159
            t0 $ trans Ts t1 $ trans Ts t2
blanchet@44734
  1160
          | _ =>
blanchet@44734
  1161
            if not (exists_subterm (fn Abs _ => true | _ => false) t) then t
blanchet@44734
  1162
            else t |> Envir.eta_contract |> do_lambdas ctxt Ts
blanchet@44734
  1163
        val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
blanchet@45676
  1164
      in t |> trans [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@44734
  1165
  end
blanchet@44727
  1166
blanchet@44868
  1167
fun do_cheaply_conceal_lambdas Ts (t1 $ t2) =
blanchet@44868
  1168
    do_cheaply_conceal_lambdas Ts t1
blanchet@44868
  1169
    $ do_cheaply_conceal_lambdas Ts t2
blanchet@44868
  1170
  | do_cheaply_conceal_lambdas Ts (Abs (_, T, t)) =
blanchet@46425
  1171
    Const (lam_lifted_poly_prefix ^ serial_string (),
blanchet@46380
  1172
           T --> fastype_of1 (T :: Ts, t))
blanchet@44868
  1173
  | do_cheaply_conceal_lambdas _ t = t
blanchet@44691
  1174
blanchet@44733
  1175
fun do_introduce_combinators ctxt Ts t =
blanchet@44727
  1176
  let val thy = Proof_Context.theory_of ctxt in
blanchet@44776
  1177
    t |> conceal_bounds Ts
blanchet@44776
  1178
      |> cterm_of thy
blanchet@44776
  1179
      |> Meson_Clausify.introduce_combinators_in_cterm
blanchet@44776
  1180
      |> prop_of |> Logic.dest_equals |> snd
blanchet@44776
  1181
      |> reveal_bounds Ts
blanchet@44727
  1182
  end
blanchet@44733
  1183
  (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@44868
  1184
  handle THM _ => t |> do_cheaply_conceal_lambdas Ts
blanchet@44733
  1185
val introduce_combinators = simple_translate_lambdas do_introduce_combinators
blanchet@44727
  1186
blanchet@46385
  1187
fun preprocess_abstractions_in_terms trans_lams facts =
blanchet@44733
  1188
  let
blanchet@44734
  1189
    val (facts, lambda_ts) =
blanchet@46385
  1190
      facts |> map (snd o snd) |> trans_lams
blanchet@44734
  1191
            |>> map2 (fn (name, (kind, _)) => fn t => (name, (kind, t))) facts
blanchet@46425
  1192
    val lam_facts =
blanchet@44734
  1193
      map2 (fn t => fn j =>
blanchet@47208
  1194
               ((lam_fact_prefix ^ Int.toString j, (Global, Eq)), (Axiom, t)))
blanchet@44734
  1195
           lambda_ts (1 upto length lambda_ts)
blanchet@46425
  1196
  in (facts, lam_facts) end
blanchet@38506
  1197
blanchet@38506
  1198
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@43224
  1199
   same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
blanchet@38506
  1200
fun freeze_term t =
blanchet@38506
  1201
  let
blanchet@45676
  1202
    fun freeze (t $ u) = freeze t $ freeze u
blanchet@45676
  1203
      | freeze (Abs (s, T, t)) = Abs (s, T, freeze t)
blanchet@45676
  1204
      | freeze (Var ((s, i), T)) =
blanchet@44734
  1205
        Free (atp_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@45676
  1206
      | freeze t = t
blanchet@45676
  1207
  in t |> exists_subterm is_Var t ? freeze end
blanchet@38506
  1208
blanchet@46964
  1209
fun presimp_prop ctxt role t =
blanchet@46070
  1210
  (let
blanchet@46070
  1211
     val thy = Proof_Context.theory_of ctxt
blanchet@46070
  1212
     val t = t |> Envir.beta_eta_contract
blanchet@46070
  1213
               |> transform_elim_prop
blanchet@46070
  1214
               |> Object_Logic.atomize_term thy
blanchet@46070
  1215
     val need_trueprop = (fastype_of t = @{typ bool})
blanchet@46070
  1216
   in
blanchet@46070
  1217
     t |> need_trueprop ? HOLogic.mk_Trueprop
blanchet@46070
  1218
       |> extensionalize_term ctxt
blanchet@46964
  1219
       |> presimplify_term ctxt
blanchet@46070
  1220
       |> HOLogic.dest_Trueprop
blanchet@46070
  1221
   end
blanchet@46070
  1222
   handle TERM _ => if role = Conjecture then @{term False} else @{term True})
blanchet@46070
  1223
  |> pair role
blanchet@43937
  1224
blanchet@47168
  1225
fun make_formula ctxt format type_enc eq_as_iff name stature kind t =
blanchet@43937
  1226
  let
blanchet@46187
  1227
    val (iformula, atomic_Ts) =
blanchet@45644
  1228
      iformula_from_prop ctxt format type_enc eq_as_iff
blanchet@45644
  1229
                         (SOME (kind <> Conjecture)) t []
blanchet@46187
  1230
      |>> close_iformula_universally
blanchet@38506
  1231
  in
blanchet@47168
  1232
    {name = name, stature = stature, kind = kind, iformula = iformula,
blanchet@46187
  1233
     atomic_types = atomic_Ts}
blanchet@38506
  1234
  end
blanchet@38506
  1235
blanchet@47168
  1236
fun make_fact ctxt format type_enc eq_as_iff ((name, stature), t) =
blanchet@45644
  1237
  case t |> make_formula ctxt format type_enc (eq_as_iff andalso format <> CNF)
blanchet@47168
  1238
                         name stature Axiom of
blanchet@45644
  1239
    formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} =>
blanchet@45644
  1240
    if s = tptp_true then NONE else SOME formula
blanchet@45644
  1241
  | formula => SOME formula
blanchet@43432
  1242
blanchet@47213
  1243
fun s_not_trueprop (@{const Trueprop} $ t) = @{const Trueprop} $ s_not t
blanchet@47213
  1244
  | s_not_trueprop t =
blanchet@47213
  1245
    if fastype_of t = @{typ bool} then s_not t
blanchet@47205
  1246
    else @{prop False} (* "t" is too meta *)
blanchet@45317
  1247
blanchet@45644
  1248
fun make_conjecture ctxt format type_enc =
blanchet@47168
  1249
  map (fn ((name, stature), (kind, t)) =>
blanchet@47213
  1250
          t |> kind = Conjecture ? s_not_trueprop
blanchet@47168
  1251
            |> make_formula ctxt format type_enc (format <> CNF) name stature
blanchet@47168
  1252
                            kind)
blanchet@38506
  1253
blanchet@43552
  1254
(** Finite and infinite type inference **)
blanchet@43552
  1255
blanchet@45676
  1256
fun tvar_footprint thy s ary =
blanchet@46382
  1257
  (case unprefix_and_unascii const_prefix s of
blanchet@45676
  1258
     SOME s =>
blanchet@45676
  1259
     s |> invert_const |> robust_const_type thy |> chop_fun ary |> fst
blanchet@45676
  1260
       |> map (fn T => Term.add_tvarsT T [] |> map fst)
blanchet@45676
  1261
   | NONE => [])
blanchet@45676
  1262
  handle TYPE _ => []
blanchet@45676
  1263
blanchet@45676
  1264
fun ghost_type_args thy s ary =
blanchet@46819
  1265
  if is_tptp_equal s then
blanchet@46819
  1266
    0 upto ary - 1
blanchet@46819
  1267
  else
blanchet@46819
  1268
    let
blanchet@46819
  1269
      val footprint = tvar_footprint thy s ary
blanchet@46819
  1270
      val eq = (s = @{const_name HOL.eq})
blanchet@46819
  1271
      fun ghosts _ [] = []
blanchet@46819
  1272
        | ghosts seen ((i, tvars) :: args) =
blanchet@46819
  1273
          ghosts (union (op =) seen tvars) args
blanchet@46819
  1274
          |> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars)
blanchet@46819
  1275
             ? cons i
blanchet@46819
  1276
    in
blanchet@46819
  1277
      if forall null footprint then
blanchet@46819
  1278
        []
blanchet@46819
  1279
      else
blanchet@46819
  1280
        0 upto length footprint - 1 ~~ footprint
blanchet@46819
  1281
        |> sort (rev_order o list_ord Term_Ord.indexname_ord o pairself snd)
blanchet@46819
  1282
        |> ghosts []
blanchet@46819
  1283
    end
blanchet@45676
  1284
blanchet@45258
  1285
type monotonicity_info =
blanchet@45258
  1286
  {maybe_finite_Ts : typ list,
blanchet@45258
  1287
   surely_finite_Ts : typ list,
blanchet@45258
  1288
   maybe_infinite_Ts : typ list,
blanchet@45258
  1289
   surely_infinite_Ts : typ list,
blanchet@45258
  1290
   maybe_nonmono_Ts : typ list}
blanchet@45258
  1291
blanchet@45256
  1292
(* These types witness that the type classes they belong to allow infinite
blanchet@45256
  1293
   models and hence that any types with these type classes is monotonic. *)
blanchet@45256
  1294
val known_infinite_types =
blanchet@45492
  1295
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}]
blanchet@45256
  1296
blanchet@47129
  1297
fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T =
blanchet@47129
  1298
  strictness <> Strict andalso is_type_surely_infinite ctxt true cached_Ts T
blanchet@43755
  1299
blanchet@43552
  1300
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
blanchet@43552
  1301
   dangerous because their "exhaust" properties can easily lead to unsound ATP
blanchet@43552
  1302
   proofs. On the other hand, all HOL infinite types can be given the same
blanchet@43552
  1303
   models in first-order logic (via Löwenheim-Skolem). *)
blanchet@43552
  1304
blanchet@45258
  1305
fun should_encode_type _ (_ : monotonicity_info) All_Types _ = true
blanchet@45258
  1306
  | should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts,
blanchet@45258
  1307
                             maybe_nonmono_Ts, ...}
blanchet@47129
  1308
                       (Noninf_Nonmono_Types (strictness, grain)) T =
blanchet@45673
  1309
    grain = Ghost_Type_Arg_Vars orelse
blanchet@45673
  1310
    (exists (type_intersect ctxt T) maybe_nonmono_Ts andalso
blanchet@45673
  1311
     not (exists (type_instance ctxt T) surely_infinite_Ts orelse
blanchet@45730
  1312
          (not (member (type_equiv ctxt) maybe_finite_Ts T) andalso
blanchet@47129
  1313
           is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts
blanchet@45673
  1314
                                           T)))
blanchet@45258
  1315
  | should_encode_type ctxt {surely_finite_Ts, maybe_infinite_Ts,
blanchet@45258
  1316
                             maybe_nonmono_Ts, ...}
blanchet@45673
  1317
                       (Fin_Nonmono_Types grain) T =
blanchet@45673
  1318
    grain = Ghost_Type_Arg_Vars orelse
blanchet@45673
  1319
    (exists (type_intersect ctxt T) maybe_nonmono_Ts andalso
blanchet@45673
  1320
     (exists (type_generalization ctxt T) surely_finite_Ts orelse
blanchet@45730
  1321
      (not (member (type_equiv ctxt) maybe_infinite_Ts T) andalso
blanchet@45673
  1322
       is_type_surely_finite ctxt T)))
blanchet@43552
  1323
  | should_encode_type _ _ _ _ = false
blanchet@43552
  1324
blanchet@45639
  1325
fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T =
blanchet@45673
  1326
    should_guard_var () andalso should_encode_type ctxt mono level T
blanchet@45258
  1327
  | should_guard_type _ _ _ _ _ = false
blanchet@43552
  1328
blanchet@45262
  1329
fun is_maybe_universal_var (IConst ((s, _), _, _)) =
blanchet@45262
  1330
    String.isPrefix bound_var_prefix s orelse
blanchet@45262
  1331
    String.isPrefix all_bound_var_prefix s
blanchet@45262
  1332
  | is_maybe_universal_var (IVar _) = true
blanchet@45262
  1333
  | is_maybe_universal_var _ = false
blanchet@43707
  1334
blanchet@46818
  1335
datatype site =
blanchet@44232
  1336
  Top_Level of bool option |
blanchet@44232
  1337
  Eq_Arg of bool option |
blanchet@44232
  1338
  Elsewhere
blanchet@43700
  1339
blanchet@46820
  1340
fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false
blanchet@46820
  1341
  | should_tag_with_type ctxt mono (Tags (_, level)) site u T =
blanchet@46820
  1342
    if granularity_of_type_level level = All_Vars then
blanchet@46820
  1343
      should_encode_type ctxt mono level T
blanchet@46820
  1344
    else
blanchet@46820
  1345
      (case (site, is_maybe_universal_var u) of
blanchet@46820
  1346
         (Eq_Arg _, true) => should_encode_type ctxt mono level T
blanchet@45676
  1347
       | _ => false)
blanchet@46820
  1348
  | should_tag_with_type _ _ _ _ _ _ = false
blanchet@43552
  1349
blanchet@45458
  1350
fun fused_type ctxt mono level =
blanchet@43835
  1351
  let
blanchet@45258
  1352
    val should_encode = should_encode_type ctxt mono level
blanchet@45458
  1353
    fun fuse 0 T = if should_encode T then T else fused_infinite_type
blanchet@45458
  1354
      | fuse ary (Type (@{type_name fun}, [T1, T2])) =
blanchet@45458
  1355
        fuse 0 T1 --> fuse (ary - 1) T2
blanchet@45458
  1356
      | fuse _ _ = raise Fail "expected function type"
blanchet@45458
  1357
  in fuse end
blanchet@43552
  1358
blanchet@45307
  1359
(** predicators and application operators **)
blanchet@41561
  1360
blanchet@43445
  1361
type sym_info =
blanchet@45700
  1362
  {pred_sym : bool, min_ary : int, max_ary : int, types : typ list,
blanchet@45700
  1363
   in_conj : bool}
blanchet@43434
  1364
blanchet@45700
  1365
fun default_sym_tab_entries type_enc =
blanchet@45700
  1366
  (make_fixed_const NONE @{const_name undefined},
blanchet@45700
  1367
       {pred_sym = false, min_ary = 0, max_ary = 0, types = [],
blanchet@45700
  1368
        in_conj = false}) ::
blanchet@45700
  1369
  ([tptp_false, tptp_true]
blanchet@45700
  1370
   |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = [],
blanchet@45700
  1371
                  in_conj = false})) @
blanchet@45700
  1372
  ([tptp_equal, tptp_old_equal]
blanchet@45700
  1373
   |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = [],
blanchet@45700
  1374
                  in_conj = false}))
blanchet@45700
  1375
  |> not (is_type_enc_higher_order type_enc)
blanchet@45700
  1376
     ? cons (prefixed_predicator_name,
blanchet@45700
  1377
             {pred_sym = true, min_ary = 1, max_ary = 1, types = [],
blanchet@45700
  1378
              in_conj = false})
blanchet@45700
  1379
blanchet@47217
  1380
datatype explicit_apply = Incomplete_Apply | Sufficient_Apply | Full_Apply
blanchet@47217
  1381
blanchet@45700
  1382
fun sym_table_for_facts ctxt type_enc explicit_apply conjs facts =
blanchet@43429
  1383
  let
blanchet@45643
  1384
    fun consider_var_ary const_T var_T max_ary =
blanchet@43905
  1385
      let
blanchet@43905
  1386
        fun iter ary T =
blanchet@45258
  1387
          if ary = max_ary orelse type_instance ctxt var_T T orelse
blanchet@45258
  1388
             type_instance ctxt T var_T then
blanchet@44051
  1389
            ary
blanchet@44051
  1390
          else
blanchet@44051
  1391
            iter (ary + 1) (range_type T)
blanchet@43905
  1392
      in iter 0 const_T end
blanchet@45262
  1393
    fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
blanchet@47217
  1394
      if explicit_apply = Sufficient_Apply andalso
blanchet@44042
  1395
         (can dest_funT T orelse T = @{typ bool}) then
blanchet@44042
  1396
        let
blanchet@44042
  1397
          val bool_vars' = bool_vars orelse body_type T = @{typ bool}
blanchet@45700
  1398
          fun repair_min_ary {pred_sym, min_ary, max_ary, types, in_conj} =
blanchet@44042
  1399
            {pred_sym = pred_sym andalso not bool_vars',
blanchet@45643
  1400
             min_ary = fold (fn T' => consider_var_ary T' T) types min_ary,
blanchet@45700
  1401
             max_ary = max_ary, types = types, in_conj = in_conj}
blanchet@44042
  1402
          val fun_var_Ts' =
blanchet@44042
  1403
            fun_var_Ts |> can dest_funT T ? insert_type ctxt I T
blanchet@44042
  1404
        in
blanchet@44042
  1405
          if bool_vars' = bool_vars andalso
blanchet@44042
  1406
             pointer_eq (fun_var_Ts', fun_var_Ts) then
blanchet@44042
  1407
            accum
blanchet@44008
  1408
          else
blanchet@45643
  1409
            ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_ary) sym_tab)
blanchet@44042
  1410
        end
blanchet@44042
  1411
      else
blanchet@44042
  1412
        accum
blanchet@45700
  1413
    fun add_fact_syms conj_fact =
blanchet@45700
  1414
      let
blanchet@45700
  1415
        fun add_iterm_syms top_level tm
blanchet@45700
  1416
                           (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
blanchet@45700
  1417
          let val (head, args) = strip_iterm_comb tm in
blanchet@45700
  1418
            (case head of
blanchet@45700
  1419
               IConst ((s, _), T, _) =>
blanchet@45700
  1420
               if String.isPrefix bound_var_prefix s orelse
blanchet@45700
  1421
                  String.isPrefix all_bound_var_prefix s then
blanchet@45700
  1422
                 add_universal_var T accum
blanchet@45700
  1423
               else if String.isPrefix exist_bound_var_prefix s then
blanchet@45700
  1424
                 accum
blanchet@45700
  1425
               else
blanchet@45700
  1426
                 let val ary = length args in
blanchet@45700
  1427
                   ((bool_vars, fun_var_Ts),
blanchet@45700
  1428
                    case Symtab.lookup sym_tab s of
blanchet@45700
  1429
                      SOME {pred_sym, min_ary, max_ary, types, in_conj} =>
blanchet@45700
  1430
                      let
blanchet@45700
  1431
                        val pred_sym =
blanchet@45700
  1432
                          pred_sym andalso top_level andalso not bool_vars
blanchet@45700
  1433
                        val types' = types |> insert_type ctxt I T
blanchet@45700
  1434
                        val in_conj = in_conj orelse conj_fact
blanchet@45700
  1435
                        val min_ary =
blanchet@47217
  1436
                          if explicit_apply = Sufficient_Apply andalso
blanchet@47217
  1437
                             not (pointer_eq (types', types)) then
blanchet@47217
  1438
                            fold (consider_var_ary T) fun_var_Ts min_ary
blanchet@47217
  1439
                          else
blanchet@45700
  1440
                            min_ary
blanchet@45700
  1441
                      in
blanchet@45700
  1442
                        Symtab.update (s, {pred_sym = pred_sym,
blanchet@45700
  1443
                                           min_ary = Int.min (ary, min_ary),
blanchet@45700
  1444
                                           max_ary = Int.max (ary, max_ary),
blanchet@45700
  1445
                                           types = types', in_conj = in_conj})
blanchet@43905
  1446
                                      sym_tab
blanchet@45700
  1447
                      end
blanchet@45700
  1448
                    | NONE =>
blanchet@45700
  1449
                      let
blanchet@45700
  1450
                        val pred_sym = top_level andalso not bool_vars
blanchet@45700
  1451
                        val min_ary =
blanchet@45700
  1452
                          case explicit_apply of
blanchet@47217
  1453
                            Incomplete_Apply => ary
blanchet@47217
  1454
                          | Sufficient_Apply =>
blanchet@47217
  1455
                            fold (consider_var_ary T) fun_var_Ts ary
blanchet@47217
  1456
                          | Full_Apply => 0
blanchet@45700
  1457
                      in
blanchet@45700
  1458
                        Symtab.update_new (s,
blanchet@45700
  1459
                            {pred_sym = pred_sym, min_ary = min_ary,
blanchet@45700
  1460
                             max_ary = ary, types = [T], in_conj = conj_fact})
blanchet@45700
  1461
                            sym_tab
blanchet@45700
  1462
                      end)
blanchet@45700
  1463
                 end
blanchet@45700
  1464
             | IVar (_, T) => add_universal_var T accum
blanchet@45700
  1465
             | IAbs ((_, T), tm) =>
blanchet@45700
  1466
               accum |> add_universal_var T |> add_iterm_syms false tm
blanchet@45700
  1467
             | _ => accum)
blanchet@45700
  1468
            |> fold (add_iterm_syms false) args
blanchet@45700
  1469
          end
blanchet@45700
  1470
      in K (add_iterm_syms true) |> formula_fold NONE |> fact_lift end
blanchet@45700
  1471
  in
blanchet@45700
  1472
    ((false, []), Symtab.empty)
blanchet@45700
  1473
    |> fold (add_fact_syms true) conjs
blanchet@45700
  1474
    |> fold (add_fact_syms false) facts
blanchet@45700
  1475
    |> snd
blanchet@45700
  1476
    |> fold Symtab.update (default_sym_tab_entries type_enc)
blanchet@45700
  1477
  end
blanchet@38506
  1478
blanchet@45643
  1479
fun min_ary_of sym_tab s =
blanchet@43429
  1480
  case Symtab.lookup sym_tab s of
blanchet@43445
  1481
    SOME ({min_ary, ...} : sym_info) => min_ary
blanchet@43429
  1482
  | NONE =>
blanchet@46382
  1483
    case unprefix_and_unascii const_prefix s of
blanchet@43418
  1484
      SOME s =>
blanchet@43441
  1485
      let val s = s |> unmangled_const_name |> invert_const in
blanchet@43807
  1486
        if s = predicator_name then 1
blanchet@43807
  1487
        else if s = app_op_name then 2
blanchet@45255
  1488
        else if s = type_guard_name then 1
blanchet@43428
  1489
        else 0
blanchet@43418
  1490
      end
blanchet@38506
  1491
    | NONE => 0
blanchet@38506
  1492
blanchet@38506
  1493
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
  1494
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
  1495
   type instantiations). If false, the constant always receives all of its
blanchet@38506
  1496
   arguments and is used as a predicate. *)
blanchet@43429
  1497
fun is_pred_sym sym_tab s =
blanchet@43429
  1498
  case Symtab.lookup sym_tab s of
blanchet@43445
  1499
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
blanchet@43445
  1500
    pred_sym andalso min_ary = max_ary
blanchet@43429
  1501
  | NONE => false
blanchet@38506
  1502
blanchet@45644
  1503
val app_op = `(make_fixed_const NONE) app_op_name
blanchet@43439
  1504
val predicator_combconst =
nik@45350
  1505
  IConst (`(make_fixed_const NONE) predicator_name, @{typ "bool => bool"}, [])
blanchet@45644
  1506
blanchet@45644
  1507
fun list_app head args = fold (curry (IApp o swap)) args head
blanchet@44730
  1508
fun predicator tm = IApp (predicator_combconst, tm)
blanchet@38506
  1509
blanchet@46816
  1510
fun firstorderize_fact thy monom_constrs format type_enc sym_tab =
blanchet@43415
  1511
  let
blanchet@45644
  1512
    fun do_app arg head =
blanchet@45644
  1513
      let
blanchet@45644
  1514
        val head_T = ityp_of head
blanchet@45644
  1515
        val (arg_T, res_T) = dest_funT head_T
blanchet@45644
  1516
        val app =
blanchet@45644
  1517
          IConst (app_op, head_T --> head_T, [arg_T, res_T])
blanchet@45645
  1518
          |> mangle_type_args_in_iterm format type_enc
blanchet@45644
  1519
      in list_app app [head, arg] end
blanchet@45644
  1520
    fun list_app_ops head args = fold do_app args head
blanchet@45644
  1521
    fun introduce_app_ops tm =
blanchet@44730
  1522
      case strip_iterm_comb tm of
blanchet@44730
  1523
        (head as IConst ((s, _), _, _), args) =>
blanchet@45644
  1524
        args |> map introduce_app_ops
blanchet@45643
  1525
             |> chop (min_ary_of sym_tab s)
blanchet@43415
  1526
             |>> list_app head
blanchet@45644
  1527
             |-> list_app_ops
blanchet@45644
  1528
      | (head, args) => list_app_ops head (map introduce_app_ops args)
blanchet@45644
  1529
    fun introduce_predicators tm =
blanchet@45644
  1530
      case strip_iterm_comb tm of
blanchet@45644
  1531
        (IConst ((s, _), _, _), _) =>
blanchet@45644
  1532
        if is_pred_sym sym_tab s then tm else predicator tm
blanchet@45644
  1533
      | _ => predicator tm
blanchet@45644
  1534
    val do_iterm =
blanchet@45644
  1535
      not (is_type_enc_higher_order type_enc)
blanchet@45644
  1536
      ? (introduce_app_ops #> introduce_predicators)
blanchet@46816
  1537
      #> filter_type_args_in_iterm thy monom_constrs type_enc
blanchet@45644
  1538
  in update_iformula (formula_map do_iterm) end
blanchet@43444
  1539
blanchet@43444
  1540
(** Helper facts **)
blanchet@43444
  1541
blanchet@45307
  1542
val not_ffalse = @{lemma "~ fFalse" by (unfold fFalse_def) fast}
blanchet@45307
  1543
val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
blanchet@45307
  1544
blanchet@44035
  1545
(* The Boolean indicates that a fairly sound type encoding is needed. *)
blanchet@43926
  1546
val helper_table =
blanchet@44035
  1547
  [(("COMBI", false), @{thms Meson.COMBI_def}),
blanchet@44035
  1548
   (("COMBK", false), @{thms Meson.COMBK_def}),
blanchet@44035
  1549
   (("COMBB", false), @{thms Meson.COMBB_def}),
blanchet@44035
  1550
   (("COMBC", false), @{thms Meson.COMBC_def}),
blanchet@44035
  1551
   (("COMBS", false), @{thms Meson.COMBS_def}),
blanchet@45307
  1552
   ((predicator_name, false), [not_ffalse, ftrue]),
blanchet@45307
  1553
   (("fFalse", false), [not_ffalse]),
blanchet@44035
  1554
   (("fFalse", true), @{thms True_or_False}),
blanchet@45307
  1555
   (("fTrue", false), [ftrue]),
blanchet@44035
  1556
   (("fTrue", true), @{thms True_or_False}),
blanchet@44035
  1557
   (("fNot", false),
blanchet@44035
  1558
    @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
blanchet@44035
  1559
           fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
blanchet@44035
  1560
   (("fconj", false),
blanchet@44035
  1561
    @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
blanchet@44035
  1562
        by (unfold fconj_def) fast+}),
blanchet@44035
  1563
   (("fdisj", false),
blanchet@44035
  1564
    @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
blanchet@44035
  1565
        by (unfold fdisj_def) fast+}),
blanchet@44035
  1566
   (("fimplies", false),
blanchet@44051
  1567
    @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
blanchet@44035
  1568
        by (unfold fimplies_def) fast+}),
nik@44537
  1569
   (("fequal", true),
nik@44537
  1570
    (* This is a lie: Higher-order equality doesn't need a sound type encoding.
nik@44537
  1571
       However, this is done so for backward compatibility: Including the
nik@44537
  1572
       equality helpers by default in Metis breaks a few existing proofs. *)
nik@44537
  1573
    @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
nik@44537
  1574
           fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}),
blanchet@44874
  1575
   (* Partial characterization of "fAll" and "fEx". A complete characterization
blanchet@44874
  1576
      would require the axiom of choice for replay with Metis. *)
blanchet@44874
  1577
   (("fAll", false), [@{lemma "~ fAll P | P x" by (auto simp: fAll_def)}]),
blanchet@44874
  1578
   (("fEx", false), [@{lemma "~ P x | fEx P" by (auto simp: fEx_def)}]),
blanchet@44035
  1579
   (("If", true), @{thms if_True if_False True_or_False})]
blanchet@44035
  1580
  |> map (apsnd (map zero_var_indexes))
blanchet@43926
  1581
blanchet@46225
  1582
fun atype_of_type_vars (Simple_Types (_, Polymorphic, _)) = SOME atype_of_types
blanchet@46225
  1583
  | atype_of_type_vars _ = NONE
blanchet@46225
  1584
blanchet@46791
  1585
fun bound_tvars type_enc sorts Ts =
blanchet@46791
  1586
  (sorts ? mk_ahorn (formulas_for_types type_enc add_sorts_on_tvar Ts))
blanchet@46187
  1587
  #> mk_aquant AForall
blanchet@46187
  1588
        (map_filter (fn TVar (x as (s, _), _) =>
blanchet@46187
  1589
                        SOME ((make_schematic_type_var x, s),
blanchet@46225
  1590
                              atype_of_type_vars type_enc)
blanchet@46187
  1591
                      | _ => NONE) Ts)
blanchet@45263
  1592
blanchet@45263
  1593
fun eq_formula type_enc atomic_Ts pred_sym tm1 tm2 =
blanchet@45263
  1594
  (if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2])
blanchet@45263
  1595
   else AAtom (ATerm (`I tptp_equal, [tm1, tm2])))
blanchet@46248
  1596
  |> close_formula_universally
blanchet@46791
  1597
  |> bound_tvars type_enc true atomic_Ts
blanchet@45263
  1598
nik@45350
  1599
val type_tag = `(make_fixed_const NONE) type_tag_name
blanchet@43971
  1600
blanchet@46172
  1601
fun type_tag_idempotence_fact format type_enc =
blanchet@43444
  1602
  let
blanchet@43444
  1603
    fun var s = ATerm (`I s, [])
blanchet@45267
  1604
    fun tag tm = ATerm (type_tag, [var "A", tm])
blanchet@45267
  1605
    val tagged_var = tag (var "X")
blanchet@43444
  1606
  in
blanchet@44000
  1607
    Formula (type_tag_idempotence_helper_name, Axiom,
blanchet@45264
  1608
             eq_formula type_enc [] false (tag tagged_var) tagged_var,
blanchet@47212
  1609
             NONE, isabelle_info format eqN)
blanchet@43444
  1610
  end
blanchet@43444
  1611
blanchet@44493
  1612
fun should_specialize_helper type_enc t =
blanchet@45348
  1613
  polymorphism_of_type_enc type_enc <> Polymorphic andalso
blanchet@44495
  1614
  level_of_type_enc type_enc <> No_Types andalso
blanchet@44495
  1615
  not (null (Term.hidden_polymorphism t))
blanchet@44000
  1616
blanchet@44729
  1617
fun helper_facts_for_sym ctxt format type_enc (s, {types, ...} : sym_info) =
blanchet@46382
  1618
  case unprefix_and_unascii const_prefix s of
blanchet@43444
  1619
    SOME mangled_s =>
blanchet@43444
  1620
    let
blanchet@43444
  1621
      val thy = Proof_Context.theory_of ctxt
blanchet@43444
  1622
      val unmangled_s = mangled_s |> unmangled_const_name
blanchet@44495
  1623
      fun dub needs_fairly_sound j k =
blanchet@47167
  1624
        unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
blanchet@47167
  1625
        (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
blanchet@47167
  1626
        (if needs_fairly_sound then typed_helper_suffix
blanchet@47167
  1627
         else untyped_helper_suffix)
blanchet@44000
  1628
      fun dub_and_inst needs_fairly_sound (th, j) =
blanchet@44495
  1629
        let val t = prop_of th in
blanchet@44495
  1630
          if should_specialize_helper type_enc t then
blanchet@44495
  1631
            map (fn T => specialize_type thy (invert_const unmangled_s, T) t)
blanchet@44495
  1632
                types
blanchet@44495
  1633
          else
blanchet@44495
  1634
            [t]
blanchet@44495
  1635
        end
blanchet@47167
  1636
        |> tag_list 1
blanchet@47169
  1637
        |> map (fn (k, t) => ((dub needs_fairly_sound j k, (Global, Eq)), t))
blanchet@44731
  1638
      val make_facts = map_filter (make_fact ctxt format type_enc false)
blanchet@44493
  1639
      val fairly_sound = is_type_enc_fairly_sound type_enc
blanchet@43444
  1640
    in
blanchet@43926
  1641
      helper_table
blanchet@44035
  1642
      |> maps (fn ((helper_s, needs_fairly_sound), ths) =>
blanchet@44000
  1643
                  if helper_s <> unmangled_s orelse
blanchet@43765
  1644
                     (needs_fairly_sound andalso not fairly_sound) then
blanchet@43444
  1645
                    []
blanchet@43444
  1646
                  else
blanchet@43444
  1647
                    ths ~~ (1 upto length ths)
blanchet@44495
  1648
                    |> maps (dub_and_inst needs_fairly_sound)
blanchet@44000
  1649
                    |> make_facts)
blanchet@43444
  1650
    end
blanchet@43444
  1651
  | NONE => []
blanchet@44729
  1652
fun helper_facts_for_sym_table ctxt format type_enc sym_tab =
blanchet@44729
  1653
  Symtab.fold_rev (append o helper_facts_for_sym ctxt format type_enc) sym_tab
blanchet@44729
  1654
                  []
blanchet@43444
  1655
blanchet@43926
  1656
(***************************************************************)
blanchet@43926
  1657
(* Type Classes Present in the Axiom or Conjecture Clauses     *)
blanchet@43926
  1658
(***************************************************************)
blanchet@43926
  1659
blanchet@43926
  1660
fun set_insert (x, s) = Symtab.update (x, ()) s
blanchet@43926
  1661
blanchet@43926
  1662
fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
blanchet@43926
  1663
blanchet@43926
  1664
(* Remove this trivial type class (FIXME: similar code elsewhere) *)
blanchet@43926
  1665
fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset
blanchet@43926
  1666
blanchet@43934
  1667
fun classes_of_terms get_Ts =
blanchet@43962
  1668
  map (map snd o get_Ts)
blanchet@43934
  1669
  #> List.foldl add_classes Symtab.empty
blanchet@43934
  1670
  #> delete_type #> Symtab.keys
blanchet@43926
  1671
wenzelm@45004
  1672
val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees
wenzelm@45004
  1673
val tvar_classes_of_terms = classes_of_terms Misc_Legacy.term_tvars
blanchet@43926
  1674
blanchet@44489
  1675
fun fold_type_constrs f (Type (s, Ts)) x =
blanchet@44489
  1676
    fold (fold_type_constrs f) Ts (f (s, x))
blanchet@44030
  1677
  | fold_type_constrs _ _ x = x
blanchet@43926
  1678
blanchet@44778
  1679
(* Type constructors used to instantiate overloaded constants are the only ones
blanchet@44778
  1680
   needed. *)
blanchet@44030
  1681
fun add_type_constrs_in_term thy =
blanchet@43926
  1682
  let
blanchet@44029
  1683
    fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
blanchet@44022
  1684
      | add (t $ u) = add t #> add u
blanchet@45602
  1685
      | add (Const x) =
blanchet@45602
  1686
        x |> robust_const_typargs thy |> fold (fold_type_constrs set_insert)
blanchet@44022
  1687
      | add (Abs (_, _, u)) = add u
blanchet@44022
  1688
      | add _ = I
blanchet@44022
  1689
  in add end
blanchet@43926
  1690
blanchet@44030
  1691
fun type_constrs_of_terms thy ts =
blanchet@44030
  1692
  Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
blanchet@43926
  1693
blanchet@46382
  1694
fun extract_lambda_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
blanchet@46382
  1695
    let val (head, args) = strip_comb t in
blanchet@46382
  1696
      (head |> dest_Const |> fst,
blanchet@46382
  1697
       fold_rev (fn t as Var ((s, _), T) =>
blanchet@46382
  1698
                    (fn u => Abs (s, T, abstract_over (t, u)))
blanchet@46382
  1699
                  | _ => raise Fail "expected Var") args u)
blanchet@46382
  1700
    end
blanchet@46382
  1701
  | extract_lambda_def _ = raise Fail "malformed lifted lambda"
blanchet@46379
  1702
blanchet@46385
  1703
fun trans_lams_from_string ctxt type_enc lam_trans =
blanchet@46385
  1704
  if lam_trans = no_lamsN then
blanchet@46385
  1705
    rpair []
blanchet@46385
  1706
  else if lam_trans = hide_lamsN then
blanchet@46385
  1707
    lift_lams ctxt type_enc ##> K []
blanchet@47193
  1708
  else if lam_trans = liftingN orelse lam_trans = lam_liftingN then
blanchet@46385
  1709
    lift_lams ctxt type_enc
blanchet@47193
  1710
  else if lam_trans = combsN then
blanchet@46385
  1711
    map (introduce_combinators ctxt) #> rpair []
blanchet@47193
  1712
  else if lam_trans = combs_and_liftingN then
blanchet@47193
  1713
    lift_lams_part_1 ctxt type_enc
blanchet@47193
  1714
    ##> maps (fn t => [t, introduce_combinators ctxt (intentionalize_def t)])
blanchet@47193
  1715
    #> lift_lams_part_2
blanchet@47196
  1716
  else if lam_trans = combs_or_liftingN then
blanchet@46425
  1717
    lift_lams_part_1 ctxt type_enc
blanchet@47196
  1718
    ##> map (fn t => case head_of (strip_qnt_body @{const_name All} t) of
blanchet@47196
  1719
                       @{term "op =::bool => bool => bool"} => t
blanchet@47196
  1720
                     | _ => introduce_combinators ctxt (intentionalize_def t))
blanchet@46425
  1721
    #> lift_lams_part_2
blanchet@46385
  1722
  else if lam_trans = keep_lamsN then
blanchet@46385
  1723
    map (Envir.eta_contract) #> rpair []
blanchet@46385
  1724
  else
blanchet@46390
  1725
    error ("Unknown lambda translation scheme: " ^ quote lam_trans ^ ".")
blanchet@46385
  1726
blanchet@46385
  1727
fun translate_formulas ctxt format prem_kind type_enc lam_trans presimp hyp_ts
blanchet@46385
  1728
                       concl_t facts =
blanchet@43444
  1729
  let
blanchet@43444
  1730
    val thy = Proof_Context.theory_of ctxt
blanchet@46385
  1731
    val trans_lams = trans_lams_from_string ctxt type_enc lam_trans
blanchet@44732
  1732
    val fact_ts = facts |> map snd
blanchet@43444
  1733
    (* Remove existing facts from the conjecture, as this can dramatically
blanchet@43444
  1734
       boost an ATP's performance (for some reason). *)
blanchet@44033
  1735
    val hyp_ts =
blanchet@44033
  1736
      hyp_ts
blanchet@44033
  1737
      |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
blanchet@44735
  1738
    val facts = facts |> map (apsnd (pair Axiom))
blanchet@44735
  1739
    val conjs =
blanchet@47213
  1740
      map (pair prem_kind) hyp_ts @ [(Conjecture, s_not_trueprop concl_t)]
blanchet@46039
  1741
      |> map (apsnd freeze_term)
blanchet@47168
  1742
      |> map2 (pair o rpair (Local, General) o string_of_int)
blanchet@47168
  1743
              (0 upto length hyp_ts)
blanchet@46425
  1744
    val ((conjs, facts), lam_facts) =
blanchet@46382
  1745
      (conjs, facts)
blanchet@46964
  1746
      |> presimp ? pairself (map (apsnd (uncurry (presimp_prop ctxt))))
blanchet@46385
  1747
      |> (if lam_trans = no_lamsN then
blanchet@46382
  1748
            rpair []
blanchet@46382
  1749
          else
blanchet@46382
  1750
            op @
blanchet@46385
  1751
            #> preprocess_abstractions_in_terms trans_lams
blanchet@46382
  1752
            #>> chop (length conjs))
blanchet@45644
  1753
    val conjs = conjs |> make_conjecture ctxt format type_enc
blanchet@44734
  1754
    val (fact_names, facts) =
blanchet@44735
  1755
      facts
blanchet@44734
  1756
      |> map_filter (fn (name, (_, t)) =>
blanchet@44734
  1757
                        make_fact ctxt format type_enc true (name, t)
blanchet@44734
  1758
                        |> Option.map (pair name))
blanchet@44732
  1759
      |> ListPair.unzip
blanchet@47203
  1760
    val lifted = lam_facts |> map (extract_lambda_def o snd o snd)
blanchet@46425
  1761
    val lam_facts =
blanchet@46425
  1762
      lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd)
blanchet@44732
  1763
    val all_ts = concl_t :: hyp_ts @ fact_ts
blanchet@43444
  1764
    val subs = tfree_classes_of_terms all_ts
blanchet@43444
  1765
    val supers = tvar_classes_of_terms all_ts
blanchet@44030
  1766
    val tycons = type_constrs_of_terms thy all_ts
blanchet@44732
  1767
    val (supers, arity_clauses) =
blanchet@44493
  1768
      if level_of_type_enc type_enc = No_Types then ([], [])
blanchet@43444
  1769
      else make_arity_clauses thy tycons supers
blanchet@44732
  1770
    val class_rel_clauses = make_class_rel_clauses thy subs supers
blanchet@43444
  1771
  in
blanchet@46379
  1772
    (fact_names |> map single, union (op =) subs supers, conjs,
blanchet@46425
  1773
     facts @ lam_facts, class_rel_clauses, arity_clauses, lifted)
blanchet@43444
  1774
  end
blanchet@43444
  1775
nik@45350
  1776
val type_guard = `(make_fixed_const NONE) type_guard_name
blanchet@43971
  1777
blanchet@45645
  1778
fun type_guard_iterm format type_enc T tm =
blanchet@45255
  1779
  IApp (IConst (type_guard, T --> @{typ bool}, [T])
blanchet@45645
  1780
        |> mangle_type_args_in_iterm format type_enc, tm)
blanchet@43444
  1781
blanchet@44282
  1782
fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
blanchet@44282
  1783
  | is_var_positively_naked_in_term name _ (ATerm ((s, _), tms)) accum =
blanchet@43841
  1784
    accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm (name, [])))
blanchet@44558
  1785
  | is_var_positively_naked_in_term _ _ _ _ = true
blanchet@45262
  1786
blanchet@46819
  1787
fun is_var_ghost_type_arg_in_term thy polym_constrs name pos tm accum =
blanchet@45673
  1788
  is_var_positively_naked_in_term name pos tm accum orelse
blanchet@45673
  1789
  let
blanchet@45673
  1790
    val var = ATerm (name, [])
blanchet@45673
  1791
    fun is_nasty_in_term (ATerm (_, [])) = false
blanchet@45673
  1792
      | is_nasty_in_term (ATerm ((s, _), tms)) =
blanchet@46819
  1793
        let
blanchet@46819
  1794
          val ary = length tms
blanchet@46819
  1795
          val polym_constr = member (op =) polym_constrs s
blanchet@46819
  1796
          val ghosts = ghost_type_args thy s ary
blanchet@46819
  1797
        in
blanchet@46819
  1798
          exists (fn (j, tm) =>
blanchet@46819
  1799
                     if polym_constr then
blanchet@46819
  1800
                       member (op =) ghosts j andalso
blanchet@46819
  1801
                       (tm = var orelse is_nasty_in_term tm)
blanchet@46819
  1802
                     else
blanchet@46819
  1803
                       tm = var andalso member (op =) ghosts j)
blanchet@46819
  1804
                 (0 upto ary - 1 ~~ tms)
blanchet@46819
  1805
          orelse (not polym_constr andalso exists is_nasty_in_term tms)
blanchet@46819
  1806
        end
blanchet@45673
  1807
      | is_nasty_in_term _ = true
blanchet@45673
  1808
  in is_nasty_in_term tm end
blanchet@45673
  1809
blanchet@46819
  1810
fun should_guard_var_in_formula thy polym_constrs level pos phi (SOME true)
blanchet@46819
  1811
                                name =
blanchet@45673
  1812
    (case granularity_of_type_level level of
blanchet@45673
  1813
       All_Vars => true
blanchet@45673
  1814
     | Positively_Naked_Vars =>
blanchet@45673
  1815
       formula_fold pos (is_var_positively_naked_in_term name) phi false
blanchet@45673
  1816
     | Ghost_Type_Arg_Vars =>
blanchet@46819
  1817
       formula_fold pos (is_var_ghost_type_arg_in_term thy polym_constrs name)
blanchet@46819
  1818
                    phi false)
blanchet@46819
  1819
  | should_guard_var_in_formula _ _ _ _ _ _ _ = true
blanchet@46819
  1820
blanchet@46819
  1821
fun always_guard_var_in_formula _ _ _ _ _ _ _ = true
blanchet@43705
  1822
blanchet@45264
  1823
fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false
blanchet@45639
  1824
  | should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T =
blanchet@45673
  1825
    granularity_of_type_level level <> All_Vars andalso
blanchet@45650
  1826
    should_encode_type ctxt mono level T
blanchet@45264
  1827
  | should_generate_tag_bound_decl _ _ _ _ _ = false
blanchet@45263
  1828
nik@44536
  1829
fun mk_aterm format type_enc name T_args args =
nik@44536
  1830
  ATerm (name, map_filter (ho_term_for_type_arg format type_enc) T_args @ args)
blanchet@43835
  1831
blanchet@46820
  1832
fun tag_with_type ctxt format mono type_enc pos T tm =
blanchet@44730
  1833
  IConst (type_tag, T --> T, [T])
blanchet@45645
  1834
  |> mangle_type_args_in_iterm format type_enc
blanchet@46820
  1835
  |> ho_term_from_iterm ctxt format mono type_enc pos
blanchet@44558
  1836
  |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm])
blanchet@44558
  1837
       | _ => raise Fail "unexpected lambda-abstraction")
blanchet@46820
  1838
and ho_term_from_iterm ctxt format mono type_enc =
blanchet@43444
  1839
  let
blanchet@46820
  1840
    fun term site u =
blanchet@43803
  1841
      let
blanchet@44730
  1842
        val (head, args) = strip_iterm_comb u
nik@44536
  1843
        val pos =
blanchet@46820
  1844
          case site of
nik@44536
  1845
            Top_Level pos => pos
nik@44536
  1846
          | Eq_Arg pos => pos
blanchet@45676
  1847
          | _ => NONE
nik@44536
  1848
        val t =
blanchet@43803
  1849
          case head of
blanchet@44730
  1850
            IConst (name as (s, _), _, T_args) =>
nik@44536
  1851
            let
blanchet@46820
  1852
              val arg_site = if is_tptp_equal s then Eq_Arg pos else Elsewhere
nik@44536
  1853
            in
blanchet@46820
  1854
              map (term arg_site) args |> mk_aterm format type_enc name T_args
nik@44536
  1855
            end
blanchet@44730
  1856
          | IVar (name, _) =>
blanchet@46820
  1857
            map (term Elsewhere) args |> mk_aterm format type_enc name []
blanchet@44730
  1858
          | IAbs ((name, T), tm) =>
blanchet@44730
  1859
            AAbs ((name, ho_type_from_typ format type_enc true 0 T),
blanchet@46820
  1860
                  term Elsewhere tm)
blanchet@44730
  1861
          | IApp _ => raise Fail "impossible \"IApp\""
blanchet@44730
  1862
        val T = ityp_of u
blanchet@43803
  1863
      in
blanchet@46820
  1864
        if should_tag_with_type ctxt mono type_enc site u T then
blanchet@46820
  1865
          tag_with_type ctxt format mono type_enc pos T t
blanchet@46818
  1866
        else
blanchet@46818
  1867
          t
blanchet@43803
  1868
      end
blanchet@46820
  1869
  in term o Top_Level end
blanchet@46818
  1870
and formula_from_iformula ctxt polym_constrs format mono type_enc
blanchet@46818
  1871
                          should_guard_var =
blanchet@43700
  1872
  let
blanchet@45673
  1873
    val thy = Proof_Context.theory_of ctxt
blanchet@45673
  1874
    val level = level_of_type_enc type_enc
blanchet@46820
  1875
    val do_term = ho_term_from_iterm ctxt format mono type_enc
blanchet@43444
  1876
    val do_bound_type =
blanchet@44493
  1877
      case type_enc of
blanchet@45673
  1878
        Simple_Types _ => fused_type ctxt mono level 0
nik@44536
  1879
        #> ho_type_from_typ format type_enc false 0 #> SOME
blanchet@43552
  1880
      | _ => K NONE
blanchet@43747
  1881
    fun do_out_of_bound_type pos phi universal (name, T) =
blanchet@45258
  1882
      if should_guard_type ctxt mono type_enc
blanchet@46819
  1883
             (fn () => should_guard_var thy polym_constrs level pos phi
blanchet@46819
  1884
                                        universal name) T then
blanchet@44730
  1885
        IVar (name, T)
blanchet@45645
  1886
        |> type_guard_iterm format type_enc T
blanchet@44232
  1887
        |> do_term pos |> AAtom |> SOME
blanchet@45264
  1888
      else if should_generate_tag_bound_decl ctxt mono type_enc universal T then
blanchet@45264
  1889
        let
blanchet@45264
  1890
          val var = ATerm (name, [])
blanchet@46820
  1891
          val tagged_var = tag_with_type ctxt format mono type_enc pos T var
blanchet@45264
  1892
        in SOME (AAtom (ATerm (`I tptp_equal, [tagged_var, var]))) end
blanchet@43444
  1893
      else
blanchet@43444
  1894
        NONE
blanchet@43747
  1895
    fun do_formula pos (AQuant (q, xs, phi)) =
blanchet@43747
  1896
        let
blanchet@43747
  1897
          val phi = phi |> do_formula pos
blanchet@43747
  1898
          val universal = Option.map (q = AExists ? not) pos
blanchet@43747
  1899
        in
blanchet@43705
  1900
          AQuant (q, xs |> map (apsnd (fn NONE => NONE
blanchet@43705
  1901
                                        | SOME T => do_bound_type T)),
blanchet@43705
  1902
                  (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
blanchet@43705
  1903
                      (map_filter
blanchet@43705
  1904
                           (fn (_, NONE) => NONE
blanchet@43705
  1905
                             | (s, SOME T) =>
blanchet@43747
  1906
                               do_out_of_bound_type pos phi universal (s, T))
blanchet@43747
  1907
                           xs)
blanchet@43705
  1908
                      phi)
blanchet@43705
  1909
        end
blanchet@43747
  1910
      | do_formula pos (AConn conn) = aconn_map pos do_formula conn
blanchet@44232
  1911
      | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
blanchet@44364
  1912
  in do_formula end
blanchet@43444
  1913
blanchet@43444
  1914
(* Each fact is given a unique fact number to avoid name clashes (e.g., because
blanchet@43444
  1915
   of monomorphization). The TPTP explicitly forbids name clashes, and some of
blanchet@43444
  1916
   the remote provers might care. *)
blanchet@46818
  1917
fun formula_line_for_fact ctxt polym_constrs format prefix encode freshen pos
blanchet@47168
  1918
        mono type_enc (j, {name, stature, kind, iformula, atomic_types}) =
blanchet@44735
  1919
  (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name, kind,
blanchet@44730
  1920
   iformula
blanchet@46818
  1921
   |> formula_from_iformula ctxt polym_constrs format mono type_enc
blanchet@46818
  1922
          should_guard_var_in_formula (if pos then SOME true else NONE)
blanchet@46248
  1923
   |> close_formula_universally
blanchet@46791
  1924
   |> bound_tvars type_enc true atomic_types,
blanchet@44364
  1925
   NONE,
blanchet@47168
  1926
   case snd stature of
blanchet@46172
  1927
     Intro => isabelle_info format introN
blanchet@46172
  1928
   | Elim => isabelle_info format elimN
blanchet@46172
  1929
   | Simp => isabelle_info format simpN
blanchet@47207
  1930
   | Eq => isabelle_info format eqN
blanchet@44364
  1931
   | _ => NONE)
blanchet@44364
  1932
  |> Formula
blanchet@43444
  1933
blanchet@46172
  1934
fun formula_line_for_class_rel_clause format type_enc
blanchet@45457
  1935
        ({name, subclass, superclass, ...} : class_rel_clause) =
blanchet@45480
  1936
  let val ty_arg = ATerm (tvar_a_name, []) in
blanchet@43448
  1937
    Formula (class_rel_clause_prefix ^ ascii_of name, Axiom,
blanchet@45480
  1938
             AConn (AImplies,
blanchet@45483
  1939
                    [type_class_formula type_enc subclass ty_arg,
blanchet@45483
  1940
                     type_class_formula type_enc superclass ty_arg])
blanchet@46700
  1941
             |> mk_aquant AForall
blanchet@46700
  1942
                          [(tvar_a_name, atype_of_type_vars type_enc)],
blanchet@47212
  1943
             NONE, isabelle_info format introN)
blanchet@43444
  1944
  end
blanchet@43444
  1945
blanchet@45483
  1946
fun formula_from_arity_atom type_enc (class, t, args) =
blanchet@45483
  1947
  ATerm (t, map (fn arg => ATerm (arg, [])) args)
blanchet@45483
  1948
  |> type_class_formula type_enc class
blanchet@43444
  1949
blanchet@46172
  1950
fun formula_line_for_arity_clause format type_enc
blanchet@46225
  1951
        ({name, prem_atoms, concl_atom} : arity_clause) =
blanchet@44366
  1952
  Formula (arity_clause_prefix ^ name, Axiom,
blanchet@45483
  1953
           mk_ahorn (map (formula_from_arity_atom type_enc) prem_atoms)
blanchet@45483
  1954
                    (formula_from_arity_atom type_enc concl_atom)
blanchet@46225
  1955
           |> mk_aquant AForall
blanchet@46225
  1956
                  (map (rpair (atype_of_type_vars type_enc)) (#3 concl_atom)),
blanchet@47212
  1957
           NONE, isabelle_info format introN)
blanchet@43444
  1958
blanchet@46818
  1959
fun formula_line_for_conjecture ctxt polym_constrs format mono type_enc
blanchet@44730
  1960
        ({name, kind, iformula, atomic_types, ...} : translated_formula) =
blanchet@43448
  1961
  Formula (conjecture_prefix ^ name, kind,
blanchet@46187
  1962
           iformula
blanchet@46818
  1963
           |> formula_from_iformula ctxt polym_constrs format mono type_enc
blanchet@46187
  1964
                  should_guard_var_in_formula (SOME false)
blanchet@46248
  1965
           |> close_formula_universally
blanchet@46791
  1966
           |> bound_tvars type_enc true atomic_types, NONE, NONE)
blanchet@43444
  1967
blanchet@45483
  1968
fun formula_line_for_free_type j phi =
blanchet@45483
  1969
  Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis, phi, NONE, NONE)
blanchet@45484
  1970
fun formula_lines_for_free_types type_enc (facts : translated_formula list) =
blanchet@43444
  1971
  let
blanchet@45483
  1972
    val phis =
blanchet@45483
  1973
      fold (union (op =)) (map #atomic_types facts) []
blanchet@45483
  1974
      |> formulas_for_types type_enc add_sorts_on_tfree
blanchet@45483
  1975
  in map2 formula_line_for_free_type (0 upto length phis - 1) phis end
blanchet@43444
  1976
blanchet@43444
  1977
(** Symbol declarations **)
blanchet@43415
  1978
blanchet@45618
  1979
fun decl_line_for_class order s =
blanchet@45459
  1980
  let val name as (s, _) = `make_type_class s in
blanchet@45480
  1981
    Decl (sym_decl_prefix ^ s, name,
blanchet@46650
  1982
          if order = First_Order then
blanchet@46650
  1983
            ATyAbs ([tvar_a_name],
blanchet@46650
  1984
                    if avoid_first_order_ghost_type_vars then
blanchet@46650
  1985
                      AFun (a_itself_atype, bool_atype)
blanchet@46650
  1986
                    else
blanchet@46650
  1987
                      bool_atype)
blanchet@45480
  1988
          else
blanchet@45618
  1989
            AFun (atype_of_types, bool_atype))
blanchet@45459
  1990
  end
blanchet@45459
  1991
blanchet@45459
  1992
fun decl_lines_for_classes type_enc classes =
blanchet@45459
  1993
  case type_enc of
blanchet@45618
  1994
    Simple_Types (order, Polymorphic, _) =>
blanchet@45618
  1995
    map (decl_line_for_class order) classes
blanchet@45459
  1996
  | _ => []
blanchet@45459
  1997
blanchet@45644
  1998
fun sym_decl_table_for_facts ctxt format type_enc sym_tab (conjs, facts) =
blanchet@43445
  1999
  let
blanchet@45700
  2000
    fun add_iterm_syms tm =
blanchet@44730
  2001
      let val (head, args) = strip_iterm_comb tm in
blanchet@43445
  2002
        (case head of
blanchet@44730
  2003
           IConst ((s, s'), T, T_args) =>
blanchet@45458
  2004
           let
blanchet@45700
  2005
             val (pred_sym, in_conj) =
blanchet@45700
  2006
               case Symtab.lookup sym_tab s of
blanchet@45724
  2007
                 SOME ({pred_sym, in_conj, ...} : sym_info) =>
blanchet@45724
  2008
                 (pred_sym, in_conj)
blanchet@45700
  2009
               | NONE => (false, false)
blanchet@45458
  2010
             val decl_sym =
blanchet@45458
  2011
               (case type_enc of
blanchet@45458
  2012
                  Guards _ => not pred_sym
blanchet@45458
  2013
                | _ => true) andalso
blanchet@45458
  2014
               is_tptp_user_symbol s
blanchet@45458
  2015
           in
blanchet@45458
  2016
             if decl_sym then
blanchet@43447
  2017
               Symtab.map_default (s, [])
blanchet@43755
  2018
                   (insert_type ctxt #3 (s', T_args, T, pred_sym, length args,
blanchet@43755
  2019
                                         in_conj))
blanchet@43445
  2020
             else
blanchet@43445
  2021
               I
blanchet@43445
  2022
           end
blanchet@45700
  2023
         | IAbs (_, tm) => add_iterm_syms tm
blanchet@43445
  2024
         | _ => I)
blanchet@45700
  2025
        #> fold add_iterm_syms args
blanchet@43445
  2026
      end
blanchet@45700
  2027
    val add_fact_syms = K add_iterm_syms |> formula_fold NONE |> fact_lift
blanchet@44837
  2028
    fun add_formula_var_types (AQuant (_, xs, phi)) =
blanchet@44837
  2029
        fold (fn (_, SOME T) => insert_type ctxt I T | _ => I) xs
blanchet@44837
  2030
        #> add_formula_var_types phi
blanchet@44837
  2031
      | add_formula_var_types (AConn (_, phis)) =
blanchet@44837
  2032
        fold add_formula_var_types phis
blanchet@44837
  2033
      | add_formula_var_types _ = I
blanchet@44837
  2034
    fun var_types () =
blanchet@44837
  2035
      if polymorphism_of_type_enc type_enc = Polymorphic then [tvar_a]
blanchet@44837
  2036
      else fold (fact_lift add_formula_var_types) (conjs @ facts) []
blanchet@44837
  2037
    fun add_undefined_const T =
blanchet@44855
  2038
      let
blanchet@44855
  2039
        val (s, s') =
blanchet@45480
  2040
          `(make_fixed_const NONE) @{const_name undefined}
blanchet@46816
  2041
          |> (case type_arg_policy [] type_enc @{const_name undefined} of
blanchet@45642
  2042
                Mangled_Type_Args => mangled_const_name format type_enc [T]
blanchet@44872
  2043
              | _ => I)
blanchet@44855
  2044
      in
blanchet@44855
  2045
        Symtab.map_default (s, [])
blanchet@44855
  2046
                           (insert_type ctxt #3 (s', [T], T, false, 0, false))
blanchet@44855
  2047
      end
blanchet@45480
  2048
    fun add_TYPE_const () =
blanchet@45480
  2049
      let val (s, s') = TYPE_name in
blanchet@45480
  2050
        Symtab.map_default (s, [])
blanchet@45480
  2051
            (insert_type ctxt #3
blanchet@45480
  2052
                         (s', [tvar_a], @{typ "'a itself"}, false, 0, false))
blanchet@45480
  2053
      end
blanchet@43568
  2054
  in
blanchet@43568
  2055
    Symtab.empty
blanchet@44493
  2056
    |> is_type_enc_fairly_sound type_enc
blanchet@45700
  2057
       ? (fold (fold add_fact_syms) [conjs, facts]
blanchet@44856
  2058
          #> (case type_enc of
blanchet@46699
  2059
                Simple_Types (First_Order, Polymorphic, _) =>
blanchet@46699
  2060
                if avoid_first_order_ghost_type_vars then add_TYPE_const ()
blanchet@46699
  2061
                else I
blanchet@45654
  2062
              | Simple_Types _ => I
blanchet@44856
  2063
              | _ => fold add_undefined_const (var_types ())))
blanchet@43568
  2064
  end
blanchet@43445
  2065
blanchet@45258
  2066
(* We add "bool" in case the helper "True_or_False" is included later. *)
blanchet@45492
  2067
fun default_mono level =
blanchet@45258
  2068
  {maybe_finite_Ts = [@{typ bool}],
blanchet@45258
  2069
   surely_finite_Ts = [@{typ bool}],
blanchet@45258
  2070
   maybe_infinite_Ts = known_infinite_types,
blanchet@45492
  2071
   surely_infinite_Ts =
blanchet@45492
  2072
     case level of
blanchet@47129
  2073
       Noninf_Nonmono_Types (Strict, _) => []
blanchet@45492
  2074
     | _ => known_infinite_types,
blanchet@45258
  2075
   maybe_nonmono_Ts = [@{typ bool}]}
blanchet@45258
  2076
blanchet@43555
  2077
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
blanchet@43555
  2078
   out with monotonicity" paper presented at CADE 2011. *)
blanchet@45258
  2079
fun add_iterm_mononotonicity_info _ _ (SOME false) _ mono = mono
blanchet@45258
  2080
  | add_iterm_mononotonicity_info ctxt level _
blanchet@45258
  2081
        (IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2))
blanchet@45258
  2082
        (mono as {maybe_finite_Ts, surely_finite_Ts, maybe_infinite_Ts,
blanchet@45258
  2083
                  surely_infinite_Ts, maybe_nonmono_Ts}) =
blanchet@45262
  2084
    if is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2] then
blanchet@45258
  2085
      case level of
blanchet@47129
  2086
        Noninf_Nonmono_Types (strictness, _) =>
blanchet@45258
  2087
        if exists (type_instance ctxt T) surely_infinite_Ts orelse
blanchet@45730
  2088
           member (type_equiv ctxt) maybe_finite_Ts T then
blanchet@45258
  2089
          mono
blanchet@47129
  2090
        else if is_type_kind_of_surely_infinite ctxt strictness
blanchet@45355
  2091
                                                surely_infinite_Ts T then
blanchet@45258
  2092
          {maybe_finite_Ts = maybe_finite_Ts,
blanchet@45258
  2093
           surely_finite_Ts = surely_finite_Ts,
blanchet@45258
  2094
           maybe_infinite_Ts = maybe_infinite_Ts,
blanchet@45258
  2095
           surely_infinite_Ts = surely_infinite_Ts |> insert_type ctxt I T,
blanchet@45258
  2096
           maybe_nonmono_Ts = maybe_nonmono_Ts}
blanchet@45258
  2097
        else
blanchet@45730
  2098
          {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv ctxt) T,
blanchet@45258
  2099
           surely_finite_Ts = surely_finite_Ts,
blanchet@45258
  2100
           maybe_infinite_Ts = maybe_infinite_Ts,
blanchet@45258
  2101
           surely_infinite_Ts = surely_infinite_Ts,
blanchet@45258
  2102
           maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type ctxt I T}
blanchet@45639
  2103
      | Fin_Nonmono_Types _ =>
blanchet@45258
  2104
        if exists (type_instance ctxt T) surely_finite_Ts orelse
blanchet@45730
  2105
           member (type_equiv ctxt) maybe_infinite_Ts T then
blanchet@45258
  2106
          mono
blanchet@45258
  2107
        else if is_type_surely_finite ctxt T then
blanchet@45258
  2108
          {maybe_finite_Ts = maybe_finite_Ts,
blanchet@45258
  2109
           surely_finite_Ts = surely_finite_Ts |> insert_type ctxt I T,
blanchet@45258
  2110
           maybe_infinite_Ts = maybe_infinite_Ts,
blanchet@45258
  2111
           surely_infinite_Ts = surely_infinite_Ts,
blanchet@45258
  2112
           maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type ctxt I T}
blanchet@45258
  2113
        else
blanchet@45258
  2114
          {maybe_finite_Ts = maybe_finite_Ts,
blanchet@45258
  2115
           surely_finite_Ts = surely_finite_Ts,
blanchet@45730
  2116
           maybe_infinite_Ts = maybe_infinite_Ts |> insert (type_equiv ctxt) T,
blanchet@45258
  2117
           surely_infinite_Ts = surely_infinite_Ts,
blanchet@45258
  2118
           maybe_nonmono_Ts = maybe_nonmono_Ts}
blanchet@45258
  2119
      | _ => mono
blanchet@45258
  2120
    else
blanchet@45258
  2121
      mono
blanchet@45258
  2122
  | add_iterm_mononotonicity_info _ _ _ _ mono = mono
blanchet@45258
  2123
fun add_fact_mononotonicity_info ctxt level
blanchet@45258
  2124
        ({kind, iformula, ...} : translated_formula) =
blanchet@43705
  2125
  formula_fold (SOME (kind <> Conjecture))
blanchet@45258
  2126
               (add_iterm_mononotonicity_info ctxt level) iformula
blanchet@45258
  2127
fun mononotonicity_info_for_facts ctxt type_enc facts =
blanchet@44493
  2128
  let val level = level_of_type_enc type_enc in
blanchet@45492
  2129
    default_mono level
blanchet@45258
  2130
    |> is_type_level_monotonicity_based level
blanchet@45258
  2131
       ? fold (add_fact_mononotonicity_info ctxt level) facts
blanchet@43700
  2132
  end
blanchet@43547
  2133
blanchet@45356
  2134
fun add_iformula_monotonic_types ctxt mono type_enc =
blanchet@45356
  2135
  let
blanchet@45356
  2136
    val level = level_of_type_enc type_enc
blanchet@45356
  2137
    val should_encode = should_encode_type ctxt mono level
blanchet@45359
  2138
    fun add_type T = not (should_encode T) ? insert_type ctxt I T
blanchet@45361
  2139
    fun add_args (IApp (tm1, tm2)) = add_args tm1 #> add_term tm2
blanchet@45361
  2140
      | add_args _ = I
blanchet@45361
  2141
    and add_term tm = add_type (ityp_of tm) #> add_args tm
blanchet@45356
  2142
  in formula_fold NONE (K add_term) end
blanchet@45356
  2143
fun add_fact_monotonic_types ctxt mono type_enc =
blanchet@45356
  2144
  add_iformula_monotonic_types ctxt mono type_enc |> fact_lift
blanchet@45356
  2145
fun monotonic_types_for_facts ctxt mono type_enc facts =
blanchet@45673
  2146
  let val level = level_of_type_enc type_enc in
blanchet@45673
  2147
    [] |> (polymorphism_of_type_enc type_enc = Polymorphic andalso
blanchet@45673
  2148
           is_type_level_monotonicity_based level andalso
blanchet@45673
  2149
           granularity_of_type_level level <> Ghost_Type_Arg_Vars)
blanchet@45673
  2150
          ? fold (add_fact_monotonic_types ctxt mono type_enc) facts
blanchet@45673
  2151
  end
blanchet@45356
  2152
blanchet@45258
  2153
fun formula_line_for_guards_mono_type ctxt format mono type_enc T =
blanchet@45255
  2154
  Formula (guards_sym_formula_prefix ^
blanchet@45255
  2155
           ascii_of (mangled_type format type_enc T),
blanchet@45255
  2156
           Axiom,
blanchet@45255
  2157
           IConst (`make_bound_var "X", T, [])
blanchet@45645
  2158
           |> type_guard_iterm format type_enc T
blanchet@45255
  2159
           |> AAtom
blanchet@46818
  2160
           |> formula_from_iformula ctxt [] format mono type_enc
blanchet@46819
  2161
                                    always_guard_var_in_formula (SOME true)
blanchet@46248
  2162
           |> close_formula_universally
blanchet@46791
  2163
           |> bound_tvars type_enc true (atomic_types_of T),
blanchet@47212
  2164
           NONE, isabelle_info format introN)
blanchet@45255
  2165
blanchet@45258
  2166
fun formula_line_for_tags_mono_type ctxt format mono type_enc T =
blanchet@45255
  2167
  let val x_var = ATerm (`make_bound_var "X", []) in
blanchet@45255
  2168
    Formula (tags_sym_formula_prefix ^
blanchet@45255
  2169
             ascii_of (mangled_type format type_enc T),
blanchet@45255
  2170
             Axiom,
blanchet@46187
  2171
             eq_formula type_enc (atomic_types_of T) false
blanchet@46820
  2172
                  (tag_with_type ctxt format mono type_enc NONE T x_var) x_var,
blanchet@47212
  2173
             NONE, isabelle_info format eqN)
blanchet@45255
  2174
  end
blanchet@45255
  2175
blanchet@45258
  2176
fun problem_lines_for_mono_types ctxt format mono type_enc Ts =
blanchet@45255
  2177
  case type_enc of
blanchet@45255
  2178
    Simple_Types _ => []
blanchet@45255
  2179
  | Guards _ =>
blanchet@45258
  2180
    map (formula_line_for_guards_mono_type ctxt format mono type_enc) Ts
blanchet@45258
  2181
  | Tags _ => map (formula_line_for_tags_mono_type ctxt format mono type_enc) Ts
blanchet@45255
  2182
blanchet@45258
  2183
fun decl_line_for_sym ctxt format mono type_enc s
blanchet@43835
  2184
                      (s', T_args, T, pred_sym, ary, _) =
blanchet@43835
  2185
  let
blanchet@45458
  2186
    val thy = Proof_Context.theory_of ctxt
blanchet@45458
  2187
    val (T, T_args) =
blanchet@45458
  2188
      if null T_args then
blanchet@45458
  2189
        (T, [])
blanchet@46382
  2190
      else case unprefix_and_unascii const_prefix s of
blanchet@45458
  2191
        SOME s' =>
blanchet@45458
  2192
        let
blanchet@45458
  2193
          val s' = s' |> invert_const
blanchet@45458
  2194
          val T = s' |> robust_const_type thy
blanchet@45458
  2195
        in (T, robust_const_typargs thy (s', T)) end
blanchet@46380
  2196
      | NONE => raise Fail "unexpected type arguments"
blanchet@43835
  2197
  in
blanchet@43839
  2198
    Decl (sym_decl_prefix ^ s, (s, s'),
blanchet@45458
  2199
          T |> fused_type ctxt mono (level_of_type_enc type_enc) ary
blanchet@45458
  2200
            |> ho_type_from_typ format type_enc pred_sym ary
blanchet@45458
  2201
            |> not (null T_args)
blanchet@45458
  2202
               ? curry ATyAbs (map (tvar_name o fst o dest_TVar) T_args))
blanchet@43835
  2203
  end
blanchet@43450
  2204
blanchet@45258
  2205
fun formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono type_enc n s
blanchet@45258
  2206
                                     j (s', T_args, T, _, ary, in_conj) =
blanchet@43450
  2207
  let
blanchet@45673
  2208
    val thy = Proof_Context.theory_of ctxt
blanchet@43580
  2209
    val (kind, maybe_negate) =
blanchet@43580
  2210
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43580
  2211
      else (Axiom, I)
blanchet@43618
  2212
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@45676
  2213
    val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  2214
    val bounds =
blanchet@44730
  2215
      bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, []))
blanchet@43450
  2216
    val bound_Ts =
blanchet@45673
  2217
      if exists (curry (op =) dummyT) T_args then
blanchet@45673
  2218
        case level_of_type_enc type_enc of
blanchet@45673
  2219
          All_Types => map SOME arg_Ts
blanchet@45673
  2220
        | level =>
blanchet@45673
  2221
          if granularity_of_type_level level = Ghost_Type_Arg_Vars then
blanchet@45673
  2222
            let val ghosts = ghost_type_args thy s ary in
blanchet@45673
  2223
              map2 (fn j => if member (op =) ghosts j then SOME else K NONE)
blanchet@45676
  2224
                   (0 upto ary - 1) arg_Ts
blanchet@45673
  2225
            end
blanchet@45673
  2226
          else
blanchet@45676
  2227
            replicate ary NONE
blanchet@45673
  2228
      else
blanchet@45676
  2229
        replicate ary NONE
blanchet@43450
  2230
  in
blanchet@44860
  2231
    Formula (guards_sym_formula_prefix ^ s ^
blanchet@43580
  2232
             (if n > 1 then "_" ^ string_of_int j else ""), kind,
blanchet@44730
  2233
             IConst ((s, s'), T, T_args)
blanchet@44730
  2234
             |> fold (curry (IApp o swap)) bounds
blanchet@45645
  2235
             |> type_guard_iterm format type_enc res_T
blanchet@43804
  2236
             |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
blanchet@46818
  2237
             |> formula_from_iformula ctxt [] format mono type_enc
blanchet@46819
  2238
                                      always_guard_var_in_formula (SOME true)
blanchet@46248
  2239
             |> close_formula_universally
blanchet@46791
  2240
             |> bound_tvars type_enc (n > 1) (atomic_types_of T)
blanchet@43580
  2241
             |> maybe_negate,
blanchet@47212
  2242
             NONE, isabelle_info format introN)
blanchet@43450
  2243
  end
blanchet@43450
  2244
blanchet@45673
  2245
fun formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono type_enc n s
blanchet@45673
  2246
        (j, (s', T_args, T, pred_sym, ary, in_conj)) =
blanchet@43700
  2247
  let
blanchet@45676
  2248
    val thy = Proof_Context.theory_of ctxt
blanchet@45676
  2249
    val level = level_of_type_enc type_enc
blanchet@45676
  2250
    val grain = granularity_of_type_level level
blanchet@43700
  2251
    val ident_base =
blanchet@45255
  2252
      tags_sym_formula_prefix ^ s ^
blanchet@43966
  2253
      (if n > 1 then "_" ^ string_of_int j else "")
blanchet@43721
  2254
    val (kind, maybe_negate) =
blanchet@43721
  2255
      if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot)
blanchet@43721
  2256
      else (Axiom, I)
blanchet@43700
  2257
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@45676
  2258
    val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  2259
    val bounds = bound_names |> map (fn name => ATerm (name, []))
nik@44536
  2260
    val cst = mk_aterm format type_enc (s, s') T_args
blanchet@46187
  2261
    val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) pred_sym
blanchet@45676
  2262
    val should_encode = should_encode_type ctxt mono level
blanchet@46820
  2263
    val tag_with = tag_with_type ctxt format mono type_enc NONE
blanchet@43700
  2264
    val add_formula_for_res =
blanchet@43700
  2265
      if should_encode res_T then
blanchet@45676
  2266
        let
blanchet@45676
  2267
          val tagged_bounds =
blanchet@45676
  2268
            if grain = Ghost_Type_Arg_Vars then
blanchet@45676
  2269
              let val ghosts = ghost_type_args thy s ary in
blanchet@45676
  2270
                map2 (fn (j, arg_T) => member (op =) ghosts j ? tag_with arg_T)
blanchet@45676
  2271
                     (0 upto ary - 1 ~~ arg_Ts) bounds
blanchet@45676
  2272
              end
blanchet@45676
  2273
            else
blanchet@45676
  2274
              bounds
blanchet@45676
  2275
        in
blanchet@45676
  2276
          cons (Formula (ident_base ^ "_res", kind,
blanchet@45676
  2277
                         eq (tag_with res_T (cst bounds)) (cst tagged_bounds),
blanchet@47212
  2278
                         NONE, isabelle_info format eqN))
blanchet@45676
  2279
        end
blanchet@43700
  2280
      else
blanchet@43700
  2281
        I
blanchet@43700
  2282
    fun add_formula_for_arg k =
blanchet@43700
  2283
      let val arg_T = nth arg_Ts k in
blanchet@43700
  2284
        if should_encode arg_T then
blanchet@43700
  2285
          case chop k bounds of
blanchet@43700
  2286
            (bounds1, bound :: bounds2) =>
blanchet@43721
  2287
            cons (Formula (ident_base ^ "_arg" ^ string_of_int (k + 1), kind,
blanchet@45255
  2288
                           eq (cst (bounds1 @ tag_with arg_T bound :: bounds2))
blanchet@45255
  2289
                              (cst bounds),
blanchet@47212
  2290
                           NONE, isabelle_info format eqN))
blanchet@43700
  2291
          | _ => raise Fail "expected nonempty tail"
blanchet@43700
  2292
        else
blanchet@43700
  2293
          I
blanchet@43700
  2294
      end
blanchet@43700
  2295
  in
blanchet@43705
  2296
    [] |> not pred_sym ? add_formula_for_res
blanchet@45676
  2297
       |> (Config.get ctxt type_tag_arguments andalso
blanchet@45676
  2298
           grain = Positively_Naked_Vars)
blanchet@45351
  2299
          ? fold add_formula_for_arg (ary - 1 downto 0)
blanchet@43700
  2300
  end
blanchet@43700
  2301
blanchet@43707
  2302
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
blanchet@43707
  2303
blanchet@46651
  2304
fun rationalize_decls ctxt (decls as decl :: (decls' as _ :: _)) =
blanchet@46651
  2305
    let
blanchet@46651
  2306
      val T = result_type_of_decl decl
blanchet@46651
  2307
              |> map_type_tvar (fn (z, _) => TVar (z, HOLogic.typeS))
blanchet@46651
  2308
    in
blanchet@46651
  2309
      if forall (type_generalization ctxt T o result_type_of_decl) decls' then
blanchet@46651
  2310
        [decl]
blanchet@46651
  2311
      else
blanchet@46651
  2312
        decls
blanchet@46651
  2313
    end
blanchet@46651
  2314
  | rationalize_decls _ decls = decls
blanchet@46651
  2315
blanchet@45258
  2316
fun problem_lines_for_sym_decls ctxt format conj_sym_kind mono type_enc
blanchet@45255
  2317
                                (s, decls) =
blanchet@44493
  2318
  case type_enc of
blanchet@46701
  2319
    Simple_Types _ => [decl_line_for_sym ctxt format mono type_enc s (hd decls)]
blanchet@45639
  2320
  | Guards (_, level) =>
blanchet@43839
  2321
    let
blanchet@46651
  2322
      val decls = decls |> rationalize_decls ctxt
blanchet@43839
  2323
      val n = length decls
blanchet@43839
  2324
      val decls =
blanchet@45258
  2325
        decls |> filter (should_encode_type ctxt mono level
blanchet@44264
  2326
                         o result_type_of_decl)
blanchet@43839
  2327
    in
blanchet@43839
  2328
      (0 upto length decls - 1, decls)
blanchet@45258
  2329
      |-> map2 (formula_line_for_guards_sym_decl ctxt format conj_sym_kind mono
blanchet@45258
  2330
                                                 type_enc n s)
blanchet@43839
  2331
    end
blanchet@45639
  2332
  | Tags (_, level) =>
blanchet@45673
  2333
    if granularity_of_type_level level = All_Vars then
blanchet@45639
  2334
      []
blanchet@45639
  2335
    else
blanchet@45639
  2336
      let val n = length decls in
blanchet@45639
  2337
        (0 upto n - 1 ~~ decls)
blanchet@45673
  2338
        |> maps (formula_lines_for_tags_sym_decl ctxt format conj_sym_kind mono
blanchet@45673
  2339
                                                 type_enc n s)
blanchet@45639
  2340
      end
blanchet@43450
  2341
blanchet@45258
  2342
fun problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono type_enc
blanchet@45356
  2343
                                     mono_Ts sym_decl_tab =
blanchet@45255
  2344
  let
blanchet@45255
  2345
    val syms = sym_decl_tab |> Symtab.dest |> sort_wrt fst
blanchet@45255
  2346
    val mono_lines =
blanchet@45258
  2347
      problem_lines_for_mono_types ctxt format mono type_enc mono_Ts
blanchet@45255
  2348
    val decl_lines =
blanchet@45255
  2349
      fold_rev (append o problem_lines_for_sym_decls ctxt format conj_sym_kind
blanchet@46818
  2350
                             mono type_enc)
blanchet@45255
  2351
               syms []
blanchet@45255
  2352
  in mono_lines @ decl_lines end
blanchet@43414
  2353
blanchet@45639
  2354
fun needs_type_tag_idempotence ctxt (Tags (poly, level)) =
blanchet@45351
  2355
    Config.get ctxt type_tag_idempotence andalso
blanchet@45639
  2356
    is_type_level_monotonicity_based level andalso
blanchet@45639
  2357
    poly <> Mangled_Monomorphic
blanchet@45351
  2358
  | needs_type_tag_idempotence _ _ = false
blanchet@43702
  2359
blanchet@43839
  2360
val implicit_declsN = "Should-be-implicit typings"
blanchet@43839
  2361
val explicit_declsN = "Explicit typings"
blanchet@41405
  2362
val factsN = "Relevant facts"
blanchet@41405
  2363
val class_relsN = "Class relationships"
blanchet@43414
  2364
val aritiesN = "Arities"
blanchet@41405
  2365
val helpersN = "Helper facts"
blanchet@41405
  2366
val conjsN = "Conjectures"
blanchet@41561
  2367
val free_typesN = "Type variables"
blanchet@41405
  2368
blanchet@46698
  2369
(* TFF allows implicit declarations of types, function symbols, and predicate
blanchet@46698
  2370
   symbols (with "$i" as the type of individuals), but some provers (e.g.,
blanchet@46698
  2371
   SNARK) require explicit declarations. The situation is similar for THF. *)
blanchet@46698
  2372
blanchet@46746
  2373
fun default_type type_enc pred_sym s =
blanchet@46698
  2374
  let
blanchet@46698
  2375
    val ind =
blanchet@46746
  2376
      case type_enc of
blanchet@46746
  2377
        Simple_Types _ =>
blanchet@46746
  2378
        if String.isPrefix type_const_prefix s then atype_of_types
blanchet@46746
  2379
        else individual_atype
blanchet@46746
  2380
      | _ => individual_atype
blanchet@46698
  2381
    fun typ 0 = if pred_sym then bool_atype else ind
blanchet@46698
  2382
      | typ ary = AFun (ind, typ (ary - 1))
blanchet@46698
  2383
  in typ end
blanchet@46698
  2384
blanchet@46698
  2385
fun nary_type_constr_type n =
blanchet@46698
  2386
  funpow n (curry AFun atype_of_types) atype_of_types
blanchet@46698
  2387
blanchet@46746
  2388
fun undeclared_syms_in_problem type_enc problem =
blanchet@46698
  2389
  let
blanchet@46698
  2390
    val declared = declared_syms_in_problem problem
blanchet@46698
  2391
    fun do_sym name ty =
blanchet@46698
  2392
      if member (op =) declared name then I else AList.default (op =) (name, ty)
blanchet@46698
  2393
    fun do_type (AType (name as (s, _), tys)) =
blanchet@46698
  2394
        is_tptp_user_symbol s
blanchet@46698
  2395
        ? do_sym name (fn () => nary_type_constr_type (length tys))
blanchet@46698
  2396
        #> fold do_type tys
blanchet@46698
  2397
      | do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2
blanchet@46698
  2398
      | do_type (ATyAbs (_, ty)) = do_type ty
blanchet@46698
  2399
    fun do_term pred_sym (ATerm (name as (s, _), tms)) =
blanchet@46698
  2400
        is_tptp_user_symbol s
blanchet@46746
  2401
        ? do_sym name (fn _ => default_type type_enc pred_sym s (length tms))
blanchet@46698
  2402
        #> fold (do_term false) tms
blanchet@46698
  2403
      | do_term _ (AAbs ((_, ty), tm)) = do_type ty #> do_term false tm
blanchet@46698
  2404
    fun do_formula (AQuant (_, xs, phi)) =
blanchet@46698
  2405
        fold do_type (map_filter snd xs) #> do_formula phi
blanchet@46698
  2406
      | do_formula (AConn (_, phis)) = fold do_formula phis
blanchet@46698
  2407
      | do_formula (AAtom tm) = do_term true tm
blanchet@46698
  2408
    fun do_problem_line (Decl (_, _, ty)) = do_type ty
blanchet@46698
  2409
      | do_problem_line (Formula (_, _, phi, _, _)) = do_formula phi
blanchet@46698
  2410
  in
blanchet@46698
  2411
    fold (fold do_problem_line o snd) problem []
blanchet@46698
  2412
    |> filter_out (is_built_in_tptp_symbol o fst o fst)
blanchet@46698
  2413
  end
blanchet@46698
  2414
blanchet@46746
  2415
fun declare_undeclared_syms_in_atp_problem type_enc problem =
blanchet@46698
  2416
  let
blanchet@46698
  2417
    val decls =
blanchet@46698
  2418
      problem
blanchet@46746
  2419
      |> undeclared_syms_in_problem type_enc
blanchet@46698
  2420
      |> sort_wrt (fst o fst)
blanchet@46698
  2421
      |> map (fn (x as (s, _), ty) => Decl (type_decl_prefix ^ s, x, ty ()))
blanchet@46698
  2422
  in (implicit_declsN, decls) :: problem end
blanchet@46698
  2423
blanchet@46816
  2424
fun exists_subdtype P =
blanchet@46816
  2425
  let
blanchet@46816
  2426
    fun ex U = P U orelse
blanchet@46816
  2427
      (case U of Datatype.DtType (_, Us) => exists ex Us | _ => false)
blanchet@46816
  2428
  in ex end
blanchet@46816
  2429
blanchet@46816
  2430
fun is_poly_constr (_, Us) =
blanchet@46816
  2431
  exists (exists_subdtype (fn Datatype.DtTFree _ => true | _ => false)) Us
blanchet@46816
  2432
blanchet@46818
  2433
fun all_constrs_of_polymorphic_datatypes thy =
blanchet@46816
  2434
  Symtab.fold (snd
blanchet@46816
  2435
               #> #descr
blanchet@46816
  2436
               #> maps (snd #> #3)
blanchet@46816
  2437
               #> (fn cs => exists is_poly_constr cs ? append cs))
blanchet@46816
  2438
              (Datatype.get_all thy) []
blanchet@46818
  2439
  |> List.partition is_poly_constr
blanchet@46818
  2440
  |> pairself (map fst)
blanchet@46816
  2441
blanchet@45674
  2442
val explicit_apply_threshold = 50
blanchet@44100
  2443
blanchet@45256
  2444
fun prepare_atp_problem ctxt format conj_sym_kind prem_kind type_enc exporter
blanchet@46385
  2445
                        lam_trans readable_names preproc hyp_ts concl_t facts =
blanchet@38506
  2446
  let
blanchet@45645
  2447
    val thy = Proof_Context.theory_of ctxt
blanchet@45275
  2448
    val type_enc = type_enc |> adjust_type_enc format
blanchet@47217
  2449
    (* Forcing explicit applications is expensive for polymorphic encodings,
blanchet@47217
  2450
       because it takes only one existential variable ranging over "'a => 'b" to
blanchet@47217
  2451
       ruin everything. Hence we do it only if there are few facts (which is
blanchet@47217
  2452
       normally the case for "metis" and the minimizer). *)
blanchet@45674
  2453
    val explicit_apply =
blanchet@47198
  2454
      if polymorphism_of_type_enc type_enc = Polymorphic andalso
blanchet@47198
  2455
         length facts >= explicit_apply_threshold then
blanchet@47217
  2456
        Incomplete_Apply
blanchet@47198
  2457
      else
blanchet@47217
  2458
        Sufficient_Apply
blanchet@46385
  2459
    val lam_trans =
blanchet@46391
  2460
      if lam_trans = keep_lamsN andalso
blanchet@46391
  2461
         not (is_type_enc_higher_order type_enc) then
blanchet@46390
  2462
        error ("Lambda translation scheme incompatible with first-order \
blanchet@44959
  2463
               \encoding.")
blanchet@44959
  2464
      else
blanchet@46385
  2465
        lam_trans
blanchet@46379
  2466
    val (fact_names, classes, conjs, facts, class_rel_clauses, arity_clauses,
blanchet@46379
  2467
         lifted) =
blanchet@46385
  2468
      translate_formulas ctxt format prem_kind type_enc lam_trans preproc hyp_ts
blanchet@46385
  2469
                         concl_t facts
blanchet@46808
  2470
    val sym_tab = sym_table_for_facts ctxt type_enc explicit_apply conjs facts
blanchet@45258
  2471
    val mono = conjs @ facts |> mononotonicity_info_for_facts ctxt type_enc
blanchet@46818
  2472
    val (polym_constrs, monom_constrs) =
blanchet@46818
  2473
      all_constrs_of_polymorphic_datatypes thy
blanchet@46818
  2474
      |>> map (make_fixed_const (SOME format))
blanchet@46816
  2475
    val firstorderize =
blanchet@46816
  2476
      firstorderize_fact thy monom_constrs format type_enc sym_tab
blanchet@45644
  2477
    val (conjs, facts) = (conjs, facts) |> pairself (map firstorderize)
blanchet@47217
  2478
    val sym_tab = sym_table_for_facts ctxt type_enc Incomplete_Apply conjs facts
blanchet@43444
  2479
    val helpers =
blanchet@45644
  2480
      sym_tab |> helper_facts_for_sym_table ctxt format type_enc
blanchet@45644
  2481
              |> map firstorderize
blanchet@45356
  2482
    val mono_Ts =
blanchet@46808
  2483
      helpers @ conjs @ facts |> monotonic_types_for_facts ctxt mono type_enc
blanchet@45459
  2484
    val class_decl_lines = decl_lines_for_classes type_enc classes
blanchet@43550
  2485
    val sym_decl_lines =
blanchet@43596
  2486
      (conjs, helpers @ facts)
blanchet@45644
  2487
      |> sym_decl_table_for_facts ctxt format type_enc sym_tab
blanchet@45258
  2488
      |> problem_lines_for_sym_decl_table ctxt format conj_sym_kind mono
blanchet@45356
  2489
                                               type_enc mono_Ts
blanchet@43750
  2490
    val helper_lines =
blanchet@43797
  2491
      0 upto length helpers - 1 ~~ helpers
blanchet@46818
  2492
      |> map (formula_line_for_fact ctxt polym_constrs format helper_prefix I
blanchet@46818
  2493
                                    false true mono type_enc)
blanchet@45351
  2494
      |> (if needs_type_tag_idempotence ctxt type_enc then
blanchet@46172
  2495
            cons (type_tag_idempotence_fact format type_enc)
blanchet@44000
  2496
          else
blanchet@44000
  2497
            I)
blanchet@43393
  2498
    (* Reordering these might confuse the proof reconstruction code or the SPASS
blanchet@43880
  2499
       FLOTTER hack. *)
blanchet@38506
  2500
    val problem =
blanchet@45459
  2501
      [(explicit_declsN, class_decl_lines @ sym_decl_lines),
blanchet@43797
  2502
       (factsN,
blanchet@46818
  2503
        map (formula_line_for_fact ctxt polym_constrs format fact_prefix
blanchet@46818
  2504
                 ascii_of (not exporter) (not exporter) mono type_enc)
blanchet@43797
  2505
            (0 upto length facts - 1 ~~ facts)),
blanchet@45457
  2506
       (class_relsN,
blanchet@46172
  2507
        map (formula_line_for_class_rel_clause format type_enc)
blanchet@46172
  2508
            class_rel_clauses),
blanchet@46172
  2509
       (aritiesN,
blanchet@46172
  2510
        map (formula_line_for_arity_clause format type_enc) arity_clauses),
blanchet@43750
  2511
       (helpersN, helper_lines),
blanchet@43803
  2512
       (conjsN,
blanchet@46818
  2513
        map (formula_line_for_conjecture ctxt polym_constrs format mono
blanchet@46818
  2514
                                         type_enc) conjs),
blanchet@44493
  2515
       (free_typesN, formula_lines_for_free_types type_enc (facts @ conjs))]
blanchet@43414
  2516
    val problem =
blanchet@43432
  2517
      problem
blanchet@43933
  2518
      |> (case format of
blanchet@43933
  2519
            CNF => ensure_cnf_problem
blanchet@43933
  2520
          | CNF_UEQ => filter_cnf_ueq_problem
blanchet@45453
  2521
          | FOF => I
blanchet@45618
  2522
          | TFF (_, TPTP_Implicit) => I
blanchet@45618
  2523
          | THF (_, TPTP_Implicit, _) => I
blanchet@46746
  2524
          | _ => declare_undeclared_syms_in_atp_problem type_enc)
blanchet@46810
  2525
    val (problem, pool) = problem |> nice_atp_problem readable_names format
blanchet@45643
  2526
    fun add_sym_ary (s, {min_ary, ...} : sym_info) =
blanchet@45644
  2527
      min_ary > 0 ? Symtab.insert (op =) (s, min_ary)
blanchet@38506
  2528
  in
blanchet@38506
  2529
    (problem,
blanchet@38506
  2530
     case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
blanchet@43620
  2531
     fact_names |> Vector.fromList,
blanchet@46379
  2532
     lifted,
blanchet@45643
  2533
     Symtab.empty |> Symtab.fold add_sym_ary sym_tab)
blanchet@38506
  2534
  end
blanchet@38506
  2535
blanchet@41561
  2536
(* FUDGE *)
blanchet@41561
  2537
val conj_weight = 0.0
blanchet@42641
  2538
val hyp_weight = 0.1
blanchet@42641
  2539
val fact_min_weight = 0.2
blanchet@41561
  2540
val fact_max_weight = 1.0
blanchet@43479
  2541
val type_info_default_weight = 0.8
blanchet@41561
  2542
blanchet@41561
  2543
fun add_term_weights weight (ATerm (s, tms)) =
nik@44535
  2544
    is_tptp_user_symbol s ? Symtab.default (s, weight)
nik@44535
  2545
    #> fold (add_term_weights weight) tms
nik@44535
  2546
  | add_term_weights weight (AAbs (_, tm)) = add_term_weights weight tm
blanchet@43448
  2547
fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) =
blanchet@43705
  2548
    formula_fold NONE (K (add_term_weights weight)) phi
blanchet@43399
  2549
  | add_problem_line_weights _ _ = I
blanchet@41561
  2550
blanchet@41561
  2551
fun add_conjectures_weights [] = I
blanchet@41561
  2552
  | add_conjectures_weights conjs =
blanchet@41561
  2553
    let val (hyps, conj) = split_last conjs in
blanchet@41561
  2554
      add_problem_line_weights conj_weight conj
blanchet@41561
  2555
      #> fold (add_problem_line_weights hyp_weight) hyps
blanchet@41561
  2556
    end
blanchet@41561
  2557
blanchet@41561
  2558
fun add_facts_weights facts =
blanchet@41561
  2559
  let
blanchet@41561
  2560
    val num_facts = length facts
blanchet@41561
  2561
    fun weight_of j =
blanchet@41561
  2562
      fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
blanchet@41561
  2563
                        / Real.fromInt num_facts
blanchet@41561
  2564
  in
blanchet@41561
  2565
    map weight_of (0 upto num_facts - 1) ~~ facts
blanchet@41561
  2566
    |> fold (uncurry add_problem_line_weights)
blanchet@41561
  2567
  end
blanchet@41561
  2568
blanchet@41561
  2569
(* Weights are from 0.0 (most important) to 1.0 (least important). *)
blanchet@41561
  2570
fun atp_problem_weights problem =
blanchet@43479
  2571
  let val get = these o AList.lookup (op =) problem in
blanchet@43479
  2572
    Symtab.empty
blanchet@43479
  2573
    |> add_conjectures_weights (get free_typesN @ get conjsN)
blanchet@43479
  2574
    |> add_facts_weights (get factsN)
blanchet@43479
  2575
    |> fold (fold (add_problem_line_weights type_info_default_weight) o get)
blanchet@43839
  2576
            [explicit_declsN, class_relsN, aritiesN]
blanchet@43479
  2577
    |> Symtab.dest
blanchet@43479
  2578
    |> sort (prod_ord Real.compare string_ord o pairself swap)
blanchet@43479
  2579
  end
blanchet@41561
  2580
blanchet@38506
  2581
end;