src/HOL/Tools/Metis/metis_generate.ML
author wenzelm
Thu, 03 Oct 2024 13:01:31 +0200
changeset 81108 92768949a923
parent 74904 cab76af373e7
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45569
diff changeset
     1
(*  Title:      HOL/Tools/Metis/metis_generate.ML
38027
505657ddb047 standardize "Author" tags
blanchet
parents: 38024
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
     3
    Author:     Kong W. Susanto, Cambridge University Computer Laboratory
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
     4
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
36393
be73a2b2443b support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents: 36378
diff changeset
     5
    Author:     Jasmin Blanchette, TU Muenchen
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     6
39494
bf7dd4902321 rename "Metis_Clauses" to "Metis_Translate" for consistency with "Sledgehammer_Translate"
blanchet
parents: 39355
diff changeset
     7
Translation of HOL to FOL for Metis.
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     8
*)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     9
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45569
diff changeset
    10
signature METIS_GENERATE =
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    11
sig
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45569
diff changeset
    12
  type type_enc = ATP_Problem_Generate.type_enc
44411
e3629929b171 change Metis's default settings if type information axioms are generated
blanchet
parents: 44409
diff changeset
    13
43159
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
    14
  datatype isa_thm =
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
    15
    Isa_Reflexive_or_Trivial |
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45510
diff changeset
    16
    Isa_Lambda_Lifted |
43159
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
    17
    Isa_Raw of thm
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
    18
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    19
  val metis_equal : string
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    20
  val metis_predicator : string
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    21
  val metis_app_op : string
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    22
  val metis_systematic_type_tag : string
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    23
  val metis_ad_hoc_type_tag : string
42098
f978caf60bbe more robust handling of variables in new Skolemizer
blanchet
parents: 41491
diff changeset
    24
  val metis_generated_var_prefix : string
43231
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    25
  val trace : bool Config.T
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    26
  val verbose : bool Config.T
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    27
  val trace_msg : Proof.context -> (unit -> string) -> unit
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    28
  val verbose_warning : Proof.context -> string -> unit
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    29
  val metis_name_table : ((string * int) * ((type_enc -> string) * bool)) list
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    30
  val reveal_old_skolem_terms : (string * term) list -> term -> term
45569
eb30a5490543 wrap lambdas earlier, to get more control over beta/eta
blanchet
parents: 45565
diff changeset
    31
  val reveal_lam_lifted : (string * term) list -> term -> term
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
    32
  val generate_metis_problem : Proof.context -> type_enc -> string -> thm list -> thm list ->
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
    33
    int Symtab.table * (Metis_Thm.thm * isa_thm) list * (unit -> (string * int) list)
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
    34
    * ((string * term) list * (string * term) list)
24310
af4af9993922 proper signature;
wenzelm
parents: 24183
diff changeset
    35
end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    36
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45569
diff changeset
    37
structure Metis_Generate : METIS_GENERATE =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    38
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    39
43092
93ec303e1917 more work on new metis that exploits the powerful new type encodings
blanchet
parents: 43091
diff changeset
    40
open ATP_Problem
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45569
diff changeset
    41
open ATP_Problem_Generate
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    42
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    43
val metis_equal = "="
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    44
val metis_predicator = "{}"
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    45
val metis_app_op = Metis_Name.toString Metis_Term.appName
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    46
val metis_systematic_type_tag =
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    47
  Metis_Name.toString Metis_Term.hasTypeFunctionName
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    48
val metis_ad_hoc_type_tag = "**"
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43003
diff changeset
    49
val metis_generated_var_prefix = "_"
17150
ce2a1aeb42aa DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents: 16976
diff changeset
    50
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61860
diff changeset
    51
val trace = Attrib.setup_config_bool \<^binding>\<open>metis_trace\<close> (K false)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61860
diff changeset
    52
val verbose = Attrib.setup_config_bool \<^binding>\<open>metis_verbose\<close> (K true)
43231
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    53
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    54
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    55
fun verbose_warning ctxt msg =
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    56
  if Config.get ctxt verbose then warning ("Metis: " ^ msg) else ()
69f22ac07395 nicely thread monomorphism verbosity in Metis and Sledgehammer
blanchet
parents: 43224
diff changeset
    57
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    58
val metis_name_table =
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    59
  [((tptp_equal, 2), (K metis_equal, false)),
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    60
   ((tptp_old_equal, 2), (K metis_equal, false)),
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    61
   ((prefixed_predicator_name, 1), (K metis_predicator, false)),
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    62
   ((prefixed_app_op_name, 2), (K metis_app_op, false)),
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
    63
   ((prefixed_type_tag_name, 2),
44782
blanchet
parents: 44768
diff changeset
    64
    (fn type_enc =>
blanchet
parents: 44768
diff changeset
    65
        if level_of_type_enc type_enc = All_Types then metis_systematic_type_tag
blanchet
parents: 44768
diff changeset
    66
        else metis_ad_hoc_type_tag, true))]
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
    67
39896
13b3a2ba9ea7 encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
blanchet
parents: 39890
diff changeset
    68
fun old_skolem_const_name i j num_T_args =
59877
a04ea4709c8d more standard Long_Name operations;
wenzelm
parents: 59582
diff changeset
    69
  Long_Name.implode (old_skolem_const_prefix :: map string_of_int [i, j, num_T_args])
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    70
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    71
fun conceal_old_skolem_terms i old_skolems t =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61860
diff changeset
    72
  if exists_Const (curry (op =) \<^const_name>\<open>Meson.skolem\<close> o fst) t then
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    73
    let
74347
f984d30cd0c3 clarified antiquotations;
wenzelm
parents: 73932
diff changeset
    74
      fun aux old_skolems (t as \<^Const_>\<open>Meson.skolem T for _\<close>) =
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    75
          let
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    76
            val (old_skolems, s) =
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    77
              if i = ~1 then
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61860
diff changeset
    78
                (old_skolems, \<^const_name>\<open>undefined\<close>)
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    79
              else
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    80
                (case AList.find (op aconv) old_skolems t of
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    81
                  s :: _ => (old_skolems, s)
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    82
                | [] =>
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    83
                  let
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    84
                    val s =
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    85
                      old_skolem_const_name i (length old_skolems) (length (Term.add_tvarsT T []))
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
    86
                  in ((s, t) :: old_skolems, s) end)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    87
          in (old_skolems, Const (s, T)) end
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    88
        | aux old_skolems (t1 $ t2) =
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    89
          let
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    90
            val (old_skolems, t1) = aux old_skolems t1
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    91
            val (old_skolems, t2) = aux old_skolems t2
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    92
          in (old_skolems, t1 $ t2) end
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    93
        | aux old_skolems (Abs (s, T, t')) =
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    94
          let val (old_skolems, t') = aux old_skolems t' in
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    95
            (old_skolems, Abs (s, T, t'))
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    96
          end
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    97
        | aux old_skolems t = (old_skolems, t)
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
    98
    in aux old_skolems t end
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
    99
  else
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
   100
    (old_skolems, t)
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
   101
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39720
diff changeset
   102
fun reveal_old_skolem_terms old_skolems =
37632
df12f798df99 move function
blanchet
parents: 37625
diff changeset
   103
  map_aterms (fn t as Const (s, _) =>
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   104
      if String.isPrefix old_skolem_const_prefix s then
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   105
        AList.lookup (op =) old_skolems s |> the
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   106
        |> map_types (map_type_tvar (K dummyT))
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   107
      else
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   108
        t
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   109
    | t => t)
37632
df12f798df99 move function
blanchet
parents: 37625
diff changeset
   110
52150
41c885784e04 handle lambda-lifted problems in Isar construction code
blanchet
parents: 52031
diff changeset
   111
fun reveal_lam_lifted lifted =
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45510
diff changeset
   112
  map_aterms (fn t as Const (s, _) =>
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   113
      if String.isPrefix lam_lifted_prefix s then
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   114
        (case AList.lookup (op =) lifted s of
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   115
          SOME t =>
74347
f984d30cd0c3 clarified antiquotations;
wenzelm
parents: 73932
diff changeset
   116
            \<^Const>\<open>Metis.lambda dummyT\<close> $
f984d30cd0c3 clarified antiquotations;
wenzelm
parents: 73932
diff changeset
   117
              map_types (map_type_tvar (K dummyT)) (reveal_lam_lifted lifted t)
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   118
        | NONE => t)
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   119
      else
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   120
        t
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   121
    | t => t)
45508
b216dc1b3630 started implementing lambda-lifting in Metis
blanchet
parents: 45043
diff changeset
   122
37577
5379f41a1322 merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet
parents: 37575
diff changeset
   123
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   124
(* ------------------------------------------------------------------------- *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   125
(* Logic maps manage the interface between HOL and first-order logic.        *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   126
(* ------------------------------------------------------------------------- *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   127
43159
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
   128
datatype isa_thm =
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
   129
  Isa_Reflexive_or_Trivial |
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45510
diff changeset
   130
  Isa_Lambda_Lifted |
43159
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
   131
  Isa_Raw of thm
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
   132
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
   133
val proxy_defs = map (fst o snd o snd) proxy_table
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53479
diff changeset
   134
fun prepare_helper ctxt =
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59877
diff changeset
   135
  Meson.make_meta_clause ctxt #> rewrite_rule ctxt (map safe_mk_meta_eq proxy_defs)
43159
29b55f292e0b added support for helpers in new Metis, so far only for polymorphic type encodings
blanchet
parents: 43132
diff changeset
   136
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   137
fun metis_term_of_atp type_enc (ATerm ((s, []), tms)) =
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   138
  if is_tptp_variable s then
43268
c0eaa8b9bff5 removed yet another hack in "make_metis" script -- respect opacity of "Metis_Name.name"
blanchet
parents: 43259
diff changeset
   139
    Metis_Term.Var (Metis_Name.fromString s)
43094
269300fb83d0 more work on new Metis
blanchet
parents: 43093
diff changeset
   140
  else
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   141
    (case AList.lookup (op =) metis_name_table (s, length tms) of
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   142
       SOME (f, swap) => (f type_enc, swap)
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   143
     | NONE => (s, false))
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   144
    |> (fn (s, swap) =>
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   145
           Metis_Term.Fn (Metis_Name.fromString s,
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   146
                          tms |> map (metis_term_of_atp type_enc)
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   147
                              |> swap ? rev))
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   148
fun metis_atom_of_atp type_enc (AAtom tm) =
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   149
    (case metis_term_of_atp type_enc tm of
43104
81d1b15aa0ae use ":" for type information (looks good in Metis's output) and handle it in new path finder
blanchet
parents: 43103
diff changeset
   150
       Metis_Term.Fn x => x
81d1b15aa0ae use ":" for type information (looks good in Metis's output) and handle it in new path finder
blanchet
parents: 43103
diff changeset
   151
     | _ => raise Fail "non CNF -- expected function")
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   152
  | metis_atom_of_atp _ _ = raise Fail "not CNF -- expected atom"
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   153
fun metis_literal_of_atp type_enc (AConn (ANot, [phi])) =
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   154
    (false, metis_atom_of_atp type_enc phi)
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   155
  | metis_literal_of_atp type_enc phi = (true, metis_atom_of_atp type_enc phi)
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   156
fun metis_literals_of_atp type_enc (AConn (AOr, phis)) =
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   157
    maps (metis_literals_of_atp type_enc) phis
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   158
  | metis_literals_of_atp type_enc phi = [metis_literal_of_atp type_enc phi]
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53479
diff changeset
   159
fun metis_axiom_of_atp ctxt type_enc clauses (Formula ((ident, _), _, phi, _, _)) =
43173
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   160
    let
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   161
      fun some isa =
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   162
        SOME (phi |> metis_literals_of_atp type_enc
44492
a330c0608da8 avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents: 44411
diff changeset
   163
                  |> Metis_LiteralSet.fromList
43173
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   164
                  |> Metis_Thm.axiom, isa)
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   165
    in
47046
62ca06cc5a99 remove two options that were found to play hardly any role
blanchet
parents: 47039
diff changeset
   166
      if String.isPrefix tags_sym_formula_prefix ident then
43173
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   167
        Isa_Reflexive_or_Trivial |> some
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   168
      else if String.isPrefix conjecture_prefix ident then
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   169
        NONE
43173
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   170
      else if String.isPrefix helper_prefix ident then
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   171
        (case (String.isSuffix typed_helper_suffix ident, space_explode "_" ident) of
43194
ef3ff8856245 fixed type helper indices in new Metis
blanchet
parents: 43193
diff changeset
   172
          (needs_fairly_sound, _ :: const :: j :: _) =>
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 69593
diff changeset
   173
          nth (AList.lookup (op =) (helper_table true) (const, needs_fairly_sound) |> the)
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   174
            (the (Int.fromString j) - 1)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53479
diff changeset
   175
          |> snd |> prepare_helper ctxt
43194
ef3ff8856245 fixed type helper indices in new Metis
blanchet
parents: 43193
diff changeset
   176
          |> Isa_Raw |> some
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   177
        | _ => raise Fail ("malformed helper identifier " ^ quote ident))
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   178
      else
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   179
        (case try (unprefix fact_prefix) ident of
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   180
          SOME s =>
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   181
          let val s = s |> space_explode "_" |> tl |> space_implode "_" in
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   182
            (case Int.fromString s of
60358
aebfbcab1eb8 clarified context;
wenzelm
parents: 59877
diff changeset
   183
              SOME j => Meson.make_meta_clause ctxt (snd (nth clauses j)) |> Isa_Raw |> some
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   184
            | NONE =>
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   185
              if String.isPrefix lam_fact_prefix (unascii_of s) then Isa_Lambda_Lifted |> some
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   186
              else raise Fail ("malformed fact identifier " ^ quote ident))
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   187
          end
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   188
        | NONE => some (Isa_Raw TrueI))
43173
b98daa96d043 don't pass "~ " to new Metis
blanchet
parents: 43160
diff changeset
   189
    end
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53479
diff changeset
   190
  | metis_axiom_of_atp _ _ _ _ = raise Fail "not CNF -- expected formula"
43092
93ec303e1917 more work on new metis that exploits the powerful new type encodings
blanchet
parents: 43091
diff changeset
   191
74347
f984d30cd0c3 clarified antiquotations;
wenzelm
parents: 73932
diff changeset
   192
fun eliminate_lam_wrappers \<^Const_>\<open>Metis.lambda _ for t\<close> = eliminate_lam_wrappers t
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   193
  | eliminate_lam_wrappers (t $ u) = eliminate_lam_wrappers t $ eliminate_lam_wrappers u
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   194
  | eliminate_lam_wrappers (Abs (s, T, t)) = Abs (s, T, eliminate_lam_wrappers t)
45569
eb30a5490543 wrap lambdas earlier, to get more control over beta/eta
blanchet
parents: 45565
diff changeset
   195
  | eliminate_lam_wrappers t = t
eb30a5490543 wrap lambdas earlier, to get more control over beta/eta
blanchet
parents: 45565
diff changeset
   196
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   197
(* Function to generate metis clauses, including comb and type clauses *)
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   198
fun generate_metis_problem ctxt type_enc lam_trans conj_clauses fact_clauses =
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43211
diff changeset
   199
  let
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   200
    val (conj_clauses, fact_clauses) =
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   201
      if is_type_enc_polymorphic type_enc then
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   202
        (conj_clauses, fact_clauses)
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   203
      else
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   204
        conj_clauses @ fact_clauses
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   205
        |> map (pair 0)
53479
f7d8224641de move metis to new monomorphizer
blanchet
parents: 52150
diff changeset
   206
        |> Monomorph.monomorph atp_schematic_consts_of ctxt
f7d8224641de move metis to new monomorphizer
blanchet
parents: 52150
diff changeset
   207
        |> chop (length conj_clauses)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 57408
diff changeset
   208
        |> apply2 (maps (map (zero_var_indexes o snd)))
47039
1b36a05a070d added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
blanchet
parents: 46409
diff changeset
   209
    (* Pretend every clause is a "simp" rule, to guide the term ordering. *)
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43211
diff changeset
   210
    val clauses =
74904
cab76af373e7 tuned metis to use map_index
desharna
parents: 74347
diff changeset
   211
      map_index (apfst (fn j => (Int.toString j, (Local, Simp)))) (conj_clauses @ fact_clauses)
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43211
diff changeset
   212
    val (old_skolems, props) =
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   213
      fold_rev (fn (name, th) => fn (old_skolems, props) =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59058
diff changeset
   214
           th |> Thm.prop_of |> Logic.strip_imp_concl
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   215
              |> conceal_old_skolem_terms (length clauses) old_skolems
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 69593
diff changeset
   216
              ||> lam_trans = liftingN ? eliminate_lam_wrappers
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   217
              ||> (fn prop => (name, prop) :: props))
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   218
         clauses ([], [])
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   219
    (*
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   220
    val _ =
45042
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44782
diff changeset
   221
      tracing ("PROPS:\n" ^
89341b897412 better type reconstruction -- prevents ill-instantiations in proof replay
blanchet
parents: 44782
diff changeset
   222
               cat_lines (map (Syntax.string_of_term ctxt o snd) props))
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   223
    *)
46365
547d1a1dcaf6 rename lambda translation schemes
blanchet
parents: 46340
diff changeset
   224
    val lam_trans = if lam_trans = combsN then no_lamsN else lam_trans
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   225
    val (atp_problem, _, lifted, sym_tab) =
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 60358
diff changeset
   226
      generate_atp_problem ctxt true CNF Hypothesis type_enc Metis lam_trans false false false []
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61860
diff changeset
   227
        \<^prop>\<open>False\<close> props
45510
96696c360b3e disable debugging output
blanchet
parents: 45508
diff changeset
   228
    (*
43295
30aaab778416 pass Metis facts and negated conjecture as facts, with (almost) correctly set localities, so that the correct encoding is used for nonmonotonic occurrences of infinite types
blanchet
parents: 43268
diff changeset
   229
    val _ = tracing ("ATP PROBLEM: " ^
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51575
diff changeset
   230
                     cat_lines (lines_of_atp_problem CNF atp_problem))
45510
96696c360b3e disable debugging output
blanchet
parents: 45508
diff changeset
   231
    *)
45508
b216dc1b3630 started implementing lambda-lifting in Metis
blanchet
parents: 45043
diff changeset
   232
    (* "rev" is for compatibility with existing proof scripts. *)
57263
2b6a96cc64c9 simplified code
blanchet
parents: 54742
diff changeset
   233
    val axioms = atp_problem
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53479
diff changeset
   234
      |> maps (map_filter (metis_axiom_of_atp ctxt type_enc clauses) o snd) |> rev
47039
1b36a05a070d added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
blanchet
parents: 46409
diff changeset
   235
    fun ord_info () = atp_problem_term_order_info atp_problem
57408
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   236
  in
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   237
    (sym_tab, axioms, ord_info, (lifted, old_skolems))
39b3562e9edc tuned whitespace and parentheses
blanchet
parents: 57263
diff changeset
   238
  end
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39494
diff changeset
   239
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   240
end;