src/HOL/Tools/Metis/metis_reconstruct.ML
author blanchet
Tue, 05 Oct 2010 11:45:10 +0200
changeset 39953 aa54f347e5e2
parent 39946 78faa9b31202
child 39958 88c9aa5666de
permissions -rw-r--r--
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39495
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/metis_reconstruct.ML
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     2
    Author:     Kong W. Susanto, Cambridge University Computer Laboratory
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     3
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     5
    Copyright   Cambridge University 2007
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     6
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     7
Proof reconstruction for Metis.
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     8
*)
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
     9
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    10
signature METIS_RECONSTRUCT =
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    11
sig
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    12
  type mode = Metis_Translate.mode
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    13
39550
f97fa74c388e merge tracing of two related modules
blanchet
parents: 39498
diff changeset
    14
  val trace : bool Unsynchronized.ref
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    15
  val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
    16
  val untyped_aconv : term -> term -> bool
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    17
  val replay_one_inference :
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    18
    Proof.context -> mode -> (string * term) list
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    19
    -> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    20
    -> (Metis_Thm.thm * thm) list
39495
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    21
end;
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    22
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    23
structure Metis_Reconstruct : METIS_RECONSTRUCT =
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    24
struct
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
    25
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    26
open Metis_Translate
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    27
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    28
val trace = Unsynchronized.ref false
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    29
fun trace_msg msg = if !trace then tracing (msg ()) else ()
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    30
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    31
datatype term_or_type = SomeTerm of term | SomeType of typ
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    32
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    33
fun terms_of [] = []
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    34
  | terms_of (SomeTerm t :: tts) = t :: terms_of tts
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    35
  | terms_of (SomeType _ :: tts) = terms_of tts;
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    36
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    37
fun types_of [] = []
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    38
  | types_of (SomeTerm (Var ((a,idx), _)) :: tts) =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    39
      if String.isPrefix "_" a then
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    40
          (*Variable generated by Metis, which might have been a type variable.*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    41
          TVar (("'" ^ a, idx), HOLogic.typeS) :: types_of tts
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    42
      else types_of tts
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    43
  | types_of (SomeTerm _ :: tts) = types_of tts
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    44
  | types_of (SomeType T :: tts) = T :: types_of tts;
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    45
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    46
fun apply_list rator nargs rands =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    47
  let val trands = terms_of rands
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    48
  in  if length trands = nargs then SomeTerm (list_comb(rator, trands))
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    49
      else raise Fail
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    50
        ("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    51
          " expected " ^ Int.toString nargs ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    52
          " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    53
  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    54
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    55
fun infer_types ctxt =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    56
  Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    57
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    58
(*We use 1 rather than 0 because variable references in clauses may otherwise conflict
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    59
  with variable constraints in the goal...at least, type inference often fails otherwise.
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    60
  SEE ALSO axiom_inf below.*)
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    61
fun mk_var (w, T) = Var ((w, 1), T)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    62
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    63
(*include the default sort, if available*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    64
fun mk_tfree ctxt w =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    65
  let val ww = "'" ^ w
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    66
  in  TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1)))  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    67
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    68
(*Remove the "apply" operator from an HO term*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    69
fun strip_happ args (Metis_Term.Fn(".",[t,u])) = strip_happ (u::args) t
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    70
  | strip_happ args x = (x, args);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    71
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    72
fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    73
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    74
fun smart_invert_const "fequal" = @{const_name HOL.eq}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    75
  | smart_invert_const s = invert_const s
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    76
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    77
fun hol_type_from_metis_term _ (Metis_Term.Var v) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    78
     (case strip_prefix_and_unascii tvar_prefix v of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    79
          SOME w => make_tvar w
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    80
        | NONE   => make_tvar v)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    81
  | hol_type_from_metis_term ctxt (Metis_Term.Fn(x, tys)) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    82
     (case strip_prefix_and_unascii type_const_prefix x of
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    83
          SOME tc => Type (smart_invert_const tc,
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    84
                           map (hol_type_from_metis_term ctxt) tys)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    85
        | NONE    =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    86
      case strip_prefix_and_unascii tfree_prefix x of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    87
          SOME tf => mk_tfree ctxt tf
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    88
        | NONE    => raise Fail ("hol_type_from_metis_term: " ^ x));
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    89
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    90
(*Maps metis terms to isabelle terms*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    91
fun hol_term_from_metis_PT ctxt fol_tm =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    92
  let val thy = ProofContext.theory_of ctxt
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    93
      val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    94
                                  Metis_Term.toString fol_tm)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    95
      fun tm_to_tt (Metis_Term.Var v) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    96
             (case strip_prefix_and_unascii tvar_prefix v of
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
    97
                  SOME w => SomeType (make_tvar w)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    98
                | NONE =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
    99
              case strip_prefix_and_unascii schematic_var_prefix v of
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   100
                  SOME w => SomeTerm (mk_var (w, HOLogic.typeT))
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   101
                | NONE   => SomeTerm (mk_var (v, HOLogic.typeT)) )
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   102
                    (*Var from Metis with a name like _nnn; possibly a type variable*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   103
        | tm_to_tt (Metis_Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   104
        | tm_to_tt (t as Metis_Term.Fn (".",_)) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   105
            let val (rator,rands) = strip_happ [] t
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   106
            in  case rator of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   107
                    Metis_Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   108
                  | _ => case tm_to_tt rator of
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   109
                             SomeTerm t => SomeTerm (list_comb(t, terms_of (map tm_to_tt rands)))
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   110
                           | _ => raise Fail "tm_to_tt: HO application"
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   111
            end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   112
        | tm_to_tt (Metis_Term.Fn (fname, args)) = applic_to_tt (fname,args)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   113
      and applic_to_tt ("=",ts) =
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   114
            SomeTerm (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   115
        | applic_to_tt (a,ts) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   116
            case strip_prefix_and_unascii const_prefix a of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   117
                SOME b =>
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   118
                  let
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   119
                    val c = smart_invert_const b
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   120
                    val ntypes = num_type_args thy c
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   121
                    val nterms = length ts - ntypes
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   122
                    val tts = map tm_to_tt ts
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   123
                    val tys = types_of (List.take(tts,ntypes))
39939
6e9aff5ee9b5 paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
blanchet
parents: 39896
diff changeset
   124
                    val t =
6e9aff5ee9b5 paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
blanchet
parents: 39896
diff changeset
   125
                      if String.isPrefix new_skolem_const_prefix c then
6e9aff5ee9b5 paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
blanchet
parents: 39896
diff changeset
   126
                        Var (new_skolem_var_from_const c,
6e9aff5ee9b5 paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
blanchet
parents: 39896
diff changeset
   127
                             Type_Infer.paramify_vars (tl tys ---> hd tys))
6e9aff5ee9b5 paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
blanchet
parents: 39896
diff changeset
   128
                      else
6e9aff5ee9b5 paramify new skolems just like old ones (cf. reveal_old_skolem_terms)
blanchet
parents: 39896
diff changeset
   129
                        Const (c, dummyT)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   130
                  in if length tys = ntypes then
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   131
                         apply_list t nterms (List.drop(tts,ntypes))
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   132
                     else
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   133
                       raise Fail ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   134
                                   " but gets " ^ Int.toString (length tys) ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   135
                                   " type arguments\n" ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   136
                                   cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   137
                                   " the terms are \n" ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   138
                                   cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   139
                     end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   140
              | NONE => (*Not a constant. Is it a type constructor?*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   141
            case strip_prefix_and_unascii type_const_prefix a of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   142
                SOME b =>
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   143
                SomeType (Type (smart_invert_const b, types_of (map tm_to_tt ts)))
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   144
              | NONE => (*Maybe a TFree. Should then check that ts=[].*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   145
            case strip_prefix_and_unascii tfree_prefix a of
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   146
                SOME b => SomeType (mk_tfree ctxt b)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   147
              | NONE => (*a fixed variable? They are Skolem functions.*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   148
            case strip_prefix_and_unascii fixed_var_prefix a of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   149
                SOME b =>
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   150
                  let val opr = Free (b, HOLogic.typeT)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   151
                  in  apply_list opr (length ts) (map tm_to_tt ts)  end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   152
              | NONE => raise Fail ("unexpected metis function: " ^ a)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   153
  in
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   154
    case tm_to_tt fol_tm of
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   155
      SomeTerm t => t
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   156
    | SomeType T => raise TYPE ("fol_tm_to_tt: Term expected", [T], [])
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   157
  end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   158
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   159
(*Maps fully-typed metis terms to isabelle terms*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   160
fun hol_term_from_metis_FT ctxt fol_tm =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   161
  let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   162
                                  Metis_Term.toString fol_tm)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   163
      fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   164
             (case strip_prefix_and_unascii schematic_var_prefix v of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   165
                  SOME w =>  mk_var(w, dummyT)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   166
                | NONE   => mk_var(v, dummyT))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   167
        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn ("=",[]), _])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   168
            Const (@{const_name HOL.eq}, HOLogic.typeT)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   169
        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (x,[]), ty])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   170
           (case strip_prefix_and_unascii const_prefix x of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   171
                SOME c => Const (smart_invert_const c, dummyT)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   172
              | NONE => (*Not a constant. Is it a fixed variable??*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   173
            case strip_prefix_and_unascii fixed_var_prefix x of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   174
                SOME v => Free (v, hol_type_from_metis_term ctxt ty)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   175
              | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   176
        | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (".",[tm1,tm2]), _])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   177
            cvt tm1 $ cvt tm2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   178
        | cvt (Metis_Term.Fn (".",[tm1,tm2])) = (*untyped application*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   179
            cvt tm1 $ cvt tm2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   180
        | cvt (Metis_Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   181
        | cvt (Metis_Term.Fn ("=", [tm1,tm2])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   182
            list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   183
        | cvt (t as Metis_Term.Fn (x, [])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   184
           (case strip_prefix_and_unascii const_prefix x of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   185
                SOME c => Const (smart_invert_const c, dummyT)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   186
              | NONE => (*Not a constant. Is it a fixed variable??*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   187
            case strip_prefix_and_unascii fixed_var_prefix x of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   188
                SOME v => Free (v, dummyT)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   189
              | NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   190
                  hol_term_from_metis_PT ctxt t))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   191
        | cvt t = (trace_msg (fn () => "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   192
            hol_term_from_metis_PT ctxt t)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   193
  in fol_tm |> cvt end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   194
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   195
fun hol_term_from_metis FT = hol_term_from_metis_FT
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   196
  | hol_term_from_metis _ = hol_term_from_metis_PT
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   197
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   198
fun hol_terms_from_fol ctxt mode old_skolems fol_tms =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   199
  let val ts = map (hol_term_from_metis mode ctxt) fol_tms
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   200
      val _ = trace_msg (fn () => "  calling type inference:")
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   201
      val _ = app (fn t => trace_msg (fn () => Syntax.string_of_term ctxt t)) ts
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   202
      val ts' = ts |> map (reveal_old_skolem_terms old_skolems)
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   203
                   |> infer_types ctxt
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   204
      val _ = app (fn t => trace_msg
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   205
                    (fn () => "  final term: " ^ Syntax.string_of_term ctxt t ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   206
                              "  of type  " ^ Syntax.string_of_typ ctxt (type_of t)))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   207
                  ts'
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   208
  in  ts'  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   209
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   210
(* ------------------------------------------------------------------------- *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   211
(* FOL step Inference Rules                                                  *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   212
(* ------------------------------------------------------------------------- *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   213
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   214
(*for debugging only*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   215
(*
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   216
fun print_thpair (fth,th) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   217
  (trace_msg (fn () => "=============================================");
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   218
   trace_msg (fn () => "Metis: " ^ Metis_Thm.toString fth);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   219
   trace_msg (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th));
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   220
*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   221
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   222
fun lookth thpairs (fth : Metis_Thm.thm) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   223
  the (AList.lookup (uncurry Metis_Thm.equal) thpairs fth)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   224
  handle Option.Option =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   225
         raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fth)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   226
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   227
fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   228
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   229
(* INFERENCE RULE: AXIOM *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   230
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   231
fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   232
    (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   233
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   234
(* INFERENCE RULE: ASSUME *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   235
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   236
val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   237
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   238
fun inst_excluded_middle thy i_atm =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   239
  let val th = EXCLUDED_MIDDLE
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   240
      val [vx] = Term.add_vars (prop_of th) []
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   241
      val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   242
  in  cterm_instantiate substs th  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   243
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   244
fun assume_inf ctxt mode old_skolems atm =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   245
  inst_excluded_middle
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   246
      (ProofContext.theory_of ctxt)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   247
      (singleton (hol_terms_from_fol ctxt mode old_skolems) (Metis_Term.Fn atm))
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   248
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   249
(* INFERENCE RULE: INSTANTIATE (SUBST). Type instantiations are ignored. Trying
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   250
   to reconstruct them admits new possibilities of errors, e.g. concerning
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   251
   sorts. Instead we try to arrange that new TVars are distinct and that types
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   252
   can be inferred from terms. *)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   253
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   254
fun inst_inf ctxt mode old_skolems thpairs fsubst th =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   255
  let val thy = ProofContext.theory_of ctxt
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   256
      val i_th = lookth thpairs th
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   257
      val i_th_vars = Term.add_vars (prop_of i_th) []
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   258
      fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   259
      fun subst_translation (x,y) =
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   260
        let val v = find_var x
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   261
            (* We call "reveal_old_skolem_terms" and "infer_types" below. *)
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   262
            val t = hol_term_from_metis mode ctxt y
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   263
        in  SOME (cterm_of thy (Var v), t)  end
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   264
        handle Option.Option =>
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   265
               (trace_msg (fn () => "\"find_var\" failed for " ^ x ^
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   266
                                    " in " ^ Display.string_of_thm ctxt i_th);
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   267
                NONE)
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   268
             | TYPE _ =>
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   269
               (trace_msg (fn () => "\"hol_term_from_metis\" failed for " ^ x ^
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   270
                                    " in " ^ Display.string_of_thm ctxt i_th);
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   271
                NONE)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   272
      fun remove_typeinst (a, t) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   273
            case strip_prefix_and_unascii schematic_var_prefix a of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   274
                SOME b => SOME (b, t)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   275
              | NONE => case strip_prefix_and_unascii tvar_prefix a of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   276
                SOME _ => NONE          (*type instantiations are forbidden!*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   277
              | NONE => SOME (a,t)    (*internal Metis var?*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   278
      val _ = trace_msg (fn () => "  isa th: " ^ Display.string_of_thm ctxt i_th)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   279
      val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   280
      val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   281
      val tms = rawtms |> map (reveal_old_skolem_terms old_skolems)
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   282
                       |> infer_types ctxt
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   283
      val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   284
      val substs' = ListPair.zip (vars, map ctm_of tms)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   285
      val _ = trace_msg (fn () =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   286
        cat_lines ("subst_translations:" ::
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   287
          (substs' |> map (fn (x, y) =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   288
            Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   289
            Syntax.string_of_term ctxt (term_of y)))));
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   290
  in cterm_instantiate substs' i_th end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   291
  handle THM (msg, _, _) =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   292
         error ("Cannot replay Metis proof in Isabelle:\n" ^ msg)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   293
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   294
(* INFERENCE RULE: RESOLVE *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   295
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   296
(* Like RSN, but we rename apart only the type variables. Vars here typically
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   297
   have an index of 1, and the use of RSN would increase this typically to 3.
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   298
   Instantiations of those Vars could then fail. See comment on "mk_var". *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   299
fun resolve_inc_tyvars thy tha i thb =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   300
  let
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   301
    val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   302
    fun aux tha thb =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   303
      case Thm.bicompose false (false, tha, nprems_of tha) i thb
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   304
           |> Seq.list_of |> distinct Thm.eq_thm of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   305
        [th] => th
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   306
      | _ => raise THM ("resolve_inc_tyvars: unique result expected", i,
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   307
                        [tha, thb])
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   308
  in
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   309
    aux tha thb
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   310
    handle TERM z =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   311
           (* The unifier, which is invoked from "Thm.bicompose", will sometimes
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   312
              refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   313
              "TERM" exception (with "add_ffpair" as first argument). We then
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   314
              perform unification of the types of variables by hand and try
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   315
              again. We could do this the first time around but this error
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   316
              occurs seldom and we don't want to break existing proofs in subtle
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   317
              ways or slow them down needlessly. *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   318
           case [] |> fold (Term.add_vars o prop_of) [tha, thb]
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   319
                   |> AList.group (op =)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   320
                   |> maps (fn ((s, _), T :: Ts) =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   321
                               map (fn T' => (Free (s, T), Free (s, T'))) Ts)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   322
                   |> rpair (Envir.empty ~1)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   323
                   |-> fold (Pattern.unify thy)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   324
                   |> Envir.type_env |> Vartab.dest
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   325
                   |> map (fn (x, (S, T)) =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   326
                              pairself (ctyp_of thy) (TVar (x, S), T)) of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   327
             [] => raise TERM z
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   328
           | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   329
  end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   330
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   331
fun mk_not (Const (@{const_name Not}, _) $ b) = b
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   332
  | mk_not b = HOLogic.mk_not b
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   333
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   334
(* Match untyped terms. *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   335
fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   336
  | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   337
  | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   338
    (a = b) (* The index is ignored, for some reason. *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   339
  | untyped_aconv (Bound i) (Bound j) = (i = j)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   340
  | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   341
  | untyped_aconv (t1 $ t2) (u1 $ u2) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   342
    untyped_aconv t1 u1 andalso untyped_aconv t2 u2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   343
  | untyped_aconv _ _ = false
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   344
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   345
(* Finding the relative location of an untyped term within a list of terms *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   346
fun literal_index lit =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   347
  let
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   348
    val lit = Envir.eta_contract lit
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   349
    fun get _ [] = raise Empty
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   350
      | get n (x :: xs) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   351
        if untyped_aconv lit (Envir.eta_contract (HOLogic.dest_Trueprop x)) then
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   352
          n
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   353
        else
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   354
          get (n+1) xs
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   355
  in get 1 end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   356
39893
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   357
(* Permute a rule's premises to move the i-th premise to the last position. *)
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   358
fun make_last i th =
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   359
  let val n = nprems_of th
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   360
  in  if 1 <= i andalso i <= n
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   361
      then Thm.permute_prems (i-1) 1 th
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   362
      else raise THM("select_literal", i, [th])
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   363
  end;
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   364
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   365
(* Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   366
   double-negations. *)
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   367
val negate_head = rewrite_rule [@{thm atomize_not}, not_not RS eq_reflection]
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   368
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   369
(* Maps the clause  [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P *)
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   370
val select_literal = negate_head oo make_last
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   371
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   372
fun resolve_inf ctxt mode old_skolems thpairs atm th1 th2 =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   373
  let
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   374
    val thy = ProofContext.theory_of ctxt
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   375
    val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   376
    val _ = trace_msg (fn () => "  isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   377
    val _ = trace_msg (fn () => "  isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   378
  in
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   379
    (* Trivial cases where one operand is type info *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   380
    if Thm.eq_thm (TrueI, i_th1) then
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   381
      i_th2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   382
    else if Thm.eq_thm (TrueI, i_th2) then
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   383
      i_th1
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   384
    else
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   385
      let
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   386
        val i_atm = singleton (hol_terms_from_fol ctxt mode old_skolems)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   387
                              (Metis_Term.Fn atm)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   388
        val _ = trace_msg (fn () => "  atom: " ^ Syntax.string_of_term ctxt i_atm)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   389
        val prems_th1 = prems_of i_th1
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   390
        val prems_th2 = prems_of i_th2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   391
        val index_th1 = literal_index (mk_not i_atm) prems_th1
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   392
              handle Empty => raise Fail "Failed to find literal in th1"
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   393
        val _ = trace_msg (fn () => "  index_th1: " ^ Int.toString index_th1)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   394
        val index_th2 = literal_index i_atm prems_th2
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   395
              handle Empty => raise Fail "Failed to find literal in th2"
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   396
        val _ = trace_msg (fn () => "  index_th2: " ^ Int.toString index_th2)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   397
    in
39893
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   398
      resolve_inc_tyvars thy (select_literal index_th1 i_th1) index_th2 i_th2
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   399
    end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   400
  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   401
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   402
(* INFERENCE RULE: REFL *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   403
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   404
val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   405
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   406
val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   407
val refl_idx = 1 + Thm.maxidx_of REFL_THM;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   408
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   409
fun refl_inf ctxt mode old_skolems t =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   410
  let val thy = ProofContext.theory_of ctxt
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   411
      val i_t = singleton (hol_terms_from_fol ctxt mode old_skolems) t
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   412
      val _ = trace_msg (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   413
      val c_t = cterm_incr_types thy refl_idx i_t
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   414
  in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   415
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   416
(* INFERENCE RULE: EQUALITY *)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   417
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   418
val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   419
val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   420
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   421
val metis_eq = Metis_Term.Fn ("=", []);
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   422
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   423
fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0  (*equality has no type arguments*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   424
  | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   425
  | get_ty_arg_size _ _ = 0;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   426
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   427
fun equality_inf ctxt mode old_skolems (pos, atm) fp fr =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   428
  let val thy = ProofContext.theory_of ctxt
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   429
      val m_tm = Metis_Term.Fn atm
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   430
      val [i_atm,i_tm] = hol_terms_from_fol ctxt mode old_skolems [m_tm, fr]
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   431
      val _ = trace_msg (fn () => "sign of the literal: " ^ Bool.toString pos)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   432
      fun replace_item_list lx 0 (_::ls) = lx::ls
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   433
        | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   434
      fun path_finder_FO tm [] = (tm, Bound 0)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   435
        | path_finder_FO tm (p::ps) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   436
            let val (tm1,args) = strip_comb tm
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   437
                val adjustment = get_ty_arg_size thy tm1
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   438
                val p' = if adjustment > p then p else p-adjustment
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   439
                val tm_p = List.nth(args,p')
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   440
                  handle Subscript =>
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   441
                         error ("Cannot replay Metis proof in Isabelle:\n" ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   442
                                "equality_inf: " ^ Int.toString p ^ " adj " ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   443
                                Int.toString adjustment ^ " term " ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   444
                                Syntax.string_of_term ctxt tm)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   445
                val _ = trace_msg (fn () => "path_finder: " ^ Int.toString p ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   446
                                      "  " ^ Syntax.string_of_term ctxt tm_p)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   447
                val (r,t) = path_finder_FO tm_p ps
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   448
            in
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   449
                (r, list_comb (tm1, replace_item_list t p' args))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   450
            end
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   451
      fun path_finder_HO tm [] = (tm, Bound 0)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   452
        | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   453
        | path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   454
        | path_finder_HO tm ps =
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   455
          raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   456
                      "equality_inf, path_finder_HO: path = " ^
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   457
                      space_implode " " (map Int.toString ps) ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   458
                      " isa-term: " ^  Syntax.string_of_term ctxt tm)
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   459
      fun path_finder_FT tm [] _ = (tm, Bound 0)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   460
        | path_finder_FT tm (0::ps) (Metis_Term.Fn ("ti", [t1, _])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   461
            path_finder_FT tm ps t1
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   462
        | path_finder_FT (t$u) (0::ps) (Metis_Term.Fn (".", [t1, _])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   463
            (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   464
        | path_finder_FT (t$u) (1::ps) (Metis_Term.Fn (".", [_, t2])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   465
            (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   466
        | path_finder_FT tm ps t =
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   467
          raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   468
                      "equality_inf, path_finder_FT: path = " ^
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   469
                      space_implode " " (map Int.toString ps) ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   470
                      " isa-term: " ^  Syntax.string_of_term ctxt tm ^
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   471
                      " fol-term: " ^ Metis_Term.toString t)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   472
      fun path_finder FO tm ps _ = path_finder_FO tm ps
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   473
        | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   474
             (*equality: not curried, as other predicates are*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   475
             if p=0 then path_finder_HO tm (0::1::ps)  (*select first operand*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   476
             else path_finder_HO tm (p::ps)        (*1 selects second operand*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   477
        | path_finder HO tm (_ :: ps) (Metis_Term.Fn ("{}", [_])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   478
             path_finder_HO tm ps      (*if not equality, ignore head to skip hBOOL*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   479
        | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   480
                            (Metis_Term.Fn ("=", [t1,t2])) =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   481
             (*equality: not curried, as other predicates are*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   482
             if p=0 then path_finder_FT tm (0::1::ps)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   483
                          (Metis_Term.Fn (".", [Metis_Term.Fn (".", [metis_eq,t1]), t2]))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   484
                          (*select first operand*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   485
             else path_finder_FT tm (p::ps)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   486
                   (Metis_Term.Fn (".", [metis_eq,t2]))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   487
                   (*1 selects second operand*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   488
        | path_finder FT tm (_ :: ps) (Metis_Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   489
             (*if not equality, ignore head to skip the hBOOL predicate*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   490
        | path_finder FT tm ps t = path_finder_FT tm ps t  (*really an error case!*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   491
      fun path_finder_lit ((nt as Const (@{const_name Not}, _)) $ tm_a) idx =
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   492
            let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   493
            in (tm, nt $ tm_rslt) end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   494
        | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   495
      val (tm_subst, body) = path_finder_lit i_atm fp
39498
e8aef7ea9cbb make "subst_translation" more robust w.r.t. type instantiations like {_1234 |-> 'a}
blanchet
parents: 39497
diff changeset
   496
      val tm_abs = Abs ("x", type_of tm_subst, body)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   497
      val _ = trace_msg (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   498
      val _ = trace_msg (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   499
      val _ = trace_msg (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   500
      val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst)  (*ill typed but gives right max*)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   501
      val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   502
      val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   503
      val eq_terms = map (pairself (cterm_of thy))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   504
        (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   505
  in  cterm_instantiate eq_terms subst'  end;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   506
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   507
val factor = Seq.hd o distinct_subgoals_tac;
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   508
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   509
fun step ctxt mode old_skolems thpairs p =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   510
  case p of
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   511
    (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   512
  | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode old_skolems f_atm
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   513
  | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) =>
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   514
    factor (inst_inf ctxt mode old_skolems thpairs f_subst f_th1)
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   515
  | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) =>
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   516
    factor (resolve_inf ctxt mode old_skolems thpairs f_atm f_th1 f_th2)
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   517
  | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode old_skolems f_tm
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   518
  | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) =>
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   519
    equality_inf ctxt mode old_skolems f_lit f_p f_r
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   520
39893
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   521
fun flexflex_first_order th =
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   522
  case Thm.tpairs_of th of
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   523
      [] => th
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   524
    | pairs =>
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   525
        let val thy = theory_of_thm th
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   526
            val (_, tenv) =
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   527
              fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty)
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   528
            val t_pairs = map Meson.term_pair_of (Vartab.dest tenv)
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   529
            val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   530
        in  th'  end
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   531
        handle THM _ => th;
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   532
39895
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   533
fun is_metis_literal_genuine (_, (s, _)) = not (String.isPrefix class_prefix s)
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   534
fun is_isabelle_literal_genuine t =
39953
aa54f347e5e2 hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
blanchet
parents: 39946
diff changeset
   535
  case t of _ $ (Const (@{const_name Meson.skolem}, _) $ _) => false | _ => true
39895
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   536
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   537
fun count p xs = fold (fn x => if p x then Integer.add 1 else I) xs 0
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   538
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   539
fun replay_one_inference ctxt mode old_skolems (fol_th, inf) thpairs =
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   540
  let
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   541
    val _ = trace_msg (fn () => "=============================================")
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   542
    val _ = trace_msg (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   543
    val _ = trace_msg (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39550
diff changeset
   544
    val th = step ctxt mode old_skolems thpairs (fol_th, inf)
39893
25a339e1ff9b move functions closer to where they're used
blanchet
parents: 39887
diff changeset
   545
             |> flexflex_first_order
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   546
    val _ = trace_msg (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   547
    val _ = trace_msg (fn () => "=============================================")
39895
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   548
    val num_metis_lits =
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   549
      fol_th |> Metis_Thm.clause |> Metis_LiteralSet.toList
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   550
             |> count is_metis_literal_genuine
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   551
    val num_isabelle_lits =
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   552
      th |> prems_of |> count is_isabelle_literal_genuine
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   553
    val _ = if num_metis_lits = num_isabelle_lits then ()
a91a84b1dfdd reintroduced code that keeps track of whether the Isabelle and Metis proofs are in sync -- generalized to work with the new skolemizer
blanchet
parents: 39893
diff changeset
   554
            else error "Cannot replay Metis proof in Isabelle: Out of sync."
39497
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   555
  in (fol_th, th) :: thpairs end
fa16349939b7 complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents: 39495
diff changeset
   556
39495
bb4fb9ffe2d1 added new "Metis_Reconstruct" module, temporarily empty
blanchet
parents:
diff changeset
   557
end;