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