src/HOL/Tools/Metis/metis_translate.ML
author blanchet
Mon, 22 Aug 2011 15:02:45 +0200
changeset 45268 b529d9501d64
parent 45256 06375952f1fa
child 45270 e3629929b171
permissions -rw-r--r--
prefer the lighter, slightly unsound monotonicity-based encodings for Metis
blanchet@40139
     1
(*  Title:      HOL/Tools/Metis/metis_translate.ML
blanchet@38261
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
blanchet@39737
     3
    Author:     Kong W. Susanto, Cambridge University Computer Laboratory
blanchet@39737
     4
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
blanchet@36393
     5
    Author:     Jasmin Blanchette, TU Muenchen
paulson@15347
     6
blanchet@39734
     7
Translation of HOL to FOL for Metis.
paulson@15347
     8
*)
paulson@15347
     9
blanchet@39734
    10
signature METIS_TRANSLATE =
wenzelm@24310
    11
sig
blanchet@44000
    12
  datatype isa_thm =
blanchet@44000
    13
    Isa_Reflexive_or_Trivial |
blanchet@44000
    14
    Isa_Raw of thm
blanchet@44000
    15
blanchet@43935
    16
  val metis_equal : string
blanchet@43935
    17
  val metis_predicator : string
blanchet@43935
    18
  val metis_app_op : string
blanchet@43947
    19
  val metis_type_tag : string
blanchet@42962
    20
  val metis_generated_var_prefix : string
blanchet@44072
    21
  val trace : bool Config.T
blanchet@44072
    22
  val verbose : bool Config.T
blanchet@44072
    23
  val trace_msg : Proof.context -> (unit -> string) -> unit
blanchet@44072
    24
  val verbose_warning : Proof.context -> string -> unit
blanchet@43945
    25
  val metis_name_table : ((string * int) * (string * bool)) list
blanchet@40067
    26
  val reveal_old_skolem_terms : (string * term) list -> term -> term
blanchet@40398
    27
  val prepare_metis_problem :
blanchet@44053
    28
    Proof.context -> string -> thm list -> thm list
blanchet@44053
    29
    -> int Symtab.table * (Metis_Thm.thm * isa_thm) list * (string * term) list
wenzelm@24310
    30
end
paulson@15347
    31
blanchet@39734
    32
structure Metis_Translate : METIS_TRANSLATE =
paulson@15347
    33
struct
paulson@15347
    34
blanchet@43933
    35
open ATP_Problem
blanchet@43926
    36
open ATP_Translate
blanchet@43926
    37
blanchet@43935
    38
val metis_equal = "="
blanchet@43935
    39
val metis_predicator = "{}"
blanchet@43935
    40
val metis_app_op = "."
blanchet@43945
    41
val metis_type_tag = ":"
blanchet@42962
    42
val metis_generated_var_prefix = "_"
blanchet@42962
    43
blanchet@44072
    44
val trace = Attrib.setup_config_bool @{binding metis_trace} (K false)
blanchet@44072
    45
val verbose = Attrib.setup_config_bool @{binding metis_verbose} (K true)
blanchet@44072
    46
blanchet@44072
    47
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@44072
    48
fun verbose_warning ctxt msg =
blanchet@44072
    49
  if Config.get ctxt verbose then warning ("Metis: " ^ msg) else ()
blanchet@44072
    50
blanchet@43935
    51
val metis_name_table =
blanchet@43945
    52
  [((tptp_equal, 2), (metis_equal, false)),
blanchet@43945
    53
   ((tptp_old_equal, 2), (metis_equal, false)),
blanchet@44015
    54
   ((prefixed_predicator_name, 1), (metis_predicator, false)),
blanchet@43971
    55
   ((prefixed_app_op_name, 2), (metis_app_op, false)),
blanchet@43971
    56
   ((prefixed_type_tag_name, 2), (metis_type_tag, true))]
blanchet@43935
    57
blanchet@40077
    58
fun old_skolem_const_name i j num_T_args =
blanchet@40077
    59
  old_skolem_const_prefix ^ Long_Name.separator ^
wenzelm@41739
    60
  (space_implode Long_Name.separator (map string_of_int [i, j, num_T_args]))
blanchet@37577
    61
blanchet@40067
    62
fun conceal_old_skolem_terms i old_skolems t =
blanchet@40134
    63
  if exists_Const (curry (op =) @{const_name Meson.skolem} o fst) t then
blanchet@37577
    64
    let
blanchet@40067
    65
      fun aux old_skolems
blanchet@40134
    66
             (t as (Const (@{const_name Meson.skolem}, Type (_, [_, T])) $ _)) =
blanchet@37577
    67
          let
blanchet@40067
    68
            val (old_skolems, s) =
blanchet@37577
    69
              if i = ~1 then
blanchet@40067
    70
                (old_skolems, @{const_name undefined})
blanchet@40067
    71
              else case AList.find (op aconv) old_skolems t of
blanchet@40067
    72
                s :: _ => (old_skolems, s)
blanchet@37577
    73
              | [] =>
blanchet@37577
    74
                let
blanchet@40077
    75
                  val s = old_skolem_const_name i (length old_skolems)
blanchet@40077
    76
                                                (length (Term.add_tvarsT T []))
blanchet@40067
    77
                in ((s, t) :: old_skolems, s) end
blanchet@40067
    78
          in (old_skolems, Const (s, T)) end
blanchet@40067
    79
        | aux old_skolems (t1 $ t2) =
blanchet@37577
    80
          let
blanchet@40067
    81
            val (old_skolems, t1) = aux old_skolems t1
blanchet@40067
    82
            val (old_skolems, t2) = aux old_skolems t2
blanchet@40067
    83
          in (old_skolems, t1 $ t2) end
blanchet@40067
    84
        | aux old_skolems (Abs (s, T, t')) =
blanchet@40067
    85
          let val (old_skolems, t') = aux old_skolems t' in
blanchet@40067
    86
            (old_skolems, Abs (s, T, t'))
blanchet@37577
    87
          end
blanchet@40067
    88
        | aux old_skolems t = (old_skolems, t)
blanchet@40067
    89
    in aux old_skolems t end
blanchet@37577
    90
  else
blanchet@40067
    91
    (old_skolems, t)
blanchet@37577
    92
blanchet@40067
    93
fun reveal_old_skolem_terms old_skolems =
blanchet@37632
    94
  map_aterms (fn t as Const (s, _) =>
blanchet@40077
    95
                 if String.isPrefix old_skolem_const_prefix s then
blanchet@40067
    96
                   AList.lookup (op =) old_skolems s |> the
blanchet@44690
    97
                   |> map_types (map_type_tvar (K dummyT))
blanchet@37632
    98
                 else
blanchet@37632
    99
                   t
blanchet@37632
   100
               | t => t)
blanchet@37632
   101
blanchet@37577
   102
blanchet@39737
   103
(* ------------------------------------------------------------------------- *)
blanchet@39737
   104
(* Logic maps manage the interface between HOL and first-order logic.        *)
blanchet@39737
   105
(* ------------------------------------------------------------------------- *)
blanchet@39737
   106
blanchet@44000
   107
datatype isa_thm =
blanchet@44000
   108
  Isa_Reflexive_or_Trivial |
blanchet@44000
   109
  Isa_Raw of thm
blanchet@44000
   110
blanchet@44000
   111
val proxy_defs = map (fst o snd o snd) proxy_table
blanchet@44000
   112
val prepare_helper =
blanchet@44000
   113
  Meson.make_meta_clause #> rewrite_rule (map safe_mk_meta_eq proxy_defs)
blanchet@44000
   114
blanchet@43933
   115
fun metis_term_from_atp (ATerm (s, tms)) =
blanchet@43935
   116
  if is_tptp_variable s then
blanchet@44109
   117
    Metis_Term.Var (Metis_Name.fromString s)
blanchet@43935
   118
  else
blanchet@44109
   119
    let
blanchet@44109
   120
      val (s, swap) = AList.lookup (op =) metis_name_table (s, length tms)
blanchet@44109
   121
                      |> the_default (s, false)
blanchet@44109
   122
    in
blanchet@44109
   123
      Metis_Term.Fn (Metis_Name.fromString s,
blanchet@44109
   124
                     tms |> map metis_term_from_atp |> swap ? rev)
blanchet@43945
   125
    end
blanchet@43945
   126
fun metis_atom_from_atp (AAtom tm) =
blanchet@43945
   127
    (case metis_term_from_atp tm of
blanchet@43945
   128
       Metis_Term.Fn x => x
blanchet@43945
   129
     | _ => raise Fail "non CNF -- expected function")
blanchet@43933
   130
  | metis_atom_from_atp _ = raise Fail "not CNF -- expected atom"
blanchet@43933
   131
fun metis_literal_from_atp (AConn (ANot, [phi])) =
blanchet@43933
   132
    (false, metis_atom_from_atp phi)
blanchet@43933
   133
  | metis_literal_from_atp phi = (true, metis_atom_from_atp phi)
blanchet@44109
   134
fun metis_literals_from_atp (AConn (AOr, phis)) =
blanchet@44109
   135
    maps metis_literals_from_atp phis
blanchet@43933
   136
  | metis_literals_from_atp phi = [metis_literal_from_atp phi]
blanchet@43933
   137
fun metis_axiom_from_atp clauses (Formula (ident, _, phi, _, _)) =
blanchet@44014
   138
    let
blanchet@44014
   139
      fun some isa =
blanchet@44014
   140
        SOME (phi |> metis_literals_from_atp |> Metis_LiteralSet.fromList
blanchet@44014
   141
                  |> Metis_Thm.axiom, isa)
blanchet@44014
   142
    in
blanchet@44014
   143
      if ident = type_tag_idempotence_helper_name orelse
blanchet@45255
   144
         String.isPrefix tags_sym_formula_prefix ident then
blanchet@44014
   145
        Isa_Reflexive_or_Trivial |> some
blanchet@44153
   146
      else if String.isPrefix conjecture_prefix ident then
blanchet@44153
   147
        NONE
blanchet@44014
   148
      else if String.isPrefix helper_prefix ident then
blanchet@44035
   149
        case (String.isSuffix typed_helper_suffix ident,
blanchet@44035
   150
              space_explode "_" ident) of
blanchet@44035
   151
          (needs_fairly_sound, _ :: const :: j :: _) =>
blanchet@44035
   152
          nth ((const, needs_fairly_sound)
blanchet@44035
   153
               |> AList.lookup (op =) helper_table |> the)
blanchet@44014
   154
              (the (Int.fromString j) - 1)
blanchet@44035
   155
          |> prepare_helper
blanchet@44035
   156
          |> Isa_Raw |> some
blanchet@44014
   157
        | _ => raise Fail ("malformed helper identifier " ^ quote ident)
blanchet@44153
   158
      else case try (unprefix fact_prefix) ident of
blanchet@44014
   159
        SOME s =>
blanchet@44153
   160
        let
blanchet@44153
   161
          val j = s |> space_explode "_" |> List.last |> Int.fromString |> the
blanchet@44153
   162
        in Meson.make_meta_clause (snd (nth clauses j)) |> Isa_Raw |> some end
blanchet@44014
   163
      | NONE => TrueI |> Isa_Raw |> some
blanchet@44014
   164
    end
blanchet@43933
   165
  | metis_axiom_from_atp _ _ = raise Fail "not CNF -- expected formula"
blanchet@43933
   166
blanchet@39737
   167
(* Function to generate metis clauses, including comb and type clauses *)
blanchet@44493
   168
fun prepare_metis_problem ctxt type_enc conj_clauses fact_clauses =
blanchet@44053
   169
  let
blanchet@45268
   170
    val type_enc = type_enc_from_string Unsound type_enc
blanchet@44153
   171
    val (conj_clauses, fact_clauses) =
blanchet@44493
   172
      if polymorphism_of_type_enc type_enc = Polymorphic then
blanchet@44153
   173
        (conj_clauses, fact_clauses)
blanchet@44153
   174
      else
blanchet@44153
   175
        conj_clauses @ fact_clauses
blanchet@44153
   176
        |> map (pair 0)
blanchet@44153
   177
        |> rpair ctxt
blanchet@44153
   178
        |-> Monomorph.monomorph atp_schematic_consts_of
blanchet@44153
   179
        |> fst |> chop (length conj_clauses)
blanchet@44153
   180
        |> pairself (maps (map (zero_var_indexes o snd)))
blanchet@44153
   181
    val num_conjs = length conj_clauses
blanchet@44053
   182
    val clauses =
blanchet@44153
   183
      map2 (fn j => pair (Int.toString j, Local))
blanchet@44153
   184
           (0 upto num_conjs - 1) conj_clauses @
blanchet@44153
   185
      (* "General" below isn't quite correct; the fact could be local. *)
blanchet@44153
   186
      map2 (fn j => pair (Int.toString (num_conjs + j), General))
blanchet@44153
   187
           (0 upto length fact_clauses - 1) fact_clauses
blanchet@44053
   188
    val (old_skolems, props) =
blanchet@44153
   189
      fold_rev (fn (name, th) => fn (old_skolems, props) =>
blanchet@44153
   190
                   th |> prop_of |> Logic.strip_imp_concl
blanchet@44153
   191
                      |> conceal_old_skolem_terms (length clauses) old_skolems
blanchet@44153
   192
                      ||> (fn prop => (name, prop) :: props))
blanchet@44153
   193
               clauses ([], [])
blanchet@44153
   194
    (*
blanchet@44153
   195
    val _ =
blanchet@44153
   196
      tracing ("PROPS:\n" ^ cat_lines (map (Syntax.string_of_term ctxt) props))
blanchet@44153
   197
    *)
blanchet@44053
   198
    val (atp_problem, _, _, _, _, _, sym_tab) =
blanchet@45256
   199
      prepare_atp_problem ctxt CNF Hypothesis Axiom type_enc false no_lambdasN
blanchet@45256
   200
                          false false [] @{prop False} props
blanchet@44153
   201
    (*
blanchet@44153
   202
    val _ = tracing ("ATP PROBLEM: " ^
blanchet@44153
   203
                     cat_lines (tptp_lines_for_atp_problem CNF atp_problem))
blanchet@44153
   204
    *)
blanchet@45268
   205
    (* "rev" is for compatibility. *)
blanchet@44053
   206
    val axioms =
blanchet@44053
   207
      atp_problem |> maps (map_filter (metis_axiom_from_atp clauses) o snd)
blanchet@44053
   208
                  |> rev
blanchet@44053
   209
  in (sym_tab, axioms, old_skolems) end
blanchet@39737
   210
paulson@15347
   211
end;