src/HOL/Tools/metis_tools.ML
changeset 33316 6a72af4e84b8
parent 33305 a103fa7c19cc
child 33317 b4534348b8fd
equal deleted inserted replaced
33315:784c1b09d485 33316:6a72af4e84b8
    61 fun fn_isa_to_met "equal" = "="
    61 fun fn_isa_to_met "equal" = "="
    62   | fn_isa_to_met x       = x;
    62   | fn_isa_to_met x       = x;
    63 
    63 
    64 fun metis_lit b c args = (b, (c, args));
    64 fun metis_lit b c args = (b, (c, args));
    65 
    65 
    66 fun hol_type_to_fol (ResClause.AtomV x) = Metis.Term.Var x
    66 fun hol_type_to_fol (Res_Clause.AtomV x) = Metis.Term.Var x
    67   | hol_type_to_fol (ResClause.AtomF x) = Metis.Term.Fn(x,[])
    67   | hol_type_to_fol (Res_Clause.AtomF x) = Metis.Term.Fn(x,[])
    68   | hol_type_to_fol (ResClause.Comp(tc,tps)) = Metis.Term.Fn(tc, map hol_type_to_fol tps);
    68   | hol_type_to_fol (Res_Clause.Comp(tc,tps)) = Metis.Term.Fn(tc, map hol_type_to_fol tps);
    69 
    69 
    70 (*These two functions insert type literals before the real literals. That is the
    70 (*These two functions insert type literals before the real literals. That is the
    71   opposite order from TPTP linkup, but maybe OK.*)
    71   opposite order from TPTP linkup, but maybe OK.*)
    72 
    72 
    73 fun hol_term_to_fol_FO tm =
    73 fun hol_term_to_fol_FO tm =
    74   case ResHolClause.strip_comb tm of
    74   case Res_HOL_Clause.strip_comb tm of
    75       (ResHolClause.CombConst(c,_,tys), tms) =>
    75       (Res_HOL_Clause.CombConst(c,_,tys), tms) =>
    76         let val tyargs = map hol_type_to_fol tys
    76         let val tyargs = map hol_type_to_fol tys
    77             val args   = map hol_term_to_fol_FO tms
    77             val args   = map hol_term_to_fol_FO tms
    78         in Metis.Term.Fn (c, tyargs @ args) end
    78         in Metis.Term.Fn (c, tyargs @ args) end
    79     | (ResHolClause.CombVar(v,_), []) => Metis.Term.Var v
    79     | (Res_HOL_Clause.CombVar(v,_), []) => Metis.Term.Var v
    80     | _ => error "hol_term_to_fol_FO";
    80     | _ => error "hol_term_to_fol_FO";
    81 
    81 
    82 fun hol_term_to_fol_HO (ResHolClause.CombVar (a, _)) = Metis.Term.Var a
    82 fun hol_term_to_fol_HO (Res_HOL_Clause.CombVar (a, _)) = Metis.Term.Var a
    83   | hol_term_to_fol_HO (ResHolClause.CombConst (a, _, tylist)) =
    83   | hol_term_to_fol_HO (Res_HOL_Clause.CombConst (a, _, tylist)) =
    84       Metis.Term.Fn (fn_isa_to_met a, map hol_type_to_fol tylist)
    84       Metis.Term.Fn (fn_isa_to_met a, map hol_type_to_fol tylist)
    85   | hol_term_to_fol_HO (ResHolClause.CombApp (tm1, tm2)) =
    85   | hol_term_to_fol_HO (Res_HOL_Clause.CombApp (tm1, tm2)) =
    86        Metis.Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]);
    86        Metis.Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]);
    87 
    87 
    88 (*The fully-typed translation, to avoid type errors*)
    88 (*The fully-typed translation, to avoid type errors*)
    89 fun wrap_type (tm, ty) = Metis.Term.Fn("ti", [tm, hol_type_to_fol ty]);
    89 fun wrap_type (tm, ty) = Metis.Term.Fn("ti", [tm, hol_type_to_fol ty]);
    90 
    90 
    91 fun hol_term_to_fol_FT (ResHolClause.CombVar(a, ty)) =
    91 fun hol_term_to_fol_FT (Res_HOL_Clause.CombVar(a, ty)) =
    92       wrap_type (Metis.Term.Var a, ty)
    92       wrap_type (Metis.Term.Var a, ty)
    93   | hol_term_to_fol_FT (ResHolClause.CombConst(a, ty, _)) =
    93   | hol_term_to_fol_FT (Res_HOL_Clause.CombConst(a, ty, _)) =
    94       wrap_type (Metis.Term.Fn(fn_isa_to_met a, []), ty)
    94       wrap_type (Metis.Term.Fn(fn_isa_to_met a, []), ty)
    95   | hol_term_to_fol_FT (tm as ResHolClause.CombApp(tm1,tm2)) =
    95   | hol_term_to_fol_FT (tm as Res_HOL_Clause.CombApp(tm1,tm2)) =
    96        wrap_type (Metis.Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
    96        wrap_type (Metis.Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
    97                   ResHolClause.type_of_combterm tm);
    97                   Res_HOL_Clause.type_of_combterm tm);
    98 
    98 
    99 fun hol_literal_to_fol FO (ResHolClause.Literal (pol, tm)) =
    99 fun hol_literal_to_fol FO (Res_HOL_Clause.Literal (pol, tm)) =
   100       let val (ResHolClause.CombConst(p,_,tys), tms) = ResHolClause.strip_comb tm
   100       let val (Res_HOL_Clause.CombConst(p,_,tys), tms) = Res_HOL_Clause.strip_comb tm
   101           val tylits = if p = "equal" then [] else map hol_type_to_fol tys
   101           val tylits = if p = "equal" then [] else map hol_type_to_fol tys
   102           val lits = map hol_term_to_fol_FO tms
   102           val lits = map hol_term_to_fol_FO tms
   103       in metis_lit pol (fn_isa_to_met p) (tylits @ lits) end
   103       in metis_lit pol (fn_isa_to_met p) (tylits @ lits) end
   104   | hol_literal_to_fol HO (ResHolClause.Literal (pol, tm)) =
   104   | hol_literal_to_fol HO (Res_HOL_Clause.Literal (pol, tm)) =
   105      (case ResHolClause.strip_comb tm of
   105      (case Res_HOL_Clause.strip_comb tm of
   106           (ResHolClause.CombConst("equal",_,_), tms) =>
   106           (Res_HOL_Clause.CombConst("equal",_,_), tms) =>
   107             metis_lit pol "=" (map hol_term_to_fol_HO tms)
   107             metis_lit pol "=" (map hol_term_to_fol_HO tms)
   108         | _ => metis_lit pol "{}" [hol_term_to_fol_HO tm])   (*hBOOL*)
   108         | _ => metis_lit pol "{}" [hol_term_to_fol_HO tm])   (*hBOOL*)
   109   | hol_literal_to_fol FT (ResHolClause.Literal (pol, tm)) =
   109   | hol_literal_to_fol FT (Res_HOL_Clause.Literal (pol, tm)) =
   110      (case ResHolClause.strip_comb tm of
   110      (case Res_HOL_Clause.strip_comb tm of
   111           (ResHolClause.CombConst("equal",_,_), tms) =>
   111           (Res_HOL_Clause.CombConst("equal",_,_), tms) =>
   112             metis_lit pol "=" (map hol_term_to_fol_FT tms)
   112             metis_lit pol "=" (map hol_term_to_fol_FT tms)
   113         | _ => metis_lit pol "{}" [hol_term_to_fol_FT tm])   (*hBOOL*);
   113         | _ => metis_lit pol "{}" [hol_term_to_fol_FT tm])   (*hBOOL*);
   114 
   114 
   115 fun literals_of_hol_thm thy mode t =
   115 fun literals_of_hol_thm thy mode t =
   116       let val (lits, types_sorts) = ResHolClause.literals_of_term thy t
   116       let val (lits, types_sorts) = Res_HOL_Clause.literals_of_term thy t
   117       in  (map (hol_literal_to_fol mode) lits, types_sorts) end;
   117       in  (map (hol_literal_to_fol mode) lits, types_sorts) end;
   118 
   118 
   119 (*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
   119 (*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
   120 fun metis_of_typeLit pos (ResClause.LTVar (s,x))  = metis_lit pos s [Metis.Term.Var x]
   120 fun metis_of_typeLit pos (Res_Clause.LTVar (s,x))  = metis_lit pos s [Metis.Term.Var x]
   121   | metis_of_typeLit pos (ResClause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])];
   121   | metis_of_typeLit pos (Res_Clause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])];
   122 
   122 
   123 fun default_sort _ (TVar _) = false
   123 fun default_sort _ (TVar _) = false
   124   | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
   124   | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
   125 
   125 
   126 fun metis_of_tfree tf =
   126 fun metis_of_tfree tf =
   130   let val thy = ProofContext.theory_of ctxt
   130   let val thy = ProofContext.theory_of ctxt
   131       val (mlits, types_sorts) =
   131       val (mlits, types_sorts) =
   132              (literals_of_hol_thm thy mode o HOLogic.dest_Trueprop o prop_of) th
   132              (literals_of_hol_thm thy mode o HOLogic.dest_Trueprop o prop_of) th
   133   in
   133   in
   134       if is_conjecture then
   134       if is_conjecture then
   135           (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), ResClause.add_typs types_sorts)
   135           (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), Res_Clause.add_typs types_sorts)
   136       else
   136       else
   137         let val tylits = ResClause.add_typs
   137         let val tylits = Res_Clause.add_typs
   138                            (filter (not o default_sort ctxt) types_sorts)
   138                            (filter (not o default_sort ctxt) types_sorts)
   139             val mtylits = if Config.get ctxt type_lits
   139             val mtylits = if Config.get ctxt type_lits
   140                           then map (metis_of_typeLit false) tylits else []
   140                           then map (metis_of_typeLit false) tylits else []
   141         in
   141         in
   142           (Metis.Thm.axiom (Metis.LiteralSet.fromList(mtylits @ mlits)), [])
   142           (Metis.Thm.axiom (Metis.LiteralSet.fromList(mtylits @ mlits)), [])
   143         end
   143         end
   144   end;
   144   end;
   145 
   145 
   146 (* ARITY CLAUSE *)
   146 (* ARITY CLAUSE *)
   147 
   147 
   148 fun m_arity_cls (ResClause.TConsLit (c,t,args)) =
   148 fun m_arity_cls (Res_Clause.TConsLit (c,t,args)) =
   149       metis_lit true (ResClause.make_type_class c) [Metis.Term.Fn(t, map Metis.Term.Var args)]
   149       metis_lit true (Res_Clause.make_type_class c) [Metis.Term.Fn(t, map Metis.Term.Var args)]
   150   | m_arity_cls (ResClause.TVarLit (c,str))     =
   150   | m_arity_cls (Res_Clause.TVarLit (c,str))     =
   151       metis_lit false (ResClause.make_type_class c) [Metis.Term.Var str];
   151       metis_lit false (Res_Clause.make_type_class c) [Metis.Term.Var str];
   152 
   152 
   153 (*TrueI is returned as the Isabelle counterpart because there isn't any.*)
   153 (*TrueI is returned as the Isabelle counterpart because there isn't any.*)
   154 fun arity_cls (ResClause.ArityClause{conclLit,premLits,...}) =
   154 fun arity_cls (Res_Clause.ArityClause{conclLit,premLits,...}) =
   155   (TrueI,
   155   (TrueI,
   156    Metis.Thm.axiom (Metis.LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
   156    Metis.Thm.axiom (Metis.LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
   157 
   157 
   158 (* CLASSREL CLAUSE *)
   158 (* CLASSREL CLAUSE *)
   159 
   159 
   160 fun m_classrel_cls subclass superclass =
   160 fun m_classrel_cls subclass superclass =
   161   [metis_lit false subclass [Metis.Term.Var "T"], metis_lit true superclass [Metis.Term.Var "T"]];
   161   [metis_lit false subclass [Metis.Term.Var "T"], metis_lit true superclass [Metis.Term.Var "T"]];
   162 
   162 
   163 fun classrel_cls (ResClause.ClassrelClause {subclass, superclass, ...}) =
   163 fun classrel_cls (Res_Clause.ClassrelClause {subclass, superclass, ...}) =
   164   (TrueI, Metis.Thm.axiom (Metis.LiteralSet.fromList (m_classrel_cls subclass superclass)));
   164   (TrueI, Metis.Thm.axiom (Metis.LiteralSet.fromList (m_classrel_cls subclass superclass)));
   165 
   165 
   166 (* ------------------------------------------------------------------------- *)
   166 (* ------------------------------------------------------------------------- *)
   167 (* FOL to HOL  (Metis to Isabelle)                                           *)
   167 (* FOL to HOL  (Metis to Isabelle)                                           *)
   168 (* ------------------------------------------------------------------------- *)
   168 (* ------------------------------------------------------------------------- *)
   207 (*Remove the "apply" operator from an HO term*)
   207 (*Remove the "apply" operator from an HO term*)
   208 fun strip_happ args (Metis.Term.Fn(".",[t,u])) = strip_happ (u::args) t
   208 fun strip_happ args (Metis.Term.Fn(".",[t,u])) = strip_happ (u::args) t
   209   | strip_happ args x = (x, args);
   209   | strip_happ args x = (x, args);
   210 
   210 
   211 fun fol_type_to_isa _ (Metis.Term.Var v) =
   211 fun fol_type_to_isa _ (Metis.Term.Var v) =
   212      (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of
   212      (case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix v of
   213           SOME w => ResReconstruct.make_tvar w
   213           SOME w => Res_Reconstruct.make_tvar w
   214         | NONE   => ResReconstruct.make_tvar v)
   214         | NONE   => Res_Reconstruct.make_tvar v)
   215   | fol_type_to_isa ctxt (Metis.Term.Fn(x, tys)) =
   215   | fol_type_to_isa ctxt (Metis.Term.Fn(x, tys)) =
   216      (case ResReconstruct.strip_prefix ResClause.tconst_prefix x of
   216      (case Res_Reconstruct.strip_prefix Res_Clause.tconst_prefix x of
   217           SOME tc => Term.Type (ResReconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
   217           SOME tc => Term.Type (Res_Reconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
   218         | NONE    =>
   218         | NONE    =>
   219       case ResReconstruct.strip_prefix ResClause.tfree_prefix x of
   219       case Res_Reconstruct.strip_prefix Res_Clause.tfree_prefix x of
   220           SOME tf => mk_tfree ctxt tf
   220           SOME tf => mk_tfree ctxt tf
   221         | NONE    => error ("fol_type_to_isa: " ^ x));
   221         | NONE    => error ("fol_type_to_isa: " ^ x));
   222 
   222 
   223 (*Maps metis terms to isabelle terms*)
   223 (*Maps metis terms to isabelle terms*)
   224 fun fol_term_to_hol_RAW ctxt fol_tm =
   224 fun fol_term_to_hol_RAW ctxt fol_tm =
   225   let val thy = ProofContext.theory_of ctxt
   225   let val thy = ProofContext.theory_of ctxt
   226       val _ = trace_msg (fn () => "fol_term_to_hol: " ^ Metis.Term.toString fol_tm)
   226       val _ = trace_msg (fn () => "fol_term_to_hol: " ^ Metis.Term.toString fol_tm)
   227       fun tm_to_tt (Metis.Term.Var v) =
   227       fun tm_to_tt (Metis.Term.Var v) =
   228              (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of
   228              (case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix v of
   229                   SOME w => Type (ResReconstruct.make_tvar w)
   229                   SOME w => Type (Res_Reconstruct.make_tvar w)
   230                 | NONE =>
   230                 | NONE =>
   231               case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of
   231               case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix v of
   232                   SOME w => Term (mk_var (w, HOLogic.typeT))
   232                   SOME w => Term (mk_var (w, HOLogic.typeT))
   233                 | NONE   => Term (mk_var (v, HOLogic.typeT)) )
   233                 | NONE   => Term (mk_var (v, HOLogic.typeT)) )
   234                     (*Var from Metis with a name like _nnn; possibly a type variable*)
   234                     (*Var from Metis with a name like _nnn; possibly a type variable*)
   235         | tm_to_tt (Metis.Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
   235         | tm_to_tt (Metis.Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
   236         | tm_to_tt (t as Metis.Term.Fn (".",_)) =
   236         | tm_to_tt (t as Metis.Term.Fn (".",_)) =
   243             end
   243             end
   244         | tm_to_tt (Metis.Term.Fn (fname, args)) = applic_to_tt (fname,args)
   244         | tm_to_tt (Metis.Term.Fn (fname, args)) = applic_to_tt (fname,args)
   245       and applic_to_tt ("=",ts) =
   245       and applic_to_tt ("=",ts) =
   246             Term (list_comb(Const ("op =", HOLogic.typeT), terms_of (map tm_to_tt ts)))
   246             Term (list_comb(Const ("op =", HOLogic.typeT), terms_of (map tm_to_tt ts)))
   247         | applic_to_tt (a,ts) =
   247         | applic_to_tt (a,ts) =
   248             case ResReconstruct.strip_prefix ResClause.const_prefix a of
   248             case Res_Reconstruct.strip_prefix Res_Clause.const_prefix a of
   249                 SOME b =>
   249                 SOME b =>
   250                   let val c = ResReconstruct.invert_const b
   250                   let val c = Res_Reconstruct.invert_const b
   251                       val ntypes = ResReconstruct.num_typargs thy c
   251                       val ntypes = Res_Reconstruct.num_typargs thy c
   252                       val nterms = length ts - ntypes
   252                       val nterms = length ts - ntypes
   253                       val tts = map tm_to_tt ts
   253                       val tts = map tm_to_tt ts
   254                       val tys = types_of (List.take(tts,ntypes))
   254                       val tys = types_of (List.take(tts,ntypes))
   255                       val ntyargs = ResReconstruct.num_typargs thy c
   255                       val ntyargs = Res_Reconstruct.num_typargs thy c
   256                   in if length tys = ntyargs then
   256                   in if length tys = ntyargs then
   257                          apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
   257                          apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
   258                      else error ("Constant " ^ c ^ " expects " ^ Int.toString ntyargs ^
   258                      else error ("Constant " ^ c ^ " expects " ^ Int.toString ntyargs ^
   259                                  " but gets " ^ Int.toString (length tys) ^
   259                                  " but gets " ^ Int.toString (length tys) ^
   260                                  " type arguments\n" ^
   260                                  " type arguments\n" ^
   261                                  cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
   261                                  cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
   262                                  " the terms are \n" ^
   262                                  " the terms are \n" ^
   263                                  cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
   263                                  cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
   264                      end
   264                      end
   265               | NONE => (*Not a constant. Is it a type constructor?*)
   265               | NONE => (*Not a constant. Is it a type constructor?*)
   266             case ResReconstruct.strip_prefix ResClause.tconst_prefix a of
   266             case Res_Reconstruct.strip_prefix Res_Clause.tconst_prefix a of
   267                 SOME b =>
   267                 SOME b =>
   268                   Type (Term.Type (ResReconstruct.invert_type_const b, types_of (map tm_to_tt ts)))
   268                   Type (Term.Type (Res_Reconstruct.invert_type_const b, types_of (map tm_to_tt ts)))
   269               | NONE => (*Maybe a TFree. Should then check that ts=[].*)
   269               | NONE => (*Maybe a TFree. Should then check that ts=[].*)
   270             case ResReconstruct.strip_prefix ResClause.tfree_prefix a of
   270             case Res_Reconstruct.strip_prefix Res_Clause.tfree_prefix a of
   271                 SOME b => Type (mk_tfree ctxt b)
   271                 SOME b => Type (mk_tfree ctxt b)
   272               | NONE => (*a fixed variable? They are Skolem functions.*)
   272               | NONE => (*a fixed variable? They are Skolem functions.*)
   273             case ResReconstruct.strip_prefix ResClause.fixed_var_prefix a of
   273             case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix a of
   274                 SOME b =>
   274                 SOME b =>
   275                   let val opr = Term.Free(b, HOLogic.typeT)
   275                   let val opr = Term.Free(b, HOLogic.typeT)
   276                   in  apply_list opr (length ts) (map tm_to_tt ts)  end
   276                   in  apply_list opr (length ts) (map tm_to_tt ts)  end
   277               | NONE => error ("unexpected metis function: " ^ a)
   277               | NONE => error ("unexpected metis function: " ^ a)
   278   in  case tm_to_tt fol_tm of Term t => t | _ => error "fol_tm_to_tt: Term expected"  end;
   278   in  case tm_to_tt fol_tm of Term t => t | _ => error "fol_tm_to_tt: Term expected"  end;
   279 
   279 
   280 (*Maps fully-typed metis terms to isabelle terms*)
   280 (*Maps fully-typed metis terms to isabelle terms*)
   281 fun fol_term_to_hol_FT ctxt fol_tm =
   281 fun fol_term_to_hol_FT ctxt fol_tm =
   282   let val _ = trace_msg (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm)
   282   let val _ = trace_msg (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm)
   283       fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
   283       fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
   284              (case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of
   284              (case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix v of
   285                   SOME w =>  mk_var(w, dummyT)
   285                   SOME w =>  mk_var(w, dummyT)
   286                 | NONE   => mk_var(v, dummyT))
   286                 | NONE   => mk_var(v, dummyT))
   287         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) =
   287         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) =
   288             Const ("op =", HOLogic.typeT)
   288             Const ("op =", HOLogic.typeT)
   289         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
   289         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
   290            (case ResReconstruct.strip_prefix ResClause.const_prefix x of
   290            (case Res_Reconstruct.strip_prefix Res_Clause.const_prefix x of
   291                 SOME c => Const (ResReconstruct.invert_const c, dummyT)
   291                 SOME c => Const (Res_Reconstruct.invert_const c, dummyT)
   292               | NONE => (*Not a constant. Is it a fixed variable??*)
   292               | NONE => (*Not a constant. Is it a fixed variable??*)
   293             case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of
   293             case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix x of
   294                 SOME v => Free (v, fol_type_to_isa ctxt ty)
   294                 SOME v => Free (v, fol_type_to_isa ctxt ty)
   295               | NONE => error ("fol_term_to_hol_FT bad constant: " ^ x))
   295               | NONE => error ("fol_term_to_hol_FT bad constant: " ^ x))
   296         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
   296         | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
   297             cvt tm1 $ cvt tm2
   297             cvt tm1 $ cvt tm2
   298         | cvt (Metis.Term.Fn (".",[tm1,tm2])) = (*untyped application*)
   298         | cvt (Metis.Term.Fn (".",[tm1,tm2])) = (*untyped application*)
   299             cvt tm1 $ cvt tm2
   299             cvt tm1 $ cvt tm2
   300         | cvt (Metis.Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
   300         | cvt (Metis.Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
   301         | cvt (Metis.Term.Fn ("=", [tm1,tm2])) =
   301         | cvt (Metis.Term.Fn ("=", [tm1,tm2])) =
   302             list_comb(Const ("op =", HOLogic.typeT), map cvt [tm1,tm2])
   302             list_comb(Const ("op =", HOLogic.typeT), map cvt [tm1,tm2])
   303         | cvt (t as Metis.Term.Fn (x, [])) =
   303         | cvt (t as Metis.Term.Fn (x, [])) =
   304            (case ResReconstruct.strip_prefix ResClause.const_prefix x of
   304            (case Res_Reconstruct.strip_prefix Res_Clause.const_prefix x of
   305                 SOME c => Const (ResReconstruct.invert_const c, dummyT)
   305                 SOME c => Const (Res_Reconstruct.invert_const c, dummyT)
   306               | NONE => (*Not a constant. Is it a fixed variable??*)
   306               | NONE => (*Not a constant. Is it a fixed variable??*)
   307             case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of
   307             case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix x of
   308                 SOME v => Free (v, dummyT)
   308                 SOME v => Free (v, dummyT)
   309               | NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x);
   309               | NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x);
   310                   fol_term_to_hol_RAW ctxt t))
   310                   fol_term_to_hol_RAW ctxt t))
   311         | cvt t = (trace_msg (fn () => "fol_term_to_hol_FT bad term: " ^ Metis.Term.toString t);
   311         | cvt t = (trace_msg (fn () => "fol_term_to_hol_FT bad term: " ^ Metis.Term.toString t);
   312             fol_term_to_hol_RAW ctxt t)
   312             fol_term_to_hol_RAW ctxt t)
   381             handle Option =>
   381             handle Option =>
   382                 (trace_msg (fn() => "List.find failed for the variable " ^ x ^
   382                 (trace_msg (fn() => "List.find failed for the variable " ^ x ^
   383                                        " in " ^ Display.string_of_thm ctxt i_th);
   383                                        " in " ^ Display.string_of_thm ctxt i_th);
   384                  NONE)
   384                  NONE)
   385       fun remove_typeinst (a, t) =
   385       fun remove_typeinst (a, t) =
   386             case ResReconstruct.strip_prefix ResClause.schematic_var_prefix a of
   386             case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix a of
   387                 SOME b => SOME (b, t)
   387                 SOME b => SOME (b, t)
   388               | NONE   => case ResReconstruct.strip_prefix ResClause.tvar_prefix a of
   388               | NONE   => case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix a of
   389                 SOME _ => NONE          (*type instantiations are forbidden!*)
   389                 SOME _ => NONE          (*type instantiations are forbidden!*)
   390               | NONE   => SOME (a,t)    (*internal Metis var?*)
   390               | NONE   => SOME (a,t)    (*internal Metis var?*)
   391       val _ = trace_msg (fn () => "  isa th: " ^ Display.string_of_thm ctxt i_th)
   391       val _ = trace_msg (fn () => "  isa th: " ^ Display.string_of_thm ctxt i_th)
   392       val substs = map_filter remove_typeinst (Metis.Subst.toList fsubst)
   392       val substs = map_filter remove_typeinst (Metis.Subst.toList fsubst)
   393       val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
   393       val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
   450       val _ = trace_msg (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
   450       val _ = trace_msg (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
   451       val c_t = cterm_incr_types thy refl_idx i_t
   451       val c_t = cterm_incr_types thy refl_idx i_t
   452   in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
   452   in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
   453 
   453 
   454 fun get_ty_arg_size _ (Const ("op =", _)) = 0  (*equality has no type arguments*)
   454 fun get_ty_arg_size _ (Const ("op =", _)) = 0  (*equality has no type arguments*)
   455   | get_ty_arg_size thy (Const (c, _)) = (ResReconstruct.num_typargs thy c handle TYPE _ => 0)
   455   | get_ty_arg_size thy (Const (c, _)) = (Res_Reconstruct.num_typargs thy c handle TYPE _ => 0)
   456   | get_ty_arg_size _ _ = 0;
   456   | get_ty_arg_size _ _ = 0;
   457 
   457 
   458 (* INFERENCE RULE: EQUALITY *)
   458 (* INFERENCE RULE: EQUALITY *)
   459 fun equality_inf ctxt mode (pos, atm) fp fr =
   459 fun equality_inf ctxt mode (pos, atm) fp fr =
   460   let val thy = ProofContext.theory_of ctxt
   460   let val thy = ProofContext.theory_of ctxt
   536       factor (resolve_inf ctxt mode thpairs f_atm f_th1 f_th2)
   536       factor (resolve_inf ctxt mode thpairs f_atm f_th1 f_th2)
   537   | step ctxt mode _ (_, Metis.Proof.Refl f_tm) = refl_inf ctxt mode f_tm
   537   | step ctxt mode _ (_, Metis.Proof.Refl f_tm) = refl_inf ctxt mode f_tm
   538   | step ctxt mode _ (_, Metis.Proof.Equality (f_lit, f_p, f_r)) =
   538   | step ctxt mode _ (_, Metis.Proof.Equality (f_lit, f_p, f_r)) =
   539       equality_inf ctxt mode f_lit f_p f_r;
   539       equality_inf ctxt mode f_lit f_p f_r;
   540 
   540 
   541 fun real_literal (_, (c, _)) = not (String.isPrefix ResClause.class_prefix c);
   541 fun real_literal (_, (c, _)) = not (String.isPrefix Res_Clause.class_prefix c);
   542 
   542 
   543 fun translate _ _ thpairs [] = thpairs
   543 fun translate _ _ thpairs [] = thpairs
   544   | translate mode ctxt thpairs ((fol_th, inf) :: infpairs) =
   544   | translate mode ctxt thpairs ((fol_th, inf) :: infpairs) =
   545       let val _ = trace_msg (fn () => "=============================================")
   545       let val _ = trace_msg (fn () => "=============================================")
   546           val _ = trace_msg (fn () => "METIS THM: " ^ Metis.Thm.toString fol_th)
   546           val _ = trace_msg (fn () => "METIS THM: " ^ Metis.Thm.toString fol_th)
   562 
   562 
   563 (* ------------------------------------------------------------------------- *)
   563 (* ------------------------------------------------------------------------- *)
   564 (* Translation of HO Clauses                                                 *)
   564 (* Translation of HO Clauses                                                 *)
   565 (* ------------------------------------------------------------------------- *)
   565 (* ------------------------------------------------------------------------- *)
   566 
   566 
   567 fun cnf_th thy th = hd (ResAxioms.cnf_axiom thy th);
   567 fun cnf_th thy th = hd (Res_Axioms.cnf_axiom thy th);
   568 
   568 
   569 val equal_imp_fequal' = cnf_th @{theory} @{thm equal_imp_fequal};
   569 val equal_imp_fequal' = cnf_th @{theory} @{thm equal_imp_fequal};
   570 val fequal_imp_equal' = cnf_th @{theory} @{thm fequal_imp_equal};
   570 val fequal_imp_equal' = cnf_th @{theory} @{thm fequal_imp_equal};
   571 
   571 
   572 val comb_I = cnf_th @{theory} ResHolClause.comb_I;
   572 val comb_I = cnf_th @{theory} Res_HOL_Clause.comb_I;
   573 val comb_K = cnf_th @{theory} ResHolClause.comb_K;
   573 val comb_K = cnf_th @{theory} Res_HOL_Clause.comb_K;
   574 val comb_B = cnf_th @{theory} ResHolClause.comb_B;
   574 val comb_B = cnf_th @{theory} Res_HOL_Clause.comb_B;
   575 val comb_C = cnf_th @{theory} ResHolClause.comb_C;
   575 val comb_C = cnf_th @{theory} Res_HOL_Clause.comb_C;
   576 val comb_S = cnf_th @{theory} ResHolClause.comb_S;
   576 val comb_S = cnf_th @{theory} Res_HOL_Clause.comb_S;
   577 
   577 
   578 fun type_ext thy tms =
   578 fun type_ext thy tms =
   579   let val subs = ResAtp.tfree_classes_of_terms tms
   579   let val subs = Res_ATP.tfree_classes_of_terms tms
   580       val supers = ResAtp.tvar_classes_of_terms tms
   580       val supers = Res_ATP.tvar_classes_of_terms tms
   581       and tycons = ResAtp.type_consts_of_terms thy tms
   581       and tycons = Res_ATP.type_consts_of_terms thy tms
   582       val (supers', arity_clauses) = ResClause.make_arity_clauses thy tycons supers
   582       val (supers', arity_clauses) = Res_Clause.make_arity_clauses thy tycons supers
   583       val classrel_clauses = ResClause.make_classrel_clauses thy subs supers'
   583       val classrel_clauses = Res_Clause.make_classrel_clauses thy subs supers'
   584   in  map classrel_cls classrel_clauses @ map arity_cls arity_clauses
   584   in  map classrel_cls classrel_clauses @ map arity_cls arity_clauses
   585   end;
   585   end;
   586 
   586 
   587 (* ------------------------------------------------------------------------- *)
   587 (* ------------------------------------------------------------------------- *)
   588 (* Logic maps manage the interface between HOL and first-order logic.        *)
   588 (* Logic maps manage the interface between HOL and first-order logic.        *)
   589 (* ------------------------------------------------------------------------- *)
   589 (* ------------------------------------------------------------------------- *)
   590 
   590 
   591 type logic_map =
   591 type logic_map =
   592   {axioms : (Metis.Thm.thm * thm) list,
   592   {axioms : (Metis.Thm.thm * thm) list,
   593    tfrees : ResClause.type_literal list};
   593    tfrees : Res_Clause.type_literal list};
   594 
   594 
   595 fun const_in_metis c (pred, tm_list) =
   595 fun const_in_metis c (pred, tm_list) =
   596   let
   596   let
   597     fun in_mterm (Metis.Term.Var _) = false
   597     fun in_mterm (Metis.Term.Var _) = false
   598       | in_mterm (Metis.Term.Fn (".", tm_list)) = exists in_mterm tm_list
   598       | in_mterm (Metis.Term.Fn (".", tm_list)) = exists in_mterm tm_list
   600   in  c = pred orelse exists in_mterm tm_list  end;
   600   in  c = pred orelse exists in_mterm tm_list  end;
   601 
   601 
   602 (*Extract TFree constraints from context to include as conjecture clauses*)
   602 (*Extract TFree constraints from context to include as conjecture clauses*)
   603 fun init_tfrees ctxt =
   603 fun init_tfrees ctxt =
   604   let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts
   604   let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts
   605   in  ResClause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end;
   605   in  Res_Clause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end;
   606 
   606 
   607 (*transform isabelle type / arity clause to metis clause *)
   607 (*transform isabelle type / arity clause to metis clause *)
   608 fun add_type_thm [] lmap = lmap
   608 fun add_type_thm [] lmap = lmap
   609   | add_type_thm ((ith, mth) :: cls) {axioms, tfrees} =
   609   | add_type_thm ((ith, mth) :: cls) {axioms, tfrees} =
   610       add_type_thm cls {axioms = (mth, ith) :: axioms,
   610       add_type_thm cls {axioms = (mth, ith) :: axioms,
   662 exception METIS of string;
   662 exception METIS of string;
   663 
   663 
   664 (* Main function to start metis prove and reconstruction *)
   664 (* Main function to start metis prove and reconstruction *)
   665 fun FOL_SOLVE mode ctxt cls ths0 =
   665 fun FOL_SOLVE mode ctxt cls ths0 =
   666   let val thy = ProofContext.theory_of ctxt
   666   let val thy = ProofContext.theory_of ctxt
   667       val th_cls_pairs = map (fn th => (Thm.get_name_hint th, ResAxioms.cnf_axiom thy th)) ths0
   667       val th_cls_pairs = map (fn th => (Thm.get_name_hint th, Res_Axioms.cnf_axiom thy th)) ths0
   668       val ths = maps #2 th_cls_pairs
   668       val ths = maps #2 th_cls_pairs
   669       val _ = trace_msg (fn () => "FOL_SOLVE: CONJECTURE CLAUSES")
   669       val _ = trace_msg (fn () => "FOL_SOLVE: CONJECTURE CLAUSES")
   670       val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
   670       val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
   671       val _ = trace_msg (fn () => "THEOREM CLAUSES")
   671       val _ = trace_msg (fn () => "THEOREM CLAUSES")
   672       val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) ths
   672       val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) ths
   673       val (mode, {axioms,tfrees}) = build_map mode ctxt cls ths
   673       val (mode, {axioms,tfrees}) = build_map mode ctxt cls ths
   674       val _ = if null tfrees then ()
   674       val _ = if null tfrees then ()
   675               else (trace_msg (fn () => "TFREE CLAUSES");
   675               else (trace_msg (fn () => "TFREE CLAUSES");
   676                     app (fn tf => trace_msg (fn _ => ResClause.tptp_of_typeLit true tf)) tfrees)
   676                     app (fn tf => trace_msg (fn _ => Res_Clause.tptp_of_typeLit true tf)) tfrees)
   677       val _ = trace_msg (fn () => "CLAUSES GIVEN TO METIS")
   677       val _ = trace_msg (fn () => "CLAUSES GIVEN TO METIS")
   678       val thms = map #1 axioms
   678       val thms = map #1 axioms
   679       val _ = app (fn th => trace_msg (fn () => Metis.Thm.toString th)) thms
   679       val _ = app (fn th => trace_msg (fn () => Metis.Thm.toString th)) thms
   680       val _ = trace_msg (fn () => "mode = " ^ string_of_mode mode)
   680       val _ = trace_msg (fn () => "mode = " ^ string_of_mode mode)
   681       val _ = trace_msg (fn () => "START METIS PROVE PROCESS")
   681       val _ = trace_msg (fn () => "START METIS PROVE PROCESS")
   713 
   713 
   714 fun metis_general_tac mode ctxt ths i st0 =
   714 fun metis_general_tac mode ctxt ths i st0 =
   715   let val _ = trace_msg (fn () =>
   715   let val _ = trace_msg (fn () =>
   716         "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths))
   716         "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths))
   717   in
   717   in
   718      if exists_type ResAxioms.type_has_empty_sort (prop_of st0)
   718      if exists_type Res_Axioms.type_has_empty_sort (prop_of st0)
   719      then (warning "Proof state contains the empty sort"; Seq.empty)
   719      then (warning "Proof state contains the empty sort"; Seq.empty)
   720      else
   720      else
   721        (Meson.MESON ResAxioms.neg_clausify
   721        (Meson.MESON Res_Axioms.neg_clausify
   722          (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1) ctxt i
   722          (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1) ctxt i
   723         THEN ResAxioms.expand_defs_tac st0) st0
   723         THEN Res_Axioms.expand_defs_tac st0) st0
   724   end
   724   end
   725   handle METIS s => (warning ("Metis: " ^ s); Seq.empty);
   725   handle METIS s => (warning ("Metis: " ^ s); Seq.empty);
   726 
   726 
   727 val metis_tac = metis_general_tac HO;
   727 val metis_tac = metis_general_tac HO;
   728 val metisF_tac = metis_general_tac FO;
   728 val metisF_tac = metis_general_tac FO;
   735   type_lits_setup #>
   735   type_lits_setup #>
   736   method @{binding metis} HO "METIS for FOL & HOL problems" #>
   736   method @{binding metis} HO "METIS for FOL & HOL problems" #>
   737   method @{binding metisF} FO "METIS for FOL problems" #>
   737   method @{binding metisF} FO "METIS for FOL problems" #>
   738   method @{binding metisFT} FT "METIS With-fully typed translation" #>
   738   method @{binding metisFT} FT "METIS With-fully typed translation" #>
   739   Method.setup @{binding finish_clausify}
   739   Method.setup @{binding finish_clausify}
   740     (Scan.succeed (K (SIMPLE_METHOD (ResAxioms.expand_defs_tac refl))))
   740     (Scan.succeed (K (SIMPLE_METHOD (Res_Axioms.expand_defs_tac refl))))
   741     "cleanup after conversion to clauses";
   741     "cleanup after conversion to clauses";
   742 
   742 
   743 end;
   743 end;