src/HOL/Tools/ATP/atp_problem_generate.ML
author blanchet
Mon, 23 Jul 2012 15:32:30 +0200
changeset 49453 3e45c98fe127
parent 49339 3ee5b5589402
child 49553 726590131ca1
permissions -rw-r--r--
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet@47148
     1
(*  Title:      HOL/Tools/ATP/atp_problem_generate.ML
blanchet@38506
     2
    Author:     Fabian Immler, TU Muenchen
blanchet@38506
     3
    Author:     Makarius
blanchet@38506
     4
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@38506
     5
blanchet@44734
     6
Translation of HOL to FOL for Metis and Sledgehammer.
blanchet@38506
     7
*)
blanchet@38506
     8
blanchet@47148
     9
signature ATP_PROBLEM_GENERATE =
blanchet@38506
    10
sig
nik@44535
    11
  type ('a, 'b) ho_term = ('a, 'b) ATP_Problem.ho_term
blanchet@43977
    12
  type connective = ATP_Problem.connective
blanchet@49150
    13
  type ('a, 'b, 'c, 'd) formula = ('a, 'b, 'c, 'd) ATP_Problem.formula
blanchet@46172
    14
  type atp_format = ATP_Problem.atp_format
blanchet@48990
    15
  type formula_role = ATP_Problem.formula_role
blanchet@38506
    16
  type 'a problem = 'a ATP_Problem.problem
blanchet@43926
    17
blanchet@49158
    18
  datatype mode = Metis | Sledgehammer | Sledgehammer_Completish | Exporter
blanchet@48961
    19
blanchet@47168
    20
  datatype scope = Global | Local | Assum | Chained
blanchet@49453
    21
  datatype status =
blanchet@49453
    22
    General | Induction | Intro | Inductive | Elim | Simp | Non_Rec_Def |
blanchet@49453
    23
    Rec_Def
blanchet@47168
    24
  type stature = scope * status
blanchet@43484
    25
blanchet@47129
    26
  datatype strictness = Strict | Non_Strict
blanchet@49198
    27
  datatype uniformity = Uniform | Non_Uniform
blanchet@49198
    28
  datatype constr_optim = With_Constr_Optim | Without_Constr_Optim
blanchet@43484
    29
  datatype type_level =
blanchet@45256
    30
    All_Types |
blanchet@49215
    31
    Undercover_Types |
blanchet@49198
    32
    Nonmono_Types of strictness * uniformity |
blanchet@49198
    33
    Const_Types of constr_optim |
blanchet@44233
    34
    No_Types
blanchet@45650
    35
  type type_enc
blanchet@43484
    36
blanchet@46384
    37
  val no_lamsN : string
blanchet@46384
    38
  val hide_lamsN : string
blanchet@47193
    39
  val liftingN : string
blanchet@47193
    40
  val combsN : string
blanchet@47193
    41
  val combs_and_liftingN : string
blanchet@47193
    42
  val combs_or_liftingN : string
blanchet@46384
    43
  val lam_liftingN : string
blanchet@46384
    44
  val keep_lamsN : string
blanchet@44367
    45
  val schematic_var_prefix : string
blanchet@44367
    46
  val fixed_var_prefix : string
blanchet@44367
    47
  val tvar_prefix : string
blanchet@44367
    48
  val tfree_prefix : string
blanchet@44367
    49
  val const_prefix : string
blanchet@44367
    50
  val type_const_prefix : string
blanchet@44367
    51
  val class_prefix : string
blanchet@46425
    52
  val lam_lifted_prefix : string
blanchet@46425
    53
  val lam_lifted_mono_prefix : string
blanchet@46425
    54
  val lam_lifted_poly_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@46425
    58
  val combinator_prefix : string
blanchet@49157
    59
  val class_decl_prefix : string
blanchet@43966
    60
  val type_decl_prefix : string
blanchet@43966
    61
  val sym_decl_prefix : string
blanchet@49157
    62
  val class_memb_prefix : string
blanchet@44860
    63
  val guards_sym_formula_prefix : string
blanchet@45255
    64
  val tags_sym_formula_prefix : string
blanchet@40445
    65
  val fact_prefix : string
blanchet@38506
    66
  val conjecture_prefix : string
blanchet@43750
    67
  val helper_prefix : string
blanchet@49157
    68
  val subclass_prefix : string
blanchet@49157
    69
  val tcon_clause_prefix : string
blanchet@43966
    70
  val tfree_clause_prefix : string
blanchet@46425
    71
  val lam_fact_prefix : string
blanchet@43750
    72
  val typed_helper_suffix : string
blanchet@43966
    73
  val untyped_helper_suffix : string
blanchet@43807
    74
  val predicator_name : string
blanchet@43807
    75
  val app_op_name : string
blanchet@45255
    76
  val type_guard_name : string
blanchet@43945
    77
  val type_tag_name : string
blanchet@47263
    78
  val native_type_prefix : string
blanchet@44015
    79
  val prefixed_predicator_name : string
blanchet@43971
    80
  val prefixed_app_op_name : string
blanchet@43971
    81
  val prefixed_type_tag_name : string
blanchet@44367
    82
  val ascii_of : string -> string
blanchet@44367
    83
  val unascii_of : string -> string
blanchet@46382
    84
  val unprefix_and_unascii : string -> string -> string option
blanchet@44000
    85
  val proxy_table : (string * (string * (thm * (string * string)))) list
blanchet@44000
    86
  val proxify_const : string -> (string * string) option
blanchet@44367
    87
  val invert_const : string -> string
blanchet@44367
    88
  val unproxify_const : string -> string
blanchet@43934
    89
  val new_skolem_var_name_from_const : string -> string
blanchet@49333
    90
  val atp_logical_consts : string list
blanchet@44089
    91
  val atp_irrelevant_consts : string list
blanchet@49339
    92
  val atp_widely_irrelevant_consts : string list
blanchet@44089
    93
  val atp_schematic_consts_of : term -> typ list Symtab.table
blanchet@44692
    94
  val is_type_enc_higher_order : type_enc -> bool
blanchet@49146
    95
  val is_type_enc_polymorphic : type_enc -> bool
blanchet@44493
    96
  val level_of_type_enc : type_enc -> type_level
blanchet@49104
    97
  val is_type_enc_sound : type_enc -> bool
blanchet@47129
    98
  val type_enc_from_string : strictness -> string -> type_enc
blanchet@46172
    99
  val adjust_type_enc : atp_format -> type_enc -> type_enc
blanchet@49317
   100
  val is_lambda_free : term -> bool
blanchet@43977
   101
  val mk_aconns :
blanchet@49150
   102
    connective -> ('a, 'b, 'c, 'd) formula list -> ('a, 'b, 'c, 'd) formula
nik@44535
   103
  val unmangled_const : string -> string * (string, 'b) ho_term list
blanchet@47220
   104
  val unmangled_const_name : string -> string list
blanchet@48961
   105
  val helper_table : ((string * bool) * (status * thm) list) list
blanchet@46385
   106
  val trans_lams_from_string :
blanchet@46385
   107
    Proof.context -> type_enc -> string -> term list -> term list * term list
blanchet@49453
   108
  val string_of_status : status -> string
blanchet@44372
   109
  val factsN : string
blanchet@40240
   110
  val prepare_atp_problem :
blanchet@48990
   111
    Proof.context -> atp_format -> formula_role -> type_enc -> mode -> string
blanchet@48927
   112
    -> bool -> bool -> bool -> term list -> term
blanchet@47168
   113
    -> ((string * stature) * term) list
blanchet@47168
   114
    -> string problem * string Symtab.table * (string * stature) list vector
blanchet@46422
   115
       * (string * term) list * int Symtab.table
blanchet@47901
   116
  val atp_problem_selection_weights : string problem -> (string * real) list
blanchet@47909
   117
  val atp_problem_term_order_info : string problem -> (string * int) list
blanchet@38506
   118
end;
blanchet@38506
   119
blanchet@47148
   120
structure ATP_Problem_Generate : ATP_PROBLEM_GENERATE =
blanchet@38506
   121
struct
blanchet@38506
   122
blanchet@43926
   123
open ATP_Util
blanchet@38506
   124
open ATP_Problem
blanchet@43926
   125
blanchet@49158
   126
datatype mode = Metis | Sledgehammer | Sledgehammer_Completish | Exporter
blanchet@48961
   127
blanchet@48638
   128
datatype scope = Global | Local | Assum | Chained
blanchet@49453
   129
datatype status =
blanchet@49453
   130
  General | Induction | Intro | Inductive | Elim | Simp | Non_Rec_Def | Rec_Def
blanchet@48638
   131
type stature = scope * status
blanchet@48638
   132
blanchet@48638
   133
datatype order =
blanchet@48638
   134
  First_Order |
blanchet@49019
   135
  Higher_Order of thf_choice
blanchet@49146
   136
datatype phantom_policy = Without_Phantom_Type_Vars | With_Phantom_Type_Vars
blanchet@49146
   137
datatype polymorphism =
blanchet@49146
   138
  Type_Class_Polymorphic |
blanchet@49146
   139
  Raw_Polymorphic of phantom_policy |
blanchet@49146
   140
  Raw_Monomorphic |
blanchet@49146
   141
  Mangled_Monomorphic
blanchet@48638
   142
datatype strictness = Strict | Non_Strict
blanchet@49198
   143
datatype uniformity = Uniform | Non_Uniform
blanchet@49198
   144
datatype constr_optim = With_Constr_Optim | Without_Constr_Optim
blanchet@48638
   145
datatype type_level =
blanchet@48638
   146
  All_Types |
blanchet@49215
   147
  Undercover_Types |
blanchet@49198
   148
  Nonmono_Types of strictness * uniformity |
blanchet@49198
   149
  Const_Types of constr_optim |
blanchet@48638
   150
  No_Types
blanchet@48638
   151
blanchet@48638
   152
datatype type_enc =
blanchet@48638
   153
  Native of order * polymorphism * type_level |
blanchet@48638
   154
  Guards of polymorphism * type_level |
blanchet@48638
   155
  Tags of polymorphism * type_level
blanchet@48638
   156
blanchet@48638
   157
fun is_type_enc_native (Native _) = true
blanchet@48638
   158
  | is_type_enc_native _ = false
blanchet@48638
   159
fun is_type_enc_higher_order (Native (Higher_Order _, _, _)) = true
blanchet@48638
   160
  | is_type_enc_higher_order _ = false
blanchet@48638
   161
blanchet@48638
   162
fun polymorphism_of_type_enc (Native (_, poly, _)) = poly
blanchet@48638
   163
  | polymorphism_of_type_enc (Guards (poly, _)) = poly
blanchet@48638
   164
  | polymorphism_of_type_enc (Tags (poly, _)) = poly
blanchet@48638
   165
blanchet@49146
   166
fun is_type_enc_polymorphic type_enc =
blanchet@49146
   167
  case polymorphism_of_type_enc type_enc of
blanchet@49146
   168
    Raw_Polymorphic _ => true
blanchet@49146
   169
  | Type_Class_Polymorphic => true
blanchet@49146
   170
  | _ => false
blanchet@49146
   171
blanchet@49216
   172
fun is_type_enc_mangling type_enc =
blanchet@49216
   173
  polymorphism_of_type_enc type_enc = Mangled_Monomorphic
blanchet@49216
   174
blanchet@48638
   175
fun level_of_type_enc (Native (_, _, level)) = level
blanchet@48638
   176
  | level_of_type_enc (Guards (_, level)) = level
blanchet@48638
   177
  | level_of_type_enc (Tags (_, level)) = level
blanchet@48638
   178
blanchet@49198
   179
fun is_type_level_uniform (Nonmono_Types (_, Non_Uniform)) = false
blanchet@49198
   180
  | is_type_level_uniform Undercover_Types = false
blanchet@49198
   181
  | is_type_level_uniform _ = true
blanchet@48638
   182
blanchet@49198
   183
fun is_type_level_sound (Const_Types _) = false
blanchet@49198
   184
  | is_type_level_sound No_Types = false
blanchet@49198
   185
  | is_type_level_sound _ = true
blanchet@49104
   186
val is_type_enc_sound = is_type_level_sound o level_of_type_enc
blanchet@48638
   187
blanchet@49107
   188
fun is_type_level_monotonicity_based (Nonmono_Types _) = true
blanchet@48638
   189
  | is_type_level_monotonicity_based _ = false
blanchet@48638
   190
blanchet@46387
   191
val no_lamsN = "no_lams" (* used internally; undocumented *)
blanchet@46384
   192
val hide_lamsN = "hide_lams"
blanchet@47193
   193
val liftingN = "lifting"
blanchet@47193
   194
val combsN = "combs"
blanchet@47193
   195
val combs_and_liftingN = "combs_and_lifting"
blanchet@47193
   196
val combs_or_liftingN = "combs_or_lifting"
blanchet@46384
   197
val keep_lamsN = "keep_lams"
blanchet@49159
   198
val lam_liftingN = "lam_lifting" (* legacy FIXME: remove *)
blanchet@44959
   199
blanchet@43926
   200
val bound_var_prefix = "B_"
blanchet@48021
   201
val all_bound_var_prefix = "A_"
blanchet@48021
   202
val exist_bound_var_prefix = "E_"
blanchet@43926
   203
val schematic_var_prefix = "V_"
blanchet@43926
   204
val fixed_var_prefix = "v_"
blanchet@43926
   205
val tvar_prefix = "T_"
blanchet@49234
   206
val tfree_prefix = "tf_"
blanchet@43926
   207
val const_prefix = "c_"
blanchet@49234
   208
val type_const_prefix = "t_"
blanchet@47263
   209
val native_type_prefix = "n_"
blanchet@43926
   210
val class_prefix = "cl_"
blanchet@43926
   211
blanchet@46380
   212
(* Freshness almost guaranteed! *)
blanchet@47220
   213
val atp_prefix = "ATP" ^ Long_Name.separator
blanchet@46380
   214
val atp_weak_prefix = "ATP:"
blanchet@46380
   215
blanchet@46425
   216
val lam_lifted_prefix = atp_weak_prefix ^ "Lam"
blanchet@46425
   217
val lam_lifted_mono_prefix = lam_lifted_prefix ^ "m"
blanchet@46425
   218
val lam_lifted_poly_prefix = lam_lifted_prefix ^ "p"
blanchet@44778
   219
blanchet@47220
   220
val skolem_const_prefix = atp_prefix ^ "Sko"
blanchet@43926
   221
val old_skolem_const_prefix = skolem_const_prefix ^ "o"
blanchet@43926
   222
val new_skolem_const_prefix = skolem_const_prefix ^ "n"
blanchet@43926
   223
blanchet@46425
   224
val combinator_prefix = "COMB"
blanchet@46425
   225
blanchet@49157
   226
val class_decl_prefix = "cl_"
blanchet@43839
   227
val type_decl_prefix = "ty_"
blanchet@43839
   228
val sym_decl_prefix = "sy_"
blanchet@49157
   229
val class_memb_prefix = "clmb_"
blanchet@44860
   230
val guards_sym_formula_prefix = "gsy_"
blanchet@45255
   231
val tags_sym_formula_prefix = "tsy_"
blanchet@47237
   232
val uncurried_alias_eq_prefix = "unc_"
blanchet@40445
   233
val fact_prefix = "fact_"
blanchet@38506
   234
val conjecture_prefix = "conj_"
blanchet@38506
   235
val helper_prefix = "help_"
blanchet@49157
   236
val subclass_prefix = "subcl_"
blanchet@49157
   237
val tcon_clause_prefix = "tcon_"
blanchet@43926
   238
val tfree_clause_prefix = "tfree_"
blanchet@38506
   239
blanchet@46425
   240
val lam_fact_prefix = "ATP.lambda_"
blanchet@43750
   241
val typed_helper_suffix = "_T"
blanchet@43750
   242
val untyped_helper_suffix = "_U"
blanchet@43750
   243
blanchet@45346
   244
val predicator_name = "pp"
blanchet@45346
   245
val app_op_name = "aa"
blanchet@45346
   246
val type_guard_name = "gg"
blanchet@45346
   247
val type_tag_name = "tt"
blanchet@43402
   248
blanchet@44015
   249
val prefixed_predicator_name = const_prefix ^ predicator_name
blanchet@43971
   250
val prefixed_app_op_name = const_prefix ^ app_op_name
blanchet@43971
   251
val prefixed_type_tag_name = const_prefix ^ type_tag_name
blanchet@43971
   252
blanchet@43926
   253
(*Escaping of special characters.
blanchet@43926
   254
  Alphanumeric characters are left unchanged.
blanchet@43926
   255
  The character _ goes to __
blanchet@43926
   256
  Characters in the range ASCII space to / go to _A to _P, respectively.
blanchet@43926
   257
  Other characters go to _nnn where nnn is the decimal ASCII code.*)
blanchet@43934
   258
val upper_a_minus_space = Char.ord #"A" - Char.ord #" "
blanchet@43926
   259
blanchet@43926
   260
fun ascii_of_char c =
blanchet@43926
   261
  if Char.isAlphaNum c then
blanchet@43926
   262
    String.str c
blanchet@43926
   263
  else if c = #"_" then
blanchet@43926
   264
    "__"
blanchet@43926
   265
  else if #" " <= c andalso c <= #"/" then
blanchet@43926
   266
    "_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space))
blanchet@43926
   267
  else
blanchet@43926
   268
    (* fixed width, in case more digits follow *)
blanchet@43926
   269
    "_" ^ stringN_of_int 3 (Char.ord c)
blanchet@43926
   270
blanchet@43926
   271
val ascii_of = String.translate ascii_of_char
blanchet@43926
   272
blanchet@43926
   273
(** Remove ASCII armoring from names in proof files **)
blanchet@43926
   274
blanchet@43926
   275
(* We don't raise error exceptions because this code can run inside a worker
blanchet@43926
   276
   thread. Also, the errors are impossible. *)
blanchet@43926
   277
val unascii_of =
blanchet@43926
   278
  let
blanchet@48990
   279
    fun un rcs [] = String.implode (rev rcs)
blanchet@43926
   280
      | un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *)
blanchet@43926
   281
        (* Three types of _ escapes: __, _A to _P, _nnn *)
blanchet@44367
   282
      | un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs
blanchet@43926
   283
      | un rcs (#"_" :: c :: cs) =
blanchet@43926
   284
        if #"A" <= c andalso c<= #"P" then
blanchet@43926
   285
          (* translation of #" " to #"/" *)
blanchet@43926
   286
          un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs
blanchet@43926
   287
        else
blanchet@44367
   288
          let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in
blanchet@43926
   289
            case Int.fromString (String.implode digits) of
blanchet@43926
   290
              SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2))
blanchet@44367
   291
            | NONE => un (c :: #"_" :: rcs) cs (* ERROR *)
blanchet@43926
   292
          end
blanchet@43926
   293
      | un rcs (c :: cs) = un (c :: rcs) cs
blanchet@43926
   294
  in un [] o String.explode end
blanchet@43926
   295
blanchet@43926
   296
(* If string s has the prefix s1, return the result of deleting it,
blanchet@43926
   297
   un-ASCII'd. *)
blanchet@46382
   298
fun unprefix_and_unascii s1 s =
blanchet@43926
   299
  if String.isPrefix s1 s then
blanchet@43926
   300
    SOME (unascii_of (String.extract (s, size s1, NONE)))
blanchet@43926
   301
  else
blanchet@43926
   302
    NONE
blanchet@43926
   303
blanchet@44000
   304
val proxy_table =
blanchet@44000
   305
  [("c_False", (@{const_name False}, (@{thm fFalse_def},
blanchet@44000
   306
       ("fFalse", @{const_name ATP.fFalse})))),
blanchet@44000
   307
   ("c_True", (@{const_name True}, (@{thm fTrue_def},
blanchet@44000
   308
       ("fTrue", @{const_name ATP.fTrue})))),
blanchet@44000
   309
   ("c_Not", (@{const_name Not}, (@{thm fNot_def},
blanchet@44000
   310
       ("fNot", @{const_name ATP.fNot})))),
blanchet@44000
   311
   ("c_conj", (@{const_name conj}, (@{thm fconj_def},
blanchet@44000
   312
       ("fconj", @{const_name ATP.fconj})))),
blanchet@44000
   313
   ("c_disj", (@{const_name disj}, (@{thm fdisj_def},
blanchet@44000
   314
       ("fdisj", @{const_name ATP.fdisj})))),
blanchet@44000
   315
   ("c_implies", (@{const_name implies}, (@{thm fimplies_def},
blanchet@44000
   316
       ("fimplies", @{const_name ATP.fimplies})))),
blanchet@44000
   317
   ("equal", (@{const_name HOL.eq}, (@{thm fequal_def},
nik@44537
   318
       ("fequal", @{const_name ATP.fequal})))),
nik@44537
   319
   ("c_All", (@{const_name All}, (@{thm fAll_def},
nik@44537
   320
       ("fAll", @{const_name ATP.fAll})))),
nik@44537
   321
   ("c_Ex", (@{const_name Ex}, (@{thm fEx_def},
nik@44537
   322
       ("fEx", @{const_name ATP.fEx}))))]
blanchet@43926
   323
blanchet@44000
   324
val proxify_const = AList.lookup (op =) proxy_table #> Option.map (snd o snd)
blanchet@43926
   325
blanchet@43926
   326
(* Readable names for the more common symbolic functions. Do not mess with the
blanchet@43926
   327
   table unless you know what you are doing. *)
blanchet@43926
   328
val const_trans_table =
blanchet@43926
   329
  [(@{type_name Product_Type.prod}, "prod"),
blanchet@43926
   330
   (@{type_name Sum_Type.sum}, "sum"),
blanchet@43926
   331
   (@{const_name False}, "False"),
blanchet@43926
   332
   (@{const_name True}, "True"),
blanchet@43926
   333
   (@{const_name Not}, "Not"),
blanchet@43926
   334
   (@{const_name conj}, "conj"),
blanchet@43926
   335
   (@{const_name disj}, "disj"),
blanchet@43926
   336
   (@{const_name implies}, "implies"),
blanchet@43926
   337
   (@{const_name HOL.eq}, "equal"),
nik@44537
   338
   (@{const_name All}, "All"),
nik@44537
   339
   (@{const_name Ex}, "Ex"),
blanchet@43926
   340
   (@{const_name If}, "If"),
blanchet@43926
   341
   (@{const_name Set.member}, "member"),
blanchet@46425
   342
   (@{const_name Meson.COMBI}, combinator_prefix ^ "I"),
blanchet@46425
   343
   (@{const_name Meson.COMBK}, combinator_prefix ^ "K"),
blanchet@46425
   344
   (@{const_name Meson.COMBB}, combinator_prefix ^ "B"),
blanchet@46425
   345
   (@{const_name Meson.COMBC}, combinator_prefix ^ "C"),
blanchet@46425
   346
   (@{const_name Meson.COMBS}, combinator_prefix ^ "S")]
blanchet@43926
   347
  |> Symtab.make
blanchet@44000
   348
  |> fold (Symtab.update o swap o snd o snd o snd) proxy_table
blanchet@43926
   349
blanchet@43926
   350
(* Invert the table of translations between Isabelle and ATPs. *)
blanchet@43926
   351
val const_trans_table_inv =
blanchet@43926
   352
  const_trans_table |> Symtab.dest |> map swap |> Symtab.make
blanchet@43926
   353
val const_trans_table_unprox =
blanchet@43926
   354
  Symtab.empty
blanchet@44000
   355
  |> fold (fn (_, (isa, (_, (_, atp)))) => Symtab.update (atp, isa)) proxy_table
blanchet@43926
   356
blanchet@43926
   357
val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
blanchet@43926
   358
val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox)
blanchet@43926
   359
blanchet@43926
   360
fun lookup_const c =
blanchet@43926
   361
  case Symtab.lookup const_trans_table c of
blanchet@43926
   362
    SOME c' => c'
blanchet@43926
   363
  | NONE => ascii_of c
blanchet@43926
   364
blanchet@44489
   365
fun ascii_of_indexname (v, 0) = ascii_of v
blanchet@44489
   366
  | ascii_of_indexname (v, i) = ascii_of v ^ "_" ^ string_of_int i
blanchet@43926
   367
blanchet@43926
   368
fun make_bound_var x = bound_var_prefix ^ ascii_of x
blanchet@45262
   369
fun make_all_bound_var x = all_bound_var_prefix ^ ascii_of x
blanchet@45262
   370
fun make_exist_bound_var x = exist_bound_var_prefix ^ ascii_of x
blanchet@43926
   371
fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
blanchet@43926
   372
fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
blanchet@43926
   373
blanchet@49157
   374
fun make_tvar (s, i) = tvar_prefix ^ ascii_of_indexname (unprefix "'" s, i)
blanchet@49157
   375
fun make_tfree s = tfree_prefix ^ ascii_of (unprefix "'" s)
blanchet@49157
   376
fun tvar_name ((x as (s, _)), _) = (make_tvar x, s)
blanchet@43926
   377
blanchet@46172
   378
(* "HOL.eq" and choice are mapped to the ATP's equivalents *)
nik@45451
   379
local
nik@45451
   380
  val choice_const = (fst o dest_Const o HOLogic.choice_const) Term.dummyT
nik@45451
   381
  fun default c = const_prefix ^ lookup_const c
nik@45451
   382
in
nik@45451
   383
  fun make_fixed_const _ @{const_name HOL.eq} = tptp_old_equal
blanchet@48638
   384
    | make_fixed_const (SOME (Native (Higher_Order THF_With_Choice, _, _))) c =
blanchet@45618
   385
      if c = choice_const then tptp_choice else default c
nik@45451
   386
    | make_fixed_const _ c = default c
nik@45451
   387
end
blanchet@43926
   388
blanchet@43926
   389
fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
blanchet@43926
   390
blanchet@49157
   391
fun make_class clas = class_prefix ^ ascii_of clas
blanchet@43926
   392
blanchet@43934
   393
fun new_skolem_var_name_from_const s =
blanchet@43934
   394
  let val ss = s |> space_explode Long_Name.separator in
blanchet@43934
   395
    nth ss (length ss - 2)
blanchet@43934
   396
  end
blanchet@43934
   397
blanchet@49333
   398
(* These are ignored anyway by the relevance filter (unless they appear in
blanchet@49333
   399
   higher-order places) but not by the monomorphizer. *)
blanchet@49333
   400
val atp_logical_consts =
blanchet@49333
   401
  [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
blanchet@49333
   402
   @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
blanchet@49333
   403
   @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
blanchet@49333
   404
blanchet@44089
   405
(* These are either simplified away by "Meson.presimplify" (most of the time) or
blanchet@44089
   406
   handled specially via "fFalse", "fTrue", ..., "fequal". *)
blanchet@44089
   407
val atp_irrelevant_consts =
blanchet@44089
   408
  [@{const_name False}, @{const_name True}, @{const_name Not},
blanchet@44089
   409
   @{const_name conj}, @{const_name disj}, @{const_name implies},
blanchet@44089
   410
   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
blanchet@44089
   411
blanchet@49333
   412
val atp_widely_irrelevant_consts = atp_logical_consts @ atp_irrelevant_consts
blanchet@44089
   413
blanchet@44099
   414
fun add_schematic_const (x as (_, T)) =
blanchet@44099
   415
  Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x
blanchet@44099
   416
val add_schematic_consts_of =
blanchet@44099
   417
  Term.fold_aterms (fn Const (x as (s, _)) =>
blanchet@49242
   418
                       not (member (op =) atp_widely_irrelevant_consts s)
blanchet@44099
   419
                       ? add_schematic_const x
blanchet@44099
   420
                      | _ => I)
blanchet@44099
   421
fun atp_schematic_consts_of t = add_schematic_consts_of t Symtab.empty
blanchet@44089
   422
blanchet@49148
   423
val tvar_a_str = "'a"
blanchet@49262
   424
val tvar_a_z = ((tvar_a_str, 100 (* arbitrary *)), HOLogic.typeS)
blanchet@49157
   425
val tvar_a = TVar tvar_a_z
blanchet@49157
   426
val tvar_a_name = tvar_name tvar_a_z
blanchet@49148
   427
val itself_name = `make_fixed_type_const @{type_name itself}
blanchet@49148
   428
val TYPE_name = `(make_fixed_const NONE) @{const_name TYPE}
blanchet@49148
   429
val tvar_a_atype = AType (tvar_a_name, [])
blanchet@49148
   430
val a_itself_atype = AType (itself_name, [tvar_a_atype])
blanchet@49148
   431
blanchet@43926
   432
(** Definitions and functions for FOL clauses and formulas for TPTP **)
blanchet@43926
   433
blanchet@49157
   434
(** Type class membership **)
blanchet@43926
   435
blanchet@49157
   436
(* In our data structures, [] exceptionally refers to the top class, not to
blanchet@49157
   437
   the empty class. *)
blanchet@44104
   438
blanchet@49157
   439
val class_of_types = the_single @{sort type}
blanchet@45483
   440
blanchet@49157
   441
fun normalize_classes cls = if member (op =) cls class_of_types then [] else cls
blanchet@43926
   442
blanchet@49157
   443
(* Arity of type constructor "s :: (arg1, ..., argN) res" *)
blanchet@49157
   444
fun make_axiom_tcon_clause (s, name, (cl, args)) =
blanchet@43926
   445
  let
blanchet@49157
   446
    val args = args |> map normalize_classes
blanchet@49148
   447
    val tvars =
blanchet@49157
   448
      1 upto length args |> map (fn j => TVar ((tvar_a_str, j), @{sort type}))
blanchet@49157
   449
  in (name, args ~~ tvars, (cl, Type (s, tvars))) end
blanchet@43926
   450
blanchet@49156
   451
(* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in
blanchet@49156
   452
   theory thy provided its arguments have the corresponding sorts. *)
blanchet@49157
   453
fun class_pairs thy tycons cls =
blanchet@49156
   454
  let
blanchet@49156
   455
    val alg = Sign.classes_of thy
blanchet@49156
   456
    fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
blanchet@49157
   457
    fun add_class tycon cl =
blanchet@49157
   458
      cons (cl, domain_sorts tycon cl)
blanchet@49156
   459
      handle Sorts.CLASS_ERROR _ => I
blanchet@49157
   460
    fun try_classes tycon = (tycon, fold (add_class tycon) cls [])
blanchet@49156
   461
  in map try_classes tycons end
blanchet@49156
   462
blanchet@49156
   463
(* Proving one (tycon, class) membership may require proving others, so
blanchet@49156
   464
   iterate. *)
blanchet@49157
   465
fun all_class_pairs _ _ [] = ([], [])
blanchet@49157
   466
  | all_class_pairs thy tycons cls =
blanchet@49156
   467
    let
blanchet@49156
   468
      fun maybe_insert_class s =
blanchet@49157
   469
        (s <> class_of_types andalso not (member (op =) cls s))
blanchet@49156
   470
        ? insert (op =) s
blanchet@49157
   471
      val pairs = class_pairs thy tycons cls
blanchet@49157
   472
      val new_cls =
blanchet@49157
   473
        [] |> fold (fold (fold (fold maybe_insert_class) o snd) o snd) pairs
blanchet@49157
   474
      val (cls', pairs') = all_class_pairs thy tycons new_cls
blanchet@49157
   475
    in (cls' @ cls, union (op =) pairs' pairs) end
blanchet@49156
   476
blanchet@49157
   477
fun tcon_clause _ _ (_, []) = []
blanchet@49157
   478
  | tcon_clause seen n (tcons, (ar as (cl, _)) :: ars) =
blanchet@49157
   479
    if cl = class_of_types then
blanchet@49157
   480
      tcon_clause seen n (tcons, ars)
blanchet@49157
   481
    else if member (op =) seen cl then
blanchet@49157
   482
      (* multiple clauses for the same (tycon, cl) pair *)
blanchet@49157
   483
      make_axiom_tcon_clause (tcons,
blanchet@49157
   484
          lookup_const tcons ^ "___" ^ ascii_of cl ^ "_" ^ string_of_int n,
blanchet@44366
   485
          ar) ::
blanchet@49157
   486
      tcon_clause seen (n + 1) (tcons, ars)
blanchet@44366
   487
    else
blanchet@49157
   488
      make_axiom_tcon_clause (tcons, lookup_const tcons ^ "___" ^ ascii_of cl,
blanchet@49157
   489
                              ar) ::
blanchet@49157
   490
      tcon_clause (cl :: seen) n (tcons, ars)
blanchet@43926
   491
blanchet@49157
   492
fun make_tcon_clauses thy tycons =
blanchet@49157
   493
  all_class_pairs thy tycons ##> maps (tcon_clause [] 1)
blanchet@43926
   494
blanchet@43926
   495
blanchet@43926
   496
(** Isabelle class relations **)
blanchet@43926
   497
blanchet@49157
   498
(* Generate a list ("sub", "supers") such that "sub" is a proper subclass of all
blanchet@49157
   499
   "supers". *)
blanchet@49157
   500
fun make_subclass_pairs thy subs supers =
blanchet@49156
   501
  let
blanchet@49157
   502
    val class_less = curry (Sorts.class_less (Sign.classes_of thy))
blanchet@49157
   503
    fun supers_of sub = (sub, filter (class_less sub) supers)
blanchet@49157
   504
  in map supers_of subs |> filter_out (null o snd) end
blanchet@43926
   505
blanchet@44730
   506
(* intermediate terms *)
blanchet@44730
   507
datatype iterm =
blanchet@49150
   508
  IConst of (string * string) * typ * typ list |
blanchet@49150
   509
  IVar of (string * string) * typ |
blanchet@44730
   510
  IApp of iterm * iterm |
blanchet@49150
   511
  IAbs of ((string * string) * typ) * iterm
blanchet@43926
   512
blanchet@44730
   513
fun ityp_of (IConst (_, T, _)) = T
blanchet@44730
   514
  | ityp_of (IVar (_, T)) = T
blanchet@44730
   515
  | ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1))
blanchet@44730
   516
  | ityp_of (IAbs ((_, T), tm)) = T --> ityp_of tm
blanchet@43926
   517
blanchet@43926
   518
(*gets the head of a combinator application, along with the list of arguments*)
blanchet@44730
   519
fun strip_iterm_comb u =
blanchet@44367
   520
  let
blanchet@44730
   521
    fun stripc (IApp (t, u), ts) = stripc (t, u :: ts)
blanchet@44367
   522
      | stripc x = x
blanchet@44367
   523
  in stripc (u, []) end
blanchet@43926
   524
blanchet@46187
   525
fun atomic_types_of T = fold_atyps (insert (op =)) T []
blanchet@43926
   526
blanchet@43926
   527
fun new_skolem_const_name s num_T_args =
blanchet@43926
   528
  [new_skolem_const_prefix, s, string_of_int num_T_args]
wenzelm@47582
   529
  |> Long_Name.implode
blanchet@43926
   530
blanchet@48947
   531
val alpha_to_beta = Logic.varifyT_global @{typ "'a => 'b"}
blanchet@48947
   532
val alpha_to_beta_to_alpha_to_beta = alpha_to_beta --> alpha_to_beta
blanchet@48947
   533
blanchet@45458
   534
fun robust_const_type thy s =
blanchet@46380
   535
  if s = app_op_name then
blanchet@48947
   536
    alpha_to_beta_to_alpha_to_beta
blanchet@46425
   537
  else if String.isPrefix lam_lifted_prefix s then
blanchet@48947
   538
    alpha_to_beta
blanchet@46380
   539
  else
blanchet@46380
   540
    (* Old Skolems throw a "TYPE" exception here, which will be caught. *)
blanchet@46380
   541
    s |> Sign.the_const_type thy
blanchet@45458
   542
blanchet@47513
   543
val robust_const_ary =
blanchet@47513
   544
  let
blanchet@47513
   545
    fun ary (Type (@{type_name fun}, [_, T])) = 1 + ary T
blanchet@47513
   546
      | ary _ = 0
blanchet@47513
   547
  in ary oo robust_const_type end
blanchet@47513
   548
blanchet@45458
   549
(* This function only makes sense if "T" is as general as possible. *)
blanchet@45458
   550
fun robust_const_typargs thy (s, T) =
blanchet@46380
   551
  if s = app_op_name then
blanchet@46380
   552
    let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end
blanchet@46380
   553
  else if String.isPrefix old_skolem_const_prefix s then
blanchet@46380
   554
    [] |> Term.add_tvarsT T |> rev |> map TVar
blanchet@46425
   555
  else if String.isPrefix lam_lifted_prefix s then
blanchet@46425
   556
    if String.isPrefix lam_lifted_poly_prefix s then
blanchet@46382
   557
      let val (T1, T2) = T |> dest_funT in [T1, T2] end
blanchet@46382
   558
    else
blanchet@46382
   559
      []
blanchet@46380
   560
  else
blanchet@46380
   561
    (s, T) |> Sign.const_typargs thy
blanchet@45458
   562
blanchet@44730
   563
(* Converts an Isabelle/HOL term (with combinators) into an intermediate term.
blanchet@44730
   564
   Also accumulates sort infomation. *)
blanchet@48638
   565
fun iterm_from_term thy type_enc bs (P $ Q) =
blanchet@43926
   566
    let
blanchet@48638
   567
      val (P', P_atomics_Ts) = iterm_from_term thy type_enc bs P
blanchet@48638
   568
      val (Q', Q_atomics_Ts) = iterm_from_term thy type_enc bs Q
blanchet@44730
   569
    in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end
blanchet@48638
   570
  | iterm_from_term thy type_enc _ (Const (c, T)) =
blanchet@48638
   571
    (IConst (`(make_fixed_const (SOME type_enc)) c, T,
blanchet@45458
   572
             robust_const_typargs thy (c, T)),
blanchet@46187
   573
     atomic_types_of T)
nik@45350
   574
  | iterm_from_term _ _ _ (Free (s, T)) =
blanchet@46380
   575
    (IConst (`make_fixed_var s, T, []), atomic_types_of T)
blanchet@48638
   576
  | iterm_from_term _ type_enc _ (Var (v as (s, _), T)) =
blanchet@43926
   577
    (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
blanchet@43926
   578
       let
blanchet@43926
   579
         val Ts = T |> strip_type |> swap |> op ::
blanchet@43926
   580
         val s' = new_skolem_const_name s (length Ts)
blanchet@48638
   581
       in IConst (`(make_fixed_const (SOME type_enc)) s', T, Ts) end
blanchet@43926
   582
     else
blanchet@46187
   583
       IVar ((make_schematic_var v, s), T), atomic_types_of T)
nik@45350
   584
  | iterm_from_term _ _ bs (Bound j) =
blanchet@46187
   585
    nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T))
blanchet@48638
   586
  | iterm_from_term thy type_enc bs (Abs (s, T, t)) =
nik@44537
   587
    let
nik@44537
   588
      fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string
nik@44537
   589
      val s = vary s
blanchet@45262
   590
      val name = `make_bound_var s
blanchet@48638
   591
      val (tm, atomic_Ts) =
blanchet@48638
   592
        iterm_from_term thy type_enc ((s, (name, T)) :: bs) t
blanchet@46187
   593
    in (IAbs ((name, T), tm), union (op =) atomic_Ts (atomic_types_of T)) end
blanchet@43926
   594
blanchet@49104
   595
(* "_query" and "_at" are for the ASCII-challenged Metis and Mirabelle. *)
blanchet@45653
   596
val queries = ["?", "_query"]
blanchet@45653
   597
val ats = ["@", "_at"]
blanchet@45653
   598
blanchet@43559
   599
fun try_unsuffixes ss s =
blanchet@43559
   600
  fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE
blanchet@43559
   601
blanchet@47129
   602
fun type_enc_from_string strictness s =
blanchet@49146
   603
  (case try (unprefix "tc_") s of
blanchet@49146
   604
     SOME s => (SOME Type_Class_Polymorphic, s)
blanchet@43484
   605
   | NONE =>
blanchet@49146
   606
       case try (unprefix "poly_") s of
blanchet@49146
   607
         (* It's still unclear whether all TFF1 implementations will support
blanchet@49146
   608
            type signatures such as "!>[A : $tType] : $o", with phantom type
blanchet@49146
   609
            variables. *)
blanchet@49146
   610
         SOME s => (SOME (Raw_Polymorphic With_Phantom_Type_Vars), s)
blanchet@49146
   611
       | NONE =>
blanchet@49146
   612
         case try (unprefix "raw_mono_") s of
blanchet@49146
   613
           SOME s => (SOME Raw_Monomorphic, s)
blanchet@49146
   614
         | NONE =>
blanchet@49146
   615
           case try (unprefix "mono_") s of
blanchet@49146
   616
             SOME s => (SOME Mangled_Monomorphic, s)
blanchet@49146
   617
           | NONE => (NONE, s))
blanchet@49107
   618
  ||> (fn s =>
blanchet@49107
   619
          case try_unsuffixes queries s of
blanchet@49107
   620
          SOME s =>
blanchet@49107
   621
          (case try_unsuffixes queries s of
blanchet@49198
   622
             SOME s => (Nonmono_Types (strictness, Non_Uniform), s)
blanchet@49198
   623
           | NONE => (Nonmono_Types (strictness, Uniform), s))
blanchet@49161
   624
         | NONE =>
blanchet@49161
   625
           (case try_unsuffixes ats s of
blanchet@49198
   626
              SOME s => (Undercover_Types, s)
blanchet@49161
   627
            | NONE => (All_Types, s)))
blanchet@45639
   628
  |> (fn (poly, (level, core)) =>
blanchet@45639
   629
         case (core, (poly, level)) of
blanchet@47263
   630
           ("native", (SOME poly, _)) =>
blanchet@45606
   631
           (case (poly, level) of
blanchet@49146
   632
              (Mangled_Monomorphic, _) =>
blanchet@49198
   633
              if is_type_level_uniform level then
blanchet@48637
   634
                Native (First_Order, Mangled_Monomorphic, level)
blanchet@45639
   635
              else
blanchet@45639
   636
                raise Same.SAME
blanchet@49146
   637
            | (Raw_Monomorphic, _) => raise Same.SAME
blanchet@49146
   638
            | (poly, All_Types) => Native (First_Order, poly, All_Types))
blanchet@47263
   639
         | ("native_higher", (SOME poly, _)) =>
blanchet@45455
   640
           (case (poly, level) of
blanchet@49146
   641
              (_, Nonmono_Types _) => raise Same.SAME
blanchet@49198
   642
            | (_, Undercover_Types) => raise Same.SAME
blanchet@45606
   643
            | (Mangled_Monomorphic, _) =>
blanchet@49198
   644
              if is_type_level_uniform level then
blanchet@48638
   645
                Native (Higher_Order THF_With_Choice, Mangled_Monomorphic,
blanchet@48638
   646
                        level)
blanchet@45639
   647
              else
blanchet@45639
   648
                raise Same.SAME
blanchet@49149
   649
            | (poly as Raw_Polymorphic _, All_Types) =>
blanchet@49149
   650
              Native (Higher_Order THF_With_Choice, poly, All_Types)
blanchet@49149
   651
            | _ => raise Same.SAME)
blanchet@45672
   652
         | ("guards", (SOME poly, _)) =>
blanchet@49149
   653
           if (poly = Mangled_Monomorphic andalso
blanchet@49198
   654
               level = Undercover_Types) orelse
blanchet@49149
   655
              poly = Type_Class_Polymorphic then
blanchet@46820
   656
             raise Same.SAME
blanchet@46820
   657
           else
blanchet@46820
   658
             Guards (poly, level)
blanchet@45672
   659
         | ("tags", (SOME poly, _)) =>
blanchet@49161
   660
           if (poly = Mangled_Monomorphic andalso
blanchet@49198
   661
               level = Undercover_Types) orelse
blanchet@49149
   662
              poly = Type_Class_Polymorphic then
blanchet@46820
   663
             raise Same.SAME
blanchet@46820
   664
           else
blanchet@46820
   665
             Tags (poly, level)
blanchet@45639
   666
         | ("args", (SOME poly, All_Types (* naja *))) =>
blanchet@49149
   667
           if poly = Type_Class_Polymorphic then raise Same.SAME
blanchet@49198
   668
           else Guards (poly, Const_Types Without_Constr_Optim)
blanchet@49198
   669
         | ("args", (SOME poly, Nonmono_Types (_, Uniform) (* naja *))) =>
blanchet@49149
   670
           if poly = Mangled_Monomorphic orelse
blanchet@49149
   671
              poly = Type_Class_Polymorphic then
blanchet@49149
   672
             raise Same.SAME
blanchet@49149
   673
           else
blanchet@49198
   674
             Guards (poly, Const_Types With_Constr_Optim)
blanchet@45639
   675
         | ("erased", (NONE, All_Types (* naja *))) =>
blanchet@49146
   676
           Guards (Raw_Polymorphic With_Phantom_Type_Vars, No_Types)
blanchet@43618
   677
         | _ => raise Same.SAME)
blanchet@45653
   678
  handle Same.SAME => error ("Unknown type encoding: " ^ quote s ^ ".")
blanchet@43484
   679
blanchet@48638
   680
fun adjust_order THF_Without_Choice (Higher_Order _) =
blanchet@48638
   681
    Higher_Order THF_Without_Choice
blanchet@48638
   682
  | adjust_order _ type_enc = type_enc
blanchet@48638
   683
blanchet@49149
   684
fun no_type_classes Type_Class_Polymorphic =
blanchet@49149
   685
    Raw_Polymorphic With_Phantom_Type_Vars
blanchet@49149
   686
  | no_type_classes poly = poly
blanchet@49149
   687
blanchet@49145
   688
fun adjust_type_enc (THF (Polymorphic, _, choice, _))
blanchet@48638
   689
                    (Native (order, poly, level)) =
blanchet@49149
   690
    Native (adjust_order choice order, no_type_classes poly, level)
blanchet@49145
   691
  | adjust_type_enc (THF (Monomorphic, _, choice, _))
blanchet@48638
   692
                         (Native (order, _, level)) =
blanchet@49019
   693
    Native (adjust_order choice order, Mangled_Monomorphic, level)
blanchet@49145
   694
  | adjust_type_enc (TFF (Monomorphic, _)) (Native (_, _, level)) =
blanchet@48637
   695
    Native (First_Order, Mangled_Monomorphic, level)
blanchet@49146
   696
  | adjust_type_enc (DFG Polymorphic) (Native (_, poly, level)) =
blanchet@49146
   697
    Native (First_Order, poly, level)
blanchet@49146
   698
  | adjust_type_enc (DFG Monomorphic) (Native (_, _, level)) =
blanchet@48637
   699
    Native (First_Order, Mangled_Monomorphic, level)
blanchet@48637
   700
  | adjust_type_enc (TFF _) (Native (_, poly, level)) =
blanchet@49149
   701
    Native (First_Order, no_type_classes poly, level)
blanchet@48637
   702
  | adjust_type_enc format (Native (_, poly, level)) =
blanchet@49149
   703
    adjust_type_enc format (Guards (no_type_classes poly, level))
blanchet@45275
   704
  | adjust_type_enc CNF_UEQ (type_enc as Guards stuff) =
blanchet@49104
   705
    (if is_type_enc_sound type_enc then Tags else Guards) stuff
blanchet@45275
   706
  | adjust_type_enc _ type_enc = type_enc
blanchet@43942
   707
blanchet@49317
   708
fun is_lambda_free t =
blanchet@47689
   709
  case t of
blanchet@49317
   710
    @{const Not} $ t1 => is_lambda_free t1
blanchet@49317
   711
  | Const (@{const_name All}, _) $ Abs (_, _, t') => is_lambda_free t'
blanchet@49317
   712
  | Const (@{const_name All}, _) $ t1 => is_lambda_free t1
blanchet@49317
   713
  | Const (@{const_name Ex}, _) $ Abs (_, _, t') => is_lambda_free t'
blanchet@49317
   714
  | Const (@{const_name Ex}, _) $ t1 => is_lambda_free t1
blanchet@49317
   715
  | @{const HOL.conj} $ t1 $ t2 => is_lambda_free t1 andalso is_lambda_free t2
blanchet@49317
   716
  | @{const HOL.disj} $ t1 $ t2 => is_lambda_free t1 andalso is_lambda_free t2
blanchet@49317
   717
  | @{const HOL.implies} $ t1 $ t2 => is_lambda_free t1 andalso is_lambda_free t2
blanchet@47689
   718
  | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@49317
   719
    is_lambda_free t1 andalso is_lambda_free t2
blanchet@47689
   720
  | _ => not (exists_subterm (fn Abs _ => true | _ => false) t)
blanchet@47689
   721
blanchet@47689
   722
fun simple_translate_lambdas do_lambdas ctxt t =
blanchet@49317
   723
  if is_lambda_free t then
blanchet@47689
   724
    t
blanchet@47689
   725
  else
blanchet@47689
   726
    let
blanchet@47689
   727
      fun trans Ts t =
blanchet@47689
   728
        case t of
blanchet@47689
   729
          @{const Not} $ t1 => @{const Not} $ trans Ts t1
blanchet@47689
   730
        | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
blanchet@47689
   731
          t0 $ Abs (s, T, trans (T :: Ts) t')
blanchet@47689
   732
        | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@47689
   733
          trans Ts (t0 $ eta_expand Ts t1 1)
blanchet@47689
   734
        | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
blanchet@47689
   735
          t0 $ Abs (s, T, trans (T :: Ts) t')
blanchet@47689
   736
        | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@47689
   737
          trans Ts (t0 $ eta_expand Ts t1 1)
blanchet@47689
   738
        | (t0 as @{const HOL.conj}) $ t1 $ t2 =>
blanchet@47689
   739
          t0 $ trans Ts t1 $ trans Ts t2
blanchet@47689
   740
        | (t0 as @{const HOL.disj}) $ t1 $ t2 =>
blanchet@47689
   741
          t0 $ trans Ts t1 $ trans Ts t2
blanchet@47689
   742
        | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
blanchet@47689
   743
          t0 $ trans Ts t1 $ trans Ts t2
blanchet@47689
   744
        | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
blanchet@47689
   745
            $ t1 $ t2 =>
blanchet@47689
   746
          t0 $ trans Ts t1 $ trans Ts t2
blanchet@47689
   747
        | _ =>
blanchet@47689
   748
          if not (exists_subterm (fn Abs _ => true | _ => false) t) then t
blanchet@47689
   749
          else t |> Envir.eta_contract |> do_lambdas ctxt Ts
blanchet@47689
   750
      val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
blanchet@47689
   751
    in t |> trans [] |> singleton (Variable.export_terms ctxt' ctxt) end
blanchet@47689
   752
blanchet@47689
   753
fun do_cheaply_conceal_lambdas Ts (t1 $ t2) =
blanchet@47689
   754
    do_cheaply_conceal_lambdas Ts t1
blanchet@47689
   755
    $ do_cheaply_conceal_lambdas Ts t2
blanchet@47689
   756
  | do_cheaply_conceal_lambdas Ts (Abs (_, T, t)) =
blanchet@47689
   757
    Const (lam_lifted_poly_prefix ^ serial_string (),
blanchet@47689
   758
           T --> fastype_of1 (T :: Ts, t))
blanchet@47689
   759
  | do_cheaply_conceal_lambdas _ t = t
blanchet@47689
   760
blanchet@47689
   761
fun concealed_bound_name j = atp_weak_prefix ^ string_of_int j
blanchet@47689
   762
fun conceal_bounds Ts t =
blanchet@47689
   763
  subst_bounds (map (Free o apfst concealed_bound_name)
blanchet@47689
   764
                    (0 upto length Ts - 1 ~~ Ts), t)
blanchet@47689
   765
fun reveal_bounds Ts =
blanchet@47689
   766
  subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
blanchet@47689
   767
                    (0 upto length Ts - 1 ~~ Ts))
blanchet@47689
   768
blanchet@47689
   769
fun do_introduce_combinators ctxt Ts t =
blanchet@47689
   770
  let val thy = Proof_Context.theory_of ctxt in
blanchet@47689
   771
    t |> conceal_bounds Ts
blanchet@47689
   772
      |> cterm_of thy
blanchet@47689
   773
      |> Meson_Clausify.introduce_combinators_in_cterm
blanchet@47689
   774
      |> prop_of |> Logic.dest_equals |> snd
blanchet@47689
   775
      |> reveal_bounds Ts
blanchet@47689
   776
  end
blanchet@47689
   777
  (* A type variable of sort "{}" will make abstraction fail. *)
blanchet@47689
   778
  handle THM _ => t |> do_cheaply_conceal_lambdas Ts
blanchet@47689
   779
val introduce_combinators = simple_translate_lambdas do_introduce_combinators
blanchet@47689
   780
blanchet@46380
   781
fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u
blanchet@46380
   782
  | constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t)
blanchet@46380
   783
  | constify_lifted (Free (x as (s, _))) =
blanchet@46425
   784
    (if String.isPrefix lam_lifted_prefix s then Const else Free) x
blanchet@46380
   785
  | constify_lifted t = t
blanchet@46380
   786
blanchet@46425
   787
fun lift_lams_part_1 ctxt type_enc =
blanchet@46439
   788
  map close_form #> rpair ctxt
blanchet@44959
   789
  #-> Lambda_Lifting.lift_lambdas
blanchet@49146
   790
          (SOME ((if is_type_enc_polymorphic type_enc then
blanchet@46435
   791
                    lam_lifted_poly_prefix
blanchet@46435
   792
                  else
blanchet@46435
   793
                    lam_lifted_mono_prefix) ^ "_a"))
blanchet@44959
   794
          Lambda_Lifting.is_quantifier
blanchet@46425
   795
  #> fst
blanchet@47689
   796
blanchet@47689
   797
fun lift_lams_part_2 ctxt (facts, lifted) =
blanchet@47689
   798
  (facts, lifted)
blanchet@47689
   799
  (* Lambda-lifting sometimes leaves some lambdas around; we need some way to get rid
blanchet@47689
   800
     of them *)
blanchet@47689
   801
  |> pairself (map (introduce_combinators ctxt))
blanchet@47689
   802
  |> pairself (map constify_lifted)
blanchet@48589
   803
  (* Requires bound variables not to clash with any schematic variables (as
blanchet@48589
   804
     should be the case right after lambda-lifting). *)
blanchet@47689
   805
  |>> map (open_form (unprefix close_form_prefix))
blanchet@47689
   806
  ||> map (open_form I)
blanchet@47689
   807
blanchet@47689
   808
fun lift_lams ctxt = lift_lams_part_2 ctxt oo lift_lams_part_1 ctxt
blanchet@44959
   809
blanchet@44959
   810
fun intentionalize_def (Const (@{const_name All}, _) $ Abs (_, _, t)) =
blanchet@44959
   811
    intentionalize_def t
blanchet@44959
   812
  | intentionalize_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
blanchet@44959
   813
    let
blanchet@44959
   814
      fun lam T t = Abs (Name.uu, T, t)
blanchet@44959
   815
      val (head, args) = strip_comb t ||> rev
blanchet@44959
   816
      val head_T = fastype_of head
blanchet@44959
   817
      val n = length args
blanchet@44959
   818
      val arg_Ts = head_T |> binder_types |> take n |> rev
blanchet@44959
   819
      val u = u |> subst_atomic (args ~~ map Bound (0 upto n - 1))
blanchet@44959
   820
    in HOLogic.eq_const head_T $ head $ fold lam arg_Ts u end
blanchet@44959
   821
  | intentionalize_def t = t
blanchet@44959
   822
blanchet@48996
   823
type ifact =
blanchet@44367
   824
  {name : string,
blanchet@47168
   825
   stature : stature,
blanchet@48990
   826
   role : formula_role,
blanchet@49150
   827
   iformula : (string * string, typ, iterm, string * string) formula,
blanchet@44367
   828
   atomic_types : typ list}
blanchet@38506
   829
blanchet@48996
   830
fun update_iformula f ({name, stature, role, iformula, atomic_types} : ifact) =
blanchet@48990
   831
  {name = name, stature = stature, role = role, iformula = f iformula,
blanchet@48996
   832
   atomic_types = atomic_types} : ifact
blanchet@43413
   833
blanchet@48996
   834
fun ifact_lift f ({iformula, ...} : ifact) = f iformula
blanchet@43429
   835
blanchet@48021
   836
fun insert_type thy get_T x xs =
blanchet@43905
   837
  let val T = get_T x in
blanchet@48021
   838
    if exists (type_instance thy T o get_T) xs then xs
blanchet@48021
   839
    else x :: filter_out (type_generalization thy T o get_T) xs
blanchet@43905
   840
  end
blanchet@43547
   841
blanchet@49217
   842
fun chop_fun 0 T = ([], T)
blanchet@49217
   843
  | chop_fun n (Type (@{type_name fun}, [dom_T, ran_T])) =
blanchet@49217
   844
    chop_fun (n - 1) ran_T |>> cons dom_T
blanchet@49217
   845
  | chop_fun _ T = ([], T)
blanchet@41384
   846
blanchet@49217
   847
fun filter_type_args thy constrs type_enc s ary T_args =
blanchet@46186
   848
  let val poly = polymorphism_of_type_enc type_enc in
blanchet@49216
   849
    if s = type_tag_name then (* ### FIXME: why not "type_guard_name" as well? *)
blanchet@49217
   850
      T_args
blanchet@45645
   851
    else case type_enc of
blanchet@49217
   852
      Native (_, Raw_Polymorphic _, _) => T_args
blanchet@49217
   853
    | Native (_, Type_Class_Polymorphic, _) => T_args
blanchet@45645
   854
    | _ =>
blanchet@49217
   855
      let
blanchet@49217
   856
        fun gen_type_args _ _ [] = []
blanchet@49217
   857
          | gen_type_args keep strip_ty T_args =
blanchet@49217
   858
            let
blanchet@49217
   859
              val U = robust_const_type thy s
blanchet@49217
   860
              val (binder_Us, body_U) = strip_ty U
blanchet@49217
   861
              val in_U_vars = fold Term.add_tvarsT binder_Us []
blanchet@49217
   862
              val out_U_vars = Term.add_tvarsT body_U []
blanchet@49217
   863
              fun filt (U_var, T) =
blanchet@49217
   864
                if keep (member (op =) in_U_vars U_var,
blanchet@49217
   865
                         member (op =) out_U_vars U_var) then
blanchet@49217
   866
                  T
blanchet@49217
   867
                else
blanchet@49217
   868
                  dummyT
blanchet@49217
   869
              val U_args = (s, U) |> robust_const_typargs thy
blanchet@49217
   870
            in map (filt o apfst dest_TVar) (U_args ~~ T_args) end
blanchet@49217
   871
            handle TYPE _ => T_args
blanchet@49217
   872
        val noninfer_type_args = gen_type_args (not o fst) (chop_fun ary)
blanchet@49217
   873
        val constr_infer_type_args = gen_type_args fst strip_type
blanchet@49217
   874
        val level = level_of_type_enc type_enc
blanchet@49217
   875
      in
blanchet@45645
   876
        if level = No_Types orelse s = @{const_name HOL.eq} orelse
blanchet@49107
   877
           (case level of Const_Types _ => s = app_op_name | _ => false) then
blanchet@49217
   878
          []
blanchet@46186
   879
        else if poly = Mangled_Monomorphic then
blanchet@49217
   880
          T_args
blanchet@49215
   881
        else if level = All_Types then
blanchet@49215
   882
          case type_enc of
blanchet@49217
   883
            Guards _ => noninfer_type_args T_args
blanchet@49217
   884
          | Tags _ => []
blanchet@49215
   885
        else if level = Undercover_Types then
blanchet@49217
   886
          noninfer_type_args T_args
blanchet@49198
   887
        else if member (op =) constrs s andalso
blanchet@49198
   888
                level <> Const_Types Without_Constr_Optim then
blanchet@49217
   889
          constr_infer_type_args T_args
blanchet@45645
   890
        else
blanchet@49217
   891
          T_args
blanchet@45645
   892
      end
blanchet@45645
   893
  end
blanchet@43088
   894
blanchet@47166
   895
val fused_infinite_type_name = "ATP.fused_inf" (* shouldn't clash *)
blanchet@45458
   896
val fused_infinite_type = Type (fused_infinite_type_name, [])
blanchet@45458
   897
blanchet@49153
   898
fun raw_ho_type_from_typ type_enc =
blanchet@43835
   899
  let
blanchet@43835
   900
    fun term (Type (s, Ts)) =
blanchet@49153
   901
      AType (case (is_type_enc_higher_order type_enc, s) of
blanchet@49153
   902
               (true, @{type_name bool}) => `I tptp_bool_type
blanchet@49153
   903
             | (true, @{type_name fun}) => `I tptp_fun_type
blanchet@49153
   904
             | _ => if s = fused_infinite_type_name andalso
blanchet@49153
   905
                       is_type_enc_native type_enc then
blanchet@49153
   906
                      `I tptp_individual_type
blanchet@49153
   907
                    else
blanchet@49153
   908
                      `make_fixed_type_const s,
blanchet@49153
   909
             map term Ts)
blanchet@49153
   910
    | term (TFree (s, _)) = AType (`make_tfree s, [])
blanchet@49157
   911
    | term (TVar z) = AType (tvar_name z, [])
blanchet@43835
   912
  in term end
blanchet@43433
   913
blanchet@49153
   914
fun ho_term_from_ho_type (AType (name, tys)) = ATerm ((name, []), map ho_term_from_ho_type tys)
blanchet@49153
   915
  | ho_term_from_ho_type _ = raise Fail "unexpected type"
blanchet@49153
   916
blanchet@49153
   917
fun ho_type_for_type_arg type_enc T =
blanchet@49153
   918
  if T = dummyT then NONE else SOME (raw_ho_type_from_typ type_enc T)
blanchet@44264
   919
blanchet@43433
   920
(* This shouldn't clash with anything else. *)
blanchet@47237
   921
val uncurried_alias_sep = "\000"
blanchet@47220
   922
val mangled_type_sep = "\001"
blanchet@47220
   923
blanchet@47237
   924
val ascii_of_uncurried_alias_sep = ascii_of uncurried_alias_sep
blanchet@43413
   925
blanchet@49153
   926
fun generic_mangled_type_name f (AType (name, [])) = f name
blanchet@49153
   927
  | generic_mangled_type_name f (AType (name, tys)) =
blanchet@43626
   928
    f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys)
blanchet@43626
   929
    ^ ")"
blanchet@49153
   930
  | generic_mangled_type_name _ _ = raise Fail "unexpected type"
blanchet@43413
   931
blanchet@48638
   932
fun mangled_type type_enc =
blanchet@49153
   933
  generic_mangled_type_name fst o raw_ho_type_from_typ type_enc
blanchet@45255
   934
blanchet@47263
   935
fun make_native_type s =
blanchet@43926
   936
  if s = tptp_bool_type orelse s = tptp_fun_type orelse
blanchet@43926
   937
     s = tptp_individual_type then
blanchet@43926
   938
    s
blanchet@43926
   939
  else
blanchet@47263
   940
    native_type_prefix ^ ascii_of s
blanchet@43926
   941
blanchet@49153
   942
fun native_ho_type_from_raw_ho_type type_enc pred_sym ary =
blanchet@43804
   943
  let
blanchet@45457
   944
    fun to_mangled_atype ty =
blanchet@47263
   945
      AType ((make_native_type (generic_mangled_type_name fst ty),
blanchet@45457
   946
              generic_mangled_type_name snd ty), [])
blanchet@49153
   947
    fun to_poly_atype (AType (name, tys)) =
blanchet@49147
   948
        AType (name, map to_poly_atype tys)
blanchet@49153
   949
      | to_poly_atype _ = raise Fail "unexpected type"
blanchet@45457
   950
    val to_atype =
blanchet@49146
   951
      if is_type_enc_polymorphic type_enc then to_poly_atype
blanchet@45457
   952
      else to_mangled_atype
blanchet@43804
   953
    fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1))
blanchet@43839
   954
    fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty
blanchet@49153
   955
      | to_fo ary (AType (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys
blanchet@49153
   956
      | to_fo _ _ = raise Fail "unexpected type"
blanchet@49153
   957
    fun to_ho (ty as AType ((s, _), tys)) =
nik@44535
   958
        if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty
blanchet@49153
   959
      | to_ho _ = raise Fail "unexpected type"
blanchet@44493
   960
  in if is_type_enc_higher_order type_enc then to_ho else to_fo ary end
blanchet@43804
   961
blanchet@49153
   962
fun native_ho_type_from_typ type_enc pred_sym ary =
blanchet@49153
   963
  native_ho_type_from_raw_ho_type type_enc pred_sym ary
blanchet@49153
   964
  o raw_ho_type_from_typ type_enc
blanchet@43804
   965
blanchet@49148
   966
(* Make atoms for sorted type variables. *)
blanchet@49148
   967
fun generic_add_sorts_on_type _ [] = I
blanchet@49148
   968
  | generic_add_sorts_on_type T (s :: ss) =
blanchet@49148
   969
    generic_add_sorts_on_type T ss
blanchet@49157
   970
    #> (if s = the_single @{sort type} then I else insert (op =) (s, T))
blanchet@49148
   971
fun add_sorts_on_tfree (T as TFree (_, S)) = generic_add_sorts_on_type T S
blanchet@49148
   972
  | add_sorts_on_tfree _ = I
blanchet@49148
   973
fun add_sorts_on_tvar (T as TVar (_, S)) = generic_add_sorts_on_type T S
blanchet@49148
   974
  | add_sorts_on_tvar _ = I
blanchet@49148
   975
blanchet@49148
   976
fun process_type_args type_enc T_args =
blanchet@49148
   977
  if is_type_enc_native type_enc then
blanchet@49153
   978
    (map (native_ho_type_from_typ type_enc false 0) T_args, [])
blanchet@49148
   979
  else
blanchet@49153
   980
    ([], map_filter (Option.map ho_term_from_ho_type
blanchet@49153
   981
                     o ho_type_for_type_arg type_enc) T_args)
blanchet@49148
   982
blanchet@49157
   983
fun class_atom type_enc (cl, T) =
blanchet@49148
   984
  let
blanchet@49157
   985
    val cl = `make_class cl
blanchet@49148
   986
    val (ty_args, tm_args) = process_type_args type_enc [T]
blanchet@49148
   987
    val tm_args =
blanchet@49148
   988
      tm_args @
blanchet@49148
   989
      (case type_enc of
blanchet@49148
   990
         Native (First_Order, Raw_Polymorphic Without_Phantom_Type_Vars, _) =>
blanchet@49148
   991
         [ATerm ((TYPE_name, ty_args), [])]
blanchet@49148
   992
       | _ => [])
blanchet@49157
   993
  in AAtom (ATerm ((cl, ty_args), tm_args)) end
blanchet@49157
   994
blanchet@49157
   995
fun class_atoms type_enc (cls, T) =
blanchet@49157
   996
  map (fn cl => class_atom type_enc (cl, T)) cls
blanchet@49157
   997
blanchet@49157
   998
fun class_membs_for_types type_enc add_sorts_on_typ Ts =
blanchet@49156
   999
  [] |> (polymorphism_of_type_enc type_enc <> Type_Class_Polymorphic andalso
blanchet@49156
  1000
         level_of_type_enc type_enc <> No_Types)
blanchet@49156
  1001
        ? fold add_sorts_on_typ Ts
blanchet@49148
  1002
blanchet@49156
  1003
fun mk_aconns c = split_last #> uncurry (fold_rev (mk_aconn c))
blanchet@49148
  1004
blanchet@49148
  1005
fun mk_ahorn [] phi = phi
blanchet@49148
  1006
  | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi])
blanchet@49148
  1007
blanchet@49148
  1008
fun mk_aquant _ [] phi = phi
blanchet@49148
  1009
  | mk_aquant q xs (phi as AQuant (q', xs', phi')) =
blanchet@49148
  1010
    if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi)
blanchet@49148
  1011
  | mk_aquant q xs phi = AQuant (q, xs, phi)
blanchet@49148
  1012
blanchet@49148
  1013
fun mk_atyquant _ [] phi = phi
blanchet@49148
  1014
  | mk_atyquant q xs (phi as ATyQuant (q', xs', phi')) =
blanchet@49148
  1015
    if q = q' then ATyQuant (q, xs @ xs', phi') else ATyQuant (q, xs, phi)
blanchet@49148
  1016
  | mk_atyquant q xs phi = ATyQuant (q, xs, phi)
blanchet@49148
  1017
blanchet@49148
  1018
fun close_universally add_term_vars phi =
blanchet@49148
  1019
  let
blanchet@49148
  1020
    fun add_formula_vars bounds (ATyQuant (_, _, phi)) =
blanchet@49148
  1021
        add_formula_vars bounds phi
blanchet@49148
  1022
      | add_formula_vars bounds (AQuant (_, xs, phi)) =
blanchet@49148
  1023
        add_formula_vars (map fst xs @ bounds) phi
blanchet@49148
  1024
      | add_formula_vars bounds (AConn (_, phis)) =
blanchet@49148
  1025
        fold (add_formula_vars bounds) phis
blanchet@49148
  1026
      | add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm
blanchet@49148
  1027
  in mk_aquant AForall (add_formula_vars [] phi []) phi end
blanchet@49148
  1028
blanchet@49148
  1029
fun add_term_vars bounds (ATerm ((name as (s, _), _), tms)) =
blanchet@49148
  1030
    (if is_tptp_variable s andalso
blanchet@49148
  1031
        not (String.isPrefix tvar_prefix s) andalso
blanchet@49148
  1032
        not (member (op =) bounds name) then
blanchet@49148
  1033
       insert (op =) (name, NONE)
blanchet@49148
  1034
     else
blanchet@49148
  1035
       I)
blanchet@49148
  1036
    #> fold (add_term_vars bounds) tms
blanchet@49148
  1037
  | add_term_vars bounds (AAbs (((name, _), tm), args)) =
blanchet@49148
  1038
    add_term_vars (name :: bounds) tm #> fold (add_term_vars bounds) args
blanchet@49148
  1039
fun close_formula_universally phi = close_universally add_term_vars phi
blanchet@49148
  1040
blanchet@49148
  1041
fun add_iterm_vars bounds (IApp (tm1, tm2)) =
blanchet@49148
  1042
    fold (add_iterm_vars bounds) [tm1, tm2]
blanchet@49148
  1043
  | add_iterm_vars _ (IConst _) = I
blanchet@49148
  1044
  | add_iterm_vars bounds (IVar (name, T)) =
blanchet@49148
  1045
    not (member (op =) bounds name) ? insert (op =) (name, SOME T)
blanchet@49148
  1046
  | add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm
blanchet@49148
  1047
blanchet@47237
  1048
fun aliased_uncurried ary (s, s') =
blanchet@47237
  1049
  (s ^ ascii_of_uncurried_alias_sep ^ string_of_int ary, s' ^ string_of_int ary)
blanchet@47237
  1050
fun unaliased_uncurried (s, s') =
blanchet@47237
  1051
  case space_explode uncurried_alias_sep s of
blanchet@47220
  1052
    [_] => (s, s')
blanchet@47220
  1053
  | [s1, s2] => (s1, unsuffix s2 s')
blanchet@47220
  1054
  | _ => raise Fail "ill-formed explicit application alias"
blanchet@47220
  1055
blanchet@47220
  1056
fun raw_mangled_const_name type_name ty_args (s, s') =
blanchet@43804
  1057
  let
blanchet@43804
  1058
    fun type_suffix f g =
blanchet@47220
  1059
      fold_rev (curry (op ^) o g o prefix mangled_type_sep o type_name f)
blanchet@47220
  1060
               ty_args ""
blanchet@43804
  1061
  in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end
blanchet@48638
  1062
fun mangled_const_name type_enc =
blanchet@49153
  1063
  map_filter (ho_type_for_type_arg type_enc)
blanchet@47220
  1064
  #> raw_mangled_const_name generic_mangled_type_name
blanchet@43413
  1065
blanchet@43413
  1066
val parse_mangled_ident =
blanchet@43413
  1067
  Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode
blanchet@43413
  1068
blanchet@43413
  1069
fun parse_mangled_type x =
blanchet@43413
  1070
  (parse_mangled_ident
blanchet@43413
  1071
   -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")")
blanchet@49147
  1072
                    [] >> (ATerm o apfst (rpair []))) x
blanchet@43413
  1073
and parse_mangled_types x =
blanchet@43413
  1074
  (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x
blanchet@43413
  1075
blanchet@43413
  1076
fun unmangled_type s =
blanchet@43413
  1077
  s |> suffix ")" |> raw_explode
blanchet@43413
  1078
    |> Scan.finite Symbol.stopper
blanchet@43413
  1079
           (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^
blanchet@43413
  1080
                                                quote s)) parse_mangled_type))
blanchet@43413
  1081
    |> fst
blanchet@43413
  1082
blanchet@47220
  1083
fun unmangled_const_name s =
blanchet@47237
  1084
  (s, s) |> unaliased_uncurried |> fst |> space_explode mangled_type_sep
blanchet@43413
  1085
fun unmangled_const s =
blanchet@47220
  1086
  let val ss = unmangled_const_name s in
blanchet@43413
  1087
    (hd ss, map unmangled_type (tl ss))
blanchet@43413
  1088
  end
blanchet@43413
  1089
blanchet@45644
  1090
fun introduce_proxies_in_iterm type_enc =
blanchet@43439
  1091
  let
blanchet@44858
  1092
    fun tweak_ho_quant ho_quant T [IAbs _] = IConst (`I ho_quant, T, [])
blanchet@44858
  1093
      | tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _]))
blanchet@44858
  1094
                       _ =
blanchet@44858
  1095
        (* Eta-expand "!!" and "??", to work around LEO-II 1.2.8 parser
blanchet@44858
  1096
           limitation. This works in conjuction with special code in
blanchet@44858
  1097
           "ATP_Problem" that uses the syntactic sugar "!" and "?" whenever
blanchet@44858
  1098
           possible. *)
blanchet@44858
  1099
        IAbs ((`I "P", p_T),
blanchet@44858
  1100
              IApp (IConst (`I ho_quant, T, []),
blanchet@44858
  1101
                    IAbs ((`I "X", x_T),
blanchet@44858
  1102
                          IApp (IConst (`I "P", p_T, []),
blanchet@44858
  1103
                                IConst (`I "X", x_T, [])))))
blanchet@44858
  1104
      | tweak_ho_quant _ _ _ = raise Fail "unexpected type for quantifier"
blanchet@44858
  1105
    fun intro top_level args (IApp (tm1, tm2)) =
blanchet@44858
  1106
        IApp (intro top_level (tm2 :: args) tm1, intro false [] tm2)
blanchet@44858
  1107
      | intro top_level args (IConst (name as (s, _), T, T_args)) =
blanchet@43441
  1108
        (case proxify_const s of
blanchet@44000
  1109
           SOME proxy_base =>
blanchet@44493
  1110
           if top_level orelse is_type_enc_higher_order type_enc then
blanchet@43841
  1111
             case (top_level, s) of
blanchet@44858
  1112
               (_, "c_False") => IConst (`I tptp_false, T, [])
blanchet@44858
  1113
             | (_, "c_True") => IConst (`I tptp_true, T, [])
blanchet@44858
  1114
             | (false, "c_Not") => IConst (`I tptp_not, T, [])
blanchet@44858
  1115
             | (false, "c_conj") => IConst (`I tptp_and, T, [])
blanchet@44858
  1116
             | (false, "c_disj") => IConst (`I tptp_or, T, [])
blanchet@44858
  1117
             | (false, "c_implies") => IConst (`I tptp_implies, T, [])
blanchet@44858
  1118
             | (false, "c_All") => tweak_ho_quant tptp_ho_forall T args
blanchet@44858
  1119
             | (false, "c_Ex") => tweak_ho_quant tptp_ho_exists T args
blanchet@43841
  1120
             | (false, s) =>
blanchet@48926
  1121
               if is_tptp_equal s then
blanchet@48926
  1122
                 if length args = 2 then
blanchet@48926
  1123
                   IConst (`I tptp_equal, T, [])
blanchet@48926
  1124
                 else
blanchet@48926
  1125
                   (* Eta-expand partially applied THF equality, because the
blanchet@48926
  1126
                      LEO-II and Satallax parsers complain about not being able to
blanchet@48926
  1127
                      infer the type of "=". *)
blanchet@48926
  1128
                   let val i_T = domain_type T in
blanchet@48926
  1129
                     IAbs ((`I "Y", i_T),
blanchet@48926
  1130
                           IAbs ((`I "Z", i_T),
blanchet@48926
  1131
                                 IApp (IApp (IConst (`I tptp_equal, T, []),
blanchet@48926
  1132
                                             IConst (`I "Y", i_T, [])),
blanchet@48926
  1133
                                       IConst (`I "Z", i_T, []))))
blanchet@48926
  1134
                   end
blanchet@44968
  1135
               else
blanchet@48926
  1136
                 IConst (name, T, [])
blanchet@44858
  1137
             | _ => IConst (name, T, [])
blanchet@43440
  1138
           else
blanchet@44858
  1139
             IConst (proxy_base |>> prefix const_prefix, T, T_args)
blanchet@46038
  1140
          | NONE => if s = tptp_choice then tweak_ho_quant tptp_choice T args
blanchet@46038
  1141
                    else IConst (name, T, T_args))
blanchet@44858
  1142
      | intro _ _ (IAbs (bound, tm)) = IAbs (bound, intro false [] tm)
blanchet@44858
  1143
      | intro _ _ tm = tm
blanchet@44858
  1144
  in intro true [] end
blanchet@43439
  1145
blanchet@48638
  1146
fun mangle_type_args_in_const type_enc (name as (s, _)) T_args =
blanchet@49218
  1147
  if String.isPrefix const_prefix s andalso is_type_enc_mangling type_enc then
blanchet@49218
  1148
    (mangled_const_name type_enc T_args name, [])
blanchet@49218
  1149
  else
blanchet@49218
  1150
    (name, T_args)
blanchet@48638
  1151
fun mangle_type_args_in_iterm type_enc =
blanchet@49216
  1152
  if is_type_enc_mangling type_enc then
blanchet@45645
  1153
    let
blanchet@45645
  1154
      fun mangle (IApp (tm1, tm2)) = IApp (mangle tm1, mangle tm2)
blanchet@45645
  1155
        | mangle (tm as IConst (_, _, [])) = tm
blanchet@47220
  1156
        | mangle (IConst (name, T, T_args)) =
blanchet@48638
  1157
          mangle_type_args_in_const type_enc name T_args
blanchet@47220
  1158
          |> (fn (name, T_args) => IConst (name, T, T_args))
blanchet@45645
  1159
        | mangle (IAbs (bound, tm)) = IAbs (bound, mangle tm)
blanchet@45645
  1160
        | mangle tm = tm
blanchet@45645
  1161
    in mangle end
blanchet@45645
  1162
  else
blanchet@45645
  1163
    I
blanchet@45645
  1164
blanchet@47220
  1165
fun filter_type_args_in_const _ _ _ _ _ [] = []
blanchet@49103
  1166
  | filter_type_args_in_const thy constrs type_enc ary s T_args =
blanchet@47220
  1167
    case unprefix_and_unascii const_prefix s of
blanchet@47220
  1168
      NONE =>
blanchet@47220
  1169
      if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then []
blanchet@47220
  1170
      else T_args
blanchet@47220
  1171
    | SOME s'' =>
blanchet@49217
  1172
      filter_type_args thy constrs type_enc (invert_const s'') ary T_args
blanchet@49103
  1173
fun filter_type_args_in_iterm thy constrs type_enc =
blanchet@38506
  1174
  let
blanchet@45645
  1175
    fun filt ary (IApp (tm1, tm2)) = IApp (filt (ary + 1) tm1, filt 0 tm2)
blanchet@45645
  1176
      | filt ary (IConst (name as (s, _), T, T_args)) =
blanchet@49103
  1177
        filter_type_args_in_const thy constrs type_enc ary s T_args
blanchet@47220
  1178
        |> (fn T_args => IConst (name, T, T_args))
blanchet@45645
  1179
      | filt _ (IAbs (bound, tm)) = IAbs (bound, filt 0 tm)
blanchet@45645
  1180
      | filt _ tm = tm
blanchet@45645
  1181
  in filt 0 end
blanchet@45644
  1182
blanchet@48920
  1183
fun iformula_from_prop ctxt type_enc iff_for_eq =
blanchet@45644
  1184
  let
blanchet@45644
  1185
    val thy = Proof_Context.theory_of ctxt
blanchet@46187
  1186
    fun do_term bs t atomic_Ts =
blanchet@48638
  1187
      iterm_from_term thy type_enc bs (Envir.eta_contract t)
blanchet@45644
  1188
      |>> (introduce_proxies_in_iterm type_enc
blanchet@48638
  1189
           #> mangle_type_args_in_iterm type_enc #> AAtom)
blanchet@46187
  1190
      ||> union (op =) atomic_Ts
blanchet@45262
  1191
    fun do_quant bs q pos s T t' =
blanchet@45262
  1192
      let
blanchet@45262
  1193
        val s = singleton (Name.variant_list (map fst bs)) s
blanchet@45262
  1194
        val universal = Option.map (q = AExists ? not) pos
blanchet@45262
  1195
        val name =
blanchet@45262
  1196
          s |> `(case universal of
blanchet@45262
  1197
                   SOME true => make_all_bound_var
blanchet@45262
  1198
                 | SOME false => make_exist_bound_var
blanchet@45262
  1199
                 | NONE => make_bound_var)
blanchet@45262
  1200
      in
blanchet@45262
  1201
        do_formula ((s, (name, T)) :: bs) pos t'
blanchet@45262
  1202
        #>> mk_aquant q [(name, SOME T)]
blanchet@46187
  1203
        ##> union (op =) (atomic_types_of T)
blanchet@38743
  1204
      end
blanchet@45262
  1205
    and do_conn bs c pos1 t1 pos2 t2 =
blanchet@45262
  1206
      do_formula bs pos1 t1 ##>> do_formula bs pos2 t2 #>> uncurry (mk_aconn c)
blanchet@45262
  1207
    and do_formula bs pos t =
blanchet@38506
  1208
      case t of
blanchet@45262
  1209
        @{const Trueprop} $ t1 => do_formula bs pos t1
blanchet@45262
  1210
      | @{const Not} $ t1 => do_formula bs (Option.map not pos) t1 #>> mk_anot
blanchet@38506
  1211
      | Const (@{const_name All}, _) $ Abs (s, T, t') =>
blanchet@45262
  1212
        do_quant bs AForall pos s T t'
blanchet@46038
  1213
      | (t0 as Const (@{const_name All}, _)) $ t1 =>
blanchet@46038
  1214
        do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
blanchet@38506
  1215
      | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
blanchet@45262
  1216
        do_quant bs AExists pos s T t'
blanchet@46038
  1217
      | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
blanchet@46038
  1218
        do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1)
blanchet@45262
  1219
      | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd pos t1 pos t2
blanchet@45262
  1220
      | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr pos t1 pos t2
blanchet@45262
  1221
      | @{const HOL.implies} $ t1 $ t2 =>
blanchet@45262
  1222
        do_conn bs AImplies (Option.map not pos) t1 pos t2
haftmann@39093
  1223
      | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
blanchet@48920
  1224
        if iff_for_eq then do_conn bs AIff NONE t1 NONE t2 else do_term bs t
blanchet@41388
  1225
      | _ => do_term bs t
blanchet@38506
  1226
  in do_formula [] end
blanchet@38506
  1227
blanchet@48022
  1228
fun presimplify_term thy t =
blanchet@48022
  1229
  if exists_Const (member (op =) Meson.presimplified_consts o fst) t then
blanchet@48022
  1230
    t |> Skip_Proof.make_thm thy
blanchet@48022
  1231
      |> Meson.presimplify
blanchet@48022
  1232
      |> prop_of
blanchet@48022
  1233
  else
blanchet@48022
  1234
    t
blanchet@38506
  1235
blanchet@46385
  1236
fun preprocess_abstractions_in_terms trans_lams facts =
blanchet@44733
  1237
  let
blanchet@44734
  1238
    val (facts, lambda_ts) =
blanchet@46385
  1239
      facts |> map (snd o snd) |> trans_lams
blanchet@48990
  1240
            |>> map2 (fn (name, (role, _)) => fn t => (name, (role, t))) facts
blanchet@46425
  1241
    val lam_facts =
blanchet@44734
  1242
      map2 (fn t => fn j =>
blanchet@49453
  1243
               ((lam_fact_prefix ^ Int.toString j,
blanchet@49453
  1244
                 (Global, Non_Rec_Def)), (Axiom, t)))
blanchet@44734
  1245
           lambda_ts (1 upto length lambda_ts)
blanchet@46425
  1246
  in (facts, lam_facts) end
blanchet@38506
  1247
blanchet@38506
  1248
(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
blanchet@43224
  1249
   same in Sledgehammer to prevent the discovery of unreplayable proofs. *)
blanchet@38506
  1250
fun freeze_term t =
blanchet@38506
  1251
  let
blanchet@45676
  1252
    fun freeze (t $ u) = freeze t $ freeze u
blanchet@45676
  1253
      | freeze (Abs (s, T, t)) = Abs (s, T, freeze t)
blanchet@45676
  1254
      | freeze (Var ((s, i), T)) =
blanchet@44734
  1255
        Free (atp_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
blanchet@45676
  1256
      | freeze t = t
blanchet@45676
  1257
  in t |> exists_subterm is_Var t ? freeze end
blanchet@38506
  1258
blanchet@48639
  1259
fun presimp_prop ctxt type_enc t =
blanchet@48584
  1260
  let
blanchet@48584
  1261
    val thy = Proof_Context.theory_of ctxt
blanchet@48584
  1262
    val t = t |> Envir.beta_eta_contract
blanchet@48584
  1263
              |> transform_elim_prop
blanchet@48584
  1264
              |> Object_Logic.atomize_term thy
blanchet@48584
  1265
    val need_trueprop = (fastype_of t = @{typ bool})
blanchet@48928
  1266
    val is_ho = is_type_enc_higher_order type_enc
blanchet@48584
  1267
  in
blanchet@48584
  1268
    t |> need_trueprop ? HOLogic.mk_Trueprop
blanchet@48969
  1269
      |> (if is_ho then unextensionalize_def
blanchet@48969
  1270
          else cong_extensionalize_term thy #> abs_extensionalize_term ctxt)
blanchet@48584
  1271
      |> presimplify_term thy
blanchet@48584
  1272
      |> HOLogic.dest_Trueprop
blanchet@48584
  1273
  end
blanchet@48584
  1274
  handle TERM _ => @{const True}
blanchet@43937
  1275
blanchet@48920
  1276
(* Satallax prefers "=" to "<=>" (for definitions) and Metis (CNF) requires "="
blanchet@48920
  1277
   for obscure technical reasons. *)
blanchet@48920
  1278
fun should_use_iff_for_eq CNF _ = false
blanchet@48920
  1279
  | should_use_iff_for_eq (THF _) format = not (is_type_enc_higher_order format)
blanchet@48920
  1280
  | should_use_iff_for_eq _ _ = true
blanchet@48920
  1281
blanchet@48990
  1282
fun make_formula ctxt format type_enc iff_for_eq name stature role t =
blanchet@43937
  1283
  let
blanchet@48920
  1284
    val iff_for_eq = iff_for_eq andalso should_use_iff_for_eq format type_enc
blanchet@46187
  1285
    val (iformula, atomic_Ts) =
blanchet@48990
  1286
      iformula_from_prop ctxt type_enc iff_for_eq (SOME (role <> Conjecture)) t
blanchet@48638
  1287
                         []
blanchet@49159
  1288
      |>> close_universally add_iterm_vars
blanchet@38506
  1289
  in
blanchet@48990
  1290
    {name = name, stature = stature, role = role, iformula = iformula,
blanchet@46187
  1291
     atomic_types = atomic_Ts}
blanchet@38506
  1292
  end
blanchet@38506
  1293
blanchet@49019
  1294
fun is_format_with_defs (THF (_, _, _, THF_With_Defs)) = true
blanchet@49019
  1295
  | is_format_with_defs _ = false
blanchet@49019
  1296
blanchet@48986
  1297
fun make_fact ctxt format type_enc iff_for_eq
blanchet@48986
  1298
              ((name, stature as (_, status)), t) =
blanchet@48986
  1299
  let
blanchet@48986
  1300
    val role =
blanchet@49453
  1301
      if is_format_with_defs format andalso status = Non_Rec_Def andalso
blanchet@49006
  1302
         is_legitimate_tptp_def t then
blanchet@48986
  1303
        Definition
blanchet@48986
  1304
      else
blanchet@48986
  1305
        Axiom
blanchet@48986
  1306
  in
blanchet@48986
  1307
    case t |> make_formula ctxt format type_enc iff_for_eq name stature role of
blanchet@48986
  1308
      formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} =>
blanchet@48986
  1309
      if s = tptp_true then NONE else SOME formula
blanchet@48986
  1310
    | formula => SOME formula
blanchet@48986
  1311
  end
blanchet@43432
  1312
blanchet@48584
  1313
fun s_not_prop (@{const Trueprop} $ t) = @{const Trueprop} $ s_not t
blanchet@48584
  1314
  | s_not_prop (@{const "==>"} $ t $ @{prop False}) = t
blanchet@48584
  1315
  | s_not_prop t = @{const "==>"} $ t $ @{prop False}
blanchet@45317
  1316
blanchet@45644
  1317
fun make_conjecture ctxt format type_enc =
blanchet@48990
  1318
  map (fn ((name, stature), (role, t)) =>
blanchet@48996
  1319
          let
blanchet@49094
  1320
            (* FIXME: The commented-out code is a hack to get decent performance
blanchet@49094
  1321
               out of LEO-II on the TPTP THF benchmarks. *)
blanchet@48996
  1322
            val role =
blanchet@49094
  1323
              if (* is_format_with_defs format andalso *)
blanchet@49006
  1324
                 role <> Conjecture andalso is_legitimate_tptp_def t then
blanchet@48996
  1325
                Definition
blanchet@48996
  1326
              else
blanchet@48996
  1327
                role
blanchet@48996
  1328
          in
blanchet@48996
  1329
            t |> role = Conjecture ? s_not
blanchet@48996
  1330
              |> make_formula ctxt format type_enc true name stature role
blanchet@48996
  1331
          end)
blanchet@38506
  1332
blanchet@43552
  1333
(** Finite and infinite type inference **)
blanchet@43552
  1334
blanchet@49215
  1335
fun tvar_footprint thy s ary =
blanchet@46382
  1336
  (case unprefix_and_unascii const_prefix s of
blanchet@45676
  1337
     SOME s =>
blanchet@49215
  1338
     let fun tvars_of T = [] |> Term.add_tvarsT T |> map fst in
blanchet@49215
  1339
       s |> invert_const |> robust_const_type thy |> chop_fun ary |> fst
blanchet@49215
  1340
         |> map tvars_of
blanchet@49215
  1341
     end
blanchet@45676
  1342
   | NONE => [])
blanchet@45676
  1343
  handle TYPE _ => []
blanchet@45676
  1344
blanchet@49215
  1345
fun type_arg_cover thy pos s ary =
blanchet@46819
  1346
  if is_tptp_equal s then
blanchet@49201
  1347
    if pos = SOME false then [] else 0 upto ary - 1
blanchet@46819
  1348
  else
blanchet@46819
  1349
    let
blanchet@49215
  1350
      val footprint = tvar_footprint thy s ary
blanchet@46819
  1351
      val eq = (s = @{const_name HOL.eq})
blanchet@49095
  1352
      fun cover _ [] = []
blanchet@49095
  1353
        | cover seen ((i, tvars) :: args) =
blanchet@49095
  1354
          cover (union (op =) seen tvars) args
blanchet@46819
  1355
          |> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars)
blanchet@46819
  1356
             ? cons i
blanchet@46819
  1357
    in
blanchet@46819
  1358
      if forall null footprint then
blanchet@46819
  1359
        []
blanchet@46819
  1360
      else
blanchet@46819
  1361
        0 upto length footprint - 1 ~~ footprint
blanchet@46819
  1362
        |> sort (rev_order o list_ord Term_Ord.indexname_ord o pairself snd)
blanchet@49095
  1363
        |> cover []
blanchet@46819
  1364
    end
blanchet@45676
  1365
blanchet@45258
  1366
type monotonicity_info =
blanchet@45258
  1367
  {maybe_finite_Ts : typ list,
blanchet@45258
  1368
   surely_infinite_Ts : typ list,
blanchet@45258
  1369
   maybe_nonmono_Ts : typ list}
blanchet@45258
  1370
blanchet@45256
  1371
(* These types witness that the type classes they belong to allow infinite
blanchet@45256
  1372
   models and hence that any types with these type classes is monotonic. *)
blanchet@45256
  1373
val known_infinite_types =
blanchet@45492
  1374
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}]
blanchet@45256
  1375
blanchet@47129
  1376
fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T =
blanchet@47129
  1377
  strictness <> Strict andalso is_type_surely_infinite ctxt true cached_Ts T
blanchet@43755
  1378
blanchet@43552
  1379
(* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are
blanchet@43552
  1380
   dangerous because their "exhaust" properties can easily lead to unsound ATP
blanchet@43552
  1381
   proofs. On the other hand, all HOL infinite types can be given the same
blanchet@43552
  1382
   models in first-order logic (via Löwenheim-Skolem). *)
blanchet@43552
  1383
blanchet@49215
  1384
fun should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts,
blanchet@49218
  1385
                             maybe_nonmono_Ts}
blanchet@49198
  1386
                       (Nonmono_Types (strictness, _)) T =
blanchet@48021
  1387
    let val thy = Proof_Context.theory_of ctxt in
blanchet@48021
  1388
      (exists (type_intersect thy T) maybe_nonmono_Ts andalso
blanchet@48021
  1389
       not (exists (type_instance thy T) surely_infinite_Ts orelse
blanchet@48021
  1390
            (not (member (type_equiv thy) maybe_finite_Ts T) andalso
blanchet@48021
  1391
             is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts
blanchet@48021
  1392
                                             T)))
blanchet@48021
  1393
    end
blanchet@49215
  1394
  | should_encode_type _ _ level _ =
blanchet@49215
  1395
    (level = All_Types orelse level = Undercover_Types)
blanchet@43552
  1396
blanchet@45639
  1397
fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T =
blanchet@45673
  1398
    should_guard_var () andalso should_encode_type ctxt mono level T
blanchet@45258
  1399
  | should_guard_type _ _ _ _ _ = false
blanchet@43552
  1400
blanchet@49215
  1401
fun is_maybe_universal_name s =
blanchet@49215
  1402
  String.isPrefix bound_var_prefix s orelse
blanchet@49215
  1403
  String.isPrefix all_bound_var_prefix s
blanchet@49215
  1404
blanchet@49215
  1405
fun is_maybe_universal_var (IConst ((s, _), _, _)) = is_maybe_universal_name s
blanchet@45262
  1406
  | is_maybe_universal_var (IVar _) = true
blanchet@45262
  1407
  | is_maybe_universal_var _ = false
blanchet@43707
  1408
blanchet@46818
  1409
datatype site =
blanchet@44232
  1410
  Top_Level of bool option |
blanchet@44232
  1411
  Eq_Arg of bool option |
blanchet@49161
  1412
  Arg of string * int * int |
blanchet@44232
  1413
  Elsewhere
blanchet@43700
  1414
blanchet@46820
  1415
fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false
blanchet@46820
  1416
  | should_tag_with_type ctxt mono (Tags (_, level)) site u T =
blanchet@49161
  1417
    let val thy = Proof_Context.theory_of ctxt in
blanchet@49198
  1418
      case level of
blanchet@49198
  1419
        Nonmono_Types (_, Non_Uniform) =>
blanchet@49161
  1420
        (case (site, is_maybe_universal_var u) of
blanchet@49161
  1421
           (Eq_Arg _, true) => should_encode_type ctxt mono level T
blanchet@49161
  1422
         | _ => false)
blanchet@49198
  1423
      | Undercover_Types =>
blanchet@49198
  1424
        (case (site, is_maybe_universal_var u) of
blanchet@49201
  1425
           (Eq_Arg pos, true) => pos <> SOME false
blanchet@49201
  1426
         | (Arg (s, j, ary), true) =>
blanchet@49215
  1427
           member (op =) (type_arg_cover thy NONE s ary) j
blanchet@49198
  1428
         | _ => false)
blanchet@49198
  1429
      | _ => should_encode_type ctxt mono level T
blanchet@49161
  1430
    end
blanchet@46820
  1431
  | should_tag_with_type _ _ _ _ _ _ = false
blanchet@43552
  1432
blanchet@45458
  1433
fun fused_type ctxt mono level =
blanchet@43835
  1434
  let
blanchet@45258
  1435
    val should_encode = should_encode_type ctxt mono level
blanchet@45458
  1436
    fun fuse 0 T = if should_encode T then T else fused_infinite_type
blanchet@45458
  1437
      | fuse ary (Type (@{type_name fun}, [T1, T2])) =
blanchet@45458
  1438
        fuse 0 T1 --> fuse (ary - 1) T2
blanchet@45458
  1439
      | fuse _ _ = raise Fail "expected function type"
blanchet@45458
  1440
  in fuse end
blanchet@43552
  1441
blanchet@45307
  1442
(** predicators and application operators **)
blanchet@41561
  1443
blanchet@43445
  1444
type sym_info =
blanchet@45700
  1445
  {pred_sym : bool, min_ary : int, max_ary : int, types : typ list,
blanchet@45700
  1446
   in_conj : bool}
blanchet@43434
  1447
blanchet@45700
  1448
fun default_sym_tab_entries type_enc =
blanchet@45700
  1449
  (make_fixed_const NONE @{const_name undefined},
blanchet@45700
  1450
       {pred_sym = false, min_ary = 0, max_ary = 0, types = [],
blanchet@45700
  1451
        in_conj = false}) ::
blanchet@45700
  1452
  ([tptp_false, tptp_true]
blanchet@45700
  1453
   |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, types = [],
blanchet@45700
  1454
                  in_conj = false})) @
blanchet@45700
  1455
  ([tptp_equal, tptp_old_equal]
blanchet@45700
  1456
   |> map (rpair {pred_sym = true, min_ary = 2, max_ary = 2, types = [],
blanchet@45700
  1457
                  in_conj = false}))
blanchet@45700
  1458
  |> not (is_type_enc_higher_order type_enc)
blanchet@45700
  1459
     ? cons (prefixed_predicator_name,
blanchet@45700
  1460
             {pred_sym = true, min_ary = 1, max_ary = 1, types = [],
blanchet@45700
  1461
              in_conj = false})
blanchet@45700
  1462
blanchet@47933
  1463
datatype app_op_level =
blanchet@47933
  1464
  Min_App_Op |
blanchet@47933
  1465
  Sufficient_App_Op |
blanchet@47933
  1466
  Sufficient_App_Op_And_Predicator |
blanchet@47933
  1467
  Full_App_Op_And_Predicator
blanchet@47217
  1468
blanchet@48947
  1469
fun add_iterm_syms_to_sym_table ctxt app_op_level conj_fact =
blanchet@43429
  1470
  let
blanchet@47513
  1471
    val thy = Proof_Context.theory_of ctxt
blanchet@45643
  1472
    fun consider_var_ary const_T var_T max_ary =
blanchet@43905
  1473
      let
blanchet@43905
  1474
        fun iter ary T =
blanchet@48021
  1475
          if ary = max_ary orelse type_instance thy var_T T orelse
blanchet@48021
  1476
             type_instance thy T var_T then
blanchet@44051
  1477
            ary
blanchet@44051
  1478
          else
blanchet@44051
  1479
            iter (ary + 1) (range_type T)
blanchet@43905
  1480
      in iter 0 const_T end
blanchet@45262
  1481
    fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
blanchet@47933
  1482
      if (app_op_level = Sufficient_App_Op andalso can dest_funT T) orelse
blanchet@47933
  1483
         (app_op_level = Sufficient_App_Op_And_Predicator andalso
blanchet@47933
  1484
          (can dest_funT T orelse T = @{typ bool})) then
blanchet@44042
  1485
        let
blanchet@47933
  1486
          val bool_vars' =
blanchet@47933
  1487
            bool_vars orelse
blanchet@47933
  1488
            (app_op_level = Sufficient_App_Op_And_Predicator andalso
blanchet@47933
  1489
             body_type T = @{typ bool})
blanchet@45700
  1490
          fun repair_min_ary {pred_sym, min_ary, max_ary, types, in_conj} =
blanchet@44042
  1491
            {pred_sym = pred_sym andalso not bool_vars',
blanchet@45643
  1492
             min_ary = fold (fn T' => consider_var_ary T' T) types min_ary,
blanchet@45700
  1493
             max_ary = max_ary, types = types, in_conj = in_conj}
blanchet@44042
  1494
          val fun_var_Ts' =
blanchet@48021
  1495
            fun_var_Ts |> can dest_funT T ? insert_type thy I T
blanchet@44042
  1496
        in
blanchet@44042
  1497
          if bool_vars' = bool_vars andalso
blanchet@44042
  1498
             pointer_eq (fun_var_Ts', fun_var_Ts) then
blanchet@44042
  1499
            accum
blanchet@44008
  1500
          else
blanchet@45643
  1501
            ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_ary) sym_tab)
blanchet@44042
  1502
        end
blanchet@44042
  1503
      else
blanchet@44042
  1504
        accum
blanchet@48947
  1505
    fun add_iterm_syms top_level tm
blanchet@48947
  1506
                       (accum as ((bool_vars, fun_var_Ts), sym_tab)) =
blanchet@48947
  1507
      let val (head, args) = strip_iterm_comb tm in
blanchet@48947
  1508
        (case head of
blanchet@48947
  1509
           IConst ((s, _), T, _) =>
blanchet@49215
  1510
           if is_maybe_universal_name s then
blanchet@48947
  1511
             add_universal_var T accum
blanchet@48947
  1512
           else if String.isPrefix exist_bound_var_prefix s then
blanchet@48947
  1513
             accum
blanchet@48947
  1514
           else
blanchet@48947
  1515
             let val ary = length args in
blanchet@48947
  1516
               ((bool_vars, fun_var_Ts),
blanchet@48947
  1517
                case Symtab.lookup sym_tab s of
blanchet@48947
  1518
                  SOME {pred_sym, min_ary, max_ary, types, in_conj} =>
blanchet@48947
  1519
                  let
blanchet@48947
  1520
                    val pred_sym =
blanchet@48947
  1521
                      pred_sym andalso top_level andalso not bool_vars
blanchet@48947
  1522
                    val types' = types |> insert_type thy I T
blanchet@48947
  1523
                    val in_conj = in_conj orelse conj_fact
blanchet@48947
  1524
                    val min_ary =
blanchet@48947
  1525
                      if (app_op_level = Sufficient_App_Op orelse
blanchet@48947
  1526
                          app_op_level = Sufficient_App_Op_And_Predicator)
blanchet@48947
  1527
                         andalso not (pointer_eq (types', types)) then
blanchet@48947
  1528
                        fold (consider_var_ary T) fun_var_Ts min_ary
blanchet@48947
  1529
                      else
blanchet@48947
  1530
                        min_ary
blanchet@48947
  1531
                  in
blanchet@48947
  1532
                    Symtab.update (s, {pred_sym = pred_sym,
blanchet@48947
  1533
                                       min_ary = Int.min (ary, min_ary),
blanchet@48947
  1534
                                       max_ary = Int.max (ary, max_ary),
blanchet@48947
  1535
                                       types = types', in_conj = in_conj})
blanchet@48947
  1536
                                  sym_tab
blanchet@48947
  1537
                  end
blanchet@48947
  1538
                | NONE =>
blanchet@48947
  1539
                  let
blanchet@48947
  1540
                    val pred_sym = top_level andalso not bool_vars
blanchet@48947
  1541
                    val ary =
blanchet@48947
  1542
                      case unprefix_and_unascii const_prefix s of
blanchet@48947
  1543
                        SOME s =>
blanchet@48947
  1544
                        (if String.isSubstring uncurried_alias_sep s then
blanchet@48947
  1545
                           ary
blanchet@48947
  1546
                         else case try (robust_const_ary thy
blanchet@48947
  1547
                                        o invert_const o hd
blanchet@48947
  1548
                                        o unmangled_const_name) s of
blanchet@48947
  1549
                           SOME ary0 => Int.min (ary0, ary)
blanchet@48947
  1550
                         | NONE => ary)
blanchet@48947
  1551
                      | NONE => ary
blanchet@48947
  1552
                    val min_ary =
blanchet@48947
  1553
                      case app_op_level of
blanchet@48947
  1554
                        Min_App_Op => ary
blanchet@48947
  1555
                      | Full_App_Op_And_Predicator => 0
blanchet@48947
  1556
                      | _ => fold (consider_var_ary T) fun_var_Ts ary
blanchet@48947
  1557
                  in
blanchet@48947
  1558
                    Symtab.update_new (s,
blanchet@48947
  1559
                        {pred_sym = pred_sym, min_ary = min_ary,
blanchet@48947
  1560
                         max_ary = ary, types = [T], in_conj = conj_fact})
blanchet@48947
  1561
                        sym_tab
blanchet@48947
  1562
                  end)
blanchet@48947
  1563
             end
blanchet@48947
  1564
         | IVar (_, T) => add_universal_var T accum
blanchet@48947
  1565
         | IAbs ((_, T), tm) =>
blanchet@48947
  1566
           accum |> add_universal_var T |> add_iterm_syms false tm
blanchet@48947
  1567
         | _ => accum)
blanchet@48947
  1568
        |> fold (add_iterm_syms false) args
blanchet@48947
  1569
      end
blanchet@48947
  1570
  in add_iterm_syms end
blanchet@48947
  1571
blanchet@48947
  1572
fun sym_table_for_facts ctxt type_enc app_op_level conjs facts =
blanchet@48947
  1573
  let
blanchet@48947
  1574
    fun add_iterm_syms conj_fact =
blanchet@48947
  1575
      add_iterm_syms_to_sym_table ctxt app_op_level conj_fact true
blanchet@45700
  1576
    fun add_fact_syms conj_fact =
blanchet@48996
  1577
      K (add_iterm_syms conj_fact) |> formula_fold NONE |> ifact_lift
blanchet@45700
  1578
  in
blanchet@45700
  1579
    ((false, []), Symtab.empty)
blanchet@45700
  1580
    |> fold (add_fact_syms true) conjs
blanchet@45700
  1581
    |> fold (add_fact_syms false) facts
blanchet@48947
  1582
    ||> fold Symtab.update (default_sym_tab_entries type_enc)
blanchet@45700
  1583
  end
blanchet@38506
  1584
blanchet@45643
  1585
fun min_ary_of sym_tab s =
blanchet@43429
  1586
  case Symtab.lookup sym_tab s of
blanchet@43445
  1587
    SOME ({min_ary, ...} : sym_info) => min_ary
blanchet@43429
  1588
  | NONE =>
blanchet@46382
  1589
    case unprefix_and_unascii const_prefix s of
blanchet@43418
  1590
      SOME s =>
blanchet@47220
  1591
      let val s = s |> unmangled_const_name |> hd |> invert_const in
blanchet@43807
  1592
        if s = predicator_name then 1
blanchet@43807
  1593
        else if s = app_op_name then 2
blanchet@45255
  1594
        else if s = type_guard_name then 1
blanchet@43428
  1595
        else 0
blanchet@43418
  1596
      end
blanchet@38506
  1597
    | NONE => 0
blanchet@38506
  1598
blanchet@38506
  1599
(* True if the constant ever appears outside of the top-level position in
blanchet@38506
  1600
   literals, or if it appears with different arities (e.g., because of different
blanchet@38506
  1601
   type instantiations). If false, the constant always receives all of its
blanchet@38506
  1602
   arguments and is used as a predicate. *)
blanchet@43429
  1603
fun is_pred_sym sym_tab s =
blanchet@43429
  1604
  case Symtab.lookup sym_tab s of
blanchet@43445
  1605
    SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) =>
blanchet@43445
  1606
    pred_sym andalso min_ary = max_ary
blanchet@43429
  1607
  | NONE => false
blanchet@38506
  1608
blanchet@48961
  1609
val fTrue_iconst =
blanchet@48961
  1610
  IConst ((const_prefix ^ "fTrue", @{const_name ATP.fTrue}), @{typ bool}, [])
blanchet@48961
  1611
val predicator_iconst =
nik@45350
  1612
  IConst (`(make_fixed_const NONE) predicator_name, @{typ "bool => bool"}, [])
blanchet@45644
  1613
blanchet@49158
  1614
fun predicatify completish tm =
blanchet@49158
  1615
  if completish then
blanchet@48961
  1616
    IApp (IApp (IConst (`I tptp_equal, @{typ "bool => bool => bool"}, []), tm),
blanchet@48961
  1617
          fTrue_iconst)
blanchet@48961
  1618
  else
blanchet@48961
  1619
    IApp (predicator_iconst, tm)
blanchet@48961
  1620
blanchet@48961
  1621
val app_op = `(make_fixed_const NONE) app_op_name
blanchet@48961
  1622
blanchet@45644
  1623
fun list_app head args = fold (curry (IApp o swap)) args head
blanchet@38506
  1624
blanchet@48638
  1625
fun mk_app_op type_enc head arg =
blanchet@43415
  1626
  let
blanchet@47220
  1627
    val head_T = ityp_of head
blanchet@47220
  1628
    val (arg_T, res_T) = dest_funT head_T
blanchet@47220
  1629
    val app =
blanchet@47220
  1630
      IConst (app_op, head_T --> head_T, [arg_T, res_T])
blanchet@48638
  1631
      |> mangle_type_args_in_iterm type_enc
blanchet@47220
  1632
  in list_app app [head, arg] end
blanchet@47220
  1633
blanchet@49103
  1634
fun firstorderize_fact thy constrs type_enc sym_tab uncurried_aliases
blanchet@49158
  1635
                       completish =
blanchet@47220
  1636
  let
blanchet@48638
  1637
    fun do_app arg head = mk_app_op type_enc head arg
blanchet@45644
  1638
    fun list_app_ops head args = fold do_app args head
blanchet@45644
  1639
    fun introduce_app_ops tm =
blanchet@47220
  1640
      let val (head, args) = tm |> strip_iterm_comb ||> map introduce_app_ops in
blanchet@47220
  1641
        case head of
blanchet@47220
  1642
          IConst (name as (s, _), T, T_args) =>
blanchet@47237
  1643
          if uncurried_aliases andalso String.isPrefix const_prefix s then
blanchet@47220
  1644
            let
blanchet@47220
  1645
              val ary = length args
blanchet@47510
  1646
              val name =
blanchet@47510
  1647
                name |> ary > min_ary_of sym_tab s ? aliased_uncurried ary
blanchet@47220
  1648
            in list_app (IConst (name, T, T_args)) args end
blanchet@47220
  1649
          else
blanchet@47220
  1650
            args |> chop (min_ary_of sym_tab s)
blanchet@47220
  1651
                 |>> list_app head |-> list_app_ops
blanchet@47220
  1652
        | _ => list_app_ops head args
blanchet@47220
  1653
      end
blanchet@45644
  1654
    fun introduce_predicators tm =
blanchet@45644
  1655
      case strip_iterm_comb tm of
blanchet@45644
  1656
        (IConst ((s, _), _, _), _) =>
blanchet@49158
  1657
        if is_pred_sym sym_tab s then tm else predicatify completish tm
blanchet@49158
  1658
      | _ => predicatify completish tm
blanchet@45644
  1659
    val do_iterm =
blanchet@45644
  1660
      not (is_type_enc_higher_order type_enc)
blanchet@45644
  1661
      ? (introduce_app_ops #> introduce_predicators)
blanchet@49103
  1662
      #> filter_type_args_in_iterm thy constrs type_enc
blanchet@45644
  1663
  in update_iformula (formula_map do_iterm) end
blanchet@43444
  1664
blanchet@43444
  1665
(** Helper facts **)
blanchet@43444
  1666
blanchet@45307
  1667
val not_ffalse = @{lemma "~ fFalse" by (unfold fFalse_def) fast}
blanchet@45307
  1668
val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
blanchet@45307
  1669
blanchet@44035
  1670
(* The Boolean indicates that a fairly sound type encoding is needed. *)
blanchet@48961
  1671
val base_helper_table =
blanchet@49453
  1672
  [(("COMBI", false), [(Non_Rec_Def, @{thm Meson.COMBI_def})]),
blanchet@49453
  1673
   (("COMBK", false), [(Non_Rec_Def, @{thm Meson.COMBK_def})]),
blanchet@49453
  1674
   (("COMBB", false), [(Non_Rec_Def, @{thm Meson.COMBB_def})]),
blanchet@49453
  1675
   (("COMBC", false), [(Non_Rec_Def, @{thm Meson.COMBC_def})]),
blanchet@49453
  1676
   (("COMBS", false), [(Non_Rec_Def, @{thm Meson.COMBS_def})]),
blanchet@48961
  1677
   ((predicator_name, false), [(General, not_ffalse), (General, ftrue)]),
blanchet@48961
  1678
   (("fFalse", false), [(General, not_ffalse)]),
blanchet@48961
  1679
   (("fFalse", true), [(General, @{thm True_or_False})]),
blanchet@48961
  1680
   (("fTrue", false), [(General, ftrue)]),
blanchet@48961
  1681
   (("fTrue", true), [(General, @{thm True_or_False})]),
blanchet@48961
  1682
   (("If", true),
blanchet@49453
  1683
    [(Non_Rec_Def, @{thm if_True}), (Non_Rec_Def, @{thm if_False}),
blanchet@48961
  1684
     (General, @{thm True_or_False})])]
blanchet@48961
  1685
blanchet@43926
  1686
val helper_table =
blanchet@48961
  1687
  base_helper_table @
blanchet@48961
  1688
  [(("fNot", false),
blanchet@44035
  1689
    @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1]
blanchet@48961
  1690
           fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}
blanchet@49453
  1691
    |> map (pair Non_Rec_Def)),
blanchet@44035
  1692
   (("fconj", false),
blanchet@44035
  1693
    @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q"
blanchet@48961
  1694
        by (unfold fconj_def) fast+}
blanchet@48961
  1695
    |> map (pair General)),
blanchet@44035
  1696
   (("fdisj", false),
blanchet@44035
  1697
    @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q"
blanchet@48961
  1698
        by (unfold fdisj_def) fast+}
blanchet@48961
  1699
    |> map (pair General)),
blanchet@44035
  1700
   (("fimplies", false),
blanchet@44051
  1701
    @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
blanchet@48961
  1702
        by (unfold fimplies_def) fast+}
blanchet@48961
  1703
    |> map (pair General)),
nik@44537
  1704
   (("fequal", true),
nik@44537
  1705
    (* This is a lie: Higher-order equality doesn't need a sound type encoding.
nik@44537
  1706
       However, this is done so for backward compatibility: Including the
nik@44537
  1707
       equality helpers by default in Metis breaks a few existing proofs. *)
nik@44537
  1708
    @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1]
blanchet@48961
  1709
           fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]}
blanchet@48961
  1710
    |> map (pair General)),
blanchet@44874
  1711
   (* Partial characterization of "fAll" and "fEx". A complete characterization
blanchet@44874
  1712
      would require the axiom of choice for replay with Metis. *)
blanchet@48961
  1713
   (("fAll", false),
blanchet@48961
  1714
    [(General, @{lemma "~ fAll P | P x" by (auto simp: fAll_def)})]),
blanchet@48961
  1715
   (("fEx", false),
blanchet@48961
  1716
    [(General, @{lemma "~ P x | fEx P" by (auto simp: fEx_def)})])]
blanchet@48961
  1717
  |> map (apsnd (map (apsnd zero_var_indexes)))
blanchet@48961
  1718
blanchet@49158
  1719
val completish_helper_table =
blanchet@48961
  1720
  base_helper_table @
blanchet@48961
  1721
  [((predicator_name, true),
blanchet@48961
  1722
    @{thms True_or_False fTrue_ne_fFalse} |> map (pair General)),
blanchet@48961
  1723
   ((app_op_name, true),
blanchet@48961
  1724
    [(General, @{lemma "EX x. ~ f x = g x | f = g" by blast})]),
blanchet@48961
  1725
   (("fconj", false),
blanchet@49453
  1726
    @{thms fconj_table fconj_laws fdisj_laws} |> map (pair Non_Rec_Def)),
blanchet@48961
  1727
   (("fdisj", false),
blanchet@49453
  1728
    @{thms fdisj_table fconj_laws fdisj_laws} |> map (pair Non_Rec_Def)),
blanchet@48961
  1729
   (("fimplies", false),
blanchet@48961
  1730
    @{thms fimplies_table fconj_laws fdisj_laws fimplies_laws}
blanchet@49453
  1731
    |> map (pair Non_Rec_Def)),
blanchet@48961
  1732
   (("fequal", false),
blanchet@49453
  1733
    (@{thms fequal_table} |> map (pair Non_Rec_Def)) @
blanchet@48961
  1734
    (@{thms fequal_laws} |> map (pair General))),
blanchet@48961
  1735
   (("fAll", false),
blanchet@49453
  1736
    @{thms fAll_table fComp_law fAll_law fEx_law} |> map (pair Non_Rec_Def)),
blanchet@48961
  1737
   (("fEx", false),
blanchet@49453
  1738
    @{thms fEx_table fComp_law fAll_law fEx_law} |> map (pair Non_Rec_Def))]
blanchet@48961
  1739
  |> map (apsnd (map (apsnd zero_var_indexes)))
blanchet@43926
  1740
blanchet@46791
  1741
fun bound_tvars type_enc sorts Ts =
blanchet@49148
  1742
  case filter is_TVar Ts of
blanchet@49148
  1743
    [] => I
blanchet@49148
  1744
  | Ts =>
blanchet@49157
  1745
    (sorts ? mk_ahorn (Ts |> class_membs_for_types type_enc add_sorts_on_tvar
blanchet@49157
  1746
                          |> map (class_atom type_enc)))
blanchet@49154
  1747
    #> (case type_enc of
blanchet@49154
  1748
          Native (_, poly, _) =>
blanchet@49157
  1749
          mk_atyquant AForall
blanchet@49157
  1750
              (map (fn TVar (z as (_, S)) =>
blanchet@49157
  1751
                       (AType (tvar_name z, []),
blanchet@49157
  1752
                        if poly = Type_Class_Polymorphic then
blanchet@49157
  1753
                          map (`make_class) (normalize_classes S)
blanchet@49157
  1754
                        else
blanchet@49157
  1755
                          [])) Ts)
blanchet@49154
  1756
        | _ =>
blanchet@49157
  1757
          mk_aquant AForall (map (fn TVar z => (tvar_name z, NONE)) Ts))
blanchet@45263
  1758
blanchet@47220
  1759
fun eq_formula type_enc atomic_Ts bounds pred_sym tm1 tm2 =
blanchet@45263
  1760
  (if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2])
blanchet@49147
  1761
   else AAtom (ATerm ((`I tptp_equal, []), [tm1, tm2])))
blanchet@47220
  1762
  |> mk_aquant AForall bounds
blanchet@46248
  1763
  |> close_formula_universally
blanchet@46791
  1764
  |> bound_tvars type_enc true atomic_Ts
blanchet@45263
  1765
blanchet@47234
  1766
val helper_rank = default_rank
blanchet@47234
  1767
val min_rank = 9 * helper_rank div 10
blanchet@47234
  1768
val max_rank = 4 * min_rank
blanchet@47234
  1769
blanchet@47234
  1770
fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n
blanchet@47234
  1771
nik@45350
  1772
val type_tag = `(make_fixed_const NONE) type_tag_name
blanchet@43971
  1773
blanchet@48961
  1774
fun could_specialize_helpers type_enc =
blanchet@49146
  1775
  not (is_type_enc_polymorphic type_enc) andalso
blanchet@48961
  1776
  level_of_type_enc type_enc <> No_Types
blanchet@44493
  1777
fun should_specialize_helper type_enc t =
blanchet@48961
  1778
  could_specialize_helpers type_enc andalso
blanchet@44495
  1779
  not (null (Term.hidden_polymorphism t))
blanchet@44000
  1780
blanchet@49158
  1781
fun add_helper_facts_for_sym ctxt format type_enc completish
blanchet@48961
  1782
                             (s, {types, ...} : sym_info) =
blanchet@46382
  1783
  case unprefix_and_unascii const_prefix s of
blanchet@43444
  1784
    SOME mangled_s =>
blanchet@43444
  1785
    let
blanchet@43444
  1786
      val thy = Proof_Context.theory_of ctxt
blanchet@47220
  1787
      val unmangled_s = mangled_s |> unmangled_const_name |> hd
blanchet@49104
  1788
      fun dub needs_sound j k =
blanchet@48961
  1789
        ascii_of unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^
blanchet@47167
  1790
        (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^
blanchet@49104
  1791
        (if needs_sound then typed_helper_suffix else untyped_helper_suffix)
blanchet@48947
  1792
      fun specialize_helper t T =
blanchet@48947
  1793
        if unmangled_s = app_op_name then
blanchet@48947
  1794
          let
blanchet@48947
  1795
            val tyenv =
blanchet@48947
  1796
              Sign.typ_match thy (alpha_to_beta, domain_type T) Vartab.empty
blanchet@48947
  1797
          in monomorphic_term tyenv t end
blanchet@48947
  1798
        else
blanchet@48947
  1799
          specialize_type thy (invert_const unmangled_s, T) t
blanchet@49104
  1800
      fun dub_and_inst needs_sound ((status, t), j) =
blanchet@48947
  1801
        (if should_specialize_helper type_enc t then
blanchet@48961
  1802
           map_filter (try (specialize_helper t)) types
blanchet@48947
  1803
         else
blanchet@48947
  1804
           [t])
blanchet@47167
  1805
        |> tag_list 1
blanchet@49104
  1806
        |> map (fn (k, t) => ((dub needs_sound j k, (Global, status)), t))
blanchet@44731
  1807
      val make_facts = map_filter (make_fact ctxt format type_enc false)
blanchet@49104
  1808
      val sound = is_type_enc_sound type_enc
blanchet@48961
  1809
      val could_specialize = could_specialize_helpers type_enc
blanchet@43444
  1810
    in
blanchet@49104
  1811
      fold (fn ((helper_s, needs_sound), ths) =>
blanchet@49104
  1812
               if (needs_sound andalso not sound) orelse
blanchet@48961
  1813
                  (helper_s <> unmangled_s andalso
blanchet@49158
  1814
                   (not completish orelse could_specialize)) then
blanchet@48681
  1815
                 I
blanchet@48681
  1816
               else
blanchet@48681
  1817
                 ths ~~ (1 upto length ths)
blanchet@49104
  1818
                 |> maps (dub_and_inst needs_sound o apfst (apsnd prop_of))
blanchet@48681
  1819
                 |> make_facts
blanchet@48681
  1820
                 |> union (op = o pairself #iformula))
blanchet@49158
  1821
           (if completish then completish_helper_table else helper_table)
blanchet@43444
  1822
    end
blanchet@48681
  1823
  | NONE => I
blanchet@49158
  1824
fun helper_facts_for_sym_table ctxt format type_enc completish sym_tab =
blanchet@49158
  1825
  Symtab.fold_rev (add_helper_facts_for_sym ctxt format type_enc completish)
blanchet@48961
  1826
                  sym_tab []
blanchet@43444
  1827
blanchet@43926
  1828
(***************************************************************)
blanchet@43926
  1829
(* Type Classes Present in the Axiom or Conjecture Clauses     *)
blanchet@43926
  1830
(***************************************************************)
blanchet@43926
  1831
blanchet@43926
  1832
fun set_insert (x, s) = Symtab.update (x, ()) s
blanchet@43926
  1833
blanchet@49157
  1834
fun add_classes (cls, cset) = List.foldl set_insert cset (flat cls)
blanchet@43926
  1835
blanchet@43934
  1836
fun classes_of_terms get_Ts =
blanchet@43962
  1837
  map (map snd o get_Ts)
blanchet@49157
  1838
  #> List.foldl add_classes Symtab.empty #> Symtab.delete_safe class_of_types
blanchet@49157
  1839
  #> Symtab.keys
blanchet@43926
  1840
wenzelm@45004
  1841
val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees
wenzelm@45004
  1842
val tvar_classes_of_terms = classes_of_terms Misc_Legacy.term_tvars
blanchet@43926
  1843
blanchet@44489
  1844
fun fold_type_constrs f (Type (s, Ts)) x =
blanchet@44489
  1845
    fold (fold_type_constrs f) Ts (f (s, x))
blanchet@44030
  1846
  | fold_type_constrs _ _ x = x
blanchet@43926
  1847
blanchet@44778
  1848
(* Type constructors used to instantiate overloaded constants are the only ones
blanchet@44778
  1849
   needed. *)
blanchet@44030
  1850
fun add_type_constrs_in_term thy =
blanchet@43926
  1851
  let
blanchet@44029
  1852
    fun add (Const (@{const_name Meson.skolem}, _) $ _) = I
blanchet@44022
  1853
      | add (t $ u) = add t #> add u
blanchet@45602
  1854
      | add (Const x) =
blanchet@45602
  1855
        x |> robust_const_typargs thy |> fold (fold_type_constrs set_insert)
blanchet@44022
  1856
      | add (Abs (_, _, u)) = add u
blanchet@44022
  1857
      | add _ = I
blanchet@44022
  1858
  in add end
blanchet@43926
  1859
blanchet@44030
  1860
fun type_constrs_of_terms thy ts =
blanchet@44030
  1861
  Symtab.keys (fold (add_type_constrs_in_term thy) ts Symtab.empty)
blanchet@43926
  1862
blanchet@46382
  1863
fun extract_lambda_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
blanchet@46382
  1864
    let val (head, args) = strip_comb t in
blanchet@46382
  1865
      (head |> dest_Const |> fst,
blanchet@46382
  1866
       fold_rev (fn t as Var ((s, _), T) =>
blanchet@46382
  1867
                    (fn u => Abs (s, T, abstract_over (t, u)))
blanchet@48947
  1868
                  | _ => raise Fail "expected \"Var\"") args u)
blanchet@46382
  1869
    end
blanchet@46382
  1870
  | extract_lambda_def _ = raise Fail "malformed lifted lambda"
blanchet@46379
  1871
blanchet@46385
  1872
fun trans_lams_from_string ctxt type_enc lam_trans =
blanchet@46385
  1873
  if lam_trans = no_lamsN then
blanchet@46385
  1874
    rpair []
blanchet@46385
  1875
  else if lam_trans = hide_lamsN then
blanchet@46385
  1876
    lift_lams ctxt type_enc ##> K []
blanchet@47193
  1877
  else if lam_trans = liftingN orelse lam_trans = lam_liftingN then
blanchet@46385
  1878
    lift_lams ctxt type_enc
blanchet@47193
  1879
  else if lam_trans = combsN then
blanchet@46385
  1880
    map (introduce_combinators ctxt) #> rpair []
blanchet@47193
  1881
  else if lam_trans = combs_and_liftingN then
blanchet@47193
  1882
    lift_lams_part_1 ctxt type_enc
blanchet@47193
  1883
    ##> maps (fn t => [t, introduce_combinators ctxt (intentionalize_def t)])
blanchet@47689
  1884
    #> lift_lams_part_2 ctxt
blanchet@47196
  1885
  else if lam_trans = combs_or_liftingN then
blanchet@46425
  1886
    lift_lams_part_1 ctxt type_enc
blanchet@47196
  1887
    ##> map (fn t => case head_of (strip_qnt_body @{const_name All} t) of
blanchet@47196
  1888
                       @{term "op =::bool => bool => bool"} => t
blanchet@47196
  1889
                     | _ => introduce_combinators ctxt (intentionalize_def t))
blanchet@47689
  1890
    #> lift_lams_part_2 ctxt
blanchet@46385
  1891
  else if lam_trans = keep_lamsN then
blanchet@46385
  1892
    map (Envir.eta_contract) #> rpair []
blanchet@46385
  1893
  else
blanchet@46390
  1894
    error ("Unknown lambda translation scheme: " ^ quote lam_trans ^ ".")
blanchet@46385
  1895
blanchet@48996
  1896
val pull_and_reorder_definitions =
blanchet@48990
  1897
  let
blanchet@48990
  1898
    fun add_consts (IApp (t, u)) = fold add_consts [t, u]
blanchet@48990
  1899
      | add_consts (IAbs (_, t)) = add_consts t
blanchet@48990
  1900
      | add_consts (IConst (name, _, _)) = insert (op =) name
blanchet@48990
  1901
      | add_consts (IVar _) = I
blanchet@48996
  1902
    fun consts_of_hs l_or_r ({iformula, ...} : ifact) =
blanchet@48990
  1903
      case iformula of
blanchet@48990
  1904
        AAtom (IApp (IApp (IConst _, t), u)) => add_consts (l_or_r (t, u)) []
blanchet@48990
  1905
      | _ => []
blanchet@48990
  1906
    (* Quadratic, but usually OK. *)
blanchet@48996
  1907
    fun reorder [] [] = []
blanchet@48996
  1908
      | reorder (fact :: skipped) [] =
blanchet@48996
  1909
        fact :: reorder [] skipped (* break cycle *)
blanchet@48996
  1910
      | reorder skipped (fact :: facts) =
blanchet@48990
  1911
        let val rhs_consts = consts_of_hs snd fact in
blanchet@49111
  1912
          if exists (exists (exists (member (op =) rhs_consts)
blanchet@49111
  1913
                     o consts_of_hs fst)) [skipped, facts] then
blanchet@48996
  1914
            reorder (fact :: skipped) facts
blanchet@48990
  1915
          else
blanchet@48996
  1916
            fact :: reorder [] (facts @ skipped)
blanchet@48990
  1917
        end
blanchet@48996
  1918
  in List.partition (curry (op =) Definition o #role) #>> reorder [] #> op @ end
blanchet@48990
  1919
blanchet@48990
  1920
fun translate_formulas ctxt prem_role format type_enc lam_trans presimp hyp_ts
blanchet@46385
  1921
                       concl_t facts =
blanchet@43444
  1922
  let
blanchet@43444
  1923
    val thy = Proof_Context.theory_of ctxt
blanchet@46385
  1924
    val trans_lams = trans_lams_from_string ctxt type_enc lam_trans
blanchet@44732
  1925
    val fact_ts = facts |> map snd
blanchet@43444
  1926
    (* Remove existing facts from the conjecture, as this can dramatically
blanchet@43444
  1927
       boost an ATP's performance (for some reason). *)
blanchet@44033
  1928
    val hyp_ts =
blanchet@44033
  1929
      hyp_ts
blanchet@44033
  1930
      |> map (fn t => if member (op aconv) fact_ts t then @{prop True} else t)
blanchet@44735
  1931
    val facts = facts |> map (apsnd (pair Axiom))
blanchet@44735
  1932
    val conjs =
blanchet@48990
  1933
      map (pair prem_role) hyp_ts @ [(Conjecture, s_not_prop concl_t)]
blanchet@46039
  1934
      |> map (apsnd freeze_term)
blanchet@47168
  1935
      |> map2 (pair o rpair (Local, General) o string_of_int)
blanchet@47168
  1936
              (0 upto length hyp_ts)
blanchet@46425
  1937
    val ((conjs, facts), lam_facts) =
blanchet@46382
  1938
      (conjs, facts)
blanchet@48639
  1939
      |> presimp ? pairself (map (apsnd (apsnd (presimp_prop ctxt type_enc))))
blanchet@46385
  1940
      |> (if lam_trans = no_lamsN then
blanchet@46382
  1941
            rpair []
blanchet@46382
  1942
          else
blanchet@46382
  1943
            op @
blanchet@46385
  1944
            #> preprocess_abstractions_in_terms trans_lams
blanchet@46382
  1945
            #>> chop (length conjs))
blanchet@48996
  1946
    val conjs =
blanchet@48996
  1947
      conjs |> make_conjecture ctxt format type_enc
blanchet@48996
  1948
            |> pull_and_reorder_definitions
blanchet@48996
  1949
    val facts =
blanchet@48996
  1950
      facts |> map_filter (fn (name, (_, t)) =>
blanchet@48996
  1951
                              make_fact ctxt format type_enc true (name, t))
blanchet@48996
  1952
            |> pull_and_reorder_definitions
blanchet@48996
  1953
    val fact_names =
blanchet@48996
  1954
      facts |> map (fn {name, stature, ...} : ifact => (name, stature))
blanchet@47203
  1955
    val lifted = lam_facts |> map (extract_lambda_def o snd o snd)
blanchet@46425
  1956
    val lam_facts =
blanchet@46425
  1957
      lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd)
blanchet@44732
  1958
    val all_ts = concl_t :: hyp_ts @ fact_ts
blanchet@43444
  1959
    val subs = tfree_classes_of_terms all_ts
blanchet@43444
  1960
    val supers = tvar_classes_of_terms all_ts
blanchet@44030
  1961
    val tycons = type_constrs_of_terms thy all_ts
blanchet@49157
  1962
    val (supers, tcon_clauses) =
blanchet@44493
  1963
      if level_of_type_enc type_enc = No_Types then ([], [])
blanchet@49157
  1964
      else make_tcon_clauses thy tycons supers
blanchet@49157
  1965
    val subclass_pairs = make_subclass_pairs thy subs supers
blanchet@43444
  1966
  in
blanchet@46379
  1967
    (fact_names |> map single, union (op =) subs supers, conjs,
blanchet@49157
  1968
     facts @ lam_facts, subclass_pairs, tcon_clauses, lifted)
blanchet@43444
  1969
  end
blanchet@43444
  1970
nik@45350
  1971
val type_guard = `(make_fixed_const NONE) type_guard_name
blanchet@43971
  1972
blanchet@48638
  1973
fun type_guard_iterm type_enc T tm =
blanchet@45255
  1974
  IApp (IConst (type_guard, T --> @{typ bool}, [T])
blanchet@48638
  1975
        |> mangle_type_args_in_iterm type_enc, tm)
blanchet@43444
  1976
blanchet@44282
  1977
fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum
blanchet@49147
  1978
  | is_var_positively_naked_in_term name _ (ATerm (((s, _), _), tms)) accum =
blanchet@49147
  1979
    accum orelse
blanchet@49147
  1980
    (is_tptp_equal s andalso member (op =) tms (ATerm ((name, []), [])))
blanchet@44558
  1981
  | is_var_positively_naked_in_term _ _ _ _ = true
blanchet@45262
  1982
blanchet@49201
  1983
fun is_var_undercover_in_term thy name pos tm accum =
blanchet@49201
  1984
  accum orelse
blanchet@45673
  1985
  let
blanchet@49147
  1986
    val var = ATerm ((name, []), [])
blanchet@49110
  1987
    fun is_undercover (ATerm (_, [])) = false
blanchet@49147
  1988
      | is_undercover (ATerm (((s, _), _), tms)) =
blanchet@46819
  1989
        let
blanchet@46819
  1990
          val ary = length tms
blanchet@49215
  1991
          val cover = type_arg_cover thy pos s ary
blanchet@46819
  1992
        in
blanchet@49102
  1993
          exists (fn (j, tm) => tm = var andalso member (op =) cover j)
blanchet@49102
  1994
                 (0 upto ary - 1 ~~ tms) orelse
blanchet@49110
  1995
          exists is_undercover tms
blanchet@46819
  1996
        end
blanchet@49110
  1997
      | is_undercover _ = true
blanchet@49110
  1998
  in is_undercover tm end
blanchet@45673
  1999
blanchet@49102
  2000
fun should_guard_var_in_formula thy level pos phi (SOME true) name =
blanchet@49198
  2001
    (case level of
blanchet@49198
  2002
       All_Types => true
blanchet@49215
  2003
     | Undercover_Types =>
blanchet@49215
  2004
       formula_fold pos (is_var_undercover_in_term thy name) phi false
blanchet@49198
  2005
     | Nonmono_Types (_, Uniform) => true
blanchet@49198
  2006
     | Nonmono_Types (_, Non_Uniform) =>
blanchet@45673
  2007
       formula_fold pos (is_var_positively_naked_in_term name) phi false
blanchet@49198
  2008
     | _ => false)
blanchet@49102
  2009
  | should_guard_var_in_formula _ _ _ _ _ _ = true
blanchet@46819
  2010
blanchet@49102
  2011
fun always_guard_var_in_formula _ _ _ _ _ _ = true
blanchet@43705
  2012
blanchet@45264
  2013
fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false
blanchet@45639
  2014
  | should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T =
blanchet@49198
  2015
    not (is_type_level_uniform level) andalso
blanchet@45650
  2016
    should_encode_type ctxt mono level T
blanchet@45264
  2017
  | should_generate_tag_bound_decl _ _ _ _ _ = false
blanchet@45263
  2018
blanchet@48638
  2019
fun mk_aterm type_enc name T_args args =
blanchet@49148
  2020
  let val (ty_args, tm_args) = process_type_args type_enc T_args in
blanchet@49148
  2021
    ATerm ((name, ty_args), tm_args @ args)
blanchet@49148
  2022
  end
blanchet@43835
  2023
blanchet@48638
  2024
fun do_bound_type ctxt mono type_enc =
blanchet@47220
  2025
  case type_enc of
blanchet@48637
  2026
    Native (_, _, level) =>
blanchet@49153
  2027
    fused_type ctxt mono level 0 #> native_ho_type_from_typ type_enc false 0
blanchet@49153
  2028
    #> SOME
blanchet@47220
  2029
  | _ => K NONE
blanchet@47220
  2030
blanchet@48638
  2031
fun tag_with_type ctxt mono type_enc pos T tm =
blanchet@44730
  2032
  IConst (type_tag, T --> T, [T])
blanchet@48638
  2033
  |> mangle_type_args_in_iterm type_enc
blanchet@48638
  2034
  |> ho_term_from_iterm ctxt mono type_enc pos
blanchet@49147
  2035
  |> (fn ATerm ((s, tys), tms) => ATerm ((s, tys), tms @ [tm])
blanchet@44558
  2036
       | _ => raise Fail "unexpected lambda-abstraction")
blanchet@48638
  2037
and ho_term_from_iterm ctxt mono type_enc pos =
blanchet@43444
  2038
  let
blanchet@46820
  2039
    fun term site u =
blanchet@43803
  2040
      let
blanchet@44730
  2041
        val (head, args) = strip_iterm_comb u
nik@44536
  2042
        val pos =
blanchet@46820
  2043
          case site of
nik@44536
  2044
            Top_Level pos => pos
nik@44536
  2045
          | Eq_Arg pos => pos
blanchet@45676
  2046
          | _ => NONE
blanchet@49215
  2047
        val T = ityp_of u
nik@44536
  2048
        val t =
blanchet@43803
  2049
          case head of
blanchet@44730
  2050
            IConst (name as (s, _), _, T_args) =>
blanchet@48024
  2051
            let
blanchet@49161
  2052
              val ary = length args
blanchet@49161
  2053
              fun arg_site j =
blanchet@49161
  2054
                if is_tptp_equal s then Eq_Arg pos else Arg (s, j, ary)
blanchet@49161
  2055
            in
blanchet@49161
  2056
              map2 (fn j => term (arg_site j)) (0 upto ary - 1) args
blanchet@49161
  2057
              |> mk_aterm type_enc name T_args
blanchet@49161
  2058
            end
blanchet@44730
  2059
          | IVar (name, _) =>
blanchet@48638
  2060
            map (term Elsewhere) args |> mk_aterm type_enc name []
blanchet@44730
  2061
          | IAbs ((name, T), tm) =>
blanchet@47689
  2062
            if is_type_enc_higher_order type_enc then
blanchet@49153
  2063
              AAbs (((name, native_ho_type_from_typ type_enc true 0 T), (* FIXME? why "true"? *)
blanchet@48926
  2064
                     term Elsewhere tm), map (term Elsewhere) args)
blanchet@47689
  2065
            else
blanchet@47689
  2066
              raise Fail "unexpected lambda-abstraction"
blanchet@44730
  2067
          | IApp _ => raise Fail "impossible \"IApp\""
blanchet@49215
  2068
        val tag = should_tag_with_type ctxt mono type_enc site u T
blanchet@49215
  2069
      in t |> tag ? tag_with_type ctxt mono type_enc pos T end
blanchet@48926
  2070
  in term (Top_Level pos) end
blanchet@49102
  2071
and formula_from_iformula ctxt mono type_enc should_guard_var =
blanchet@43700
  2072
  let
blanchet@45673
  2073
    val thy = Proof_Context.theory_of ctxt
blanchet@45673
  2074
    val level = level_of_type_enc type_enc
blanchet@48638
  2075
    val do_term = ho_term_from_iterm ctxt mono type_enc
blanchet@43747
  2076
    fun do_out_of_bound_type pos phi universal (name, T) =
blanchet@45258
  2077
      if should_guard_type ctxt mono type_enc
blanchet@49102
  2078
             (fn () => should_guard_var thy level pos phi universal name) T then
blanchet@44730
  2079
        IVar (name, T)
blanchet@48638
  2080
        |> type_guard_iterm type_enc T
blanchet@44232
  2081
        |> do_term pos |> AAtom |> SOME
blanchet@45264
  2082
      else if should_generate_tag_bound_decl ctxt mono type_enc universal T then
blanchet@45264
  2083
        let
blanchet@49147
  2084
          val var = ATerm ((name, []), [])
blanchet@48638
  2085
          val tagged_var = tag_with_type ctxt mono type_enc pos T var
blanchet@49147
  2086
        in SOME (AAtom (ATerm ((`I tptp_equal, []), [tagged_var, var]))) end
blanchet@43444
  2087
      else
blanchet@43444
  2088
        NONE
blanchet@49148
  2089
    fun do_formula pos (ATyQuant (q, xs, phi)) =
blanchet@49153
  2090
        ATyQuant (q, map (apfst (native_ho_type_from_typ type_enc false 0)) xs,
blanchet@49148
  2091
                  do_formula pos phi)
blanchet@49148
  2092
      | do_formula pos (AQuant (q, xs, phi)) =
blanchet@43747
  2093
        let
blanchet@43747
  2094
          val phi = phi |> do_formula pos
blanchet@43747
  2095
          val universal = Option.map (q = AExists ? not) pos
blanchet@48638
  2096
          val do_bound_type = do_bound_type ctxt mono type_enc
blanchet@43747
  2097
        in
blanchet@43705
  2098
          AQuant (q, xs |> map (apsnd (fn NONE => NONE
blanchet@43705
  2099
                                        | SOME T => do_bound_type T)),
blanchet@43705
  2100
                  (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd))
blanchet@43705
  2101
                      (map_filter
blanchet@43705
  2102
                           (fn (_, NONE) => NONE
blanchet@43705
  2103
                             | (s, SOME T) =>
blanchet@43747
  2104
                               do_out_of_bound_type pos phi universal (s, T))
blanchet@43747
  2105
                           xs)
blanchet@43705
  2106
                      phi)
blanchet@43705
  2107
        end
blanchet@43747
  2108
      | do_formula pos (AConn conn) = aconn_map pos do_formula conn
blanchet@44232
  2109
      | do_formula pos (AAtom tm) = AAtom (do_term pos tm)
blanchet@44364
  2110
  in do_formula end
blanchet@43444
  2111
blanchet@49453
  2112
fun string_of_status General = ""
blanchet@49453
  2113
  | string_of_status Induction = inductionN
blanchet@49453
  2114
  | string_of_status Intro = introN
blanchet@49453
  2115
  | string_of_status Inductive = inductiveN
blanchet@49453
  2116
  | string_of_status Elim = elimN
blanchet@49453
  2117
  | string_of_status Simp = simpN
blanchet@49453
  2118
  | string_of_status Non_Rec_Def = non_rec_defN
blanchet@49453
  2119
  | string_of_status Rec_Def = rec_defN
blanchet@49453
  2120
blanchet@43444
  2121
(* Each fact is given a unique fact number to avoid name clashes (e.g., because
blanchet@43444
  2122
   of monomorphization). The TPTP explicitly forbids name clashes, and some of
blanchet@43444
  2123
   the remote provers might care. *)
blanchet@49161
  2124
fun line_for_fact ctxt prefix encode freshen pos mono type_enc rank
blanchet@49453
  2125
        (j, {name, stature = (_, status), role, iformula, atomic_types}) =
blanchet@49453
  2126
  Formula (prefix ^ (if freshen then string_of_int j ^ "_" else "") ^
blanchet@49453
  2127
           encode name,
blanchet@49453
  2128
           role,
blanchet@49453
  2129
           iformula
blanchet@49453
  2130
           |> formula_from_iformula ctxt mono type_enc
blanchet@49453
  2131
                  should_guard_var_in_formula (if pos then SOME true else NONE)
blanchet@49453
  2132
           |> close_formula_universally
blanchet@49453
  2133
           |> bound_tvars type_enc true atomic_types,
blanchet@49453
  2134
           NONE, isabelle_info (string_of_status status) (rank j))
blanchet@43444
  2135
blanchet@49161
  2136
fun lines_for_subclass type_enc sub super =
blanchet@49157
  2137
  Formula (subclass_prefix ^ ascii_of sub ^ "___" ^ ascii_of super, Axiom,
blanchet@49148
  2138
           AConn (AImplies,
blanchet@49157
  2139
                  [sub, super] |> map (fn s => class_atom type_enc (s, tvar_a)))
blanchet@49148
  2140
           |> bound_tvars type_enc false [tvar_a],
blanchet@49148
  2141
           NONE, isabelle_info inductiveN helper_rank)
blanchet@49148
  2142
blanchet@49161
  2143
fun lines_for_subclass_pair type_enc (sub, supers) =
blanchet@49157
  2144
  if polymorphism_of_type_enc type_enc = Type_Class_Polymorphic then
blanchet@49157
  2145
    [Class_Decl (class_decl_prefix ^ ascii_of sub, `make_class sub,
blanchet@49157
  2146
                 map (`make_class) supers)]
blanchet@49157
  2147
  else
blanchet@49161
  2148
    map (lines_for_subclass type_enc sub) supers
blanchet@49157
  2149
blanchet@49161
  2150
fun line_for_tcon_clause type_enc (name, prems, (cl, T)) =
blanchet@49157
  2151
  if polymorphism_of_type_enc type_enc = Type_Class_Polymorphic then
blanchet@49157
  2152
    Class_Memb (class_memb_prefix ^ name,
blanchet@49157
  2153
                map (fn (cls, T) =>
blanchet@49157
  2154
                        (T |> dest_TVar |> tvar_name, map (`make_class) cls))
blanchet@49157
  2155
                    prems,
blanchet@49157
  2156
                native_ho_type_from_typ type_enc false 0 T, `make_class cl)
blanchet@49157
  2157
  else
blanchet@49157
  2158
    Formula (tcon_clause_prefix ^ name, Axiom,
blanchet@49157
  2159
             mk_ahorn (maps (class_atoms type_enc) prems)
blanchet@49157
  2160
                      (class_atom type_enc (cl, T))
blanchet@49157
  2161
             |> bound_tvars type_enc true (snd (dest_Type T)),
blanchet@48019
  2162
             NONE, isabelle_info inductiveN helper_rank)
blanchet@43444
  2163
blanchet@49161
  2164
fun line_for_conjecture ctxt mono type_enc
blanchet@49161
  2165
                        ({name, role, iformula, atomic_types, ...} : ifact) =
blanchet@48990
  2166
  Formula (conjecture_prefix ^ name, role,
blanchet@46187
  2167
           iformula
blanchet@49102
  2168
           |> formula_from_iformula ctxt mono type_enc
blanchet@46187
  2169
                  should_guard_var_in_formula (SOME false)
blanchet@46248
  2170
           |> close_formula_universally
blanchet@47234
  2171
           |> bound_tvars type_enc true atomic_types, NONE, [])
blanchet@43444
  2172
blanchet@49161
  2173
fun lines_for_free_types type_enc (facts : ifact list) =
blanchet@49200
  2174
  if is_type_enc_polymorphic type_enc then
blanchet@49200
  2175
    let
blanchet@49200
  2176
      val type_classes =
blanchet@49200
  2177
        (polymorphism_of_type_enc type_enc = Type_Class_Polymorphic)
blanchet@49200
  2178
      fun line j (cl, T) =
blanchet@49200
  2179
        if type_classes then
blanchet@49200
  2180
          Class_Memb (class_memb_prefix ^ string_of_int j, [],
blanchet@49200
  2181
                      native_ho_type_from_typ type_enc false 0 T,
blanchet@49200
  2182
                      `make_class cl)
blanchet@49200
  2183
        else
blanchet@49200
  2184
          Formula (tfree_clause_prefix ^ string_of_int j, Hypothesis,
blanchet@49200
  2185
                   class_atom type_enc (cl, T), NONE, [])
blanchet@49200
  2186
      val membs =
blanchet@49200
  2187
        fold (union (op =)) (map #atomic_types facts) []
blanchet@49200
  2188
        |> class_membs_for_types type_enc add_sorts_on_tfree
blanchet@49200
  2189
    in map2 line (0 upto length membs - 1) membs end
blanchet@49200
  2190
  else
blanchet@49200
  2191
    []
blanchet@43444
  2192
blanchet@43444
  2193
(** Symbol declarations **)
blanchet@43415
  2194
blanchet@49156
  2195
fun decl_line_for_class phantoms s =
blanchet@49157
  2196
  let val name as (s, _) = `make_class s in
blanchet@49152
  2197
    Sym_Decl (sym_decl_prefix ^ s, name,
blanchet@49152
  2198
              APi ([tvar_a_name],
blanchet@49152
  2199
                   if phantoms = Without_Phantom_Type_Vars then
blanchet@49152
  2200
                     AFun (a_itself_atype, bool_atype)
blanchet@49152
  2201
                   else
blanchet@49152
  2202
                     bool_atype))
blanchet@45459
  2203
  end
blanchet@45459
  2204
blanchet@45459
  2205
fun decl_lines_for_classes type_enc classes =
blanchet@45459
  2206
  case type_enc of
blanchet@49156
  2207
    Native (_, Raw_Polymorphic phantoms, _) =>
blanchet@49156
  2208
    map (decl_line_for_class phantoms) classes
blanchet@45459
  2209
  | _ => []
blanchet@45459
  2210
blanchet@48638
  2211
fun sym_decl_table_for_facts thy type_enc sym_tab (conjs, facts, extra_tms) =
blanchet@43445
  2212
  let
blanchet@45700
  2213
    fun add_iterm_syms tm =
blanchet@44730
  2214
      let val (head, args) = strip_iterm_comb tm in
blanchet@43445
  2215
        (case head of
blanchet@44730
  2216
           IConst ((s, s'), T, T_args) =>
blanchet@45458
  2217
           let
blanchet@45700
  2218
             val (pred_sym, in_conj) =
blanchet@45700
  2219
               case Symtab.lookup sym_tab s of
blanchet@45724
  2220
                 SOME ({pred_sym, in_conj, ...} : sym_info) =>
blanchet@45724
  2221
                 (pred_sym, in_conj)
blanchet@45700
  2222
               | NONE => (false, false)
blanchet@45458
  2223
             val decl_sym =
blanchet@45458
  2224
               (case type_enc of
blanchet@45458
  2225
                  Guards _ => not pred_sym
blanchet@45458
  2226
                | _ => true) andalso
blanchet@45458
  2227
               is_tptp_user_symbol s
blanchet@45458
  2228
           in
blanchet@45458
  2229
             if decl_sym then
blanchet@43447
  2230
               Symtab.map_default (s, [])
blanchet@48021
  2231
                   (insert_type thy #3 (s', T_args, T, pred_sym, length args,
blanchet@48021
  2232
                                        in_conj))
blanchet@43445
  2233
             else
blanchet@43445
  2234
               I
blanchet@43445
  2235
           end
blanchet@45700
  2236
         | IAbs (_, tm) => add_iterm_syms tm
blanchet@43445
  2237
         | _ => I)
blanchet@45700
  2238
        #> fold add_iterm_syms args
blanchet@43445
  2239
      end
blanchet@48996
  2240
    val add_fact_syms = K add_iterm_syms |> formula_fold NONE |> ifact_lift
blanchet@49148
  2241
    fun add_formula_var_types (ATyQuant (_, _, phi)) = add_formula_var_types phi
blanchet@49148
  2242
      | add_formula_var_types (AQuant (_, xs, phi)) =
blanchet@48021
  2243
        fold (fn (_, SOME T) => insert_type thy I T | _ => I) xs
blanchet@44837
  2244
        #> add_formula_var_types phi
blanchet@44837
  2245
      | add_formula_var_types (AConn (_, phis)) =
blanchet@44837
  2246
        fold add_formula_var_types phis
blanchet@44837
  2247
      | add_formula_var_types _ = I
blanchet@44837
  2248
    fun var_types () =
blanchet@49146
  2249
      if is_type_enc_polymorphic type_enc then [tvar_a]
blanchet@48996
  2250
      else fold (ifact_lift add_formula_var_types) (conjs @ facts) []
blanchet@44837
  2251
    fun add_undefined_const T =
blanchet@44855
  2252
      let
blanchet@49215
  2253
        (* FIXME: make sure type arguments are filtered / clean up code *)
blanchet@44855
  2254
        val (s, s') =
blanchet@45480
  2255
          `(make_fixed_const NONE) @{const_name undefined}
blanchet@49216
  2256
          |> (is_type_enc_mangling type_enc ? mangled_const_name type_enc [T])
blanchet@44855
  2257
      in
blanchet@44855
  2258
        Symtab.map_default (s, [])
blanchet@48021
  2259
                           (insert_type thy #3 (s', [T], T, false, 0, false))
blanchet@44855
  2260
      end
blanchet@45480
  2261
    fun add_TYPE_const () =
blanchet@45480
  2262
      let val (s, s') = TYPE_name in
blanchet@45480
  2263
        Symtab.map_default (s, [])
blanchet@48021
  2264
            (insert_type thy #3
blanchet@45480
  2265
                         (s', [tvar_a], @{typ "'a itself"}, false, 0, false))
blanchet@45480
  2266
      end
blanchet@43568
  2267
  in
blanchet@43568
  2268
    Symtab.empty
blanchet@49104
  2269
    |> is_type_enc_sound type_enc
blanchet@45700
  2270
       ? (fold (fold add_fact_syms) [conjs, facts]
blanchet@47220
  2271
          #> fold add_iterm_syms extra_tms
blanchet@44856
  2272
          #> (case type_enc of
blanchet@49146
  2273
                Native (First_Order, Raw_Polymorphic phantoms, _) =>
blanchet@49146
  2274
                phantoms = Without_Phantom_Type_Vars ? add_TYPE_const ()
blanchet@48637
  2275
              | Native _ => I
blanchet@44856
  2276
              | _ => fold add_undefined_const (var_types ())))
blanchet@43568
  2277
  end
blanchet@43445
  2278
blanchet@45258
  2279
(* We add "bool" in case the helper "True_or_False" is included later. *)
blanchet@49158
  2280
fun default_mono level completish =
blanchet@45258
  2281
  {maybe_finite_Ts = [@{typ bool}],
blanchet@45492
  2282
   surely_infinite_Ts =
blanchet@45492
  2283
     case level of
blanchet@49107
  2284
       Nonmono_Types (Strict, _) => []
blanchet@45492
  2285
     | _ => known_infinite_types,
blanchet@49158
  2286
   maybe_nonmono_Ts = [if completish then tvar_a else @{typ bool}]}
blanchet@45258
  2287
blanchet@43555
  2288
(* This inference is described in section 2.3 of Claessen et al.'s "Sorting it
blanchet@43555
  2289
   out with monotonicity" paper presented at CADE 2011. *)
blanchet@45258
  2290
fun add_iterm_mononotonicity_info _ _ (SOME false) _ mono = mono
blanchet@45258
  2291
  | add_iterm_mononotonicity_info ctxt level _
blanchet@45258
  2292
        (IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2))
blanchet@49118
  2293
        (mono as {maybe_finite_Ts, surely_infinite_Ts, maybe_nonmono_Ts}) =
blanchet@48021
  2294
    let val thy = Proof_Context.theory_of ctxt in
blanchet@48021
  2295
      if is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2] then
blanchet@48021
  2296
        case level of
blanchet@49107
  2297
          Nonmono_Types (strictness, _) =>
blanchet@48021
  2298
          if exists (type_instance thy T) surely_infinite_Ts orelse
blanchet@48021
  2299
             member (type_equiv thy) maybe_finite_Ts T then
blanchet@48021
  2300
            mono
blanchet@48021
  2301
          else if is_type_kind_of_surely_infinite ctxt strictness
blanchet@48021
  2302
                                                  surely_infinite_Ts T then
blanchet@48021
  2303
            {maybe_finite_Ts = maybe_finite_Ts,
blanchet@48021
  2304
             surely_infinite_Ts = surely_infinite_Ts |> insert_type thy I T,
blanchet@48021
  2305
             maybe_nonmono_Ts = maybe_nonmono_Ts}
blanchet@48021
  2306
          else
blanchet@48021
  2307
            {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv thy) T,
blanchet@48021
  2308
             surely_infinite_Ts = surely_infinite_Ts,
blanchet@48021
  2309
             maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T}
blanchet@48021
  2310
        | _ => mono
blanchet@48021
  2311
      else
blanchet@48021
  2312
        mono
blanchet@48021
  2313
    end
blanchet@45258
  2314
  | add_iterm_mononotonicity_info _ _ _ _ mono = mono
blanchet@48996
  2315
fun add_fact_mononotonicity_info ctxt level ({role, iformula, ...} : ifact) =
blanchet@48990
  2316
  formula_fold (SOME (role <> Conjecture))
blanchet@45258
  2317
               (add_iterm_mononotonicity_info ctxt level) iformula
blanchet@49158
  2318
fun mononotonicity_info_for_facts ctxt type_enc completish facts =
blanchet@44493
  2319
  let val level = level_of_type_enc type_enc in
blanchet@49158
  2320
    default_mono level completish
blanchet@45258
  2321
    |> is_type_level_monotonicity_based level
blanchet@45258
  2322
       ? fold (add_fact_mononotonicity_info ctxt level) facts
blanchet@43700
  2323
  end
blanchet@43547
  2324
blanchet@45356
  2325
fun add_iformula_monotonic_types ctxt mono type_enc =
blanchet@45356
  2326
  let
blanchet@48021
  2327
    val thy = Proof_Context.theory_of ctxt
blanchet@45356
  2328
    val level = level_of_type_enc type_enc
blanchet@45356
  2329
    val should_encode = should_encode_type ctxt mono level
blanchet@48021
  2330
    fun add_type T = not (should_encode T) ? insert_type thy I T
blanchet@45361
  2331
    fun add_args (IApp (tm1, tm2)) = add_args tm1 #> add_term tm2
blanchet@45361
  2332
      | add_args _ = I
blanchet@45361
  2333
    and add_term tm = add_type (ityp_of tm) #> add_args tm
blanchet@45356
  2334
  in formula_fold NONE (K add_term) end
blanchet@45356
  2335
fun add_fact_monotonic_types ctxt mono type_enc =
blanchet@48996
  2336
  add_iformula_monotonic_types ctxt mono type_enc |> ifact_lift
blanchet@45356
  2337
fun monotonic_types_for_facts ctxt mono type_enc facts =
blanchet@45673
  2338
  let val level = level_of_type_enc type_enc in
blanchet@49146
  2339
    [] |> (is_type_enc_polymorphic type_enc andalso
blanchet@49198
  2340
           is_type_level_monotonicity_based level)
blanchet@45673
  2341
          ? fold (add_fact_monotonic_types ctxt mono type_enc) facts
blanchet@45673
  2342
  end
blanchet@45356
  2343
blanchet@49161
  2344
fun line_for_guards_mono_type ctxt mono type_enc T =
blanchet@45255
  2345
  Formula (guards_sym_formula_prefix ^
blanchet@48638
  2346
           ascii_of (mangled_type type_enc T),
blanchet@45255
  2347
           Axiom,
blanchet@45255
  2348
           IConst (`make_bound_var "X", T, [])
blanchet@48638
  2349
           |> type_guard_iterm type_enc T
blanchet@45255
  2350
           |> AAtom
blanchet@49102
  2351
           |> formula_from_iformula ctxt mono type_enc
blanchet@46819
  2352
                                    always_guard_var_in_formula (SOME true)
blanchet@46248
  2353
           |> close_formula_universally
blanchet@46791
  2354
           |> bound_tvars type_enc true (atomic_types_of T),
blanchet@48019
  2355
           NONE, isabelle_info inductiveN helper_rank)
blanchet@45255
  2356
blanchet@49161
  2357
fun line_for_tags_mono_type ctxt mono type_enc T =
blanchet@49147
  2358
  let val x_var = ATerm ((`make_bound_var "X", []), []) in
blanchet@45255
  2359
    Formula (tags_sym_formula_prefix ^
blanchet@48638
  2360
             ascii_of (mangled_type type_enc T),
blanchet@45255
  2361
             Axiom,
blanchet@47220
  2362
             eq_formula type_enc (atomic_types_of T) [] false
blanchet@48638
  2363
                  (tag_with_type ctxt mono type_enc NONE T x_var) x_var,
blanchet@49453
  2364
             NONE, isabelle_info non_rec_defN helper_rank)
blanchet@45255
  2365
  end
blanchet@45255
  2366
blanchet@49155
  2367
fun lines_for_mono_types ctxt mono type_enc Ts =
blanchet@45255
  2368
  case type_enc of
blanchet@48637
  2369
    Native _ => []
blanchet@49161
  2370
  | Guards _ => map (line_for_guards_mono_type ctxt mono type_enc) Ts
blanchet@49161
  2371
  | Tags _ => map (line_for_tags_mono_type ctxt mono type_enc) Ts
blanchet@45255
  2372
blanchet@48638
  2373
fun decl_line_for_sym ctxt mono type_enc s
blanchet@43835
  2374
                      (s', T_args, T, pred_sym, ary, _) =
blanchet@43835
  2375
  let
blanchet@45458
  2376
    val thy = Proof_Context.theory_of ctxt
blanchet@45458
  2377
    val (T, T_args) =
blanchet@45458
  2378
      if null T_args then
blanchet@45458
  2379
        (T, [])
blanchet@46382
  2380
      else case unprefix_and_unascii const_prefix s of
blanchet@45458
  2381
        SOME s' =>
blanchet@45458
  2382
        let
blanchet@45458
  2383
          val s' = s' |> invert_const
blanchet@45458
  2384
          val T = s' |> robust_const_type thy
blanchet@45458
  2385
        in (T, robust_const_typargs thy (s', T)) end
blanchet@46380
  2386
      | NONE => raise Fail "unexpected type arguments"
blanchet@43835
  2387
  in
blanchet@49152
  2388
    Sym_Decl (sym_decl_prefix ^ s, (s, s'),
blanchet@49152
  2389
              T |> fused_type ctxt mono (level_of_type_enc type_enc) ary
blanchet@49153
  2390
                |> native_ho_type_from_typ type_enc pred_sym ary
blanchet@49152
  2391
                |> not (null T_args)
blanchet@49157
  2392
                   ? curry APi (map (tvar_name o dest_TVar) T_args))
blanchet@43835
  2393
  end
blanchet@43450
  2394
blanchet@48990
  2395
fun honor_conj_sym_role in_conj =
blanchet@48927
  2396
  if in_conj then (Hypothesis, I) else (Axiom, I)
blanchet@47220
  2397
blanchet@49161
  2398
fun line_for_guards_sym_decl ctxt mono type_enc n s j
blanchet@49161
  2399
                             (s', T_args, T, _, ary, in_conj) =
blanchet@43450
  2400
  let
blanchet@45673
  2401
    val thy = Proof_Context.theory_of ctxt
blanchet@48990
  2402
    val (role, maybe_negate) = honor_conj_sym_role in_conj
blanchet@43618
  2403
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@45676
  2404
    val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
blanchet@43700
  2405
    val bounds =
blanchet@44730
  2406
      bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, []))
blanchet@43450
  2407
    val bound_Ts =
blanchet@45673
  2408
      if exists (curry (op =) dummyT) T_args then
blanchet@49215
  2409
        let val cover = type_arg_cover thy NONE s ary in
blanchet@49095
  2410
          map2 (fn j => if member (op =) cover j then SOME else K NONE)
blanchet@49095
  2411
               (0 upto ary - 1) arg_Ts
blanchet@49095
  2412
        end
blanchet@45673
  2413
      else
blanchet@45676
  2414
        replicate ary NONE
blanchet@43450
  2415
  in
blanchet@44860
  2416
    Formula (guards_sym_formula_prefix ^ s ^
blanchet@48990
  2417
             (if n > 1 then "_" ^ string_of_int j else ""), role,
blanchet@44730
  2418
             IConst ((s, s'), T, T_args)
blanchet@44730
  2419
             |> fold (curry (IApp o swap)) bounds
blanchet@48638
  2420
             |> type_guard_iterm type_enc res_T
blanchet@43804
  2421
             |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts)
blanchet@49102
  2422
             |> formula_from_iformula ctxt mono type_enc
blanchet@46819
  2423
                                      always_guard_var_in_formula (SOME true)
blanchet@46248
  2424
             |> close_formula_universally
blanchet@46791
  2425
             |> bound_tvars type_enc (n > 1) (atomic_types_of T)
blanchet@43580
  2426
             |> maybe_negate,
blanchet@48019
  2427
             NONE, isabelle_info inductiveN helper_rank)
blanchet@43450
  2428
  end
blanchet@43450
  2429
blanchet@49161
  2430
fun lines_for_tags_sym_decl ctxt mono type_enc n s
blanchet@49161
  2431
                            (j, (s', T_args, T, pred_sym, ary, in_conj)) =
blanchet@43700
  2432
  let
blanchet@45676
  2433
    val thy = Proof_Context.theory_of ctxt
blanchet@45676
  2434
    val level = level_of_type_enc type_enc
blanchet@49198
  2435
    val ident =
blanchet@45255
  2436
      tags_sym_formula_prefix ^ s ^
blanchet@43966
  2437
      (if n > 1 then "_" ^ string_of_int j else "")
blanchet@48990
  2438
    val (role, maybe_negate) = honor_conj_sym_role in_conj
blanchet@43700
  2439
    val (arg_Ts, res_T) = chop_fun ary T
blanchet@45676
  2440
    val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int)
blanchet@49147
  2441
    val bounds = bound_names |> map (fn name => ATerm ((name, []), []))
blanchet@48638
  2442
    val cst = mk_aterm type_enc (s, s') T_args
blanchet@47220
  2443
    val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) [] pred_sym
blanchet@48638
  2444
    val tag_with = tag_with_type ctxt mono type_enc NONE
blanchet@49215
  2445
    fun formula c =
blanchet@49215
  2446
      [Formula (ident, role, eq (tag_with res_T c) c, NONE, isabelle_info
blanchet@49453
  2447
                non_rec_defN helper_rank)]
blanchet@49215
  2448
  in
blanchet@49215
  2449
    if pred_sym orelse not (should_encode_type ctxt mono level res_T) then
blanchet@49215
  2450
      []
blanchet@49215
  2451
    else if level = Undercover_Types then
blanchet@49215
  2452
      let
blanchet@49215
  2453
        val cover = type_arg_cover thy NONE s ary
blanchet@49215
  2454
        fun maybe_tag (j, arg_T) = member (op =) cover j ? tag_with arg_T
blanchet@49215
  2455
        val bounds = bounds |> map2 maybe_tag (0 upto ary - 1 ~~ arg_Ts)
blanchet@49215
  2456
      in formula (cst bounds) end
blanchet@49215
  2457
    else
blanchet@49215
  2458
      formula (cst bounds)
blanchet@49215
  2459
  end
blanchet@43700
  2460
blanchet@43707
  2461
fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd
blanchet@43707
  2462
blanchet@48021
  2463
fun rationalize_decls thy (decls as decl :: (decls' as _ :: _)) =
blanchet@46651
  2464
    let
blanchet@46651
  2465
      val T = result_type_of_decl decl
blanchet@46651
  2466
              |> map_type_tvar (fn (z, _) => TVar (z, HOLogic.typeS))
blanchet@46651
  2467
    in
blanchet@48021
  2468
      if forall (type_generalization thy T o result_type_of_decl) decls' then
blanchet@46651
  2469
        [decl]
blanchet@46651
  2470
      else
blanchet@46651
  2471
        decls
blanchet@46651
  2472
    end
blanchet@46651
  2473
  | rationalize_decls _ decls = decls
blanchet@46651
  2474
blanchet@49155
  2475
fun lines_for_sym_decls ctxt mono type_enc (s, decls) =
blanchet@44493
  2476
  case type_enc of
blanchet@48638
  2477
    Native _ => [decl_line_for_sym ctxt mono type_enc s (hd decls)]
blanchet@45639
  2478
  | Guards (_, level) =>
blanchet@43839
  2479
    let
blanchet@48021
  2480
      val thy = Proof_Context.theory_of ctxt
blanchet@48021
  2481
      val decls = decls |> rationalize_decls thy
blanchet@43839
  2482
      val n = length decls
blanchet@43839
  2483
      val decls =
blanchet@45258
  2484
        decls |> filter (should_encode_type ctxt mono level
blanchet@44264
  2485
                         o result_type_of_decl)
blanchet@43839
  2486
    in
blanchet@43839
  2487
      (0 upto length decls - 1, decls)
blanchet@49161
  2488
      |-> map2 (line_for_guards_sym_decl ctxt mono type_enc n s)
blanchet@43839
  2489
    end
blanchet@45639
  2490
  | Tags (_, level) =>
blanchet@49198
  2491
    if is_type_level_uniform level then
blanchet@45639
  2492
      []
blanchet@45639
  2493
    else
blanchet@45639
  2494
      let val n = length decls in
blanchet@45639
  2495
        (0 upto n - 1 ~~ decls)
blanchet@49161
  2496
        |> maps (lines_for_tags_sym_decl ctxt mono type_enc n s)
blanchet@45639
  2497
      end
blanchet@43450
  2498
blanchet@49155
  2499
fun lines_for_sym_decl_table ctxt mono type_enc mono_Ts sym_decl_tab =
blanchet@45255
  2500
  let
blanchet@45255
  2501
    val syms = sym_decl_tab |> Symtab.dest |> sort_wrt fst
blanchet@49155
  2502
    val mono_lines = lines_for_mono_types ctxt mono type_enc mono_Ts
blanchet@49253
  2503
    val decl_lines = maps (lines_for_sym_decls ctxt mono type_enc) syms
blanchet@45255
  2504
  in mono_lines @ decl_lines end
blanchet@43414
  2505
blanchet@47220
  2506
fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2)
blanchet@47220
  2507
blanchet@49103
  2508
fun do_uncurried_alias_lines_for_sym ctxt constrs mono type_enc sym_tab0 sym_tab
blanchet@49103
  2509
                                     base_s0 types in_conj =
blanchet@47220
  2510
  let
blanchet@47220
  2511
    fun do_alias ary =
blanchet@47220
  2512
      let
blanchet@47220
  2513
        val thy = Proof_Context.theory_of ctxt
blanchet@48990
  2514
        val (role, maybe_negate) = honor_conj_sym_role in_conj
blanchet@48638
  2515
        val base_name = base_s0 |> `(make_fixed_const (SOME type_enc))
blanchet@47220
  2516
        val T = case types of [T] => T | _ => robust_const_type thy base_s0
blanchet@47220
  2517
        val T_args = robust_const_typargs thy (base_s0, T)
blanchet@47220
  2518
        val (base_name as (base_s, _), T_args) =
blanchet@48638
  2519
          mangle_type_args_in_const type_enc base_name T_args
blanchet@47230
  2520
        val base_ary = min_ary_of sym_tab0 base_s
blanchet@47220
  2521
        fun do_const name = IConst (name, T, T_args)
blanchet@49103
  2522
        val filter_ty_args = filter_type_args_in_iterm thy constrs type_enc
blanchet@48638
  2523
        val ho_term_of = ho_term_from_iterm ctxt mono type_enc (SOME true)
blanchet@47220
  2524
        val name1 as (s1, _) =
blanchet@47237
  2525
          base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1)
blanchet@47237
  2526
        val name2 as (s2, _) = base_name |> aliased_uncurried ary
blanchet@47220
  2527
        val (arg_Ts, _) = chop_fun ary T
blanchet@47220
  2528
        val bound_names =
blanchet@47220
  2529
          1 upto ary |> map (`I o make_bound_var o string_of_int)
blanchet@47220
  2530
        val bounds = bound_names ~~ arg_Ts
blanchet@47220
  2531
        val (first_bounds, last_bound) =
blanchet@47220
  2532
          bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last
blanchet@47220
  2533
        val tm1 =
blanchet@48638
  2534
          mk_app_op type_enc (list_app (do_const name1) first_bounds) last_bound
blanchet@47265
  2535
          |> filter_ty_args
blanchet@47265
  2536
        val tm2 =
blanchet@47265
  2537
          list_app (do_const name2) (first_bounds @ [last_bound])
blanchet@47265
  2538
          |> filter_ty_args
blanchet@48638
  2539
        val do_bound_type = do_bound_type ctxt mono type_enc
blanchet@47220
  2540
        val eq =
blanchet@47220
  2541
          eq_formula type_enc (atomic_types_of T)
blanchet@47220
  2542
                     (map (apsnd do_bound_type) bounds) false
blanchet@47265
  2543
                     (ho_term_of tm1) (ho_term_of tm2)
blanchet@47220
  2544
      in
blanchet@47220
  2545
        ([tm1, tm2],
blanchet@48990
  2546
         [Formula (uncurried_alias_eq_prefix ^ s2, role, eq |> maybe_negate,
blanchet@49453
  2547
                   NONE, isabelle_info non_rec_defN helper_rank)])
blanchet@47220
  2548
        |> (if ary - 1 = base_ary orelse Symtab.defined sym_tab s1 then I
blanchet@47220
  2549
            else pair_append (do_alias (ary - 1)))
blanchet@47220
  2550
      end
blanchet@47220
  2551
  in do_alias end
blanchet@49103
  2552
fun uncurried_alias_lines_for_sym ctxt constrs mono type_enc sym_tab0
blanchet@48927
  2553
        sym_tab (s, {min_ary, types, in_conj, ...} : sym_info) =
blanchet@47220
  2554
  case unprefix_and_unascii const_prefix s of
blanchet@47220
  2555
    SOME mangled_s =>
blanchet@47237
  2556
    if String.isSubstring uncurried_alias_sep mangled_s then
blanchet@47220
  2557
      let
blanchet@47220
  2558
        val base_s0 = mangled_s |> unmangled_const_name |> hd |> invert_const
blanchet@47220
  2559
      in
blanchet@49103
  2560
        do_uncurried_alias_lines_for_sym ctxt constrs mono type_enc sym_tab0
blanchet@49103
  2561
                                         sym_tab base_s0 types in_conj min_ary
blanchet@47220
  2562
      end
blanchet@47220
  2563
    else
blanchet@47220
  2564
      ([], [])
blanchet@47220
  2565
  | NONE => ([], [])
blanchet@49103
  2566
fun uncurried_alias_lines_for_sym_table ctxt constrs mono type_enc
blanchet@48927
  2567
                                        uncurried_aliases sym_tab0 sym_tab =
blanchet@47220
  2568
  ([], [])
blanchet@47237
  2569
  |> uncurried_aliases
blanchet@47230
  2570
     ? Symtab.fold_rev
blanchet@47230
  2571
           (pair_append
blanchet@49103
  2572
            o uncurried_alias_lines_for_sym ctxt constrs mono type_enc sym_tab0
blanchet@49103
  2573
                                            sym_tab) sym_tab
blanchet@47220
  2574
blanchet@43839
  2575
val implicit_declsN = "Should-be-implicit typings"
blanchet@43839
  2576
val explicit_declsN = "Explicit typings"
blanchet@47237
  2577
val uncurried_alias_eqsN = "Uncurried aliases"
blanchet@41405
  2578
val factsN = "Relevant facts"
blanchet@49157
  2579
val subclassesN = "Subclasses"
blanchet@49157
  2580
val tconsN = "Type constructors"
blanchet@41405
  2581
val helpersN = "Helper facts"
blanchet@41405
  2582
val conjsN = "Conjectures"
blanchet@49157
  2583
val free_typesN = "Free types"
blanchet@41405
  2584
blanchet@46698
  2585
(* TFF allows implicit declarations of types, function symbols, and predicate
blanchet@46698
  2586
   symbols (with "$i" as the type of individuals), but some provers (e.g.,
blanchet@46698
  2587
   SNARK) require explicit declarations. The situation is similar for THF. *)
blanchet@46698
  2588
blanchet@49156
  2589
fun default_type pred_sym =
blanchet@46698
  2590
  let
blanchet@49151
  2591
    fun typ 0 0 = if pred_sym then bool_atype else individual_atype
blanchet@49151
  2592
      | typ 0 tm_ary = AFun (individual_atype, typ 0 (tm_ary - 1))
blanchet@49148
  2593
      | typ ty_ary tm_ary = APi (replicate ty_ary tvar_a_name, typ 0 tm_ary)
blanchet@46698
  2594
  in typ end
blanchet@46698
  2595
blanchet@49155
  2596
fun undeclared_in_problem problem =
blanchet@46698
  2597
  let
blanchet@49152
  2598
    fun do_sym (name as (s, _)) value =
blanchet@49152
  2599
      if is_tptp_user_symbol s then Symtab.default (s, (name, value)) else I
blanchet@49157
  2600
    fun do_class name = apfst (apfst (do_sym name ()))
blanchet@47220
  2601
    fun do_type (AType (name, tys)) =
blanchet@49156
  2602
        apfst (apsnd (do_sym name (length tys))) #> fold do_type tys
blanchet@46698
  2603
      | do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2
blanchet@49148
  2604
      | do_type (APi (_, ty)) = do_type ty
blanchet@49156
  2605
    fun do_term pred_sym (ATerm ((name, tys), tms)) =
blanchet@49152
  2606
        apsnd (do_sym name
blanchet@49156
  2607
                   (fn _ => default_type pred_sym (length tys) (length tms)))
blanchet@49147
  2608
        #> fold do_type tys #> fold (do_term false) tms
blanchet@48926
  2609
      | do_term _ (AAbs (((_, ty), tm), args)) =
blanchet@48926
  2610
        do_type ty #> do_term false tm #> fold (do_term false) args
blanchet@49156
  2611
    fun do_formula (ATyQuant (_, xs, phi)) =
blanchet@49157
  2612
        fold (do_type o fst) xs #> fold (fold do_class o snd) xs
blanchet@49156
  2613
        #> do_formula phi
blanchet@49148
  2614
      | do_formula (AQuant (_, xs, phi)) =
blanchet@46698
  2615
        fold do_type (map_filter snd xs) #> do_formula phi
blanchet@46698
  2616
      | do_formula (AConn (_, phis)) = fold do_formula phis
blanchet@46698
  2617
      | do_formula (AAtom tm) = do_term true tm
blanchet@49157
  2618
    fun do_line (Class_Decl (_, _, cls)) = fold do_class cls
blanchet@49157
  2619
      | do_line (Type_Decl _) = I
blanchet@49155
  2620
      | do_line (Sym_Decl (_, _, ty)) = do_type ty
blanchet@49157
  2621
      | do_line (Class_Memb (_, xs, ty, cl)) =
blanchet@49157
  2622
        fold (fold do_class o snd) xs #> do_type ty #> do_class cl
blanchet@49155
  2623
      | do_line (Formula (_, _, phi, _, _)) = do_formula phi
blanchet@49156
  2624
    val ((cls, tys), syms) = declared_in_atp_problem problem
blanchet@46698
  2625
  in
blanchet@49156
  2626
    ((Symtab.empty, Symtab.empty), Symtab.empty)
blanchet@49156
  2627
    |>> apfst (fold (fn (s, _) => Symtab.default (s, (("", ""), ()))) cls)
blanchet@49156
  2628
    |>> apsnd (fold (fn (s, _) => Symtab.default (s, (("", ""), 0))) tys)
blanchet@49152
  2629
    ||> fold (fn (s, _) => Symtab.default (s, (("", ""), K tvar_a_atype))) syms
blanchet@49155
  2630
    |> fold (fold do_line o snd) problem
blanchet@46698
  2631
  end
blanchet@46698
  2632
blanchet@49156
  2633
fun declare_undeclared_in_problem heading problem =
blanchet@46698
  2634
  let
blanchet@49156
  2635
    val ((cls, tys), syms) = undeclared_in_problem problem
blanchet@46698
  2636
    val decls =
blanchet@48021
  2637
      Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
blanchet@49157
  2638
                    | (s, (cls, ())) =>
blanchet@49157
  2639
                      cons (Class_Decl (class_decl_prefix ^ s, cls, [])))
blanchet@49156
  2640
                  cls [] @
blanchet@49156
  2641
      Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
blanchet@49157
  2642
                    | (s, (ty, ary)) =>
blanchet@49157
  2643
                      cons (Type_Decl (type_decl_prefix ^ s, ty, ary)))
blanchet@49156
  2644
                  tys [] @
blanchet@49152
  2645
      Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *)
blanchet@48021
  2646
                    | (s, (sym, ty)) =>
blanchet@49157
  2647
                      cons (Sym_Decl (sym_decl_prefix ^ s, sym, ty ())))
blanchet@49152
  2648
                  syms []
blanchet@49156
  2649
  in (heading, decls) :: problem end
blanchet@46698
  2650
blanchet@49156
  2651
fun is_poly_constr (Datatype.DtTFree _) = true
blanchet@49156
  2652
  | is_poly_constr (Datatype.DtType (_, Us)) = exists is_poly_constr Us
blanchet@49156
  2653
  | is_poly_constr _ = false
blanchet@46816
  2654
blanchet@46818
  2655
fun all_constrs_of_polymorphic_datatypes thy =
blanchet@49103
  2656
  Symtab.fold (snd #> #descr #> maps (#3 o snd)
blanchet@49103
  2657
               #> (fn cs => exists (exists is_poly_constr o snd) cs
blanchet@49103
  2658
                            ? append (map fst cs)))
blanchet@49103
  2659
               (Datatype.get_all thy) []
blanchet@46816
  2660
blanchet@47933
  2661
val app_op_and_predicator_threshold = 50
blanchet@44100
  2662
blanchet@48990
  2663
fun prepare_atp_problem ctxt format prem_role type_enc mode lam_trans
blanchet@48927
  2664
                        uncurried_aliases readable_names preproc hyp_ts concl_t
blanchet@48927
  2665
                        facts =
blanchet@38506
  2666
  let
blanchet@45645
  2667
    val thy = Proof_Context.theory_of ctxt
blanchet@45275
  2668
    val type_enc = type_enc |> adjust_type_enc format
blanchet@47217
  2669
    (* Forcing explicit applications is expensive for polymorphic encodings,
blanchet@47217
  2670
       because it takes only one existential variable ranging over "'a => 'b" to
blanchet@47217
  2671
       ruin everything. Hence we do it only if there are few facts (which is
blanchet@47217
  2672
       normally the case for "metis" and the minimizer). *)
blanchet@47220
  2673
    val app_op_level =
blanchet@49158
  2674
      if mode = Sledgehammer_Completish then
blanchet@48961
  2675
        Full_App_Op_And_Predicator
blanchet@48961
  2676
      else if length facts + length hyp_ts
blanchet@48961
  2677
              > app_op_and_predicator_threshold then
blanchet@49146
  2678
        if is_type_enc_polymorphic type_enc then Min_App_Op
blanchet@48926
  2679
        else Sufficient_App_Op
blanchet@47198
  2680
      else
blanchet@47933
  2681
        Sufficient_App_Op_And_Predicator
blanchet@48961
  2682
    val exporter = (mode = Exporter)
blanchet@49158
  2683
    val completish = (mode = Sledgehammer_Completish)
blanchet@46385
  2684
    val lam_trans =
blanchet@46391
  2685
      if lam_trans = keep_lamsN andalso
blanchet@46391
  2686
         not (is_type_enc_higher_order type_enc) then
blanchet@49111
  2687
        liftingN
blanchet@44959
  2688
      else
blanchet@46385
  2689
        lam_trans
blanchet@49157
  2690
    val (fact_names, classes, conjs, facts, subclass_pairs, tcon_clauses,
blanchet@46379
  2691
         lifted) =
blanchet@48990
  2692
      translate_formulas ctxt prem_role format type_enc lam_trans preproc hyp_ts
blanchet@46385
  2693
                         concl_t facts
blanchet@48947
  2694
    val (_, sym_tab0) =
blanchet@48947
  2695
      sym_table_for_facts ctxt type_enc app_op_level conjs facts
blanchet@49120
  2696
    val mono =
blanchet@49158
  2697
      conjs @ facts |> mononotonicity_info_for_facts ctxt type_enc completish
blanchet@49103
  2698
    val constrs = all_constrs_of_polymorphic_datatypes thy
blanchet@47228
  2699
    fun firstorderize in_helper =
blanchet@49103
  2700
      firstorderize_fact thy constrs type_enc sym_tab0
blanchet@49158
  2701
          (uncurried_aliases andalso not in_helper) completish
blanchet@47228
  2702
    val (conjs, facts) = (conjs, facts) |> pairself (map (firstorderize false))
blanchet@48947
  2703
    val (ho_stuff, sym_tab) =
blanchet@48947
  2704
      sym_table_for_facts ctxt type_enc Min_App_Op conjs facts
blanchet@48947
  2705
    val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) =
blanchet@49103
  2706
      uncurried_alias_lines_for_sym_table ctxt constrs mono type_enc
blanchet@48947
  2707
                                          uncurried_aliases sym_tab0 sym_tab
blanchet@48947
  2708
    val (_, sym_tab) =
blanchet@48947
  2709
      (ho_stuff, sym_tab)
blanchet@48947
  2710
      |> fold (add_iterm_syms_to_sym_table ctxt Min_App_Op false false)
blanchet@48947
  2711
              uncurried_alias_eq_tms
blanchet@43444
  2712
    val helpers =
blanchet@49158
  2713
      sym_tab |> helper_facts_for_sym_table ctxt format type_enc completish
blanchet@47228
  2714
              |> map (firstorderize true)
blanchet@45356
  2715
    val mono_Ts =
blanchet@46808
  2716
      helpers @ conjs @ facts |> monotonic_types_for_facts ctxt mono type_enc
blanchet@45459
  2717
    val class_decl_lines = decl_lines_for_classes type_enc classes
blanchet@43550
  2718
    val sym_decl_lines =
blanchet@47237
  2719
      (conjs, helpers @ facts, uncurried_alias_eq_tms)
blanchet@48638
  2720
      |> sym_decl_table_for_facts thy type_enc sym_tab
blanchet@49155
  2721
      |> lines_for_sym_decl_table ctxt mono type_enc mono_Ts
blanchet@47234
  2722
    val num_facts = length facts
blanchet@47234
  2723
    val fact_lines =
blanchet@49161
  2724
      map (line_for_fact ctxt fact_prefix ascii_of (not exporter)
blanchet@49102
  2725
               (not exporter) mono type_enc (rank_of_fact_num num_facts))
blanchet@47234
  2726
          (0 upto num_facts - 1 ~~ facts)
blanchet@49161
  2727
    val subclass_lines = maps (lines_for_subclass_pair type_enc) subclass_pairs
blanchet@49161
  2728
    val tcon_lines = map (line_for_tcon_clause type_enc) tcon_clauses
blanchet@43750
  2729
    val helper_lines =
blanchet@43797
  2730
      0 upto length helpers - 1 ~~ helpers
blanchet@49161
  2731
      |> map (line_for_fact ctxt helper_prefix I false true mono type_enc
blanchet@49161
  2732
                            (K default_rank))
blanchet@49161
  2733
    val free_type_lines = lines_for_free_types type_enc (facts @ conjs)
blanchet@49161
  2734
    val conj_lines = map (line_for_conjecture ctxt mono type_enc) conjs
blanchet@49020
  2735
    (* Reordering these might confuse the proof reconstruction code. *)
blanchet@38506
  2736
    val problem =
blanchet@45459
  2737
      [(explicit_declsN, class_decl_lines @ sym_decl_lines),
blanchet@47237
  2738
       (uncurried_alias_eqsN, uncurried_alias_eq_lines),
blanchet@47234
  2739
       (factsN, fact_lines),
blanchet@49157
  2740
       (subclassesN, subclass_lines),
blanchet@49157
  2741
       (tconsN, tcon_lines),
blanchet@43750
  2742
       (helpersN, helper_lines),
blanchet@49156
  2743
       (free_typesN, free_type_lines),
blanchet@49161
  2744
       (conjsN, conj_lines)]
blanchet@43414
  2745
    val problem =
blanchet@49151
  2746
      problem |> (case format of
blanchet@49151
  2747
                    CNF => ensure_cnf_problem
blanchet@49151
  2748
                  | CNF_UEQ => filter_cnf_ueq_problem
blanchet@49151
  2749
                  | FOF => I
blanchet@49151
  2750
                  | TFF (_, TPTP_Implicit) => I
blanchet@49151
  2751
                  | THF (_, TPTP_Implicit, _, _) => I
blanchet@49156
  2752
                  | _ => declare_undeclared_in_problem implicit_declsN)
blanchet@46810
  2753
    val (problem, pool) = problem |> nice_atp_problem readable_names format
blanchet@45643
  2754
    fun add_sym_ary (s, {min_ary, ...} : sym_info) =
blanchet@45644
  2755
      min_ary > 0 ? Symtab.insert (op =) (s, min_ary)
blanchet@38506
  2756
  in
blanchet@49111
  2757
    (problem, pool |> Option.map snd |> the_default Symtab.empty,
blanchet@49111
  2758
     fact_names |> Vector.fromList, lifted,
blanchet@45643
  2759
     Symtab.empty |> Symtab.fold add_sym_ary sym_tab)
blanchet@38506
  2760
  end
blanchet@38506
  2761
blanchet@41561
  2762
(* FUDGE *)
blanchet@41561
  2763
val conj_weight = 0.0
blanchet@42641
  2764
val hyp_weight = 0.1
blanchet@42641
  2765
val fact_min_weight = 0.2
blanchet@41561
  2766
val fact_max_weight = 1.0
blanchet@43479
  2767
val type_info_default_weight = 0.8
blanchet@41561
  2768
blanchet@47270
  2769
(* Weights are from 0.0 (most important) to 1.0 (least important). *)
blanchet@47901
  2770
fun atp_problem_selection_weights problem =
blanchet@41561
  2771
  let
blanchet@49147
  2772
    fun add_term_weights weight (ATerm ((s, _), tms)) =
blanchet@47270
  2773
        is_tptp_user_symbol s ? Symtab.default (s, weight)
blanchet@47270
  2774
        #> fold (add_term_weights weight) tms
blanchet@48926
  2775
      | add_term_weights weight (AAbs ((_, tm), args)) =
blanchet@48926
  2776
        add_term_weights weight tm #> fold (add_term_weights weight) args
blanchet@47270
  2777
    fun add_line_weights weight (Formula (_, _, phi, _, _)) =
blanchet@47270
  2778
        formula_fold NONE (K (add_term_weights weight)) phi
blanchet@47270
  2779
      | add_line_weights _ _ = I
blanchet@47270
  2780
    fun add_conjectures_weights [] = I
blanchet@47270
  2781
      | add_conjectures_weights conjs =
blanchet@47270
  2782
        let val (hyps, conj) = split_last conjs in
blanchet@47270
  2783
          add_line_weights conj_weight conj
blanchet@47270
  2784
          #> fold (add_line_weights hyp_weight) hyps
blanchet@47270
  2785
        end
blanchet@47270
  2786
    fun add_facts_weights facts =
blanchet@47270
  2787
      let
blanchet@47270
  2788
        val num_facts = length facts
blanchet@47270
  2789
        fun weight_of j =
blanchet@47270
  2790
          fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j
blanchet@47270
  2791
                            / Real.fromInt num_facts
blanchet@47270
  2792
      in
blanchet@47270
  2793
        map weight_of (0 upto num_facts - 1) ~~ facts
blanchet@47270
  2794
        |> fold (uncurry add_line_weights)
blanchet@47270
  2795
      end
blanchet@47270
  2796
    val get = these o AList.lookup (op =) problem
blanchet@41561
  2797
  in
blanchet@43479
  2798
    Symtab.empty
blanchet@43479
  2799
    |> add_conjectures_weights (get free_typesN @ get conjsN)
blanchet@43479
  2800
    |> add_facts_weights (get factsN)
blanchet@47270
  2801
    |> fold (fold (add_line_weights type_info_default_weight) o get)
blanchet@49157
  2802
            [explicit_declsN, subclassesN, tconsN]
blanchet@43479
  2803
    |> Symtab.dest
blanchet@43479
  2804
    |> sort (prod_ord Real.compare string_ord o pairself swap)
blanchet@43479
  2805
  end
blanchet@41561
  2806
blanchet@47933
  2807
(* Ugly hack: may make innocent victims (collateral damage) *)
blanchet@47933
  2808
fun may_be_app s args = String.isPrefix app_op_name s andalso length args = 2
blanchet@47933
  2809
fun may_be_predicator s =
blanchet@47933
  2810
  member (op =) [predicator_name, prefixed_predicator_name] s
blanchet@47933
  2811
blanchet@49147
  2812
fun strip_predicator (tm as ATerm ((s, _), [tm'])) =
blanchet@47933
  2813
    if may_be_predicator s then tm' else tm
blanchet@47933
  2814
  | strip_predicator tm = tm
blanchet@47933
  2815
blanchet@49147
  2816
fun make_head_roll (ATerm ((s, _), tms)) =
blanchet@47933
  2817
    if may_be_app s tms then make_head_roll (hd tms) ||> append (tl tms)
blanchet@47933
  2818
    else (s, tms)
blanchet@47903
  2819
  | make_head_roll _ = ("", [])
blanchet@47271
  2820
blanchet@49148
  2821
fun strip_up_to_predicator (ATyQuant (_, _, phi)) = strip_up_to_predicator phi
blanchet@49148
  2822
  | strip_up_to_predicator (AQuant (_, _, phi)) = strip_up_to_predicator phi
blanchet@47933
  2823
  | strip_up_to_predicator (AConn (_, phis)) = maps strip_up_to_predicator phis
blanchet@47933
  2824
  | strip_up_to_predicator (AAtom tm) = [strip_predicator tm]
blanchet@47933
  2825
blanchet@49148
  2826
fun strip_ahorn_etc (ATyQuant (_, _, phi)) = strip_ahorn_etc phi
blanchet@49148
  2827
  | strip_ahorn_etc (AQuant (_, _, phi)) = strip_ahorn_etc phi
blanchet@47933
  2828
  | strip_ahorn_etc (AConn (AImplies, [phi1, phi2])) =
blanchet@47933
  2829
    strip_ahorn_etc phi2 |>> append (strip_up_to_predicator phi1)
blanchet@47933
  2830
  | strip_ahorn_etc phi = ([], hd (strip_up_to_predicator phi))
blanchet@47933
  2831
blanchet@49148
  2832
fun strip_iff_etc (ATyQuant (_, _, phi)) = strip_iff_etc phi
blanchet@49148
  2833
  | strip_iff_etc (AQuant (_, _, phi)) = strip_iff_etc phi
blanchet@47933
  2834
  | strip_iff_etc (AConn (AIff, [phi1, phi2])) =
blanchet@47933
  2835
    pairself strip_up_to_predicator (phi1, phi2)
blanchet@47933
  2836
  | strip_iff_etc _ = ([], [])
blanchet@47933
  2837
blanchet@47901
  2838
val max_term_order_weight = 2147483647
blanchet@47278
  2839
blanchet@47909
  2840
fun atp_problem_term_order_info problem =
blanchet@47270
  2841
  let
blanchet@47909
  2842
    fun add_edge s s' =
blanchet@47909
  2843
      Graph.default_node (s, ())
blanchet@47909
  2844
      #> Graph.default_node (s', ())
blanchet@47909
  2845
      #> Graph.add_edge_acyclic (s, s')
blanchet@49147
  2846
    fun add_term_deps head (ATerm ((s, _), args)) =
blanchet@47933
  2847
        if is_tptp_user_symbol head then
blanchet@47933
  2848
          (if is_tptp_user_symbol s then perhaps (try (add_edge s head)) else I)
blanchet@47933
  2849
          #> fold (add_term_deps head) args
blanchet@47933
  2850
        else
blanchet@47933
  2851
          I
blanchet@48926
  2852
      | add_term_deps head (AAbs ((_, tm), args)) =
blanchet@48926
  2853
        add_term_deps head tm #> fold (add_term_deps head) args
blanchet@47933
  2854
    fun add_intro_deps pred (Formula (_, role, phi, _, info)) =
blanchet@47933
  2855
        if pred (role, info) then
blanchet@47933
  2856
          let val (hyps, concl) = strip_ahorn_etc phi in
blanchet@47933
  2857
            case make_head_roll concl of
blanchet@47933
  2858
              (head, args as _ :: _) => fold (add_term_deps head) (hyps @ args)
blanchet@47933
  2859
            | _ => I
blanchet@47933
  2860
          end
blanchet@47933
  2861
        else
blanchet@47933
  2862
          I
blanchet@47933
  2863
      | add_intro_deps _ _ = I
blanchet@49147
  2864
    fun add_atom_eq_deps (SOME true) (ATerm ((s, _), [lhs as _, rhs])) =
blanchet@47271
  2865
        if is_tptp_equal s then
blanchet@47909
  2866
          case make_head_roll lhs of
blanchet@47933
  2867
            (head, args as _ :: _) => fold (add_term_deps head) (rhs :: args)
blanchet@47909
  2868
          | _ => I
blanchet@47271
  2869
        else
blanchet@47271
  2870
          I
blanchet@47933
  2871
      | add_atom_eq_deps _ _ = I
blanchet@47933
  2872
    fun add_eq_deps pred (Formula (_, role, phi, _, info)) =
blanchet@47910
  2873
        if pred (role, info) then
blanchet@47933
  2874
          case strip_iff_etc phi of
blanchet@47933
  2875
            ([lhs], rhs) =>
blanchet@47933
  2876
            (case make_head_roll lhs of
blanchet@47933
  2877
               (head, args as _ :: _) => fold (add_term_deps head) (rhs @ args)
blanchet@47933
  2878
             | _ => I)
blanchet@47933
  2879
          | _ => formula_fold (SOME (role <> Conjecture)) add_atom_eq_deps phi
blanchet@47270
  2880
        else
blanchet@47270
  2881
          I
blanchet@47933
  2882
      | add_eq_deps _ _ = I
blanchet@48986
  2883
    fun has_status status (_, info) = extract_isabelle_status info = SOME status
blanchet@47910
  2884
    fun is_conj (role, _) = (role = Conjecture orelse role = Hypothesis)
blanchet@47270
  2885
    val graph =
blanchet@47910
  2886
      Graph.empty
blanchet@49453
  2887
      |> fold (fold (add_eq_deps (has_status non_rec_defN)) o snd) problem
blanchet@49453
  2888
      |> fold (fold (add_eq_deps (has_status rec_defN orf has_status simpN
blanchet@49453
  2889
                                  orf is_conj)) o snd) problem
blanchet@48019
  2890
      |> fold (fold (add_intro_deps (has_status inductiveN)) o snd) problem
blanchet@47933
  2891
      |> fold (fold (add_intro_deps (has_status introN)) o snd) problem
blanchet@47901
  2892
    fun next_weight w = if w + w <= max_term_order_weight then w + w else w + 1
blanchet@47270
  2893
    fun add_weights _ [] = I
blanchet@47270
  2894
      | add_weights weight syms =
blanchet@47270
  2895
        fold (AList.update (op =) o rpair weight) syms
blanchet@47270
  2896
        #> add_weights (next_weight weight)
blanchet@47270
  2897
               (fold (union (op =) o Graph.immediate_succs graph) syms [])
blanchet@47270
  2898
  in
blanchet@47909
  2899
    (* Sorting is not just for aesthetics: It specifies the precedence order
blanchet@47909
  2900
       for the term ordering (KBO or LPO), from smaller to larger values. *)
blanchet@47274
  2901
    [] |> add_weights 1 (Graph.minimals graph) |> sort (int_ord o pairself snd)
blanchet@47270
  2902
  end
blanchet@47270
  2903
blanchet@38506
  2904
end;