src/HOL/Tools/Metis/metis_translate.ML
author blanchet
Wed, 08 Jun 2011 08:47:43 +0200
changeset 44100 30c141dc22d6
parent 44089 69375eaa9016
child 44109 c0eaa8b9bff5
permissions -rw-r--r--
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
     1 (*  Title:      HOL/Tools/Metis/metis_translate.ML
     2     Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
     3     Author:     Kong W. Susanto, Cambridge University Computer Laboratory
     4     Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
     5     Author:     Jasmin Blanchette, TU Muenchen
     6 
     7 Translation of HOL to FOL for Metis.
     8 *)
     9 
    10 signature METIS_TRANSLATE =
    11 sig
    12   type type_literal = ATP_Translate.type_literal
    13 
    14   datatype isa_thm =
    15     Isa_Reflexive_or_Trivial |
    16     Isa_Raw of thm
    17 
    18   val metis_equal : string
    19   val metis_predicator : string
    20   val metis_app_op : string
    21   val metis_type_tag : string
    22   val metis_generated_var_prefix : string
    23   val trace : bool Config.T
    24   val verbose : bool Config.T
    25   val trace_msg : Proof.context -> (unit -> string) -> unit
    26   val verbose_warning : Proof.context -> string -> unit
    27   val metis_name_table : ((string * int) * (string * bool)) list
    28   val reveal_old_skolem_terms : (string * term) list -> term -> term
    29   val prepare_metis_problem :
    30     Proof.context -> string -> thm list -> thm list
    31     -> int Symtab.table * (Metis_Thm.thm * isa_thm) list * (string * term) list
    32 end
    33 
    34 structure Metis_Translate : METIS_TRANSLATE =
    35 struct
    36 
    37 open ATP_Problem
    38 open ATP_Translate
    39 
    40 val metis_equal = "="
    41 val metis_predicator = "{}"
    42 val metis_app_op = "."
    43 val metis_type_tag = ":"
    44 val metis_generated_var_prefix = "_"
    45 
    46 val trace = Attrib.setup_config_bool @{binding metis_trace} (K false)
    47 val verbose = Attrib.setup_config_bool @{binding metis_verbose} (K true)
    48 
    49 fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
    50 fun verbose_warning ctxt msg =
    51   if Config.get ctxt verbose then warning ("Metis: " ^ msg) else ()
    52 
    53 val metis_name_table =
    54   [((tptp_equal, 2), (metis_equal, false)),
    55    ((tptp_old_equal, 2), (metis_equal, false)),
    56    ((prefixed_predicator_name, 1), (metis_predicator, false)),
    57    ((prefixed_app_op_name, 2), (metis_app_op, false)),
    58    ((prefixed_type_tag_name, 2), (metis_type_tag, true))]
    59 
    60 fun old_skolem_const_name i j num_T_args =
    61   old_skolem_const_prefix ^ Long_Name.separator ^
    62   (space_implode Long_Name.separator (map string_of_int [i, j, num_T_args]))
    63 
    64 fun conceal_old_skolem_terms i old_skolems t =
    65   if exists_Const (curry (op =) @{const_name Meson.skolem} o fst) t then
    66     let
    67       fun aux old_skolems
    68              (t as (Const (@{const_name Meson.skolem}, Type (_, [_, T])) $ _)) =
    69           let
    70             val (old_skolems, s) =
    71               if i = ~1 then
    72                 (old_skolems, @{const_name undefined})
    73               else case AList.find (op aconv) old_skolems t of
    74                 s :: _ => (old_skolems, s)
    75               | [] =>
    76                 let
    77                   val s = old_skolem_const_name i (length old_skolems)
    78                                                 (length (Term.add_tvarsT T []))
    79                 in ((s, t) :: old_skolems, s) end
    80           in (old_skolems, Const (s, T)) end
    81         | aux old_skolems (t1 $ t2) =
    82           let
    83             val (old_skolems, t1) = aux old_skolems t1
    84             val (old_skolems, t2) = aux old_skolems t2
    85           in (old_skolems, t1 $ t2) end
    86         | aux old_skolems (Abs (s, T, t')) =
    87           let val (old_skolems, t') = aux old_skolems t' in
    88             (old_skolems, Abs (s, T, t'))
    89           end
    90         | aux old_skolems t = (old_skolems, t)
    91     in aux old_skolems t end
    92   else
    93     (old_skolems, t)
    94 
    95 fun reveal_old_skolem_terms old_skolems =
    96   map_aterms (fn t as Const (s, _) =>
    97                  if String.isPrefix old_skolem_const_prefix s then
    98                    AList.lookup (op =) old_skolems s |> the
    99                    |> map_types Type_Infer.paramify_vars
   100                  else
   101                    t
   102                | t => t)
   103 
   104 
   105 (* ------------------------------------------------------------------------- *)
   106 (* Logic maps manage the interface between HOL and first-order logic.        *)
   107 (* ------------------------------------------------------------------------- *)
   108 
   109 datatype isa_thm =
   110   Isa_Reflexive_or_Trivial |
   111   Isa_Raw of thm
   112 
   113 val proxy_defs = map (fst o snd o snd) proxy_table
   114 val prepare_helper =
   115   Meson.make_meta_clause #> rewrite_rule (map safe_mk_meta_eq proxy_defs)
   116 
   117 fun metis_name_from_atp s ary =
   118   AList.lookup (op =) metis_name_table (s, ary) |> the_default (s, false)
   119 fun metis_term_from_atp (ATerm (s, tms)) =
   120   if is_tptp_variable s then
   121     Metis_Term.Var s
   122   else
   123     let val (s, swap) = metis_name_from_atp s (length tms) in
   124       Metis_Term.Fn (s, tms |> map metis_term_from_atp |> swap ? rev)
   125     end
   126 fun metis_atom_from_atp (AAtom tm) =
   127     (case metis_term_from_atp tm of
   128        Metis_Term.Fn x => x
   129      | _ => raise Fail "non CNF -- expected function")
   130   | metis_atom_from_atp _ = raise Fail "not CNF -- expected atom"
   131 fun metis_literal_from_atp (AConn (ANot, [phi])) =
   132     (false, metis_atom_from_atp phi)
   133   | metis_literal_from_atp phi = (true, metis_atom_from_atp phi)
   134 fun metis_literals_from_atp (AConn (AOr, [phi1, phi2])) =
   135     uncurry (union (op =)) (pairself metis_literals_from_atp (phi1, phi2))
   136   | metis_literals_from_atp phi = [metis_literal_from_atp phi]
   137 fun metis_axiom_from_atp clauses (Formula (ident, _, phi, _, _)) =
   138     let
   139       fun some isa =
   140         SOME (phi |> metis_literals_from_atp |> Metis_LiteralSet.fromList
   141                   |> Metis_Thm.axiom, isa)
   142     in
   143       if ident = type_tag_idempotence_helper_name orelse
   144          String.isPrefix lightweight_tags_sym_formula_prefix ident then
   145         Isa_Reflexive_or_Trivial |> some
   146       else if String.isPrefix helper_prefix ident then
   147         case (String.isSuffix typed_helper_suffix ident,
   148               space_explode "_" ident) of
   149           (needs_fairly_sound, _ :: const :: j :: _) =>
   150           nth ((const, needs_fairly_sound)
   151                |> AList.lookup (op =) helper_table |> the)
   152               (the (Int.fromString j) - 1)
   153           |> prepare_helper
   154           |> Isa_Raw |> some
   155         | _ => raise Fail ("malformed helper identifier " ^ quote ident)
   156       else case try (unprefix conjecture_prefix) ident of
   157         SOME s =>
   158         let val j = the (Int.fromString s) in
   159           if j = length clauses then NONE
   160           else Meson.make_meta_clause (nth clauses j) |> Isa_Raw |> some
   161         end
   162       | NONE => TrueI |> Isa_Raw |> some
   163     end
   164   | metis_axiom_from_atp _ _ = raise Fail "not CNF -- expected formula"
   165 
   166 (* Function to generate metis clauses, including comb and type clauses *)
   167 fun prepare_metis_problem ctxt type_sys conj_clauses fact_clauses =
   168   let
   169     val type_sys = type_sys_from_string type_sys
   170     val clauses =
   171       conj_clauses @ fact_clauses
   172       |> (if polymorphism_of_type_sys type_sys = Polymorphic then
   173             I
   174           else
   175             map (pair 0)
   176             #> rpair ctxt
   177             #-> Monomorph.monomorph atp_schematic_consts_of
   178             #> fst #> maps (map (zero_var_indexes o snd)))
   179     val (old_skolems, props) =
   180       fold_rev (fn clause => fn (old_skolems, props) =>
   181                    clause |> prop_of |> Logic.strip_imp_concl
   182                           |> conceal_old_skolem_terms (length clauses)
   183                                                       old_skolems
   184                           ||> (fn prop => prop :: props))
   185            clauses ([], [])
   186 (*
   187 val _ = tracing ("PROPS:\n" ^ cat_lines (map (Syntax.string_of_term ctxt) props))
   188 *)
   189     val (atp_problem, _, _, _, _, _, sym_tab) =
   190       prepare_atp_problem ctxt CNF Hypothesis Axiom type_sys false false props
   191                           @{prop False} []
   192 (*
   193 val _ = tracing ("ATP PROBLEM: " ^ cat_lines (tptp_lines_for_atp_problem CNF atp_problem))
   194 *)
   195     (* "rev" is for compatibility *)
   196     val axioms =
   197       atp_problem |> maps (map_filter (metis_axiom_from_atp clauses) o snd)
   198                   |> rev
   199   in (sym_tab, axioms, old_skolems) end
   200 
   201 end;