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