blanchet@40358: (* Title: HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML blanchet@38506: Author: Fabian Immler, TU Muenchen blanchet@38506: Author: Makarius blanchet@38506: Author: Jasmin Blanchette, TU Muenchen blanchet@38506: blanchet@39734: Translation of HOL to FOL for Sledgehammer. blanchet@38506: *) blanchet@38506: blanchet@40249: signature SLEDGEHAMMER_ATP_TRANSLATE = blanchet@38506: sig blanchet@43088: type 'a fo_term = 'a ATP_Problem.fo_term blanchet@43580: type formula_kind = ATP_Problem.formula_kind blanchet@38506: type 'a problem = 'a ATP_Problem.problem blanchet@43511: type locality = Sledgehammer_Filter.locality blanchet@43484: blanchet@43484: datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic blanchet@43484: datatype type_level = blanchet@43484: All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types blanchet@43484: blanchet@43484: datatype type_system = blanchet@43587: Simple_Types of type_level | blanchet@43484: Preds of polymorphism * type_level | blanchet@43484: Tags of polymorphism * type_level blanchet@43484: blanchet@40358: type translated_formula blanchet@38506: blanchet@43517: val readable_names : bool Config.T blanchet@40445: val fact_prefix : string blanchet@38506: val conjecture_prefix : string blanchet@43439: val predicator_base : string blanchet@43415: val explicit_app_base : string blanchet@43420: val type_pred_base : string blanchet@43433: val tff_type_prefix : string blanchet@43484: val type_sys_from_string : string -> type_system blanchet@43484: val polymorphism_of_type_sys : type_system -> polymorphism blanchet@43484: val level_of_type_sys : type_system -> type_level blanchet@43484: val is_type_sys_virtually_sound : type_system -> bool blanchet@43484: val is_type_sys_fairly_sound : type_system -> bool blanchet@41384: val num_atp_type_args : theory -> type_system -> string -> int blanchet@43413: val unmangled_const : string -> string * string fo_term list blanchet@41336: val translate_atp_fact : blanchet@43511: Proof.context -> bool -> (string * locality) * thm blanchet@43511: -> translated_formula option * ((string * locality) * thm) blanchet@40240: val prepare_atp_problem : blanchet@43580: Proof.context -> formula_kind -> formula_kind -> type_system -> bool blanchet@43580: -> term list -> term blanchet@41339: -> (translated_formula option * ((string * 'a) * thm)) list blanchet@43412: -> string problem * string Symtab.table * int * int blanchet@43412: * (string * 'a) list vector blanchet@41561: val atp_problem_weights : string problem -> (string * real) list blanchet@38506: end; blanchet@38506: blanchet@41388: structure Sledgehammer_ATP_Translate : SLEDGEHAMMER_ATP_TRANSLATE = blanchet@38506: struct blanchet@38506: blanchet@38506: open ATP_Problem blanchet@39734: open Metis_Translate blanchet@38506: open Sledgehammer_Util blanchet@43511: open Sledgehammer_Filter blanchet@43511: blanchet@43511: (* experimental *) blanchet@43511: val generate_useful_info = false blanchet@38506: blanchet@43439: (* Readable names are often much shorter, especially if types are mangled in blanchet@43460: names. Also, the logic for generating legal SNARK sort names is only blanchet@43460: implemented for readable names. Finally, readable names are, well, more blanchet@43460: readable. For these reason, they are enabled by default. *) blanchet@43517: val readable_names = blanchet@43517: Attrib.setup_config_bool @{binding sledgehammer_atp_readable_names} (K true) blanchet@43439: blanchet@43414: val type_decl_prefix = "type_" blanchet@43414: val sym_decl_prefix = "sym_" blanchet@40445: val fact_prefix = "fact_" blanchet@38506: val conjecture_prefix = "conj_" blanchet@38506: val helper_prefix = "help_" blanchet@43414: val class_rel_clause_prefix = "crel_"; blanchet@38506: val arity_clause_prefix = "arity_" blanchet@40156: val tfree_prefix = "tfree_" blanchet@38506: blanchet@43439: val predicator_base = "hBOOL" blanchet@43415: val explicit_app_base = "hAPP" blanchet@43413: val type_pred_base = "is" blanchet@43433: val tff_type_prefix = "ty_" blanchet@43402: blanchet@43433: fun make_tff_type s = tff_type_prefix ^ ascii_of s blanchet@43402: blanchet@38506: (* Freshness almost guaranteed! *) blanchet@38506: val sledgehammer_weak_prefix = "Sledgehammer:" blanchet@38506: blanchet@43484: datatype polymorphism = Polymorphic | Monomorphic | Mangled_Monomorphic blanchet@43484: datatype type_level = blanchet@43484: All_Types | Nonmonotonic_Types | Finite_Types | Const_Arg_Types | No_Types blanchet@43484: blanchet@43484: datatype type_system = blanchet@43587: Simple_Types of type_level | blanchet@43484: Preds of polymorphism * type_level | blanchet@43484: Tags of polymorphism * type_level blanchet@43484: blanchet@43559: fun try_unsuffixes ss s = blanchet@43559: fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE blanchet@43559: blanchet@43484: fun type_sys_from_string s = blanchet@43587: (case try (unprefix "poly_") s of blanchet@43587: SOME s => (SOME Polymorphic, s) blanchet@43484: | NONE => blanchet@43484: case try (unprefix "mono_") s of blanchet@43587: SOME s => (SOME Monomorphic, s) blanchet@43587: | NONE => blanchet@43587: case try (unprefix "mangled_") s of blanchet@43587: SOME s => (SOME Mangled_Monomorphic, s) blanchet@43587: | NONE => (NONE, s)) blanchet@43484: ||> (fn s => blanchet@43559: (* "_query" and "_bang" are for the ASCII-challenged Mirabelle. *) blanchet@43559: case try_unsuffixes ["?", "_query"] s of blanchet@43484: SOME s => (Nonmonotonic_Types, s) blanchet@43484: | NONE => blanchet@43559: case try_unsuffixes ["!", "_bang"] s of blanchet@43484: SOME s => (Finite_Types, s) blanchet@43484: | NONE => (All_Types, s)) blanchet@43587: |> (fn (poly, (level, core)) => blanchet@43587: case (core, (poly, level)) of blanchet@43587: ("simple_types", (NONE, level)) => Simple_Types level blanchet@43587: | ("preds", (SOME poly, level)) => Preds (poly, level) blanchet@43587: | ("tags", (SOME poly, level)) => Tags (poly, level) blanchet@43587: | ("args", (SOME poly, All_Types (* naja *))) => blanchet@43587: Preds (poly, Const_Arg_Types) blanchet@43587: | ("erased", (NONE, All_Types (* naja *))) => blanchet@43587: Preds (Polymorphic, No_Types) blanchet@43484: | _ => error ("Unknown type system: " ^ quote s ^ ".")) blanchet@43484: blanchet@43587: fun polymorphism_of_type_sys (Simple_Types _) = Mangled_Monomorphic blanchet@43484: | polymorphism_of_type_sys (Preds (poly, _)) = poly blanchet@43484: | polymorphism_of_type_sys (Tags (poly, _)) = poly blanchet@43484: blanchet@43587: fun level_of_type_sys (Simple_Types level) = level blanchet@43484: | level_of_type_sys (Preds (_, level)) = level blanchet@43484: | level_of_type_sys (Tags (_, level)) = level blanchet@43484: blanchet@43557: fun is_type_level_virtually_sound level = blanchet@43557: level = All_Types orelse level = Nonmonotonic_Types blanchet@43484: val is_type_sys_virtually_sound = blanchet@43484: is_type_level_virtually_sound o level_of_type_sys blanchet@43484: blanchet@43484: fun is_type_level_fairly_sound level = blanchet@43484: is_type_level_virtually_sound level orelse level = Finite_Types blanchet@43484: val is_type_sys_fairly_sound = is_type_level_fairly_sound o level_of_type_sys blanchet@43484: blanchet@43557: fun is_type_level_partial level = blanchet@43557: level = Nonmonotonic_Types orelse level = Finite_Types blanchet@43557: blanchet@43444: fun formula_map f (AQuant (q, xs, phi)) = AQuant (q, xs, formula_map f phi) blanchet@43444: | formula_map f (AConn (c, phis)) = AConn (c, map (formula_map f) phis) blanchet@43444: | formula_map f (AAtom tm) = AAtom (f tm) blanchet@43444: blanchet@43550: fun formula_fold pos f = blanchet@43547: let blanchet@43547: fun aux pos (AQuant (_, _, phi)) = aux pos phi blanchet@43550: | aux pos (AConn (ANot, [phi])) = aux (Option.map not pos) phi blanchet@43547: | aux pos (AConn (AImplies, [phi1, phi2])) = blanchet@43550: aux (Option.map not pos) phi1 #> aux pos phi2 blanchet@43550: | aux pos (AConn (AAnd, phis)) = fold (aux pos) phis blanchet@43550: | aux pos (AConn (AOr, phis)) = fold (aux pos) phis blanchet@43550: | aux _ (AConn (_, phis)) = fold (aux NONE) phis blanchet@43547: | aux pos (AAtom tm) = f pos tm blanchet@43550: in aux (SOME pos) end blanchet@43444: blanchet@40358: type translated_formula = blanchet@38991: {name: string, blanchet@43511: locality: locality, blanchet@43396: kind: formula_kind, blanchet@43433: combformula: (name, typ, combterm) formula, blanchet@43433: atomic_types: typ list} blanchet@38506: blanchet@43511: fun update_combformula f ({name, locality, kind, combformula, atomic_types} blanchet@43511: : translated_formula) = blanchet@43511: {name = name, locality = locality, kind = kind, combformula = f combformula, blanchet@43433: atomic_types = atomic_types} : translated_formula blanchet@43413: blanchet@43429: fun fact_lift f ({combformula, ...} : translated_formula) = f combformula blanchet@43429: blanchet@43443: val boring_consts = [explicit_app_base, @{const_name Metis.fequal}] blanchet@43443: blanchet@43443: fun should_omit_type_args type_sys s = blanchet@43460: s <> type_pred_base andalso s <> type_tag_name andalso blanchet@43460: (s = @{const_name HOL.eq} orelse level_of_type_sys type_sys = No_Types orelse blanchet@43460: (case type_sys of blanchet@43460: Tags (_, All_Types) => true blanchet@43460: | _ => polymorphism_of_type_sys type_sys <> Mangled_Monomorphic andalso blanchet@43460: member (op =) boring_consts s)) blanchet@43547: blanchet@43460: datatype type_arg_policy = No_Type_Args | Explicit_Type_Args | Mangled_Type_Args blanchet@41384: blanchet@43460: fun general_type_arg_policy type_sys = blanchet@43460: if level_of_type_sys type_sys = No_Types then blanchet@43460: No_Type_Args blanchet@43460: else if polymorphism_of_type_sys type_sys = Mangled_Monomorphic then blanchet@43460: Mangled_Type_Args blanchet@43460: else blanchet@43460: Explicit_Type_Args blanchet@43434: blanchet@43395: fun type_arg_policy type_sys s = blanchet@43443: if should_omit_type_args type_sys s then No_Type_Args blanchet@43434: else general_type_arg_policy type_sys blanchet@43088: blanchet@41384: fun num_atp_type_args thy type_sys s = blanchet@43428: if type_arg_policy type_sys s = Explicit_Type_Args then num_type_args thy s blanchet@43428: else 0 blanchet@41384: blanchet@43224: fun atp_type_literals_for_types type_sys kind Ts = blanchet@43460: if level_of_type_sys type_sys = No_Types then blanchet@43224: [] blanchet@43224: else blanchet@43224: Ts |> type_literals_for_types blanchet@43224: |> filter (fn TyLitVar _ => kind <> Conjecture blanchet@43224: | TyLitFree _ => kind = Conjecture) blanchet@41385: blanchet@43580: fun mk_anot phi = AConn (ANot, [phi]) blanchet@38506: fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2]) blanchet@43405: fun mk_aconns c phis = blanchet@43405: let val (phis', phi') = split_last phis in blanchet@43405: fold_rev (mk_aconn c) phis' phi' blanchet@43405: end blanchet@38506: fun mk_ahorn [] phi = phi blanchet@43405: | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi]) blanchet@43393: fun mk_aquant _ [] phi = phi blanchet@43393: | mk_aquant q xs (phi as AQuant (q', xs', phi')) = blanchet@43393: if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi) blanchet@43393: | mk_aquant q xs phi = AQuant (q, xs, phi) blanchet@38506: blanchet@43393: fun close_universally atom_vars phi = blanchet@41393: let blanchet@41393: fun formula_vars bounds (AQuant (_, xs, phi)) = blanchet@43397: formula_vars (map fst xs @ bounds) phi blanchet@41393: | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis blanchet@43393: | formula_vars bounds (AAtom tm) = blanchet@43397: union (op =) (atom_vars tm [] blanchet@43397: |> filter_out (member (op =) bounds o fst)) blanchet@43393: in mk_aquant AForall (formula_vars [] phi []) phi end blanchet@43393: blanchet@43402: fun combterm_vars (CombApp (tm1, tm2)) = fold combterm_vars [tm1, tm2] blanchet@43393: | combterm_vars (CombConst _) = I blanchet@43445: | combterm_vars (CombVar (name, T)) = insert (op =) (name, SOME T) blanchet@43545: fun close_combformula_universally phi = close_universally combterm_vars phi blanchet@43393: blanchet@43393: fun term_vars (ATerm (name as (s, _), tms)) = blanchet@43402: is_atp_variable s ? insert (op =) (name, NONE) blanchet@43397: #> fold term_vars tms blanchet@43545: fun close_formula_universally phi = close_universally term_vars phi blanchet@41393: blanchet@43433: fun fo_term_from_typ (Type (s, Ts)) = blanchet@43433: ATerm (`make_fixed_type_const s, map fo_term_from_typ Ts) blanchet@43433: | fo_term_from_typ (TFree (s, _)) = blanchet@43433: ATerm (`make_fixed_type_var s, []) blanchet@43433: | fo_term_from_typ (TVar ((x as (s, _)), _)) = blanchet@43433: ATerm ((make_schematic_type_var x, s), []) blanchet@43433: blanchet@43433: (* This shouldn't clash with anything else. *) blanchet@43413: val mangled_type_sep = "\000" blanchet@43413: blanchet@43433: fun generic_mangled_type_name f (ATerm (name, [])) = f name blanchet@43433: | generic_mangled_type_name f (ATerm (name, tys)) = blanchet@43433: f name ^ "(" ^ commas (map (generic_mangled_type_name f) tys) ^ ")" blanchet@43433: val mangled_type_name = blanchet@43433: fo_term_from_typ blanchet@43433: #> (fn ty => (make_tff_type (generic_mangled_type_name fst ty), blanchet@43433: generic_mangled_type_name snd ty)) blanchet@43413: blanchet@43445: fun generic_mangled_type_suffix f g Ts = blanchet@43413: fold_rev (curry (op ^) o g o prefix mangled_type_sep blanchet@43445: o generic_mangled_type_name f) Ts "" blanchet@43433: fun mangled_const_name T_args (s, s') = blanchet@43433: let val ty_args = map fo_term_from_typ T_args in blanchet@43433: (s ^ generic_mangled_type_suffix fst ascii_of ty_args, blanchet@43433: s' ^ generic_mangled_type_suffix snd I ty_args) blanchet@43433: end blanchet@43413: blanchet@43413: val parse_mangled_ident = blanchet@43413: Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode blanchet@43413: blanchet@43413: fun parse_mangled_type x = blanchet@43413: (parse_mangled_ident blanchet@43413: -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")") blanchet@43413: [] >> ATerm) x blanchet@43413: and parse_mangled_types x = blanchet@43413: (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x blanchet@43413: blanchet@43413: fun unmangled_type s = blanchet@43413: s |> suffix ")" |> raw_explode blanchet@43413: |> Scan.finite Symbol.stopper blanchet@43413: (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^ blanchet@43413: quote s)) parse_mangled_type)) blanchet@43413: |> fst blanchet@43413: blanchet@43432: val unmangled_const_name = space_explode mangled_type_sep #> hd blanchet@43413: fun unmangled_const s = blanchet@43413: let val ss = space_explode mangled_type_sep s in blanchet@43413: (hd ss, map unmangled_type (tl ss)) blanchet@43413: end blanchet@43413: blanchet@43545: fun introduce_proxies tm = blanchet@43439: let blanchet@43439: fun aux top_level (CombApp (tm1, tm2)) = blanchet@43439: CombApp (aux top_level tm1, aux false tm2) blanchet@43445: | aux top_level (CombConst (name as (s, s'), T, T_args)) = blanchet@43441: (case proxify_const s of blanchet@43439: SOME proxy_base => blanchet@43439: if top_level then blanchet@43439: (case s of blanchet@43439: "c_False" => (tptp_false, s') blanchet@43439: | "c_True" => (tptp_true, s') blanchet@43439: | _ => name, []) blanchet@43440: else blanchet@43445: (proxy_base |>> prefix const_prefix, T_args) blanchet@43445: | NONE => (name, T_args)) blanchet@43445: |> (fn (name, T_args) => CombConst (name, T, T_args)) blanchet@43439: | aux _ tm = tm blanchet@43545: in aux true tm end blanchet@43439: blanchet@43433: fun combformula_from_prop thy eq_as_iff = blanchet@38506: let blanchet@43439: fun do_term bs t atomic_types = blanchet@41388: combterm_from_term thy bs (Envir.eta_contract t) blanchet@43439: |>> (introduce_proxies #> AAtom) blanchet@43439: ||> union (op =) atomic_types blanchet@38506: fun do_quant bs q s T t' = blanchet@38743: let val s = Name.variant (map fst bs) s in blanchet@38743: do_formula ((s, T) :: bs) t' blanchet@43433: #>> mk_aquant q [(`make_bound_var s, SOME T)] blanchet@38743: end blanchet@38506: and do_conn bs c t1 t2 = blanchet@38506: do_formula bs t1 ##>> do_formula bs t2 blanchet@43402: #>> uncurry (mk_aconn c) blanchet@38506: and do_formula bs t = blanchet@38506: case t of blanchet@43402: @{const Not} $ t1 => do_formula bs t1 #>> mk_anot blanchet@38506: | Const (@{const_name All}, _) $ Abs (s, T, t') => blanchet@38506: do_quant bs AForall s T t' blanchet@38506: | Const (@{const_name Ex}, _) $ Abs (s, T, t') => blanchet@38506: do_quant bs AExists s T t' haftmann@39028: | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2 haftmann@39028: | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2 haftmann@39019: | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2 haftmann@39093: | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 => blanchet@41388: if eq_as_iff then do_conn bs AIff t1 t2 else do_term bs t blanchet@41388: | _ => do_term bs t blanchet@38506: in do_formula [] end blanchet@38506: blanchet@38841: val presimplify_term = prop_of o Meson.presimplify oo Skip_Proof.make_thm blanchet@38506: wenzelm@41739: fun concealed_bound_name j = sledgehammer_weak_prefix ^ string_of_int j blanchet@38506: fun conceal_bounds Ts t = blanchet@38506: subst_bounds (map (Free o apfst concealed_bound_name) blanchet@38506: (0 upto length Ts - 1 ~~ Ts), t) blanchet@38506: fun reveal_bounds Ts = blanchet@38506: subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j)) blanchet@38506: (0 upto length Ts - 1 ~~ Ts)) blanchet@38506: blanchet@43612: fun extensionalize_term ctxt t = blanchet@43612: let val thy = Proof_Context.theory_of ctxt in blanchet@43612: t |> cterm_of thy |> Meson.extensionalize_conv ctxt blanchet@43612: |> prop_of |> Logic.dest_equals |> snd blanchet@43612: end blanchet@38831: blanchet@38506: fun introduce_combinators_in_term ctxt kind t = wenzelm@43232: let val thy = Proof_Context.theory_of ctxt in blanchet@38716: if Meson.is_fol_term thy t then blanchet@38716: t blanchet@38716: else blanchet@38716: let blanchet@38716: fun aux Ts t = blanchet@38716: case t of blanchet@38716: @{const Not} $ t1 => @{const Not} $ aux Ts t1 blanchet@38716: | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') => blanchet@38716: t0 $ Abs (s, T, aux (T :: Ts) t') blanchet@38890: | (t0 as Const (@{const_name All}, _)) $ t1 => blanchet@38890: aux Ts (t0 $ eta_expand Ts t1 1) blanchet@38716: | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') => blanchet@38716: t0 $ Abs (s, T, aux (T :: Ts) t') blanchet@38890: | (t0 as Const (@{const_name Ex}, _)) $ t1 => blanchet@38890: aux Ts (t0 $ eta_expand Ts t1 1) haftmann@39028: | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 haftmann@39028: | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 haftmann@39019: | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 haftmann@39093: | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _]))) blanchet@38716: $ t1 $ t2 => blanchet@38716: t0 $ aux Ts t1 $ aux Ts t2 blanchet@38716: | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then blanchet@38716: t blanchet@38716: else blanchet@38716: t |> conceal_bounds Ts blanchet@38716: |> Envir.eta_contract blanchet@38716: |> cterm_of thy blanchet@40071: |> Meson_Clausify.introduce_combinators_in_cterm blanchet@38716: |> prop_of |> Logic.dest_equals |> snd blanchet@38716: |> reveal_bounds Ts blanchet@39616: val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single blanchet@38716: in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end blanchet@38716: handle THM _ => blanchet@38716: (* A type variable of sort "{}" will make abstraction fail. *) blanchet@38836: if kind = Conjecture then HOLogic.false_const blanchet@38836: else HOLogic.true_const blanchet@38716: end blanchet@38506: blanchet@38506: (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the blanchet@43224: same in Sledgehammer to prevent the discovery of unreplayable proofs. *) blanchet@38506: fun freeze_term t = blanchet@38506: let blanchet@38506: fun aux (t $ u) = aux t $ aux u blanchet@38506: | aux (Abs (s, T, t)) = Abs (s, T, aux t) blanchet@38506: | aux (Var ((s, i), T)) = blanchet@38506: Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T) blanchet@38506: | aux t = t blanchet@38506: in t |> exists_subterm is_Var t ? aux end blanchet@38506: blanchet@40445: (* making fact and conjecture formulas *) blanchet@43511: fun make_formula ctxt eq_as_iff presimp name loc kind t = blanchet@38506: let wenzelm@43232: val thy = Proof_Context.theory_of ctxt blanchet@38831: val t = t |> Envir.beta_eta_contract blanchet@38890: |> transform_elim_term blanchet@41459: |> Object_Logic.atomize_term thy blanchet@43434: val need_trueprop = (fastype_of t = @{typ bool}) blanchet@38890: val t = t |> need_trueprop ? HOLogic.mk_Trueprop blanchet@43607: |> Raw_Simplifier.rewrite_term thy blanchet@43607: (Meson.unfold_set_const_simps ctxt) [] blanchet@43612: |> extensionalize_term ctxt blanchet@38506: |> presimp ? presimplify_term thy blanchet@38506: |> perhaps (try (HOLogic.dest_Trueprop)) blanchet@38506: |> introduce_combinators_in_term ctxt kind blanchet@38836: |> kind <> Axiom ? freeze_term blanchet@43433: val (combformula, atomic_types) = blanchet@43433: combformula_from_prop thy eq_as_iff t [] blanchet@38506: in blanchet@43511: {name = name, locality = loc, kind = kind, combformula = combformula, blanchet@43433: atomic_types = atomic_types} blanchet@38506: end blanchet@38506: blanchet@43511: fun make_fact ctxt keep_trivial eq_as_iff presimp ((name, loc), t) = blanchet@43511: case (keep_trivial, make_formula ctxt eq_as_iff presimp name loc Axiom t) of blanchet@42861: (false, {combformula = AAtom (CombConst (("c_True", _), _, _)), ...}) => blanchet@42861: NONE blanchet@42861: | (_, formula) => SOME formula blanchet@43432: blanchet@43580: fun make_conjecture ctxt prem_kind ts = blanchet@38836: let val last = length ts - 1 in blanchet@43580: map2 (fn j => fn t => blanchet@43580: let blanchet@43580: val (kind, maybe_negate) = blanchet@43580: if j = last then blanchet@43580: (Conjecture, I) blanchet@43580: else blanchet@43580: (prem_kind, blanchet@43580: if prem_kind = Conjecture then update_combformula mk_anot blanchet@43580: else I) blanchet@43580: in blanchet@43580: make_formula ctxt true true (string_of_int j) Chained kind t blanchet@43580: |> maybe_negate blanchet@43580: end) blanchet@38836: (0 upto last) ts blanchet@38836: end blanchet@38506: blanchet@43552: (** Finite and infinite type inference **) blanchet@43552: blanchet@43552: (* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are blanchet@43552: dangerous because their "exhaust" properties can easily lead to unsound ATP blanchet@43552: proofs. On the other hand, all HOL infinite types can be given the same blanchet@43552: models in first-order logic (via Löwenheim-Skolem). *) blanchet@43552: blanchet@43552: fun should_encode_type _ _ All_Types _ = true blanchet@43552: | should_encode_type ctxt _ Finite_Types T = is_type_surely_finite ctxt T blanchet@43552: | should_encode_type _ nonmono_Ts Nonmonotonic_Types T = blanchet@43552: exists (curry Type.raw_instance T) nonmono_Ts blanchet@43552: | should_encode_type _ _ _ _ = false blanchet@43552: blanchet@43552: fun should_predicate_on_type ctxt nonmono_Ts (Preds (_, level)) T = blanchet@43552: should_encode_type ctxt nonmono_Ts level T blanchet@43552: | should_predicate_on_type _ _ _ _ = false blanchet@43552: blanchet@43552: fun should_tag_with_type ctxt nonmono_Ts (Tags (_, level)) T = blanchet@43552: should_encode_type ctxt nonmono_Ts level T blanchet@43552: | should_tag_with_type _ _ _ _ = false blanchet@43552: blanchet@43552: val homo_infinite_T = @{typ ind} (* any infinite type *) blanchet@43552: blanchet@43552: fun homogenized_type ctxt nonmono_Ts level T = blanchet@43552: if should_encode_type ctxt nonmono_Ts level T then T else homo_infinite_T blanchet@43552: blanchet@43444: (** "hBOOL" and "hAPP" **) blanchet@41561: blanchet@43445: type sym_info = blanchet@43434: {pred_sym : bool, min_ary : int, max_ary : int, typ : typ option} blanchet@43434: blanchet@43445: fun add_combterm_syms_to_table explicit_apply = blanchet@43429: let blanchet@43429: fun aux top_level tm = blanchet@43429: let val (head, args) = strip_combterm_comb tm in blanchet@43429: (case head of blanchet@43434: CombConst ((s, _), T, _) => blanchet@43429: if String.isPrefix bound_var_prefix s then blanchet@43429: I blanchet@43429: else blanchet@43434: let val ary = length args in blanchet@43429: Symtab.map_default blanchet@43429: (s, {pred_sym = true, blanchet@43434: min_ary = if explicit_apply then 0 else ary, blanchet@43434: max_ary = 0, typ = SOME T}) blanchet@43434: (fn {pred_sym, min_ary, max_ary, typ} => blanchet@43429: {pred_sym = pred_sym andalso top_level, blanchet@43434: min_ary = Int.min (ary, min_ary), blanchet@43434: max_ary = Int.max (ary, max_ary), blanchet@43434: typ = if typ = SOME T then typ else NONE}) blanchet@43429: end blanchet@43429: | _ => I) blanchet@43429: #> fold (aux false) args blanchet@43429: end blanchet@43429: in aux true end blanchet@43545: fun add_fact_syms_to_table explicit_apply = blanchet@43550: fact_lift (formula_fold true (K (add_combterm_syms_to_table explicit_apply))) blanchet@38506: blanchet@43546: val default_sym_table_entries : (string * sym_info) list = blanchet@43434: [("equal", {pred_sym = true, min_ary = 2, max_ary = 2, typ = NONE}), blanchet@43439: (make_fixed_const predicator_base, blanchet@43434: {pred_sym = true, min_ary = 1, max_ary = 1, typ = NONE})] @ blanchet@43439: ([tptp_false, tptp_true] blanchet@43434: |> map (rpair {pred_sym = true, min_ary = 0, max_ary = 0, typ = NONE})) blanchet@41388: blanchet@43415: fun sym_table_for_facts explicit_apply facts = blanchet@43439: Symtab.empty |> fold Symtab.default default_sym_table_entries blanchet@43445: |> fold (add_fact_syms_to_table explicit_apply) facts blanchet@38506: blanchet@43429: fun min_arity_of sym_tab s = blanchet@43429: case Symtab.lookup sym_tab s of blanchet@43445: SOME ({min_ary, ...} : sym_info) => min_ary blanchet@43429: | NONE => blanchet@43429: case strip_prefix_and_unascii const_prefix s of blanchet@43418: SOME s => blanchet@43441: let val s = s |> unmangled_const_name |> invert_const in blanchet@43439: if s = predicator_base then 1 blanchet@43418: else if s = explicit_app_base then 2 blanchet@43418: else if s = type_pred_base then 1 blanchet@43428: else 0 blanchet@43418: end blanchet@38506: | NONE => 0 blanchet@38506: blanchet@38506: (* True if the constant ever appears outside of the top-level position in blanchet@38506: literals, or if it appears with different arities (e.g., because of different blanchet@38506: type instantiations). If false, the constant always receives all of its blanchet@38506: arguments and is used as a predicate. *) blanchet@43429: fun is_pred_sym sym_tab s = blanchet@43429: case Symtab.lookup sym_tab s of blanchet@43445: SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) => blanchet@43445: pred_sym andalso min_ary = max_ary blanchet@43429: | NONE => false blanchet@38506: blanchet@43439: val predicator_combconst = blanchet@43439: CombConst (`make_fixed_const predicator_base, @{typ "bool => bool"}, []) blanchet@43439: fun predicator tm = CombApp (predicator_combconst, tm) blanchet@38506: blanchet@43439: fun introduce_predicators_in_combterm sym_tab tm = blanchet@43413: case strip_combterm_comb tm of blanchet@43413: (CombConst ((s, _), _, _), _) => blanchet@43439: if is_pred_sym sym_tab s then tm else predicator tm blanchet@43439: | _ => predicator tm blanchet@38506: blanchet@43415: fun list_app head args = fold (curry (CombApp o swap)) args head blanchet@38506: blanchet@43415: fun explicit_app arg head = blanchet@43415: let blanchet@43433: val head_T = combtyp_of head blanchet@43563: val (arg_T, res_T) = dest_funT head_T blanchet@43415: val explicit_app = blanchet@43433: CombConst (`make_fixed_const explicit_app_base, head_T --> head_T, blanchet@43563: [arg_T, res_T]) blanchet@43415: in list_app explicit_app [head, arg] end blanchet@43415: fun list_explicit_app head args = fold explicit_app args head blanchet@43415: blanchet@43436: fun introduce_explicit_apps_in_combterm sym_tab = blanchet@43415: let blanchet@43415: fun aux tm = blanchet@43415: case strip_combterm_comb tm of blanchet@43415: (head as CombConst ((s, _), _, _), args) => blanchet@43415: args |> map aux blanchet@43428: |> chop (min_arity_of sym_tab s) blanchet@43415: |>> list_app head blanchet@43415: |-> list_explicit_app blanchet@43415: | (head, args) => list_explicit_app head (map aux args) blanchet@43415: in aux end blanchet@43415: blanchet@43571: fun impose_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys = blanchet@43444: let blanchet@43444: fun aux (CombApp tmp) = CombApp (pairself aux tmp) blanchet@43445: | aux (CombConst (name as (s, _), T, T_args)) = blanchet@43571: let blanchet@43571: val level = level_of_type_sys type_sys blanchet@43571: val (T, T_args) = blanchet@43571: (* Aggressively merge most "hAPPs" if the type system is unsound blanchet@43571: anyway, by distinguishing overloads only on the homogenized blanchet@43571: result type. *) blanchet@43571: if s = const_prefix ^ explicit_app_base andalso blanchet@43591: length T_args = 2 andalso blanchet@43571: not (is_type_sys_virtually_sound type_sys) then blanchet@43571: T_args |> map (homogenized_type ctxt nonmono_Ts level) blanchet@43571: |> (fn Ts => let val T = hd Ts --> nth Ts 1 in blanchet@43571: (T --> T, tl Ts) blanchet@43571: end) blanchet@43571: else blanchet@43571: (T, T_args) blanchet@43571: in blanchet@43571: (case strip_prefix_and_unascii const_prefix s of blanchet@43571: NONE => (name, T_args) blanchet@43571: | SOME s'' => blanchet@43571: let val s'' = invert_const s'' in blanchet@43571: case type_arg_policy type_sys s'' of blanchet@43571: No_Type_Args => (name, []) blanchet@43571: | Explicit_Type_Args => (name, T_args) blanchet@43571: | Mangled_Type_Args => (mangled_const_name T_args name, []) blanchet@43571: end) blanchet@43571: |> (fn (name, T_args) => CombConst (name, T, T_args)) blanchet@43571: end blanchet@43444: | aux tm = tm blanchet@43444: in aux end blanchet@43444: blanchet@43571: fun repair_combterm ctxt nonmono_Ts type_sys sym_tab = blanchet@43436: introduce_explicit_apps_in_combterm sym_tab blanchet@43439: #> introduce_predicators_in_combterm sym_tab blanchet@43571: #> impose_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys blanchet@43571: fun repair_fact ctxt nonmono_Ts type_sys sym_tab = blanchet@43571: update_combformula (formula_map blanchet@43571: (repair_combterm ctxt nonmono_Ts type_sys sym_tab)) blanchet@43444: blanchet@43444: (** Helper facts **) blanchet@43444: blanchet@43444: fun ti_ti_helper_fact () = blanchet@43444: let blanchet@43444: fun var s = ATerm (`I s, []) blanchet@43460: fun tag tm = ATerm (`make_fixed_const type_tag_name, [var "X", tm]) blanchet@43444: in blanchet@43483: Formula (helper_prefix ^ "ti_ti", Axiom, blanchet@43444: AAtom (ATerm (`I "equal", [tag (tag (var "Y")), tag (var "Y")])) blanchet@43444: |> close_formula_universally, NONE, NONE) blanchet@43444: end blanchet@43444: blanchet@43445: fun helper_facts_for_sym ctxt type_sys (s, {typ, ...} : sym_info) = blanchet@43444: case strip_prefix_and_unascii const_prefix s of blanchet@43444: SOME mangled_s => blanchet@43444: let blanchet@43444: val thy = Proof_Context.theory_of ctxt blanchet@43444: val unmangled_s = mangled_s |> unmangled_const_name blanchet@43450: fun dub_and_inst c needs_some_types (th, j) = blanchet@43450: ((c ^ "_" ^ string_of_int j ^ (if needs_some_types then "T" else ""), blanchet@43511: Chained), blanchet@43444: let val t = th |> prop_of in blanchet@43460: t |> (general_type_arg_policy type_sys = Mangled_Type_Args andalso blanchet@43444: not (null (Term.hidden_polymorphism t))) blanchet@43444: ? (case typ of blanchet@43444: SOME T => specialize_type thy (invert_const unmangled_s, T) blanchet@43444: | NONE => I) blanchet@43444: end) blanchet@43444: fun make_facts eq_as_iff = blanchet@43444: map_filter (make_fact ctxt false eq_as_iff false) blanchet@43460: val has_some_types = is_type_sys_fairly_sound type_sys blanchet@43444: in blanchet@43444: metis_helpers blanchet@43450: |> maps (fn (metis_s, (needs_some_types, ths)) => blanchet@43444: if metis_s <> unmangled_s orelse blanchet@43460: (needs_some_types andalso not has_some_types) then blanchet@43444: [] blanchet@43444: else blanchet@43444: ths ~~ (1 upto length ths) blanchet@43450: |> map (dub_and_inst mangled_s needs_some_types) blanchet@43450: |> make_facts (not needs_some_types)) blanchet@43444: end blanchet@43444: | NONE => [] blanchet@43444: fun helper_facts_for_sym_table ctxt type_sys sym_tab = blanchet@43444: Symtab.fold_rev (append o helper_facts_for_sym ctxt type_sys) sym_tab [] blanchet@43444: blanchet@43444: fun translate_atp_fact ctxt keep_trivial = blanchet@43444: `(make_fact ctxt keep_trivial true true o apsnd prop_of) blanchet@43444: blanchet@43580: fun translate_formulas ctxt prem_kind type_sys hyp_ts concl_t rich_facts = blanchet@43444: let blanchet@43444: val thy = Proof_Context.theory_of ctxt blanchet@43444: val fact_ts = map (prop_of o snd o snd) rich_facts blanchet@43444: val (facts, fact_names) = blanchet@43444: rich_facts blanchet@43444: |> map_filter (fn (NONE, _) => NONE blanchet@43444: | (SOME fact, (name, _)) => SOME (fact, name)) blanchet@43444: |> ListPair.unzip blanchet@43444: (* Remove existing facts from the conjecture, as this can dramatically blanchet@43444: boost an ATP's performance (for some reason). *) blanchet@43444: val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts) blanchet@43444: val goal_t = Logic.list_implies (hyp_ts, concl_t) blanchet@43444: val all_ts = goal_t :: fact_ts blanchet@43444: val subs = tfree_classes_of_terms all_ts blanchet@43444: val supers = tvar_classes_of_terms all_ts blanchet@43444: val tycons = type_consts_of_terms thy all_ts blanchet@43580: val conjs = make_conjecture ctxt prem_kind (hyp_ts @ [concl_t]) blanchet@43444: val (supers', arity_clauses) = blanchet@43460: if level_of_type_sys type_sys = No_Types then ([], []) blanchet@43444: else make_arity_clauses thy tycons supers blanchet@43444: val class_rel_clauses = make_class_rel_clauses thy subs supers' blanchet@43444: in blanchet@43444: (fact_names |> map single, (conjs, facts, class_rel_clauses, arity_clauses)) blanchet@43444: end blanchet@43444: blanchet@43444: fun fo_literal_from_type_literal (TyLitVar (class, name)) = blanchet@43444: (true, ATerm (class, [ATerm (name, [])])) blanchet@43444: | fo_literal_from_type_literal (TyLitFree (class, name)) = blanchet@43444: (true, ATerm (class, [ATerm (name, [])])) blanchet@43444: blanchet@43444: fun formula_from_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot blanchet@43444: blanchet@43571: fun type_pred_combatom ctxt nonmono_Ts type_sys T tm = blanchet@43444: CombApp (CombConst (`make_fixed_const type_pred_base, T --> @{typ bool}, [T]), blanchet@43444: tm) blanchet@43571: |> impose_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys blanchet@43444: |> AAtom blanchet@43444: blanchet@43550: fun formula_from_combformula ctxt nonmono_Ts type_sys = blanchet@43444: let blanchet@43460: fun tag_with_type type_sys T tm = blanchet@43460: CombConst (`make_fixed_const type_tag_name, T --> T, [T]) blanchet@43571: |> impose_type_arg_policy_in_combterm ctxt nonmono_Ts type_sys blanchet@43460: |> do_term true blanchet@43460: |> (fn ATerm (s, tms) => ATerm (s, tms @ [tm])) blanchet@43460: and do_term top_level u = blanchet@43444: let blanchet@43444: val (head, args) = strip_combterm_comb u blanchet@43445: val (x, T_args) = blanchet@43444: case head of blanchet@43445: CombConst (name, _, T_args) => (name, T_args) blanchet@43444: | CombVar (name, _) => (name, []) blanchet@43444: | CombApp _ => raise Fail "impossible \"CombApp\"" blanchet@43445: val t = ATerm (x, map fo_term_from_typ T_args @ blanchet@43444: map (do_term false) args) blanchet@43445: val T = combtyp_of u blanchet@43444: in blanchet@43550: t |> (if not top_level andalso blanchet@43550: should_tag_with_type ctxt nonmono_Ts type_sys T then blanchet@43460: tag_with_type type_sys T blanchet@43444: else blanchet@43444: I) blanchet@43444: end blanchet@43444: val do_bound_type = blanchet@43552: case type_sys of blanchet@43587: Simple_Types level => blanchet@43552: SOME o mangled_type_name o homogenized_type ctxt nonmono_Ts level blanchet@43552: | _ => K NONE blanchet@43444: fun do_out_of_bound_type (s, T) = blanchet@43550: if should_predicate_on_type ctxt nonmono_Ts type_sys T then blanchet@43571: type_pred_combatom ctxt nonmono_Ts type_sys T (CombVar (s, T)) blanchet@43444: |> do_formula |> SOME blanchet@43444: else blanchet@43444: NONE blanchet@43444: and do_formula (AQuant (q, xs, phi)) = blanchet@43444: AQuant (q, xs |> map (apsnd (fn NONE => NONE blanchet@43445: | SOME T => do_bound_type T)), blanchet@43444: (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd)) blanchet@43444: (map_filter blanchet@43444: (fn (_, NONE) => NONE blanchet@43445: | (s, SOME T) => do_out_of_bound_type (s, T)) xs) blanchet@43444: (do_formula phi)) blanchet@43444: | do_formula (AConn (c, phis)) = AConn (c, map do_formula phis) blanchet@43444: | do_formula (AAtom tm) = AAtom (do_term true tm) blanchet@43444: in do_formula end blanchet@43444: blanchet@43592: fun bound_atomic_types type_sys Ts = blanchet@43592: mk_ahorn (map (formula_from_fo_literal o fo_literal_from_type_literal) blanchet@43592: (atp_type_literals_for_types type_sys Axiom Ts)) blanchet@43592: blanchet@43550: fun formula_for_fact ctxt nonmono_Ts type_sys blanchet@43444: ({combformula, atomic_types, ...} : translated_formula) = blanchet@43592: combformula blanchet@43592: |> close_combformula_universally blanchet@43592: |> formula_from_combformula ctxt nonmono_Ts type_sys blanchet@43592: |> bound_atomic_types type_sys atomic_types blanchet@43444: |> close_formula_universally blanchet@43444: blanchet@43511: fun useful_isabelle_info s = SOME (ATerm ("[]", [ATerm ("isabelle_" ^ s, [])])) blanchet@43511: blanchet@43444: (* Each fact is given a unique fact number to avoid name clashes (e.g., because blanchet@43444: of monomorphization). The TPTP explicitly forbids name clashes, and some of blanchet@43444: the remote provers might care. *) blanchet@43550: fun formula_line_for_fact ctxt prefix nonmono_Ts type_sys blanchet@43511: (j, formula as {name, locality, kind, ...}) = blanchet@43550: Formula (prefix ^ (if polymorphism_of_type_sys type_sys = Polymorphic then "" blanchet@43550: else string_of_int j ^ "_") ^ blanchet@43518: ascii_of name, blanchet@43550: kind, formula_for_fact ctxt nonmono_Ts type_sys formula, NONE, blanchet@43511: if generate_useful_info then blanchet@43511: case locality of blanchet@43511: Intro => useful_isabelle_info "intro" blanchet@43511: | Elim => useful_isabelle_info "elim" blanchet@43511: | Simp => useful_isabelle_info "simp" blanchet@43511: | _ => NONE blanchet@43511: else blanchet@43511: NONE) blanchet@43444: blanchet@43444: fun formula_line_for_class_rel_clause (ClassRelClause {name, subclass, blanchet@43444: superclass, ...}) = blanchet@43444: let val ty_arg = ATerm (`I "T", []) in blanchet@43448: Formula (class_rel_clause_prefix ^ ascii_of name, Axiom, blanchet@43444: AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])), blanchet@43444: AAtom (ATerm (superclass, [ty_arg]))]) blanchet@43444: |> close_formula_universally, NONE, NONE) blanchet@43444: end blanchet@43444: blanchet@43444: fun fo_literal_from_arity_literal (TConsLit (c, t, args)) = blanchet@43444: (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)])) blanchet@43444: | fo_literal_from_arity_literal (TVarLit (c, sort)) = blanchet@43444: (false, ATerm (c, [ATerm (sort, [])])) blanchet@43444: blanchet@43444: fun formula_line_for_arity_clause (ArityClause {name, conclLit, premLits, blanchet@43444: ...}) = blanchet@43448: Formula (arity_clause_prefix ^ ascii_of name, Axiom, blanchet@43444: mk_ahorn (map (formula_from_fo_literal o apfst not blanchet@43444: o fo_literal_from_arity_literal) premLits) blanchet@43444: (formula_from_fo_literal blanchet@43444: (fo_literal_from_arity_literal conclLit)) blanchet@43444: |> close_formula_universally, NONE, NONE) blanchet@43444: blanchet@43550: fun formula_line_for_conjecture ctxt nonmono_Ts type_sys blanchet@43444: ({name, kind, combformula, ...} : translated_formula) = blanchet@43448: Formula (conjecture_prefix ^ name, kind, blanchet@43550: formula_from_combformula ctxt nonmono_Ts type_sys blanchet@43444: (close_combformula_universally combformula) blanchet@43444: |> close_formula_universally, NONE, NONE) blanchet@43444: blanchet@43444: fun free_type_literals type_sys ({atomic_types, ...} : translated_formula) = blanchet@43444: atomic_types |> atp_type_literals_for_types type_sys Conjecture blanchet@43444: |> map fo_literal_from_type_literal blanchet@43444: blanchet@43444: fun formula_line_for_free_type j lit = blanchet@43448: Formula (tfree_prefix ^ string_of_int j, Hypothesis, blanchet@43444: formula_from_fo_literal lit, NONE, NONE) blanchet@43444: fun formula_lines_for_free_types type_sys facts = blanchet@43444: let blanchet@43444: val litss = map (free_type_literals type_sys) facts blanchet@43444: val lits = fold (union (op =)) litss [] blanchet@43444: in map2 formula_line_for_free_type (0 upto length lits - 1) lits end blanchet@43444: blanchet@43444: (** Symbol declarations **) blanchet@43415: blanchet@43547: fun insert_type get_T x xs = blanchet@43547: let val T = get_T x in blanchet@43547: if exists (curry Type.raw_instance T o get_T) xs then xs blanchet@43547: else x :: filter_out ((fn T' => Type.raw_instance (T', T)) o get_T) xs blanchet@43547: end blanchet@43547: blanchet@43445: fun should_declare_sym type_sys pred_sym s = blanchet@43413: not (String.isPrefix bound_var_prefix s) andalso s <> "equal" andalso blanchet@43516: not (String.isPrefix "$" s) andalso blanchet@43587: ((case type_sys of Simple_Types _ => true | _ => false) orelse not pred_sym) blanchet@43413: blanchet@43568: fun sym_decl_table_for_facts type_sys repaired_sym_tab (conjs, facts) = blanchet@43445: let blanchet@43568: fun add_combterm in_conj tm = blanchet@43445: let val (head, args) = strip_combterm_comb tm in blanchet@43445: (case head of blanchet@43445: CombConst ((s, s'), T, T_args) => blanchet@43445: let val pred_sym = is_pred_sym repaired_sym_tab s in blanchet@43445: if should_declare_sym type_sys pred_sym s then blanchet@43447: Symtab.map_default (s, []) blanchet@43568: (insert_type #3 (s', T_args, T, pred_sym, length args, blanchet@43568: in_conj)) blanchet@43445: else blanchet@43445: I blanchet@43445: end blanchet@43445: | _ => I) blanchet@43568: #> fold (add_combterm in_conj) args blanchet@43445: end blanchet@43568: fun add_fact in_conj = blanchet@43568: fact_lift (formula_fold true (K (add_combterm in_conj))) blanchet@43568: in blanchet@43568: Symtab.empty blanchet@43568: |> is_type_sys_fairly_sound type_sys blanchet@43568: ? (fold (add_fact true) conjs #> fold (add_fact false) facts) blanchet@43568: end blanchet@43445: blanchet@43547: fun is_var_or_bound_var (CombConst ((s, _), _, _)) = blanchet@43547: String.isPrefix bound_var_prefix s blanchet@43547: | is_var_or_bound_var (CombVar _) = true blanchet@43547: | is_var_or_bound_var _ = false blanchet@43547: blanchet@43555: (* This inference is described in section 2.3 of Claessen et al.'s "Sorting it blanchet@43555: out with monotonicity" paper presented at CADE 2011. *) blanchet@43550: fun add_combterm_nonmonotonic_types _ (SOME false) _ = I blanchet@43550: | add_combterm_nonmonotonic_types ctxt _ blanchet@43550: (CombApp (CombApp (CombConst (("equal", _), Type (_, [T, _]), _), tm1), blanchet@43550: tm2)) = blanchet@43550: (exists is_var_or_bound_var [tm1, tm2] andalso blanchet@43550: not (is_type_surely_infinite ctxt T)) ? insert_type I T blanchet@43550: | add_combterm_nonmonotonic_types _ _ _ = I blanchet@43550: fun add_fact_nonmonotonic_types ctxt ({kind, combformula, ...} blanchet@43550: : translated_formula) = blanchet@43550: formula_fold (kind <> Conjecture) (add_combterm_nonmonotonic_types ctxt) blanchet@43550: combformula blanchet@43550: fun add_nonmonotonic_types_for_facts ctxt type_sys facts = blanchet@43550: level_of_type_sys type_sys = Nonmonotonic_Types blanchet@43596: ? (fold (add_fact_nonmonotonic_types ctxt) facts blanchet@43596: (* in case helper "True_or_False" is included *) blanchet@43596: #> insert_type I @{typ bool}) blanchet@43547: blanchet@43445: fun n_ary_strip_type 0 T = ([], T) blanchet@43445: | n_ary_strip_type n (Type (@{type_name fun}, [dom_T, ran_T])) = blanchet@43445: n_ary_strip_type (n - 1) ran_T |>> cons dom_T blanchet@43445: | n_ary_strip_type _ _ = raise Fail "unexpected non-function" blanchet@43445: blanchet@43568: fun result_type_of_decl (_, _, T, _, ary, _) = n_ary_strip_type ary T |> snd blanchet@43450: blanchet@43568: fun decl_line_for_sym s (s', _, T, pred_sym, ary, _) = blanchet@43450: let val (arg_Ts, res_T) = n_ary_strip_type ary T in blanchet@43483: Decl (sym_decl_prefix ^ s, (s, s'), map mangled_type_name arg_Ts, blanchet@43450: if pred_sym then `I tptp_tff_bool_type else mangled_type_name res_T) blanchet@43450: end blanchet@43450: blanchet@43463: fun is_polymorphic_type T = fold_atyps (fn TVar _ => K true | _ => I) T false blanchet@43463: blanchet@43580: fun formula_line_for_sym_decl ctxt conj_sym_kind nonmono_Ts type_sys n s j blanchet@43568: (s', T_args, T, _, ary, in_conj) = blanchet@43450: let blanchet@43580: val (kind, maybe_negate) = blanchet@43580: if in_conj then (conj_sym_kind, conj_sym_kind = Conjecture ? mk_anot) blanchet@43580: else (Axiom, I) blanchet@43450: val (arg_Ts, res_T) = n_ary_strip_type ary T blanchet@43450: val bound_names = blanchet@43450: 1 upto length arg_Ts |> map (`I o make_bound_var o string_of_int) blanchet@43450: val bound_tms = blanchet@43450: bound_names ~~ arg_Ts |> map (fn (name, T) => CombConst (name, T, [])) blanchet@43450: val bound_Ts = blanchet@43463: arg_Ts |> map (fn T => if n > 1 orelse is_polymorphic_type T then SOME T blanchet@43463: else NONE) blanchet@43450: in blanchet@43483: Formula (sym_decl_prefix ^ s ^ blanchet@43580: (if n > 1 then "_" ^ string_of_int j else ""), kind, blanchet@43450: CombConst ((s, s'), T, T_args) blanchet@43450: |> fold (curry (CombApp o swap)) bound_tms blanchet@43571: |> type_pred_combatom ctxt nonmono_Ts type_sys res_T blanchet@43450: |> mk_aquant AForall (bound_names ~~ bound_Ts) blanchet@43550: |> formula_from_combformula ctxt nonmono_Ts type_sys blanchet@43592: |> n > 1 ? bound_atomic_types type_sys (atyps_of T) blanchet@43580: |> close_formula_universally blanchet@43580: |> maybe_negate, blanchet@43450: NONE, NONE) blanchet@43450: end blanchet@43450: blanchet@43580: fun problem_lines_for_sym_decls ctxt conj_sym_kind nonmono_Ts type_sys blanchet@43580: (s, decls) = blanchet@43552: case type_sys of blanchet@43587: Simple_Types _ => map (decl_line_for_sym s) decls blanchet@43552: | _ => blanchet@43445: let blanchet@43450: val decls = blanchet@43450: case decls of blanchet@43450: decl :: (decls' as _ :: _) => blanchet@43463: let val T = result_type_of_decl decl in blanchet@43463: if forall ((fn T' => Type.raw_instance (T', T)) blanchet@43463: o result_type_of_decl) decls' then blanchet@43463: [decl] blanchet@43463: else blanchet@43463: decls blanchet@43463: end blanchet@43450: | _ => decls blanchet@43450: val n = length decls blanchet@43450: val decls = blanchet@43550: decls |> filter (should_predicate_on_type ctxt nonmono_Ts type_sys blanchet@43450: o result_type_of_decl) blanchet@43445: in blanchet@43580: (0 upto length decls - 1, decls) blanchet@43580: |-> map2 (formula_line_for_sym_decl ctxt conj_sym_kind nonmono_Ts type_sys blanchet@43580: n s) blanchet@43445: end blanchet@43450: blanchet@43580: fun problem_lines_for_sym_decl_table ctxt conj_sym_kind nonmono_Ts type_sys blanchet@43580: sym_decl_tab = blanchet@43580: Symtab.fold_rev (append o problem_lines_for_sym_decls ctxt conj_sym_kind blanchet@43580: nonmono_Ts type_sys) blanchet@43445: sym_decl_tab [] blanchet@43410: blanchet@43414: fun add_tff_types_in_formula (AQuant (_, xs, phi)) = blanchet@43414: union (op =) (map_filter snd xs) #> add_tff_types_in_formula phi blanchet@43414: | add_tff_types_in_formula (AConn (_, phis)) = blanchet@43414: fold add_tff_types_in_formula phis blanchet@43414: | add_tff_types_in_formula (AAtom _) = I blanchet@43414: blanchet@43433: fun add_tff_types_in_problem_line (Decl (_, _, arg_Ts, res_T)) = blanchet@43433: union (op =) (res_T :: arg_Ts) blanchet@43448: | add_tff_types_in_problem_line (Formula (_, _, phi, _, _)) = blanchet@43414: add_tff_types_in_formula phi blanchet@43414: blanchet@43414: fun tff_types_in_problem problem = blanchet@43414: fold (fold add_tff_types_in_problem_line o snd) problem [] blanchet@43414: blanchet@43416: fun decl_line_for_tff_type (s, s') = blanchet@43439: Decl (type_decl_prefix ^ ascii_of s, (s, s'), [], `I tptp_tff_type_of_types) blanchet@43414: blanchet@43414: val type_declsN = "Types" blanchet@43415: val sym_declsN = "Symbol types" blanchet@41405: val factsN = "Relevant facts" blanchet@41405: val class_relsN = "Class relationships" blanchet@43414: val aritiesN = "Arities" blanchet@41405: val helpersN = "Helper facts" blanchet@41405: val conjsN = "Conjectures" blanchet@41561: val free_typesN = "Type variables" blanchet@41405: blanchet@41405: fun offset_of_heading_in_problem _ [] j = j blanchet@41405: | offset_of_heading_in_problem needle ((heading, lines) :: problem) j = blanchet@41405: if heading = needle then j blanchet@41405: else offset_of_heading_in_problem needle problem (j + length lines) blanchet@41405: blanchet@43580: fun prepare_atp_problem ctxt conj_sym_kind prem_kind type_sys explicit_apply blanchet@43580: hyp_ts concl_t facts = blanchet@38506: let blanchet@41561: val (fact_names, (conjs, facts, class_rel_clauses, arity_clauses)) = blanchet@43580: translate_formulas ctxt prem_kind type_sys hyp_ts concl_t facts blanchet@43434: val sym_tab = conjs @ facts |> sym_table_for_facts explicit_apply blanchet@43552: val nonmono_Ts = blanchet@43552: [] |> fold (add_nonmonotonic_types_for_facts ctxt type_sys) [facts, conjs] blanchet@43571: val repair = repair_fact ctxt nonmono_Ts type_sys sym_tab blanchet@43552: val (conjs, facts) = (conjs, facts) |> pairself (map repair) blanchet@43550: val repaired_sym_tab = conjs @ facts |> sym_table_for_facts false blanchet@43444: val helpers = blanchet@43552: repaired_sym_tab |> helper_facts_for_sym_table ctxt type_sys |> map repair blanchet@43550: val sym_decl_lines = blanchet@43596: (conjs, helpers @ facts) blanchet@43550: |> sym_decl_table_for_facts type_sys repaired_sym_tab blanchet@43580: |> problem_lines_for_sym_decl_table ctxt conj_sym_kind nonmono_Ts type_sys blanchet@43393: (* Reordering these might confuse the proof reconstruction code or the SPASS blanchet@43393: Flotter hack. *) blanchet@38506: val problem = blanchet@43432: [(sym_declsN, sym_decl_lines), blanchet@43550: (factsN, map (formula_line_for_fact ctxt fact_prefix nonmono_Ts type_sys) blanchet@43051: (0 upto length facts - 1 ~~ facts)), blanchet@43416: (class_relsN, map formula_line_for_class_rel_clause class_rel_clauses), blanchet@43416: (aritiesN, map formula_line_for_arity_clause arity_clauses), blanchet@43550: (helpersN, map (formula_line_for_fact ctxt helper_prefix nonmono_Ts blanchet@43550: type_sys) blanchet@43434: (0 upto length helpers - 1 ~~ helpers) blanchet@43450: |> (case type_sys of blanchet@43460: Tags (Polymorphic, level) => blanchet@43557: is_type_level_partial level blanchet@43460: ? cons (ti_ti_helper_fact ()) blanchet@43450: | _ => I)), blanchet@43550: (conjsN, map (formula_line_for_conjecture ctxt nonmono_Ts type_sys) blanchet@43550: conjs), blanchet@43416: (free_typesN, formula_lines_for_free_types type_sys (facts @ conjs))] blanchet@43414: val problem = blanchet@43432: problem blanchet@43552: |> (case type_sys of blanchet@43587: Simple_Types _ => blanchet@43432: cons (type_declsN, blanchet@43432: map decl_line_for_tff_type (tff_types_in_problem problem)) blanchet@43552: | _ => I) blanchet@43517: val (problem, pool) = blanchet@43517: problem |> nice_atp_problem (Config.get ctxt readable_names) blanchet@38506: in blanchet@38506: (problem, blanchet@38506: case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty, blanchet@43456: offset_of_heading_in_problem conjsN problem 0, blanchet@43412: offset_of_heading_in_problem factsN problem 0, blanchet@41405: fact_names |> Vector.fromList) blanchet@38506: end blanchet@38506: blanchet@41561: (* FUDGE *) blanchet@41561: val conj_weight = 0.0 blanchet@42641: val hyp_weight = 0.1 blanchet@42641: val fact_min_weight = 0.2 blanchet@41561: val fact_max_weight = 1.0 blanchet@43479: val type_info_default_weight = 0.8 blanchet@41561: blanchet@41561: fun add_term_weights weight (ATerm (s, tms)) = blanchet@43599: (not (is_atp_variable s) andalso s <> "equal" andalso blanchet@43599: not (String.isPrefix "$" s)) ? Symtab.default (s, weight) blanchet@41561: #> fold (add_term_weights weight) tms blanchet@43448: fun add_problem_line_weights weight (Formula (_, _, phi, _, _)) = blanchet@43550: formula_fold true (K (add_term_weights weight)) phi blanchet@43399: | add_problem_line_weights _ _ = I blanchet@41561: blanchet@41561: fun add_conjectures_weights [] = I blanchet@41561: | add_conjectures_weights conjs = blanchet@41561: let val (hyps, conj) = split_last conjs in blanchet@41561: add_problem_line_weights conj_weight conj blanchet@41561: #> fold (add_problem_line_weights hyp_weight) hyps blanchet@41561: end blanchet@41561: blanchet@41561: fun add_facts_weights facts = blanchet@41561: let blanchet@41561: val num_facts = length facts blanchet@41561: fun weight_of j = blanchet@41561: fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j blanchet@41561: / Real.fromInt num_facts blanchet@41561: in blanchet@41561: map weight_of (0 upto num_facts - 1) ~~ facts blanchet@41561: |> fold (uncurry add_problem_line_weights) blanchet@41561: end blanchet@41561: blanchet@41561: (* Weights are from 0.0 (most important) to 1.0 (least important). *) blanchet@41561: fun atp_problem_weights problem = blanchet@43479: let val get = these o AList.lookup (op =) problem in blanchet@43479: Symtab.empty blanchet@43479: |> add_conjectures_weights (get free_typesN @ get conjsN) blanchet@43479: |> add_facts_weights (get factsN) blanchet@43479: |> fold (fold (add_problem_line_weights type_info_default_weight) o get) blanchet@43479: [sym_declsN, class_relsN, aritiesN] blanchet@43479: |> Symtab.dest blanchet@43479: |> sort (prod_ord Real.compare string_ord o pairself swap) blanchet@43479: end blanchet@41561: blanchet@38506: end;