diff --git a/src/HOL/Tools/ATP/atp_problem_generate.ML b/src/HOL/Tools/ATP/atp_problem_generate.ML --- a/src/HOL/Tools/ATP/atp_problem_generate.ML +++ b/src/HOL/Tools/ATP/atp_problem_generate.ML @@ -1,2929 +1,2930 @@ (* Title: HOL/Tools/ATP/atp_problem_generate.ML Author: Fabian Immler, TU Muenchen Author: Makarius Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, UniBw Muenchen Translation of HOL to FOL for Metis and Sledgehammer. *) signature ATP_PROBLEM_GENERATE = sig type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term type atp_connective = ATP_Problem.atp_connective type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula type atp_format = ATP_Problem.atp_format type atp_formula_role = ATP_Problem.atp_formula_role type 'a atp_problem = 'a ATP_Problem.atp_problem datatype mode = Metis | Sledgehammer | Sledgehammer_Completish of int | Exporter | Translator datatype scope = Global | Local | Assum | Chained datatype status = General | Induction | Intro | Inductive | Elim | Simp | Non_Rec_Def | Rec_Def type stature = scope * status datatype strictness = Strict | Non_Strict datatype uniformity = Uniform | Non_Uniform datatype ctr_optim = With_Ctr_Optim | Without_Ctr_Optim datatype type_level = All_Types | Undercover_Types | Nonmono_Types of strictness * uniformity | Const_Types of ctr_optim | No_Types type type_enc val no_lamsN : string val opaque_liftingN : string val liftingN : string val opaque_combsN : string val combsN : string val combs_and_liftingN : string val combs_or_liftingN : string val keep_lamsN : string val schematic_var_prefix : string val fixed_var_prefix : string val tvar_prefix : string val tfree_prefix : string val const_prefix : string val type_const_prefix : string val class_prefix : string val lam_lifted_prefix : string val lam_lifted_mono_prefix : string val lam_lifted_poly_prefix : string val skolem_const_prefix : string val old_skolem_const_prefix : string val new_skolem_const_prefix : string val combinator_prefix : string val class_decl_prefix : string val type_decl_prefix : string val sym_decl_prefix : string val datatype_decl_prefix : string val class_memb_prefix : string val guards_sym_formula_prefix : string val tags_sym_formula_prefix : string val fact_prefix : string val conjecture_prefix : string val helper_prefix : string val subclass_prefix : string val tcon_clause_prefix : string val tfree_clause_prefix : string val lam_fact_prefix : string val typed_helper_suffix : string val untyped_helper_suffix : string val predicator_name : string val app_op_name : string val type_guard_name : string val type_tag_name : string val native_type_prefix : string val prefixed_predicator_name : string val prefixed_app_op_name : string val prefixed_type_tag_name : string val ascii_of : string -> string val unascii_of : string -> string val unprefix_and_unascii : string -> string -> string option val proxy_table : (string * (string * (thm * (string * string)))) list val proxify_const : string -> (string * string) option val invert_const : string -> string val unproxify_const : string -> string val new_skolem_var_name_of_const : string -> string val atp_logical_consts : string list val atp_irrelevant_consts : string list val atp_widely_irrelevant_consts : string list val is_irrelevant_const : string -> bool val is_widely_irrelevant_const : string -> bool val atp_schematic_consts_of : term -> typ list Symtab.table val is_type_enc_higher_order : type_enc -> bool val is_type_enc_polymorphic : type_enc -> bool val level_of_type_enc : type_enc -> type_level val is_type_enc_sound : type_enc -> bool val type_enc_of_string : strictness -> string -> type_enc val adjust_type_enc : atp_format -> type_enc -> type_enc val is_first_order_lambda_free : term -> bool val do_cheaply_conceal_lambdas : typ list -> term -> term val mk_aconns : atp_connective -> ('a, 'b, 'c, 'd) atp_formula list -> ('a, 'b, 'c, 'd) atp_formula val unmangled_type : string -> (string, 'a) ATP_Problem.atp_term val unmangled_const : string -> string * (string, 'b) atp_term list val unmangled_const_name : string -> string list val helper_table : bool -> ((string * bool) * (status * thm) list) list val trans_lams_of_string : Proof.context -> type_enc -> string -> term list -> term list * term list val string_of_status : status -> string val factsN : string val generate_atp_problem : Proof.context -> bool -> atp_format -> atp_formula_role -> type_enc -> mode -> string -> bool -> bool -> bool -> term list -> term -> ((string * stature) * term) list -> string atp_problem * string Symtab.table * (string * term) list * int Symtab.table val atp_problem_selection_weights : string atp_problem -> (string * real) list val atp_problem_term_order_info : string atp_problem -> (string * int) list end; structure ATP_Problem_Generate : ATP_PROBLEM_GENERATE = struct open ATP_Util open ATP_Problem datatype mode = Metis | Sledgehammer | Sledgehammer_Completish of int | Exporter | Translator datatype scope = Global | Local | Assum | Chained datatype status = General | Induction | Intro | Inductive | Elim | Simp | Non_Rec_Def | Rec_Def type stature = scope * status datatype order = First_Order | Higher_Order of thf_flavor datatype phantom_policy = Without_Phantom_Type_Vars | With_Phantom_Type_Vars datatype polymorphism = Type_Class_Polymorphic | Raw_Polymorphic of phantom_policy | Raw_Monomorphic | Mangled_Monomorphic datatype strictness = Strict | Non_Strict datatype uniformity = Uniform | Non_Uniform datatype ctr_optim = With_Ctr_Optim | Without_Ctr_Optim datatype type_level = All_Types | Undercover_Types | Nonmono_Types of strictness * uniformity | Const_Types of ctr_optim | No_Types datatype type_enc = Native of order * fool * polymorphism * type_level | Guards of polymorphism * type_level | Tags of polymorphism * type_level (* not clear whether ATPs prefer to have their negative variables tagged *) val tag_neg_vars = false fun is_type_enc_native (Native _) = true | is_type_enc_native _ = false fun is_type_enc_full_higher_order (Native (Higher_Order THF_Lambda_Free, _, _, _)) = false | is_type_enc_full_higher_order (Native (Higher_Order _, _, _, _)) = true | is_type_enc_full_higher_order _ = false fun is_type_enc_fool (Native (_, With_FOOL, _, _)) = true | is_type_enc_fool _ = false fun is_type_enc_higher_order (Native (Higher_Order _, _, _, _)) = true | is_type_enc_higher_order _ = false fun polymorphism_of_type_enc (Native (_, _, poly, _)) = poly | polymorphism_of_type_enc (Guards (poly, _)) = poly | polymorphism_of_type_enc (Tags (poly, _)) = poly fun is_type_enc_polymorphic type_enc = (case polymorphism_of_type_enc type_enc of Raw_Polymorphic _ => true | Type_Class_Polymorphic => true | _ => false) fun is_type_enc_mangling type_enc = polymorphism_of_type_enc type_enc = Mangled_Monomorphic fun level_of_type_enc (Native (_, _, _, level)) = level | level_of_type_enc (Guards (_, level)) = level | level_of_type_enc (Tags (_, level)) = level fun is_type_level_uniform (Nonmono_Types (_, Non_Uniform)) = false | is_type_level_uniform Undercover_Types = false | is_type_level_uniform _ = true fun is_type_level_sound (Const_Types _) = false | is_type_level_sound No_Types = false | is_type_level_sound _ = true val is_type_enc_sound = is_type_level_sound o level_of_type_enc fun is_type_level_monotonicity_based (Nonmono_Types _) = true | is_type_level_monotonicity_based _ = false val no_lamsN = "no_lams" (* used internally; undocumented *) val opaque_liftingN = "opaque_lifting" val liftingN = "lifting" val opaque_combsN = "opaque_combs" val combsN = "combs" val combs_and_liftingN = "combs_and_lifting" val combs_or_liftingN = "combs_or_lifting" val keep_lamsN = "keep_lams" val bound_var_prefix = "B_" val all_bound_var_prefix = "A_" val exist_bound_var_prefix = "E_" val schematic_var_prefix = "V_" val fixed_var_prefix = "v_" val tvar_prefix = "T_" val tfree_prefix = "tf_" val const_prefix = "c_" val type_const_prefix = "t_" val native_type_prefix = "n_" val class_prefix = "cl_" (* Freshness almost guaranteed! *) val atp_prefix = "ATP" ^ Long_Name.separator val atp_weak_prefix = "ATP:" val atp_weak_suffix = ":ATP" val lam_lifted_prefix = atp_weak_prefix ^ "Lam" val lam_lifted_mono_prefix = lam_lifted_prefix ^ "m" val lam_lifted_poly_prefix = lam_lifted_prefix ^ "p" val skolem_const_prefix = atp_prefix ^ "Sko" val old_skolem_const_prefix = skolem_const_prefix ^ "o" val new_skolem_const_prefix = skolem_const_prefix ^ "n" val combinator_prefix = "COMB" val class_decl_prefix = "cl_" val type_decl_prefix = "ty_" val sym_decl_prefix = "sy_" val datatype_decl_prefix = "dt_" val class_memb_prefix = "cm_" val guards_sym_formula_prefix = "gsy_" val tags_sym_formula_prefix = "tsy_" val uncurried_alias_eq_prefix = "unc_" val fact_prefix = "fact_" val conjecture_prefix = "conj_" val helper_prefix = "help_" val subclass_prefix = "subcl_" val tcon_clause_prefix = "tcon_" val tfree_clause_prefix = "tfree_" val lam_fact_prefix = "ATP.lambda_" val typed_helper_suffix = "_T" val untyped_helper_suffix = "_U" val predicator_name = "pp" val app_op_name = "aa" val type_guard_name = "gg" val type_tag_name = "tt" val prefixed_predicator_name = const_prefix ^ predicator_name val prefixed_app_op_name = const_prefix ^ app_op_name val prefixed_type_tag_name = const_prefix ^ type_tag_name (*Escaping of special characters. Alphanumeric characters are left unchanged. The character _ goes to __. Characters in the range ASCII space to / go to _A to _P, respectively. Other characters go to _nnn where nnn is the decimal ASCII code. *) val upper_a_minus_space = Char.ord #"A" - Char.ord #" " fun ascii_of_char c = if Char.isAlphaNum c then String.str c else if c = #"_" then "__" else if #" " <= c andalso c <= #"/" then "_" ^ String.str (Char.chr (Char.ord c + upper_a_minus_space)) else (* fixed width, in case more digits follow *) "_" ^ stringN_of_int 3 (Char.ord c) val ascii_of = String.translate ascii_of_char (** Remove ASCII armoring from names in proof files **) (* We don't raise error exceptions because this code can run inside a worker thread. Also, the errors are impossible. *) val unascii_of = let fun un rcs [] = String.implode (rev rcs) | un rcs [#"_"] = un (#"_" :: rcs) [] (* ERROR *) (* Three types of _ escapes: __, _A to _P, _nnn *) | un rcs (#"_" :: #"_" :: cs) = un (#"_" :: rcs) cs | un rcs (#"_" :: c :: cs) = if #"A" <= c andalso c<= #"P" then (* translation of #" " to #"/" *) un (Char.chr (Char.ord c - upper_a_minus_space) :: rcs) cs else let val digits = List.take (c :: cs, 3) handle General.Subscript => [] in (case Int.fromString (String.implode digits) of SOME n => un (Char.chr n :: rcs) (List.drop (cs, 2)) | NONE => un (c :: #"_" :: rcs) cs (* ERROR *)) end | un rcs (c :: cs) = un (c :: rcs) cs in un [] o String.explode end (* If string s has the prefix s1, return the result of deleting it, un-ASCII'd. *) fun unprefix_and_unascii s1 s = if String.isPrefix s1 s then SOME (unascii_of (String.extract (s, size s1, NONE))) else NONE val proxy_table = [("c_False", (\<^const_name>\False\, (@{thm fFalse_def}, ("fFalse", \<^const_name>\fFalse\)))), ("c_True", (\<^const_name>\True\, (@{thm fTrue_def}, ("fTrue", \<^const_name>\fTrue\)))), ("c_Not", (\<^const_name>\Not\, (@{thm fNot_def}, ("fNot", \<^const_name>\fNot\)))), ("c_conj", (\<^const_name>\conj\, (@{thm fconj_def}, ("fconj", \<^const_name>\fconj\)))), ("c_disj", (\<^const_name>\disj\, (@{thm fdisj_def}, ("fdisj", \<^const_name>\fdisj\)))), ("c_implies", (\<^const_name>\implies\, (@{thm fimplies_def}, ("fimplies", \<^const_name>\fimplies\)))), ("equal", (\<^const_name>\HOL.eq\, (@{thm fequal_def}, ("fequal", \<^const_name>\fequal\)))), ("c_All", (\<^const_name>\All\, (@{thm fAll_def}, ("fAll", \<^const_name>\fAll\)))), ("c_Ex", (\<^const_name>\Ex\, (@{thm fEx_def}, ("fEx", \<^const_name>\fEx\))))] val proxify_const = AList.lookup (op =) proxy_table #> Option.map (snd o snd) (* Readable names for the more common symbolic functions. Do not mess with the table unless you know what you are doing. *) val const_trans_table = [(\<^const_name>\False\, "False"), (\<^const_name>\True\, "True"), (\<^const_name>\Not\, "Not"), (\<^const_name>\conj\, "conj"), (\<^const_name>\disj\, "disj"), (\<^const_name>\implies\, "implies"), (\<^const_name>\HOL.eq\, "equal"), (\<^const_name>\All\, "All"), (\<^const_name>\Ex\, "Ex"), (\<^const_name>\If\, "If"), (\<^const_name>\Set.member\, "member"), (\<^const_name>\Meson.COMBI\, combinator_prefix ^ "I"), (\<^const_name>\Meson.COMBK\, combinator_prefix ^ "K"), (\<^const_name>\Meson.COMBB\, combinator_prefix ^ "B"), (\<^const_name>\Meson.COMBC\, combinator_prefix ^ "C"), (\<^const_name>\Meson.COMBS\, combinator_prefix ^ "S")] |> Symtab.make |> fold (Symtab.update o swap o snd o snd o snd) proxy_table (* Invert the table of translations between Isabelle and ATPs. *) val const_trans_table_inv = const_trans_table |> Symtab.dest |> map swap |> Symtab.make val const_trans_table_unprox = Symtab.empty |> fold (fn (_, (isa, (_, (_, atp)))) => Symtab.update (atp, isa)) proxy_table val invert_const = perhaps (Symtab.lookup const_trans_table_inv) val unproxify_const = perhaps (Symtab.lookup const_trans_table_unprox) fun lookup_const c = (case Symtab.lookup const_trans_table c of SOME c' => c' | NONE => ascii_of c) fun ascii_of_indexname (v, 0) = ascii_of v | ascii_of_indexname (v, i) = ascii_of v ^ "_" ^ string_of_int i fun make_bound_var x = bound_var_prefix ^ ascii_of x fun make_all_bound_var x = all_bound_var_prefix ^ ascii_of x fun make_exist_bound_var x = exist_bound_var_prefix ^ ascii_of x fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v fun make_fixed_var x = fixed_var_prefix ^ ascii_of x fun make_tvar (s, i) = tvar_prefix ^ ascii_of_indexname (unquote_tvar s, i) fun make_tfree s = tfree_prefix ^ ascii_of (unquote_tvar s) fun tvar_name ((x as (s, _)), _) = (make_tvar x, s) (* "HOL.eq" and choice are mapped to the ATP's equivalents *) local val choice_const = (fst o dest_Const o HOLogic.choice_const) dummyT fun default c = const_prefix ^ lookup_const c in fun make_fixed_const _ \<^const_name>\HOL.eq\ = tptp_old_equal | make_fixed_const (SOME (Native (Higher_Order THF_With_Choice, _, _, _))) c = if c = choice_const then tptp_choice else default c | make_fixed_const _ c = default c end fun make_fixed_type_const c = type_const_prefix ^ lookup_const c fun make_class clas = class_prefix ^ ascii_of clas fun new_skolem_var_name_of_const s = let val ss = Long_Name.explode s in nth ss (length ss - 2) end (* These are ignored anyway by the relevance filter (unless they appear in higher-order places) but not by the monomorphizer. *) val atp_logical_consts = [\<^const_name>\Pure.prop\, \<^const_name>\Pure.conjunction\, \<^const_name>\Pure.all\, \<^const_name>\Pure.imp\, \<^const_name>\Pure.eq\, \<^const_name>\Trueprop\, \<^const_name>\All\, \<^const_name>\Ex\, \<^const_name>\Ex1\, \<^const_name>\Ball\, \<^const_name>\Bex\] (* These are either simplified away by "Meson.presimplify" (most of the time) or handled specially via "fFalse", "fTrue", ..., "fequal". *) val atp_irrelevant_consts = [\<^const_name>\False\, \<^const_name>\True\, \<^const_name>\Not\, \<^const_name>\conj\, \<^const_name>\disj\, \<^const_name>\implies\, \<^const_name>\HOL.eq\, \<^const_name>\If\, \<^const_name>\Let\] val atp_widely_irrelevant_consts = atp_logical_consts @ atp_irrelevant_consts val atp_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_irrelevant_consts) val atp_widely_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_widely_irrelevant_consts) val is_irrelevant_const = Symtab.defined atp_irrelevant_const_tab val is_widely_irrelevant_const = Symtab.defined atp_widely_irrelevant_const_tab fun add_schematic_const (x as (_, T)) = Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x val add_schematic_consts_of = Term.fold_aterms (fn Const (x as (s, _)) => not (member (op =) atp_widely_irrelevant_consts s) ? add_schematic_const x | _ => I) fun atp_schematic_consts_of t = add_schematic_consts_of t Symtab.empty val tvar_a_str = "'a" val tvar_a_z = ((tvar_a_str, 0), \<^sort>\type\) val tvar_a = TVar tvar_a_z val tvar_a_name = tvar_name tvar_a_z val itself_name = `make_fixed_type_const \<^type_name>\itself\ val TYPE_name = `(make_fixed_const NONE) \<^const_name>\Pure.type\ val tvar_a_atype = AType ((tvar_a_name, []), []) val a_itself_atype = AType ((itself_name, []), [tvar_a_atype]) (** Definitions and functions for FOL clauses and formulas for TPTP **) (** Type class membership **) (* In our data structures, [] exceptionally refers to the top class, not to the empty class. *) val class_of_types = the_single \<^sort>\type\ fun normalize_classes cls = if member (op =) cls class_of_types then [] else cls (* Arity of type constructor "s :: (arg1, ..., argN) res" *) fun make_axiom_tcon_clause (s, name, (cl, args)) = let val args = args |> map normalize_classes val tvars = 1 upto length args |> map (fn j => TVar ((tvar_a_str, j), \<^sort>\type\)) in (name, args ~~ tvars, (cl, Type (s, tvars))) end (* Generate all pairs (tycon, class, sorts) such that tycon belongs to class in theory thy provided its arguments have the corresponding sorts. *) fun class_pairs thy tycons cls = let val alg = Sign.classes_of thy fun domain_sorts tycon = Sorts.mg_domain alg tycon o single fun add_class tycon cl = cons (cl, domain_sorts tycon cl) handle Sorts.CLASS_ERROR _ => I fun try_classes tycon = (tycon, fold (add_class tycon) cls []) in map try_classes tycons end (* Proving one (tycon, class) membership may require proving others, so iterate. *) fun all_class_pairs _ _ _ [] = ([], []) | all_class_pairs thy tycons old_cls cls = let val old_cls' = cls @ old_cls fun maybe_insert_class s = not (member (op =) old_cls' s) ? insert (op =) s val pairs = class_pairs thy tycons cls val new_cls = fold (fold (fold (fold maybe_insert_class) o snd) o snd) pairs [] val (cls', pairs') = all_class_pairs thy tycons old_cls' new_cls in (cls' @ cls, union (op =) pairs' pairs) end fun tcon_clause _ _ [] = [] | tcon_clause seen n ((_, []) :: rest) = tcon_clause seen n rest | tcon_clause seen n ((tcons, (ar as (cl, _)) :: ars) :: rest) = if cl = class_of_types then tcon_clause seen n ((tcons, ars) :: rest) else if member (op =) seen cl then (* multiple clauses for the same (tycon, cl) pair *) make_axiom_tcon_clause (tcons, lookup_const tcons ^ "___" ^ ascii_of cl ^ "_" ^ string_of_int n, ar) :: tcon_clause seen (n + 1) ((tcons, ars) :: rest) else make_axiom_tcon_clause (tcons, lookup_const tcons ^ "___" ^ ascii_of cl, ar) :: tcon_clause (cl :: seen) n ((tcons, ars) :: rest) fun make_tcon_clauses thy tycons = all_class_pairs thy tycons [] ##> tcon_clause [] 1 (** Isabelle class relations **) (* Generate a list ("sub", "supers") such that "sub" is a proper subclass of all "supers". *) fun make_subclass_pairs thy subs supers = let val class_less = curry (Sorts.class_less (Sign.classes_of thy)) fun supers_of sub = (sub, filter (class_less sub) supers) in map supers_of subs |> filter_out (null o snd) end (* intermediate terms *) datatype iterm = IConst of (string * string) * typ * typ list | IVar of (string * string) * typ | IApp of iterm * iterm | IAbs of ((string * string) * typ) * iterm fun ityp_of (IConst (_, T, _)) = T | ityp_of (IVar (_, T)) = T | ityp_of (IApp (t1, _)) = snd (dest_funT (ityp_of t1)) | ityp_of (IAbs ((_, T), tm)) = T --> ityp_of tm (*gets the head of a combinator application, along with the list of arguments*) fun strip_iterm_comb u = let fun stripc (IApp (t, u), ts) = stripc (t, u :: ts) | stripc x = x in stripc (u, []) end fun atomic_types_of T = fold_atyps (insert (op =)) T [] fun new_skolem_const_name s num_T_args = [new_skolem_const_prefix, s, string_of_int num_T_args] |> Long_Name.implode val alpha_to_beta = Logic.varifyT_global \<^typ>\'a => 'b\ val alpha_to_beta_to_alpha_to_beta = alpha_to_beta --> alpha_to_beta fun robust_const_type thy s = if s = app_op_name then alpha_to_beta_to_alpha_to_beta else if String.isPrefix lam_lifted_prefix s then alpha_to_beta else (* Old Skolems throw a "TYPE" exception here, which will be caught. *) s |> Sign.the_const_type thy fun ary_of (Type (\<^type_name>\fun\, [_, T])) = 1 + ary_of T | ary_of _ = 0 (* This function only makes sense if "T" is as general as possible. *) fun robust_const_type_args thy (s, T) = if s = app_op_name then let val (T1, T2) = T |> domain_type |> dest_funT in [T1, T2] end else if String.isPrefix old_skolem_const_prefix s then [] |> Term.add_tvarsT T |> rev |> map TVar else if String.isPrefix lam_lifted_prefix s then if String.isPrefix lam_lifted_poly_prefix s then let val (T1, T2) = T |> dest_funT in [T1, T2] end else [] else (s, T) |> Sign.const_typargs thy (* Converts an Isabelle/HOL term (with combinators) into an intermediate term. Also accumulates sort infomation. *) fun iterm_of_term thy type_enc bs (P $ Q) = let val (P', P_atomics_Ts) = iterm_of_term thy type_enc bs P val (Q', Q_atomics_Ts) = iterm_of_term thy type_enc bs Q in (IApp (P', Q'), union (op =) P_atomics_Ts Q_atomics_Ts) end | iterm_of_term thy type_enc _ (Const (c, T)) = (IConst (`(make_fixed_const (SOME type_enc)) c, T, robust_const_type_args thy (c, T)), atomic_types_of T) | iterm_of_term _ _ _ (Free (s, T)) = (IConst (`make_fixed_var s, T, []), atomic_types_of T) | iterm_of_term _ type_enc _ (Var (v as (s, _), T)) = (if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then let val Ts = T |> strip_type |> swap |> op :: val s' = new_skolem_const_name s (length Ts) in IConst (`(make_fixed_const (SOME type_enc)) s', T, Ts) end else IVar ((make_schematic_var v, s), T), atomic_types_of T) | iterm_of_term _ _ bs (Bound j) = nth bs j |> (fn (_, (name, T)) => (IConst (name, T, []), atomic_types_of T)) | iterm_of_term thy type_enc bs (Abs (s, T, t)) = let fun vary s = s |> AList.defined (op =) bs s ? vary o Symbol.bump_string val s = vary s val name = `make_bound_var s val (tm, atomic_Ts) = iterm_of_term thy type_enc ((s, (name, T)) :: bs) t in (IAbs ((name, T), tm), union (op =) atomic_Ts (atomic_types_of T)) end (* "_query" and "_at" are for the ASCII-challenged Metis and Mirabelle. *) val queries = ["?", "_query"] val ats = ["@", "_at"] fun try_unsuffixes ss s = fold (fn s' => fn NONE => try (unsuffix s') s | some => some) ss NONE fun type_enc_of_string strictness s = let val (poly, s) = (case try (unprefix "tc_") s of SOME s => (SOME Type_Class_Polymorphic, s) | NONE => (case try (unprefix "poly_") s of SOME s => (SOME (Raw_Polymorphic With_Phantom_Type_Vars), s) | NONE => (case try (unprefix "ml_poly_") s of SOME s => (SOME (Raw_Polymorphic Without_Phantom_Type_Vars), s) | NONE => (case try (unprefix "raw_mono_") s of SOME s => (SOME Raw_Monomorphic, s) | NONE => (case try (unprefix "mono_") s of SOME s => (SOME Mangled_Monomorphic, s) | NONE => (NONE, s)))))) val (level, s) = case try_unsuffixes queries s of SOME s => (case try_unsuffixes queries s of SOME s => (Nonmono_Types (strictness, Non_Uniform), s) | NONE => (Nonmono_Types (strictness, Uniform), s)) | NONE => (case try_unsuffixes ats s of SOME s => (Undercover_Types, s) | NONE => (All_Types, s)) fun native_of_string s = let val (fool, core) = (case try (unsuffix "_fool") s of SOME s => (With_FOOL, s) | NONE => (Without_FOOL, s)) in (case (core, poly) of ("native", SOME poly) => (case (poly, level) of (Mangled_Monomorphic, _) => if is_type_level_uniform level then Native (First_Order, fool, Mangled_Monomorphic, level) else raise Same.SAME | (Raw_Monomorphic, _) => raise Same.SAME | (poly, All_Types) => Native (First_Order, fool, poly, All_Types)) | ("native_higher", SOME poly) => (case (poly, level) of (_, Nonmono_Types _) => raise Same.SAME | (_, Undercover_Types) => raise Same.SAME | (Mangled_Monomorphic, _) => if is_type_level_uniform level then Native (Higher_Order THF_With_Choice, fool, Mangled_Monomorphic, level) else raise Same.SAME | (poly as Raw_Polymorphic _, All_Types) => Native (Higher_Order THF_With_Choice, fool, poly, All_Types) | _ => raise Same.SAME)) end fun nonnative_of_string core = (case (core, poly, level) of ("guards", SOME poly, _) => if (poly = Mangled_Monomorphic andalso level = Undercover_Types) orelse poly = Type_Class_Polymorphic then raise Same.SAME else Guards (poly, level) | ("tags", SOME poly, _) => if (poly = Mangled_Monomorphic andalso level = Undercover_Types) orelse poly = Type_Class_Polymorphic then raise Same.SAME else Tags (poly, level) | ("args", SOME poly, All_Types (* naja *)) => if poly = Type_Class_Polymorphic then raise Same.SAME else Guards (poly, Const_Types Without_Ctr_Optim) | ("args", SOME poly, Nonmono_Types (_, Uniform) (* naja *)) => if poly = Mangled_Monomorphic orelse poly = Type_Class_Polymorphic then raise Same.SAME else Guards (poly, Const_Types With_Ctr_Optim) | ("erased", NONE, All_Types (* naja *)) => Guards (Raw_Polymorphic With_Phantom_Type_Vars, No_Types) | _ => raise Same.SAME) in if String.isPrefix "native" s then native_of_string s else nonnative_of_string s end handle Same.SAME => error ("Unknown type encoding: " ^ quote s) fun min_hologic THF_Lambda_Free _ = THF_Lambda_Free | min_hologic _ THF_Lambda_Free = THF_Lambda_Free | min_hologic THF_Without_Choice _ = THF_Without_Choice | min_hologic _ THF_Without_Choice = THF_Without_Choice | min_hologic _ _ = THF_With_Choice fun adjust_hologic hologic (Higher_Order hologic') = Higher_Order (min_hologic hologic hologic') | adjust_hologic _ type_enc = type_enc fun adjust_fool Without_FOOL _ = Without_FOOL | adjust_fool _ fool = fool fun no_type_classes Type_Class_Polymorphic = Raw_Polymorphic With_Phantom_Type_Vars | no_type_classes poly = poly fun adjust_type_enc (THF (fool, Polymorphic, hologic)) (Native (order, fool', poly, level)) = Native (adjust_hologic hologic order, adjust_fool fool fool', no_type_classes poly, level) | adjust_type_enc (THF (fool, Monomorphic, hologic)) (Native (order, fool', _, level)) = Native (adjust_hologic hologic order, adjust_fool fool fool', Mangled_Monomorphic, level) | adjust_type_enc (TFF (fool, Monomorphic)) (Native (_, fool', _, level)) = Native (First_Order, adjust_fool fool fool', Mangled_Monomorphic, level) | adjust_type_enc (DFG Polymorphic) (Native (_, _, poly, level)) = Native (First_Order, Without_FOOL, poly, level) | adjust_type_enc (DFG Monomorphic) (Native (_, _, _, level)) = Native (First_Order, Without_FOOL, Mangled_Monomorphic, level) | adjust_type_enc (TFF (fool, _)) (Native (_, fool', poly, level)) = Native (First_Order, adjust_fool fool fool', no_type_classes poly, level) | adjust_type_enc format (Native (_, _, poly, level)) = adjust_type_enc format (Guards (no_type_classes poly, level)) | adjust_type_enc CNF_UEQ (type_enc as Guards stuff) = (if is_type_enc_sound type_enc then Tags else Guards) stuff | adjust_type_enc _ type_enc = type_enc fun is_first_order_lambda_free t = (case t of \<^const>\Not\ $ t1 => is_first_order_lambda_free t1 | Const (\<^const_name>\All\, _) $ Abs (_, _, t') => is_first_order_lambda_free t' | Const (\<^const_name>\All\, _) $ t1 => is_first_order_lambda_free t1 | Const (\<^const_name>\Ex\, _) $ Abs (_, _, t') => is_first_order_lambda_free t' | Const (\<^const_name>\Ex\, _) $ t1 => is_first_order_lambda_free t1 | \<^const>\HOL.conj\ $ t1 $ t2 => is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2 | \<^const>\HOL.disj\ $ t1 $ t2 => is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2 | \<^const>\HOL.implies\ $ t1 $ t2 => is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2 | Const (\<^const_name>\HOL.eq\, Type (_, [\<^typ>\bool\, _])) $ t1 $ t2 => is_first_order_lambda_free t1 andalso is_first_order_lambda_free t2 | _ => not (exists_subterm (fn Abs _ => true | _ => false) t)) fun simple_translate_lambdas do_lambdas ctxt type_enc t = if is_first_order_lambda_free t then t else let fun trans_first_order Ts t = (case t of \<^const>\Not\ $ t1 => \<^const>\Not\ $ trans_first_order Ts t1 | (t0 as Const (\<^const_name>\All\, _)) $ Abs (s, T, t') => t0 $ Abs (s, T, trans_first_order (T :: Ts) t') | (t0 as Const (\<^const_name>\All\, _)) $ t1 => trans_first_order Ts (t0 $ eta_expand Ts t1 1) | (t0 as Const (\<^const_name>\Ex\, _)) $ Abs (s, T, t') => t0 $ Abs (s, T, trans_first_order (T :: Ts) t') | (t0 as Const (\<^const_name>\Ex\, _)) $ t1 => trans_first_order Ts (t0 $ eta_expand Ts t1 1) | (t0 as \<^const>\HOL.conj\) $ t1 $ t2 => t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2 | (t0 as \<^const>\HOL.disj\) $ t1 $ t2 => t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2 | (t0 as \<^const>\HOL.implies\) $ t1 $ t2 => t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2 | (t0 as Const (\<^const_name>\HOL.eq\, Type (_, [\<^typ>\bool\, _]))) $ t1 $ t2 => t0 $ trans_first_order Ts t1 $ trans_first_order Ts t2 | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then t else t |> Envir.eta_contract |> do_lambdas ctxt Ts) fun trans_fool Ts t = (case t of (t0 as Const (\<^const_name>\All\, _)) $ Abs (s, T, t') => t0 $ Abs (s, T, trans_fool (T :: Ts) t') | (t0 as Const (\<^const_name>\All\, _)) $ t1 => trans_fool Ts (t0 $ eta_expand Ts t1 1) | (t0 as Const (\<^const_name>\Ex\, _)) $ Abs (s, T, t') => t0 $ Abs (s, T, trans_fool (T :: Ts) t') | (t0 as Const (\<^const_name>\Ex\, _)) $ t1 => trans_fool Ts (t0 $ eta_expand Ts t1 1) | t1 $ t2 => trans_fool Ts t1 $ trans_fool Ts t2 | Abs _ => t |> Envir.eta_contract |> do_lambdas ctxt Ts | _ => t) val trans = if is_type_enc_fool type_enc then trans_fool else trans_first_order val (t, ctxt') = yield_singleton (Variable.import_terms true) t ctxt in t |> trans [] |> singleton (Variable.export_terms ctxt' ctxt) end fun do_cheaply_conceal_lambdas Ts (t1 $ t2) = do_cheaply_conceal_lambdas Ts t1 $ do_cheaply_conceal_lambdas Ts t2 | do_cheaply_conceal_lambdas Ts (Abs (_, T, t)) = Const (lam_lifted_poly_prefix ^ serial_string (), T --> fastype_of1 (T :: Ts, t)) | do_cheaply_conceal_lambdas _ t = t fun concealed_bound_name j = atp_weak_prefix ^ string_of_int j fun conceal_bounds Ts t = subst_bounds (map (Free o apfst concealed_bound_name) (0 upto length Ts - 1 ~~ Ts), t) fun reveal_bounds Ts = subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j)) (0 upto length Ts - 1 ~~ Ts)) fun do_introduce_combinators ctxt Ts t = (t |> conceal_bounds Ts |> Thm.cterm_of ctxt |> Meson_Clausify.introduce_combinators_in_cterm ctxt |> Thm.prop_of |> Logic.dest_equals |> snd |> reveal_bounds Ts) (* A type variable of sort "{}" will make abstraction fail. *) handle THM _ => t |> do_cheaply_conceal_lambdas Ts val introduce_combinators = simple_translate_lambdas do_introduce_combinators fun constify_lifted (t $ u) = constify_lifted t $ constify_lifted u | constify_lifted (Abs (s, T, t)) = Abs (s, T, constify_lifted t) | constify_lifted (Free (x as (s, _))) = (if String.isPrefix lam_lifted_prefix s then Const else Free) x | constify_lifted t = t fun lift_lams_part_1 ctxt type_enc = map hol_close_form #> rpair ctxt #-> Lambda_Lifting.lift_lambdas (SOME ((if is_type_enc_polymorphic type_enc then lam_lifted_poly_prefix else lam_lifted_mono_prefix) ^ "_a")) Lambda_Lifting.is_quantifier #> fst fun lift_lams_part_2 ctxt type_enc (facts, lifted) = (facts, lifted) (* Lambda-lifting sometimes leaves some lambdas around; we need some way to get rid of them *) |> apply2 (map (introduce_combinators ctxt type_enc)) |> apply2 (map constify_lifted) (* Requires bound variables not to clash with any schematic variables (as should be the case right after lambda-lifting). *) |>> map (hol_open_form (unprefix hol_close_form_prefix)) ||> map (hol_open_form I) fun lift_lams ctxt type_enc = lift_lams_part_2 ctxt type_enc o lift_lams_part_1 ctxt type_enc fun intentionalize_def (Const (\<^const_name>\All\, _) $ Abs (_, _, t)) = intentionalize_def t | intentionalize_def (Const (\<^const_name>\HOL.eq\, _) $ t $ u) = let fun lam T t = Abs (Name.uu, T, t) val (head, args) = strip_comb t ||> rev val head_T = fastype_of head val n = length args val arg_Ts = head_T |> binder_types |> take n |> rev val u = u |> subst_atomic (args ~~ map Bound (0 upto n - 1)) in HOLogic.eq_const head_T $ head $ fold lam arg_Ts u end | intentionalize_def t = t type ifact = {name : string, stature : stature, role : atp_formula_role, iformula : (string * string, typ, iterm, string * string) atp_formula, atomic_types : typ list} fun update_iformula f ({name, stature, role, iformula, atomic_types} : ifact) = {name = name, stature = stature, role = role, iformula = f iformula, atomic_types = atomic_types} : ifact fun ifact_lift f ({iformula, ...} : ifact) = f iformula fun insert_type thy get_T x xs = let val T = get_T x in if exists (type_instance thy T o get_T) xs then xs else x :: filter_out (type_generalization thy T o get_T) xs end fun chop_fun 0 T = ([], T) | chop_fun n (Type (\<^type_name>\fun\, [dom_T, ran_T])) = chop_fun (n - 1) ran_T |>> cons dom_T | chop_fun _ T = ([], T) fun filter_type_args thy ctrss type_enc s ary T_args = let val poly = polymorphism_of_type_enc type_enc in if s = type_tag_name then (* FIXME: why not "type_guard_name" as well? *) T_args else (case type_enc of Native (_, _, Raw_Polymorphic _, _) => T_args | Native (_, _, Type_Class_Polymorphic, _) => T_args | _ => let fun gen_type_args _ _ [] = [] | gen_type_args keep strip_ty T_args = let val U = robust_const_type thy s val (binder_Us, body_U) = strip_ty U val in_U_vars = fold Term.add_tvarsT binder_Us [] val out_U_vars = Term.add_tvarsT body_U [] fun filt (U_var, T) = if keep (member (op =) in_U_vars U_var, member (op =) out_U_vars U_var) then T else dummyT val U_args = (s, U) |> robust_const_type_args thy in map (filt o apfst dest_TVar) (U_args ~~ T_args) end handle TYPE _ => T_args fun is_always_ctr (s', T') = s' = s andalso type_equiv thy (T', robust_const_type thy s') val noninfer_type_args = gen_type_args (not o fst) (chop_fun ary) val ctr_infer_type_args = gen_type_args fst strip_type val level = level_of_type_enc type_enc in if level = No_Types orelse s = \<^const_name>\HOL.eq\ orelse (case level of Const_Types _ => s = app_op_name | _ => false) then [] else if poly = Mangled_Monomorphic then T_args else if level = All_Types then (case type_enc of Guards _ => noninfer_type_args T_args | Tags _ => []) else if level = Undercover_Types then noninfer_type_args T_args else if level <> Const_Types Without_Ctr_Optim andalso exists (exists is_always_ctr) ctrss then ctr_infer_type_args T_args else T_args end) end fun raw_atp_type_of_typ type_enc = let fun term (Type (s, Ts)) = AType ((if s = \<^type_name>\fun\ andalso is_type_enc_higher_order type_enc then `I tptp_fun_type else if s = \<^type_name>\bool\ andalso (is_type_enc_full_higher_order type_enc orelse is_type_enc_fool type_enc) then `I tptp_bool_type else `make_fixed_type_const s, []), map term Ts) | term (TFree (s, _)) = AType ((`make_tfree s, []), []) | term (TVar z) = AType ((tvar_name z, []), []) in term end fun atp_term_of_atp_type (AType ((name, _), tys)) = ATerm ((name, []), map atp_term_of_atp_type tys) | atp_term_of_atp_type _ = raise Fail "unexpected type" fun atp_type_of_type_arg type_enc T = if T = dummyT then NONE else SOME (raw_atp_type_of_typ type_enc T) (* This shouldn't clash with anything else. *) val uncurried_alias_sep = "\000" val mangled_type_sep = "\001" val ascii_of_uncurried_alias_sep = ascii_of uncurried_alias_sep fun generic_mangled_type_name f (AType ((name, _), [])) = f name | generic_mangled_type_name f (AType ((name, _), tys)) = f name ^ "(" ^ space_implode "," (map (generic_mangled_type_name f) tys) ^ ")" | generic_mangled_type_name _ _ = raise Fail "unexpected type" fun mangled_type type_enc = generic_mangled_type_name fst o raw_atp_type_of_typ type_enc fun make_native_type s = if s = tptp_bool_type orelse s = tptp_fun_type orelse s = tptp_individual_type then s else native_type_prefix ^ ascii_of s fun native_atp_type_of_raw_atp_type type_enc pred_sym ary = let fun to_mangled_atype ty = AType (((make_native_type (generic_mangled_type_name fst ty), generic_mangled_type_name snd ty), []), []) fun to_poly_atype (AType ((name, clss), tys)) = AType ((name, clss), map to_poly_atype tys) | to_poly_atype _ = raise Fail "unexpected type" val to_atype = if is_type_enc_polymorphic type_enc then to_poly_atype else to_mangled_atype fun to_afun f1 f2 tys = AFun (f1 (hd tys), f2 (nth tys 1)) fun to_ho (ty as AType (((s, _), _), tys)) = if s = tptp_fun_type then to_afun to_ho to_ho tys else to_atype ty | to_ho _ = raise Fail "unexpected type" fun to_lfho (ty as AType (((s, _), _), tys)) = if s = tptp_fun_type then to_afun to_ho to_lfho tys else if pred_sym then bool_atype else to_atype ty | to_lfho _ = raise Fail "unexpected type" fun to_fo 0 ty = if pred_sym then bool_atype else to_atype ty | to_fo ary (AType (_, tys)) = to_afun to_atype (to_fo (ary - 1)) tys | to_fo _ _ = raise Fail "unexpected type" in if is_type_enc_full_higher_order type_enc then to_ho else if is_type_enc_higher_order type_enc then to_lfho else to_fo ary end fun native_atp_type_of_typ type_enc pred_sym ary = native_atp_type_of_raw_atp_type type_enc pred_sym ary o raw_atp_type_of_typ type_enc (* Make atoms for sorted type variables. *) fun generic_add_sorts_on_type _ [] = I | generic_add_sorts_on_type T (s :: ss) = generic_add_sorts_on_type T ss #> (if s = the_single \<^sort>\type\ then I else insert (op =) (s, T)) fun add_sorts_on_tfree (T as TFree (_, S)) = generic_add_sorts_on_type T S | add_sorts_on_tfree _ = I fun add_sorts_on_tvar (T as TVar (_, S)) = generic_add_sorts_on_type T S | add_sorts_on_tvar _ = I fun process_type_args type_enc T_args = if is_type_enc_native type_enc then (map (native_atp_type_of_typ type_enc false 0) T_args, []) else ([], map_filter (Option.map atp_term_of_atp_type o atp_type_of_type_arg type_enc) T_args) fun class_atom type_enc (cl, T) = let val cl = `make_class cl val (ty_args, tm_args) = process_type_args type_enc [T] val tm_args = tm_args @ (case type_enc of Native (_, _, Raw_Polymorphic Without_Phantom_Type_Vars, _) => [ATerm ((TYPE_name, ty_args), [])] | _ => []) in AAtom (ATerm ((cl, ty_args), tm_args)) end fun class_atoms type_enc (cls, T) = map (fn cl => class_atom type_enc (cl, T)) cls fun class_membs_of_types type_enc add_sorts_on_typ Ts = [] |> level_of_type_enc type_enc <> No_Types ? fold add_sorts_on_typ Ts fun mk_aconns c = split_last #> uncurry (fold_rev (mk_aconn c)) fun mk_ahorn [] phi = phi | mk_ahorn phis psi = AConn (AImplies, [mk_aconns AAnd phis, psi]) fun mk_aquant _ [] phi = phi | mk_aquant q xs (phi as AQuant (q', xs', phi')) = if q = q' then AQuant (q, xs @ xs', phi') else AQuant (q, xs, phi) | mk_aquant q xs phi = AQuant (q, xs, phi) fun mk_atyquant _ [] phi = phi | mk_atyquant q xs (phi as ATyQuant (q', xs', phi')) = if q = q' then ATyQuant (q, xs @ xs', phi') else ATyQuant (q, xs, phi) | mk_atyquant q xs phi = ATyQuant (q, xs, phi) fun close_universally add_term_vars phi = let fun add_formula_vars bounds (ATyQuant (_, _, phi)) = add_formula_vars bounds phi | add_formula_vars bounds (AQuant (_, xs, phi)) = add_formula_vars (map fst xs @ bounds) phi | add_formula_vars bounds (AConn (_, phis)) = fold (add_formula_vars bounds) phis | add_formula_vars bounds (AAtom tm) = add_term_vars bounds tm in mk_aquant AForall (rev (add_formula_vars [] phi [])) phi end fun add_term_vars bounds (ATerm ((name as (s, _), _), tms)) = (if is_tptp_variable s andalso not (String.isPrefix tvar_prefix s) andalso not (member (op =) bounds name) then insert (op =) (name, NONE) else I) #> fold (add_term_vars bounds) tms | add_term_vars bounds (AAbs (((name, _), tm), args)) = add_term_vars (name :: bounds) tm #> fold (add_term_vars bounds) args fun close_formula_universally phi = close_universally add_term_vars phi fun add_iterm_vars bounds (IApp (tm1, tm2)) = fold (add_iterm_vars bounds) [tm1, tm2] | add_iterm_vars _ (IConst _) = I | add_iterm_vars bounds (IVar (name, T)) = not (member (op =) bounds name) ? insert (op =) (name, SOME T) | add_iterm_vars bounds (IAbs (_, tm)) = add_iterm_vars bounds tm fun aliased_uncurried ary (s, s') = (s ^ ascii_of_uncurried_alias_sep ^ string_of_int ary, s' ^ string_of_int ary) fun unaliased_uncurried (s, s') = (case space_explode uncurried_alias_sep s of [_] => (s, s') | [s1, s2] => (s1, unsuffix s2 s') | _ => raise Fail "ill-formed explicit application alias") fun raw_mangled_const_name type_name ty_args (s, s') = let fun type_suffix f g = fold_rev (prefix o g o prefix mangled_type_sep o type_name f) ty_args "" in (s ^ type_suffix fst ascii_of, s' ^ type_suffix snd I) end fun mangled_const_name type_enc = map_filter (atp_type_of_type_arg type_enc) #> raw_mangled_const_name generic_mangled_type_name val parse_mangled_ident = Scan.many1 (not o member (op =) ["(", ")", ","]) >> implode fun parse_mangled_type x = (parse_mangled_ident -- Scan.optional ($$ "(" |-- Scan.optional parse_mangled_types [] --| $$ ")") [] >> (ATerm o apfst (rpair []))) x and parse_mangled_types x = (parse_mangled_type ::: Scan.repeat ($$ "," |-- parse_mangled_type)) x fun unmangled_type s = s |> suffix ")" |> raw_explode |> Scan.finite Symbol.stopper (Scan.error (!! (fn _ => raise Fail ("unrecognized mangled type " ^ quote s)) parse_mangled_type)) |> fst fun unmangled_const_name s = (s, s) |> unaliased_uncurried |> fst |> space_explode mangled_type_sep fun unmangled_const s = let val ss = unmangled_const_name s in (hd ss, map unmangled_type (tl ss)) end val unmangled_invert_const = invert_const o hd o unmangled_const_name fun vars_of_iterm vars (IConst ((s, _), _, _)) = insert (op =) s vars | vars_of_iterm vars (IVar ((s, _), _)) = insert (op =) s vars | vars_of_iterm vars (IApp (tm1, tm2)) = union (op =) (vars_of_iterm vars tm1) (vars_of_iterm vars tm2) | vars_of_iterm vars (IAbs (_, tm)) = vars_of_iterm vars tm fun generate_unique_name gen unique n = let val x = gen n in if unique x then x else generate_unique_name gen unique (n + 1) end fun eta_expand_quantifier_body (tm as IAbs _) = tm | eta_expand_quantifier_body tm = let (* We accumulate all variables because E 2.5 does not support variable shadowing. *) val vars = vars_of_iterm [] tm val x = generate_unique_name (fn n => "X" ^ (if n = 0 then "" else string_of_int n)) (fn name => not (exists (equal name) vars)) 0 val T = domain_type (ityp_of tm) in IAbs ((`I x, T), IApp (tm, IConst (`I x, T, []))) end fun introduce_builtins_and_proxies_in_iterm type_enc = let val is_fool = is_type_enc_fool type_enc fun tweak_ho_quant ho_quant (T as Type (_, [p_T as Type (_, [x_T, _]), _])) [] = (* Eta-expand "!!" and "??", to work around LEO-II, Leo-III, and Satallax parser limitations. This works in conjunction with special code in "ATP_Problem" that uses the syntactic sugar "!" and "?" whenever possible. *) IAbs ((`I "P", p_T), IApp (IConst (`I ho_quant, T, []), IAbs ((`I "X", x_T), IApp (IConst (`I "P", p_T, []), IConst (`I "X", x_T, []))))) | tweak_ho_quant ho_quant T _ = IConst (`I ho_quant, T, []) fun tweak_ho_equal T argc = if argc = 2 then IConst (`I tptp_equal, T, []) else (* Eta-expand partially applied THF equality, because the LEO-II and Satallax parsers complain about not being able to infer the type of "=". *) let val i_T = domain_type T in IAbs ((`I "Y", i_T), IAbs ((`I "Z", i_T), IApp (IApp (IConst (`I tptp_equal, T, []), IConst (`I "Y", i_T, [])), IConst (`I "Z", i_T, [])))) end fun intro top_level args (IApp (tm1, tm2)) = let val tm1' = intro top_level (tm2 :: args) tm1 val tm2' = intro false [] tm2 val tm2'' = (case tm1' of IConst ((s, _), _, _) => if s = tptp_ho_forall orelse s = tptp_ho_exists then eta_expand_quantifier_body tm2' else tm2' | _ => tm2') in IApp (tm1', tm2'') end | intro top_level args (IConst (name as (s, _), T, T_args)) = if is_fool andalso s = "c_If" andalso (length args = 3 orelse is_type_enc_higher_order type_enc) then IConst (`I tptp_ite, T, []) else (case proxify_const s of SOME proxy_base => let val argc = length args fun plain_const () = IConst (name, T, []) fun proxy_const () = IConst (proxy_base |>> prefix const_prefix, T, T_args) fun handle_fool card x = if card = argc then x else proxy_const () in if top_level then (case s of "c_False" => IConst (`I tptp_false, T, []) | "c_True" => IConst (`I tptp_true, T, []) | _ => plain_const ()) else if is_type_enc_full_higher_order type_enc then (case s of "c_False" => IConst (`I tptp_false, T, []) | "c_True" => IConst (`I tptp_true, T, []) | "c_Not" => IConst (`I tptp_not, T, []) | "c_conj" => IConst (`I tptp_and, T, []) | "c_disj" => IConst (`I tptp_or, T, []) | "c_implies" => IConst (`I tptp_implies, T, []) | "c_All" => tweak_ho_quant tptp_ho_forall T args | "c_Ex" => tweak_ho_quant tptp_ho_exists T args | s => if is_tptp_equal s then tweak_ho_equal T argc else plain_const ()) else if is_fool then (case s of "c_False" => IConst (`I tptp_false, T, []) | "c_True" => IConst (`I tptp_true, T, []) | "c_Not" => handle_fool 1 (IConst (`I tptp_not, T, [])) | "c_conj" => handle_fool 2 (IConst (`I tptp_and, T, [])) | "c_disj" => handle_fool 2 (IConst (`I tptp_or, T, [])) | "c_implies" => handle_fool 2 (IConst (`I tptp_implies, T, [])) | "c_All" => handle_fool 1 (tweak_ho_quant tptp_ho_forall T args) | "c_Ex" => handle_fool 1 (tweak_ho_quant tptp_ho_exists T args) | s => if is_tptp_equal s then handle_fool 2 (IConst (`I tptp_equal, T, [])) else plain_const ()) else proxy_const () end | NONE => if s = tptp_choice then tweak_ho_quant tptp_choice T args else IConst (name, T, T_args)) | intro _ _ (IAbs (bound, tm)) = IAbs (bound, intro false [] tm) | intro _ _ tm = tm in intro true [] end fun mangle_type_args_in_const type_enc (name as (s, _)) T_args = if String.isPrefix const_prefix s andalso is_type_enc_mangling type_enc then (mangled_const_name type_enc T_args name, []) else (name, T_args) fun mangle_type_args_in_iterm type_enc = if is_type_enc_mangling type_enc then let fun mangle (IApp (tm1, tm2)) = IApp (mangle tm1, mangle tm2) | mangle (tm as IConst (_, _, [])) = tm | mangle (IConst (name, T, T_args)) = mangle_type_args_in_const type_enc name T_args |> (fn (name, T_args) => IConst (name, T, T_args)) | mangle (IAbs (bound, tm)) = IAbs (bound, mangle tm) | mangle tm = tm in mangle end else I fun filter_type_args_in_const _ _ _ _ _ [] = [] | filter_type_args_in_const thy ctrss type_enc ary s T_args = (case unprefix_and_unascii const_prefix s of NONE => if level_of_type_enc type_enc = No_Types orelse s = tptp_choice then [] else T_args | SOME s'' => filter_type_args thy ctrss type_enc (unmangled_invert_const s'') ary T_args) fun filter_type_args_in_iterm thy ctrss type_enc = let fun filt ary (IApp (tm1, tm2)) = IApp (filt (ary + 1) tm1, filt 0 tm2) | filt ary (IConst (name as (s, _), T, T_args)) = filter_type_args_in_const thy ctrss type_enc ary s T_args |> (fn T_args => IConst (name, T, T_args)) | filt _ (IAbs (bound, tm)) = IAbs (bound, filt 0 tm) | filt _ tm = tm in filt 0 end fun iformula_of_prop ctxt type_enc iff_for_eq = let val thy = Proof_Context.theory_of ctxt fun do_term bs t atomic_Ts = iterm_of_term thy type_enc bs (Envir.eta_contract t) |>> (introduce_builtins_and_proxies_in_iterm type_enc #> mangle_type_args_in_iterm type_enc #> AAtom) ||> union (op =) atomic_Ts fun do_quant bs q pos s T t' = let val s = singleton (Name.variant_list (map fst bs)) s val universal = Option.map (q = AExists ? not) pos val name = s |> `(case universal of SOME true => make_all_bound_var | SOME false => make_exist_bound_var | NONE => make_bound_var) in do_formula ((s, (name, T)) :: bs) pos t' #>> mk_aquant q [(name, SOME T)] ##> union (op =) (atomic_types_of T) end and do_conn bs c pos1 t1 pos2 t2 = do_formula bs pos1 t1 ##>> do_formula bs pos2 t2 #>> uncurry (mk_aconn c) and do_formula bs pos t = (case t of \<^const>\Trueprop\ $ t1 => do_formula bs pos t1 | \<^const>\Not\ $ t1 => do_formula bs (Option.map not pos) t1 #>> mk_anot | Const (\<^const_name>\All\, _) $ Abs (s, T, t') => do_quant bs AForall pos s T t' | (t0 as Const (\<^const_name>\All\, _)) $ t1 => do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1) | Const (\<^const_name>\Ex\, _) $ Abs (s, T, t') => do_quant bs AExists pos s T t' | (t0 as Const (\<^const_name>\Ex\, _)) $ t1 => do_formula bs pos (t0 $ eta_expand (map (snd o snd) bs) t1 1) | \<^const>\HOL.conj\ $ t1 $ t2 => do_conn bs AAnd pos t1 pos t2 | \<^const>\HOL.disj\ $ t1 $ t2 => do_conn bs AOr pos t1 pos t2 | \<^const>\HOL.implies\ $ t1 $ t2 => do_conn bs AImplies (Option.map not pos) t1 pos t2 | Const (\<^const_name>\HOL.eq\, Type (_, [\<^typ>\bool\, _])) $ t1 $ t2 => if iff_for_eq then do_conn bs AIff NONE t1 NONE t2 else do_term bs t | _ => do_term bs t) in do_formula [] end -fun presimplify_term ctxt t = +fun presimplify_term simp_options ctxt t = if exists_Const (member (op =) Meson.presimplified_consts o fst) t then t |> Skip_Proof.make_thm (Proof_Context.theory_of ctxt) - |> Meson.presimplify ctxt + |> Meson.presimplify simp_options ctxt |> Thm.prop_of else t fun preprocess_abstractions_in_terms trans_lams facts = let val (facts, lambda_ts) = facts |> map (snd o snd) |> trans_lams |>> map2 (fn (name, (role, _)) => fn t => (name, (role, t))) facts val lam_facts = map2 (fn t => fn j => ((lam_fact_prefix ^ Int.toString j, (Global, Non_Rec_Def)), (Axiom, t))) lambda_ts (1 upto length lambda_ts) in (facts, lam_facts) end (* Metis's use of "resolve_tac" freezes the schematic variables. We simulate this in Sledgehammer to prevent the discovery of unreplayable proofs. *) fun freeze_term t = let (* Freshness is desirable for completeness, but not for soundness. *) fun indexed_name (s, i) = s ^ "_" ^ string_of_int i ^ atp_weak_suffix fun freeze (t $ u) = freeze t $ freeze u | freeze (Abs (s, T, t)) = Abs (s, T, freeze t) | freeze (Var (x, T)) = Free (indexed_name x, T) | freeze t = t fun freeze_tvar (x, S) = TFree (indexed_name x, S) in t |> exists_subterm is_Var t ? freeze |> exists_type (exists_subtype is_TVar) t ? map_types (map_type_tvar freeze_tvar) end -fun presimp_prop ctxt type_enc t = +fun presimp_prop simp_options ctxt type_enc t = let val t = t |> Envir.beta_eta_contract |> transform_elim_prop |> Object_Logic.atomize_term ctxt val need_trueprop = (fastype_of t = \<^typ>\bool\) val is_ho = is_type_enc_full_higher_order type_enc in t |> need_trueprop ? HOLogic.mk_Trueprop |> (if is_ho then unextensionalize_def else cong_extensionalize_term ctxt #> abs_extensionalize_term ctxt) - |> presimplify_term ctxt + |> presimplify_term simp_options ctxt |> HOLogic.dest_Trueprop end handle TERM _ => \<^const>\True\ (* Satallax prefers "=" to "<=>" (for definitions) and Metis (CNF) requires "=" for technical reasons. *) fun should_use_iff_for_eq CNF _ = false | should_use_iff_for_eq (THF _) format = not (is_type_enc_full_higher_order format) | should_use_iff_for_eq _ _ = true fun make_formula ctxt format type_enc iff_for_eq name stature role t = let val iff_for_eq = iff_for_eq andalso should_use_iff_for_eq format type_enc val (iformula, atomic_Ts) = iformula_of_prop ctxt type_enc iff_for_eq (SOME (role <> Conjecture)) t [] |>> close_universally add_iterm_vars in {name = name, stature = stature, role = role, iformula = iformula, atomic_types = atomic_Ts} end fun make_fact ctxt format type_enc iff_for_eq ((name, stature), t) = (case make_formula ctxt format type_enc iff_for_eq name stature Axiom t of formula as {iformula = AAtom (IConst ((s, _), _, _)), ...} => if s = tptp_true then NONE else SOME formula | formula => SOME formula) fun make_conjecture ctxt format type_enc = map (fn ((name, stature), (role, t)) => let val t' = t |> role = Conjecture ? s_not in make_formula ctxt format type_enc true name stature role t' end) (** Finite and infinite type inference **) fun tvar_footprint thy s ary = (case unprefix_and_unascii const_prefix s of SOME s => let fun tvars_of T = [] |> Term.add_tvarsT T |> map fst in s |> unmangled_invert_const |> robust_const_type thy |> chop_fun ary |> fst |> map tvars_of end | NONE => []) handle TYPE _ => [] fun type_arg_cover thy pos s ary = if is_tptp_equal s then if pos = SOME false then [] else 0 upto ary - 1 else let val footprint = tvar_footprint thy s ary val eq = (s = \<^const_name>\HOL.eq\) fun cover _ [] = [] | cover seen ((i, tvars) :: args) = cover (union (op =) seen tvars) args |> (eq orelse exists (fn tvar => not (member (op =) seen tvar)) tvars) ? cons i in if forall null footprint then [] else 0 upto length footprint - 1 ~~ footprint |> sort (rev_order o list_ord Term_Ord.indexname_ord o apply2 snd) |> cover [] end type monotonicity_info = {maybe_finite_Ts : typ list, surely_infinite_Ts : typ list, maybe_nonmono_Ts : typ list} (* These types witness that the type classes they belong to allow infinite models and hence that any types with these type classes is monotonic. *) val known_infinite_types = [\<^typ>\nat\, HOLogic.intT, HOLogic.realT, \<^typ>\nat => bool\] fun is_type_kind_of_surely_infinite ctxt strictness cached_Ts T = strictness <> Strict andalso is_type_surely_infinite ctxt true cached_Ts T (* Finite types such as "unit", "bool", "bool * bool", and "bool => bool" are dangerous because their "exhaust" properties can easily lead to unsound ATP proofs. On the other hand, all HOL infinite types can be given the same models in first-order logic (via Loewenheim-Skolem). *) fun should_encode_type ctxt {maybe_finite_Ts, surely_infinite_Ts, maybe_nonmono_Ts} (Nonmono_Types (strictness, _)) T = let val thy = Proof_Context.theory_of ctxt in (exists (type_intersect thy T) maybe_nonmono_Ts andalso not (exists (type_instance thy T) surely_infinite_Ts orelse (not (member (type_equiv thy) maybe_finite_Ts T) andalso is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts T))) end | should_encode_type _ _ level _ = (level = All_Types orelse level = Undercover_Types) fun should_guard_type ctxt mono (Guards (_, level)) should_guard_var T = should_guard_var () andalso should_encode_type ctxt mono level T | should_guard_type _ _ _ _ _ = false fun is_maybe_universal_name s = String.isPrefix bound_var_prefix s orelse String.isPrefix all_bound_var_prefix s fun is_maybe_universal_var (IConst ((s, _), _, _)) = is_maybe_universal_name s | is_maybe_universal_var (IVar _) = true | is_maybe_universal_var _ = false datatype site = Top_Level of bool option | Eq_Arg of bool option | Arg of string * int * int | Elsewhere fun should_tag_with_type _ _ _ (Top_Level _) _ _ = false | should_tag_with_type ctxt mono (Tags (_, level)) site u T = let val thy = Proof_Context.theory_of ctxt in (case level of Nonmono_Types (_, Non_Uniform) => (case (site, is_maybe_universal_var u) of (Eq_Arg pos, true) => (pos <> SOME false orelse tag_neg_vars) andalso should_encode_type ctxt mono level T | _ => false) | Undercover_Types => (case (site, is_maybe_universal_var u) of (Eq_Arg pos, true) => pos <> SOME false | (Arg (s, j, ary), true) => member (op =) (type_arg_cover thy NONE s ary) j | _ => false) | _ => should_encode_type ctxt mono level T) end | should_tag_with_type _ _ _ _ _ _ = false (** predicators and application operators **) type sym_info = {pred_sym : bool, min_ary : int, max_ary : int, types : typ list, in_conj : bool} fun default_sym_tab_entries type_enc = let fun mk_sym_info pred n = {pred_sym = pred, min_ary = n, max_ary = n, types = [], in_conj = false} in (make_fixed_const NONE \<^const_name>\undefined\, mk_sym_info false 0) :: (map (rpair (mk_sym_info true 0)) [tptp_false, tptp_true]) @ (map (rpair (mk_sym_info true 1)) tptp_unary_builtins) @ (map (rpair (mk_sym_info true 2)) (tptp_old_equal :: tptp_binary_builtins)) |> not (is_type_enc_fool type_enc orelse is_type_enc_full_higher_order type_enc) ? cons (prefixed_predicator_name, mk_sym_info true 1) end datatype app_op_level = Min_App_Op | Sufficient_App_Op | Sufficient_App_Op_And_Predicator | Full_App_Op_And_Predicator fun add_iterm_syms_to_sym_table ctxt app_op_level conj_fact = let val thy = Proof_Context.theory_of ctxt fun consider_var_ary const_T var_T max_ary = let fun iter ary T = if ary = max_ary orelse type_instance thy var_T T orelse type_instance thy T var_T then ary else iter (ary + 1) (range_type T) in iter 0 const_T end fun add_universal_var T (accum as ((bool_vars, fun_var_Ts), sym_tab)) = if (app_op_level = Sufficient_App_Op andalso can dest_funT T) orelse (app_op_level = Sufficient_App_Op_And_Predicator andalso (can dest_funT T orelse T = \<^typ>\bool\)) then let val bool_vars' = bool_vars orelse (app_op_level = Sufficient_App_Op_And_Predicator andalso body_type T = \<^typ>\bool\) fun repair_min_ary {pred_sym, min_ary, max_ary, types, in_conj} = {pred_sym = pred_sym andalso not bool_vars', min_ary = fold (fn T' => consider_var_ary T' T) types min_ary, max_ary = max_ary, types = types, in_conj = in_conj} val fun_var_Ts' = fun_var_Ts |> can dest_funT T ? insert_type thy I T in if bool_vars' = bool_vars andalso fun_var_Ts' = fun_var_Ts then accum else ((bool_vars', fun_var_Ts'), Symtab.map (K repair_min_ary) sym_tab) end else accum fun add_iterm_syms top_level tm (accum as ((bool_vars, fun_var_Ts), sym_tab)) = let val (head, args) = strip_iterm_comb tm in (case head of IConst ((s, _), T, _) => if is_maybe_universal_name s then add_universal_var T accum else if String.isPrefix exist_bound_var_prefix s then accum else let val ary = length args in ((bool_vars, fun_var_Ts), (case Symtab.lookup sym_tab s of SOME {pred_sym, min_ary, max_ary, types, in_conj} => let val pred_sym = pred_sym andalso top_level andalso not bool_vars val types' = types |> insert_type thy I T val in_conj = in_conj orelse conj_fact val min_ary = if (app_op_level = Sufficient_App_Op orelse app_op_level = Sufficient_App_Op_And_Predicator) andalso types' <> types then fold (consider_var_ary T) fun_var_Ts min_ary else min_ary in Symtab.update (s, {pred_sym = pred_sym, min_ary = Int.min (ary, min_ary), max_ary = Int.max (ary, max_ary), types = types', in_conj = in_conj}) sym_tab end | NONE => let val max_ary = (case unprefix_and_unascii const_prefix s of SOME s => (if String.isSubstring uncurried_alias_sep s then ary else (case try (ary_of o robust_const_type thy o unmangled_invert_const) s of SOME ary0 => Int.min (ary0, ary) | NONE => ary)) | NONE => ary) val pred_sym = top_level andalso max_ary = ary andalso not bool_vars val min_ary = (case app_op_level of Min_App_Op => max_ary | Full_App_Op_And_Predicator => 0 | _ => fold (consider_var_ary T) fun_var_Ts max_ary) in Symtab.update_new (s, {pred_sym = pred_sym, min_ary = min_ary, max_ary = max_ary, types = [T], in_conj = conj_fact}) sym_tab end)) end | IVar (_, T) => add_universal_var T accum | IAbs ((_, T), tm) => accum |> add_universal_var T |> add_iterm_syms false tm | _ => accum) |> fold (add_iterm_syms false) args end in add_iterm_syms end fun sym_table_of_facts ctxt type_enc app_op_level conjs facts = let fun add_iterm_syms conj_fact = add_iterm_syms_to_sym_table ctxt app_op_level conj_fact true fun add_fact_syms conj_fact = ifact_lift (formula_fold NONE (K (add_iterm_syms conj_fact))) in ((false, []), Symtab.empty) |> fold (add_fact_syms true) conjs |> fold (add_fact_syms false) facts ||> fold Symtab.update (default_sym_tab_entries type_enc) end fun min_ary_of sym_tab s = (case Symtab.lookup sym_tab s of SOME ({min_ary, ...} : sym_info) => min_ary | NONE => (case unprefix_and_unascii const_prefix s of SOME s => let val s = s |> unmangled_invert_const in if s = predicator_name then 1 else if s = app_op_name then 2 else if s = type_guard_name then 1 else 0 end | NONE => 0)) (* True if the constant ever appears outside of the top-level position in literals, or if it appears with different arities (e.g., because of different type instantiations). If false, the constant always receives all of its arguments and is used as a predicate. *) fun is_pred_sym sym_tab s = (case Symtab.lookup sym_tab s of SOME ({pred_sym, min_ary, max_ary, ...} : sym_info) => pred_sym andalso min_ary = max_ary | NONE => false) val fTrue_iconst = IConst ((const_prefix ^ "fTrue", \<^const_name>\fTrue\), \<^typ>\bool\, []) val predicator_iconst = IConst (`(make_fixed_const NONE) predicator_name, \<^typ>\bool => bool\, []) fun predicatify completish tm = if completish > 1 then IApp (IApp (IConst (`I tptp_equal, \<^typ>\bool => bool => bool\, []), tm), fTrue_iconst) else IApp (predicator_iconst, tm) val app_op = `(make_fixed_const NONE) app_op_name fun list_app head args = fold (curry (IApp o swap)) args head fun mk_app_op type_enc head arg = let val head_T = ityp_of head val (arg_T, res_T) = dest_funT head_T val app = IConst (app_op, head_T --> head_T, [arg_T, res_T]) |> mangle_type_args_in_iterm type_enc in list_app app [head, arg] end fun firstorderize_fact thy ctrss type_enc uncurried_aliases completish sym_tab = let fun do_app arg head = mk_app_op type_enc head arg fun list_app_ops (head, args) = fold do_app args head fun introduce_app_ops tm = let val (head, args) = tm |> strip_iterm_comb ||> map introduce_app_ops in (case head of IConst (name as (s, _), T, T_args) => let val min_ary = min_ary_of sym_tab s val ary = if uncurried_aliases andalso String.isPrefix const_prefix s then let val ary = length args (* In polymorphic native type encodings, it is impossible to declare a fully polymorphic symbol that takes more arguments than its signature (even though such concrete instances, where a type variable is instantiated by a function type, are possible.) *) val official_ary = if is_type_enc_polymorphic type_enc then (case unprefix_and_unascii const_prefix s of SOME s' => (case try (ary_of o robust_const_type thy) (invert_const s') of SOME ary => ary | NONE => min_ary) | NONE => min_ary) else 1000000000 (* irrealistically big arity *) in Int.min (ary, official_ary) end else min_ary val head = if ary = min_ary then head else IConst (aliased_uncurried ary name, T, T_args) in args |> chop ary |>> list_app head |> list_app_ops end | _ => list_app_ops (head, args)) end fun introduce_predicators tm = (case strip_iterm_comb tm of (IConst ((s, _), _, _), _) => if is_pred_sym sym_tab s then tm else predicatify completish tm | _ => predicatify completish tm) val is_ho = is_type_enc_higher_order type_enc val is_full_ho = is_type_enc_full_higher_order type_enc val is_fool = is_type_enc_fool type_enc val do_iterm = (not is_ho ? introduce_app_ops) #> (not (is_full_ho orelse is_fool) ? introduce_predicators) #> filter_type_args_in_iterm thy ctrss type_enc in update_iformula (formula_map do_iterm) end (** Helper facts **) val not_ffalse = @{lemma "\ fFalse" by (unfold fFalse_def) fast} val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast} (* The Boolean indicates that a fairly sound type encoding is needed. *) fun helper_table with_combs = (if with_combs then [(("COMBI", false), [(Non_Rec_Def, @{thm Meson.COMBI_def})]), (("COMBK", false), [(Non_Rec_Def, @{thm Meson.COMBK_def})]), (("COMBB", false), [(Non_Rec_Def, @{thm Meson.COMBB_def})]), (("COMBC", false), [(Non_Rec_Def, @{thm Meson.COMBC_def})]), (("COMBS", false), [(Non_Rec_Def, @{thm Meson.COMBS_def})])] else []) @ [((predicator_name, false), [(General, not_ffalse), (General, ftrue)]), (("fFalse", false), [(General, not_ffalse)]), (("fFalse", true), [(General, @{thm True_or_False})]), (("fTrue", false), [(General, ftrue)]), (("fTrue", true), [(General, @{thm True_or_False})]), (("If", true), [(Non_Rec_Def, @{thm if_True}), (Non_Rec_Def, @{thm if_False}), (General, @{thm True_or_False})]), (("fNot", false), @{thms fNot_def [THEN Meson.iff_to_disjD, THEN conjunct1] fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]} |> map (pair Non_Rec_Def)), (("fconj", false), @{lemma "\ P \ \ Q \ fconj P Q" "\ fconj P Q \ P" "\ fconj P Q \ Q" by (unfold fconj_def) fast+} |> map (pair General)), (("fdisj", false), @{lemma "\ P \ fdisj P Q" "\ Q \ fdisj P Q" "\ fdisj P Q \ P \ Q" by (unfold fdisj_def) fast+} |> map (pair General)), (("fimplies", false), @{lemma "P \ fimplies P Q" "\ Q \ fimplies P Q" "\ fimplies P Q \ \ P \ Q" by (unfold fimplies_def) fast+} |> map (pair General)), (("fequal", true), (* This is a lie: Higher-order equality doesn't need a sound type encoding. However, this is done so for backward compatibility: Including the equality helpers by default in Metis breaks a few existing proofs. *) @{thms fequal_def [THEN Meson.iff_to_disjD, THEN conjunct1] fequal_def [THEN Meson.iff_to_disjD, THEN conjunct2]} |> map (pair General)), (* Partial characterization of "fAll" and "fEx". A complete characterization would require the axiom of choice for replay with Metis. *) (("fAll", false), [(General, @{lemma "\ fAll P \ P x" by (auto simp: fAll_def)})]), (("fEx", false), [(General, @{lemma "\ P x \ fEx P" by (auto simp: fEx_def)})])] |> map (apsnd (map (apsnd zero_var_indexes))) fun completish_helper_table with_combs = helper_table with_combs @ [((predicator_name, true), @{thms True_or_False fTrue_ne_fFalse} |> map (pair General)), ((app_op_name, true), [(General, @{lemma "\x. \ f x = g x \ f = g" by blast}), (General, @{lemma "\p. (p x \ p y) \ x = y" by blast})]), (("fconj", false), @{thms fconj_table fconj_laws fdisj_laws} |> map (pair Non_Rec_Def)), (("fdisj", false), @{thms fdisj_table fconj_laws fdisj_laws} |> map (pair Non_Rec_Def)), (("fimplies", false), @{thms fimplies_table fconj_laws fdisj_laws fimplies_laws} |> map (pair Non_Rec_Def)), (("fequal", false), (@{thms fequal_table} |> map (pair Non_Rec_Def)) @ (@{thms fequal_laws} |> map (pair General))), (("fAll", false), @{thms fAll_table fComp_law fAll_law fEx_law} |> map (pair Non_Rec_Def)), (("fEx", false), @{thms fEx_table fComp_law fAll_law fEx_law} |> map (pair Non_Rec_Def))] |> map (apsnd (map (apsnd zero_var_indexes))) fun bound_tvars type_enc sorts Ts = (case filter is_TVar Ts of [] => I | Ts => ((sorts andalso polymorphism_of_type_enc type_enc <> Type_Class_Polymorphic) ? mk_ahorn (Ts |> class_membs_of_types type_enc add_sorts_on_tvar |> map (class_atom type_enc))) #> (case type_enc of Native (_, _, Type_Class_Polymorphic, _) => mk_atyquant AForall (map (fn TVar (z as (_, S)) => (AType ((tvar_name z, []), []), map (`make_class) (normalize_classes S) )) Ts) | Native (_, _, Raw_Polymorphic _, _) => mk_atyquant AForall (map (fn TVar (z as _) => (AType ((tvar_name z, []), []), [])) Ts) | _ => mk_aquant AForall (map (fn TVar z => (tvar_name z, NONE)) Ts))) fun eq_formula type_enc atomic_Ts bounds pred_sym tm1 tm2 = (if pred_sym then AConn (AIff, [AAtom tm1, AAtom tm2]) else AAtom (ATerm ((`I tptp_equal, []), [tm1, tm2]))) |> mk_aquant AForall bounds |> close_formula_universally |> bound_tvars type_enc true atomic_Ts val helper_rank = default_rank val min_rank = 9 * helper_rank div 10 val max_rank = 4 * min_rank fun rank_of_fact_num n j = min_rank + (max_rank - min_rank) * j div n val type_tag = `(make_fixed_const NONE) type_tag_name fun could_specialize_helpers type_enc = not (is_type_enc_polymorphic type_enc) andalso level_of_type_enc type_enc <> No_Types fun should_specialize_helper type_enc t = could_specialize_helpers type_enc andalso not (null (Term.hidden_polymorphism t)) fun add_helper_facts_of_sym ctxt format type_enc lam_trans completish (s, {types, ...} : sym_info) = (case unprefix_and_unascii const_prefix s of SOME mangled_s => let val thy = Proof_Context.theory_of ctxt val unmangled_s = mangled_s |> unmangled_const_name |> hd fun dub needs_sound j k = ascii_of unmangled_s ^ "_" ^ string_of_int j ^ "_" ^ string_of_int k ^ (if mangled_s = unmangled_s then "" else "_" ^ ascii_of mangled_s) ^ (if needs_sound then typed_helper_suffix else untyped_helper_suffix) fun specialize_helper t T = if unmangled_s = app_op_name then let val tyenv = Sign.typ_match thy (alpha_to_beta, domain_type T) Vartab.empty in Envir.subst_term_types tyenv t end else specialize_type thy (invert_const unmangled_s, T) t fun dub_and_inst needs_sound ((status, t), j) = (if should_specialize_helper type_enc t then map_filter (try (specialize_helper t)) types else [t]) |> tag_list 1 |> map (fn (k, t) => ((dub needs_sound j k, (Global, status)), t)) val make_facts = map_filter (make_fact ctxt format type_enc false) val sound = is_type_enc_sound type_enc val could_specialize = could_specialize_helpers type_enc val with_combs = lam_trans <> opaque_combsN in fold (fn ((helper_s, needs_sound), ths) => if (needs_sound andalso not sound) orelse (helper_s <> unmangled_s andalso (completish < 3 orelse could_specialize)) then I else ths ~~ (1 upto length ths) |> maps (dub_and_inst needs_sound o apfst (apsnd Thm.prop_of)) |> make_facts |> union (op = o apply2 #iformula)) ((if completish >= 3 then completish_helper_table else helper_table) with_combs) end | NONE => I) fun helper_facts_of_sym_table ctxt format type_enc lam_trans completish sym_tab = Symtab.fold_rev (add_helper_facts_of_sym ctxt format type_enc lam_trans completish) sym_tab [] (***************************************************************) (* Type Classes Present in the Axiom or Conjecture Clauses *) (***************************************************************) fun set_insert (x, s) = Symtab.update (x, ()) s fun add_classes (cls, cset) = List.foldl set_insert cset (flat cls) fun classes_of_terms get_Ts = map (map snd o get_Ts) #> List.foldl add_classes Symtab.empty #> Symtab.delete_safe class_of_types #> Symtab.keys val tfree_classes_of_terms = classes_of_terms Misc_Legacy.term_tfrees val tvar_classes_of_terms = classes_of_terms Misc_Legacy.term_tvars fun fold_type_ctrs f (Type (s, Ts)) x = fold (fold_type_ctrs f) Ts (f (s, x)) | fold_type_ctrs _ _ x = x (* Type constructors used to instantiate overloaded constants are the only ones needed. *) fun add_type_ctrs_in_term thy = let fun add (Const (\<^const_name>\Meson.skolem\, _) $ _) = I | add (t $ u) = add t #> add u | add (Const x) = x |> robust_const_type_args thy |> fold (fold_type_ctrs set_insert) | add (Abs (_, _, u)) = add u | add _ = I in add end fun type_ctrs_of_terms thy ts = Symtab.keys (fold (add_type_ctrs_in_term thy) ts Symtab.empty) fun trans_lams_of_string ctxt type_enc lam_trans = if lam_trans = no_lamsN then rpair [] else if lam_trans = opaque_liftingN then lift_lams ctxt type_enc ##> K [] else if lam_trans = liftingN then lift_lams ctxt type_enc else if lam_trans = opaque_combsN orelse lam_trans = combsN then map (introduce_combinators ctxt type_enc) #> rpair [] else if lam_trans = combs_and_liftingN then lift_lams_part_1 ctxt type_enc ##> maps (fn t => [t, introduce_combinators ctxt type_enc (intentionalize_def t)]) #> lift_lams_part_2 ctxt type_enc else if lam_trans = combs_or_liftingN then lift_lams_part_1 ctxt type_enc ##> map (fn t => (case head_of (strip_qnt_body \<^const_name>\All\ t) of \<^term>\(=) ::bool => bool => bool\ => t | _ => introduce_combinators ctxt type_enc (intentionalize_def t))) #> lift_lams_part_2 ctxt type_enc else if lam_trans = keep_lamsN then map (Envir.eta_contract) #> rpair [] else error ("Unknown lambda translation scheme: " ^ quote lam_trans) val pull_and_reorder_definitions = let fun add_consts (IApp (t, u)) = fold add_consts [t, u] | add_consts (IAbs (_, t)) = add_consts t | add_consts (IConst (name, _, _)) = insert (op =) name | add_consts (IVar _) = I fun consts_of_hs l_or_r ({iformula, ...} : ifact) = (case iformula of AAtom (IApp (IApp (IConst _, t), u)) => add_consts (l_or_r (t, u)) [] | _ => []) (* Quadratic, but usually OK. *) fun reorder [] [] = [] | reorder (fact :: skipped) [] = fact :: reorder [] skipped (* break cycle *) | reorder skipped (fact :: facts) = let val rhs_consts = consts_of_hs snd fact in if exists (exists (exists (member (op =) rhs_consts) o consts_of_hs fst)) [skipped, facts] then reorder (fact :: skipped) facts else fact :: reorder [] (facts @ skipped) end in List.partition (curry (op =) Definition o #role) #>> reorder [] #> op @ end fun s_not_prop (\<^const>\Trueprop\ $ t) = \<^const>\Trueprop\ $ s_not t | s_not_prop (\<^const>\Pure.imp\ $ t $ \<^prop>\False\) = t | s_not_prop t = \<^const>\Pure.imp\ $ t $ \<^prop>\False\ -fun translate_formulas ctxt prem_role format type_enc lam_trans presimp hyp_ts concl_t facts = +fun translate_formulas simp_options ctxt prem_role format type_enc lam_trans presimp hyp_ts concl_t facts = let val thy = Proof_Context.theory_of ctxt val trans_lams = trans_lams_of_string ctxt type_enc lam_trans val fact_ts = facts |> map snd (* Remove existing facts from the conjecture, as this can dramatically boost an ATP's performance (for some reason). *) val hyp_ts = hyp_ts |> map (fn t => if member (op aconv) fact_ts t then \<^prop>\True\ else t) val hyp_ts = map freeze_term hyp_ts; val concl_t = freeze_term concl_t; val facts = facts |> map (apsnd (pair Axiom)) val conjs = map (pair prem_role) hyp_ts @ [(Conjecture, s_not_prop concl_t)] |> map2 (pair o rpair (Local, General) o string_of_int) (0 upto length hyp_ts) val ((conjs, facts), lam_facts) = (conjs, facts) - |> presimp ? apply2 (map (apsnd (apsnd (presimp_prop ctxt type_enc)))) + |> presimp ? apply2 (map (apsnd (apsnd (presimp_prop simp_options ctxt type_enc)))) |> (if lam_trans = no_lamsN then rpair [] else op @ #> preprocess_abstractions_in_terms trans_lams #>> chop (length conjs)) val conjs = conjs |> make_conjecture ctxt format type_enc |> pull_and_reorder_definitions val facts = facts |> map_filter (fn (name, (_, t)) => make_fact ctxt format type_enc true (name, t)) |> pull_and_reorder_definitions val lifted = lam_facts |> map (extract_lambda_def dest_Const o snd o snd) val lam_facts = lam_facts |> map_filter (make_fact ctxt format type_enc true o apsnd snd) val all_ts = concl_t :: hyp_ts @ fact_ts val subs = tfree_classes_of_terms all_ts val supers = tvar_classes_of_terms all_ts val tycons = type_ctrs_of_terms thy all_ts val (supers, tcon_clauses) = if level_of_type_enc type_enc = No_Types then ([], []) else make_tcon_clauses thy tycons supers val subclass_pairs = make_subclass_pairs thy subs supers in (union (op =) subs supers, conjs, facts @ lam_facts, subclass_pairs, tcon_clauses, lifted) end val type_guard = `(make_fixed_const NONE) type_guard_name fun type_guard_iterm type_enc T tm = IApp (IConst (type_guard, T --> \<^typ>\bool\, [T]) |> mangle_type_args_in_iterm type_enc, tm) fun is_var_positively_naked_in_term _ (SOME false) _ accum = accum | is_var_positively_naked_in_term name _ (ATerm (((s, _), _), tms)) accum = accum orelse (is_tptp_equal s andalso member (op =) tms (ATerm ((name, []), []))) | is_var_positively_naked_in_term _ _ _ _ = true fun is_var_undercover_in_term thy name pos tm accum = accum orelse let val var = ATerm ((name, []), []) fun is_undercover (ATerm (_, [])) = false | is_undercover (ATerm (((s, _), _), tms)) = let val ary = length tms val cover = type_arg_cover thy pos s ary in exists (fn (j, tm) => tm = var andalso member (op =) cover j) (0 upto ary - 1 ~~ tms) orelse exists is_undercover tms end | is_undercover _ = true in is_undercover tm end fun should_guard_var_in_formula thy level pos phi (SOME true) name = (case level of All_Types => true | Undercover_Types => formula_fold pos (is_var_undercover_in_term thy name) phi false | Nonmono_Types (_, Uniform) => true | Nonmono_Types (_, Non_Uniform) => formula_fold pos (is_var_positively_naked_in_term name) phi false | _ => false) | should_guard_var_in_formula _ _ _ _ _ _ = true fun always_guard_var_in_formula _ _ _ _ _ _ = true fun should_generate_tag_bound_decl _ _ _ (SOME true) _ = false | should_generate_tag_bound_decl ctxt mono (Tags (_, level)) _ T = not (is_type_level_uniform level) andalso should_encode_type ctxt mono level T | should_generate_tag_bound_decl _ _ _ _ _ = false fun mk_aterm type_enc name T_args args = let val (ty_args, tm_args) = process_type_args type_enc T_args in ATerm ((name, ty_args), tm_args @ args) end fun do_bound_type type_enc = (case type_enc of Native _ => native_atp_type_of_typ type_enc false 0 #> SOME | _ => K NONE) fun tag_with_type ctxt mono type_enc pos T tm = IConst (type_tag, T --> T, [T]) |> mangle_type_args_in_iterm type_enc |> atp_term_of_iterm ctxt mono type_enc pos |> (fn ATerm ((s, tys), tms) => ATerm ((s, tys), tms @ [tm]) | _ => raise Fail "unexpected lambda-abstraction") and atp_term_of_iterm ctxt mono type_enc pos = let fun term site u = let val (head, args) = strip_iterm_comb u val pos = (case site of Top_Level pos => pos | Eq_Arg pos => pos | _ => NONE) val T = ityp_of u val t = (case head of IConst (name as (s, _), _, T_args) => let val ary = length args fun arg_site j = if is_tptp_equal s then Eq_Arg pos else Arg (s, j, ary) in map2 (fn j => term (arg_site j)) (0 upto ary - 1) args |> mk_aterm type_enc name T_args end | IVar (name, _) => map (term Elsewhere) args |> mk_aterm type_enc name [] | IAbs ((name, T), tm) => if is_type_enc_higher_order type_enc then AAbs (((name, native_atp_type_of_typ type_enc false 0 T), term Elsewhere tm), map (term Elsewhere) args) else raise Fail "unexpected lambda-abstraction" | IApp _ => raise Fail "impossible \"IApp\"") val tag = should_tag_with_type ctxt mono type_enc site u T in t |> tag ? tag_with_type ctxt mono type_enc pos T end in term (Top_Level pos) end and formula_of_iformula ctxt mono type_enc should_guard_var = let val thy = Proof_Context.theory_of ctxt val level = level_of_type_enc type_enc val do_term = atp_term_of_iterm ctxt mono type_enc fun do_out_of_bound_type pos phi universal (name, T) = if should_guard_type ctxt mono type_enc (fn () => should_guard_var thy level pos phi universal name) T then IVar (name, T) |> type_guard_iterm type_enc T |> do_term pos |> AAtom |> SOME else if should_generate_tag_bound_decl ctxt mono type_enc universal T then let val var = ATerm ((name, []), []) val tagged_var = tag_with_type ctxt mono type_enc pos T var in SOME (AAtom (ATerm ((`I tptp_equal, []), [tagged_var, var]))) end else NONE fun do_formula pos (ATyQuant (q, xs, phi)) = ATyQuant (q, map (apfst (native_atp_type_of_typ type_enc false 0)) xs, do_formula pos phi) | do_formula pos (AQuant (q, xs, phi)) = let val phi = phi |> do_formula pos val universal = Option.map (q = AExists ? not) pos in AQuant (q, xs |> map (apsnd (fn NONE => NONE | SOME T => do_bound_type type_enc T)), (if q = AForall then mk_ahorn else fold_rev (mk_aconn AAnd)) (map_filter (fn (_, NONE) => NONE | (s, SOME T) => do_out_of_bound_type pos phi universal (s, T)) xs) phi) end | do_formula pos (AConn conn) = aconn_map pos do_formula conn | do_formula pos (AAtom tm) = AAtom (do_term pos tm) in do_formula end fun string_of_status General = "" | string_of_status Induction = inductionN | string_of_status Intro = introN | string_of_status Inductive = inductiveN | string_of_status Elim = elimN | string_of_status Simp = simpN | string_of_status Non_Rec_Def = non_rec_defN | string_of_status Rec_Def = rec_defN (* Each fact is given a unique fact number to avoid name clashes (e.g., because of monomorphization). The TPTP forbids name clashes, and some of the remote provers might care. *) fun line_of_fact ctxt generate_info prefix encode alt freshen pos mono type_enc rank (j, {name, stature = (_, status), role, iformula, atomic_types}) = Formula ((prefix ^ (if freshen then string_of_int j ^ "_" else "") ^ encode name, alt name), role, iformula |> formula_of_iformula ctxt mono type_enc should_guard_var_in_formula (if pos then SOME true else NONE) |> close_formula_universally |> bound_tvars type_enc true atomic_types, NONE, isabelle_info generate_info (string_of_status status) (rank j)) fun lines_of_subclass generate_info type_enc sub super = Formula ((subclass_prefix ^ ascii_of sub ^ "___" ^ ascii_of super, ""), Axiom, AConn (AImplies, [sub, super] |> map (fn s => class_atom type_enc (s, tvar_a))) |> bound_tvars type_enc false [tvar_a], NONE, isabelle_info generate_info inductiveN helper_rank) fun lines_of_subclass_pair generate_info type_enc (sub, supers) = if polymorphism_of_type_enc type_enc = Type_Class_Polymorphic then [Class_Decl (class_decl_prefix ^ ascii_of sub, `make_class sub, map (`make_class) supers)] else map (lines_of_subclass generate_info type_enc sub) supers fun line_of_tcon_clause generate_info type_enc (name, prems, (cl, T)) = if polymorphism_of_type_enc type_enc = Type_Class_Polymorphic then Class_Memb (class_memb_prefix ^ name, map (fn (cls, T) => (T |> dest_TVar |> tvar_name, map (`make_class) cls)) prems, native_atp_type_of_typ type_enc false 0 T, `make_class cl) else Formula ((tcon_clause_prefix ^ name, ""), Axiom, mk_ahorn (maps (class_atoms type_enc) prems) (class_atom type_enc (cl, T)) |> bound_tvars type_enc true (snd (dest_Type T)), NONE, isabelle_info generate_info inductiveN helper_rank) fun line_of_conjecture ctxt mono type_enc ({name, role, iformula, atomic_types, ...} : ifact) = Formula ((conjecture_prefix ^ name, ""), role, iformula |> formula_of_iformula ctxt mono type_enc should_guard_var_in_formula (SOME false) |> close_formula_universally |> bound_tvars type_enc true atomic_types, NONE, []) fun lines_of_free_types type_enc (facts : ifact list) = if is_type_enc_polymorphic type_enc then let val type_classes = (polymorphism_of_type_enc type_enc = Type_Class_Polymorphic) fun line j (cl, T) = if type_classes then Class_Memb (class_memb_prefix ^ string_of_int j, [], native_atp_type_of_typ type_enc false 0 T, `make_class cl) else Formula ((tfree_clause_prefix ^ string_of_int j, ""), Hypothesis, class_atom type_enc (cl, T), NONE, []) val membs = fold (union (op =)) (map #atomic_types facts) [] |> class_membs_of_types type_enc add_sorts_on_tfree in map2 line (0 upto length membs - 1) membs end else [] (** Symbol declarations **) fun decl_line_of_class phantoms s = let val name as (s, _) = `make_class s in Sym_Decl (sym_decl_prefix ^ s, name, APi ([tvar_a_name], if phantoms = Without_Phantom_Type_Vars then AFun (a_itself_atype, bool_atype) else bool_atype)) end fun decl_lines_of_classes type_enc = (case type_enc of Native (_, _, Raw_Polymorphic phantoms, _) => map (decl_line_of_class phantoms) | _ => K []) fun sym_decl_table_of_facts thy type_enc sym_tab (conjs, facts, extra_tms) = let fun add_iterm_syms tm = let val (head, args) = strip_iterm_comb tm in (case head of IConst ((s, s'), T, T_args) => let val (pred_sym, in_conj) = (case Symtab.lookup sym_tab s of SOME ({pred_sym, in_conj, ...} : sym_info) => (pred_sym, in_conj) | NONE => (false, false)) val decl_sym = (case type_enc of Guards _ => not pred_sym | _ => true) andalso is_tptp_user_symbol s in if decl_sym then Symtab.map_default (s, []) (insert_type thy #3 (s', T_args, T, pred_sym, length args, in_conj)) else I end | IAbs (_, tm) => add_iterm_syms tm | _ => I) #> fold add_iterm_syms args end val add_fact_syms = ifact_lift (formula_fold NONE (K add_iterm_syms)) fun add_formula_var_types (ATyQuant (_, _, phi)) = add_formula_var_types phi | add_formula_var_types (AQuant (_, xs, phi)) = fold (fn (_, SOME T) => insert_type thy I T | _ => I) xs #> add_formula_var_types phi | add_formula_var_types (AConn (_, phis)) = fold add_formula_var_types phis | add_formula_var_types _ = I fun var_types () = if is_type_enc_polymorphic type_enc then [tvar_a] else fold (ifact_lift add_formula_var_types) (conjs @ facts) [] fun add_undefined_const T = let (* FIXME: make sure type arguments are filtered / clean up code *) val (s, s') = `(make_fixed_const NONE) \<^const_name>\undefined\ |> (is_type_enc_mangling type_enc ? mangled_const_name type_enc [T]) in Symtab.map_default (s, []) (insert_type thy #3 (s', [T], T, false, 0, false)) end fun add_TYPE_const () = let val (s, s') = TYPE_name in Symtab.map_default (s, []) (insert_type thy #3 (s', [tvar_a], \<^typ>\'a itself\, false, 0, false)) end in Symtab.empty |> is_type_enc_sound type_enc ? (fold (fold add_fact_syms) [conjs, facts] #> fold add_iterm_syms extra_tms #> (case type_enc of Native (_, _, Raw_Polymorphic phantoms, _) => phantoms = Without_Phantom_Type_Vars ? add_TYPE_const () | Native _ => I | _ => fold add_undefined_const (var_types ()))) end (* We add "bool" in case the helper "True_or_False" is included later. *) fun default_mono level completish = {maybe_finite_Ts = [\<^typ>\bool\], surely_infinite_Ts = (case level of Nonmono_Types (Strict, _) => [] | _ => known_infinite_types), maybe_nonmono_Ts = [if completish >= 3 then tvar_a else \<^typ>\bool\]} (* This inference is described in section 4 of Blanchette et al., "Encoding monomorphic and polymorphic types", TACAS 2013. *) fun add_iterm_mononotonicity_info ctxt level polarity tm (mono as {maybe_finite_Ts, surely_infinite_Ts, maybe_nonmono_Ts}) = let val thy = Proof_Context.theory_of ctxt fun update_mono T mono = (case level of Nonmono_Types (strictness, _) => if exists (type_instance thy T) surely_infinite_Ts orelse member (type_equiv thy) maybe_finite_Ts T then mono else if is_type_kind_of_surely_infinite ctxt strictness surely_infinite_Ts T then {maybe_finite_Ts = maybe_finite_Ts, surely_infinite_Ts = surely_infinite_Ts |> insert_type thy I T, maybe_nonmono_Ts = maybe_nonmono_Ts} else {maybe_finite_Ts = maybe_finite_Ts |> insert (type_equiv thy) T, surely_infinite_Ts = surely_infinite_Ts, maybe_nonmono_Ts = maybe_nonmono_Ts |> insert_type thy I T} | _ => mono) fun update_mono_rec (IConst ((_, s'), Type (_, [T, _]), _)) = if String.isPrefix \<^const_name>\fequal\ s' then update_mono T else I | update_mono_rec (IApp (tm1, tm2)) = fold update_mono_rec [tm1, tm2] | update_mono_rec (IAbs (_, tm)) = update_mono_rec tm | update_mono_rec _ = I in mono |> (case tm of IApp (IApp (IConst ((s, _), Type (_, [T, _]), _), tm1), tm2) => ((polarity <> SOME false andalso is_tptp_equal s andalso exists is_maybe_universal_var [tm1, tm2]) ? update_mono T) #> fold update_mono_rec [tm1, tm2] | _ => update_mono_rec tm) end fun add_fact_mononotonicity_info ctxt level ({role, iformula, ...} : ifact) = formula_fold (SOME (role <> Conjecture)) (add_iterm_mononotonicity_info ctxt level) iformula fun mononotonicity_info_of_facts ctxt type_enc completish facts = let val level = level_of_type_enc type_enc in default_mono level completish |> is_type_level_monotonicity_based level ? fold (add_fact_mononotonicity_info ctxt level) facts end fun fold_arg_types f (IApp (tm1, tm2)) = fold_arg_types f tm1 #> fold_term_types f tm2 | fold_arg_types _ _ = I and fold_term_types f tm = f (ityp_of tm) #> fold_arg_types f tm fun add_iformula_monotonic_types ctxt mono type_enc = let val thy = Proof_Context.theory_of ctxt val level = level_of_type_enc type_enc val should_encode = should_encode_type ctxt mono level fun add_type T = not (should_encode T) ? insert_type thy I T in formula_fold NONE (K (fold_term_types add_type)) end fun add_fact_monotonic_types ctxt mono type_enc = ifact_lift (add_iformula_monotonic_types ctxt mono type_enc) fun monotonic_types_of_facts ctxt mono type_enc facts = let val level = level_of_type_enc type_enc in [] |> (is_type_enc_polymorphic type_enc andalso is_type_level_monotonicity_based level) ? fold (add_fact_monotonic_types ctxt mono type_enc) facts end fun line_of_guards_mono_type ctxt generate_info mono type_enc T = Formula ((guards_sym_formula_prefix ^ ascii_of (mangled_type type_enc T), ""), Axiom, IConst (`make_bound_var "X", T, []) |> type_guard_iterm type_enc T |> AAtom |> formula_of_iformula ctxt mono type_enc always_guard_var_in_formula (SOME true) |> close_formula_universally |> bound_tvars type_enc true (atomic_types_of T), NONE, isabelle_info generate_info inductiveN helper_rank) fun line_of_tags_mono_type ctxt generate_info mono type_enc T = let val x_var = ATerm ((`make_bound_var "X", []), []) in Formula ((tags_sym_formula_prefix ^ ascii_of (mangled_type type_enc T), ""), Axiom, eq_formula type_enc (atomic_types_of T) [] false (tag_with_type ctxt mono type_enc NONE T x_var) x_var, NONE, isabelle_info generate_info non_rec_defN helper_rank) end fun lines_of_mono_types ctxt generate_info mono type_enc = (case type_enc of Native _ => K [] | Guards _ => map (line_of_guards_mono_type ctxt generate_info mono type_enc) | Tags _ => map (line_of_tags_mono_type ctxt generate_info mono type_enc)) fun decl_line_of_sym ctxt type_enc s (s', T_args, T, pred_sym, ary, _) = let val thy = Proof_Context.theory_of ctxt val (T, T_args) = if null T_args then (T, []) else (case unprefix_and_unascii const_prefix s of SOME s' => let val s' = s' |> unmangled_invert_const val T = s' |> robust_const_type thy in (T, robust_const_type_args thy (s', T)) end | NONE => raise Fail "unexpected type arguments") in Sym_Decl (sym_decl_prefix ^ s, (s, s'), T |> native_atp_type_of_typ type_enc pred_sym ary |> not (null T_args) ? curry APi (map (tvar_name o dest_TVar) T_args)) end fun honor_conj_sym_role in_conj = (if in_conj then Hypothesis else Axiom, I) fun line_of_guards_sym_decl ctxt generate_info mono type_enc n s j (s', T_args, T, _, ary, in_conj) = let val thy = Proof_Context.theory_of ctxt val (role, maybe_negate) = honor_conj_sym_role in_conj val (arg_Ts, res_T) = chop_fun ary T val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int) val bounds = bound_names ~~ arg_Ts |> map (fn (name, T) => IConst (name, T, [])) val bound_Ts = (case level_of_type_enc type_enc of All_Types => if null T_args then replicate ary NONE else map SOME arg_Ts | Undercover_Types => let val cover = type_arg_cover thy NONE s ary in map2 (fn j => if member (op =) cover j then SOME else K NONE) (0 upto ary - 1) arg_Ts end | _ => replicate ary NONE) in Formula ((guards_sym_formula_prefix ^ s ^ (if n > 1 then "_" ^ string_of_int j else ""), ""), role, IConst ((s, s'), T, T_args) |> fold (curry (IApp o swap)) bounds |> type_guard_iterm type_enc res_T |> AAtom |> mk_aquant AForall (bound_names ~~ bound_Ts) |> formula_of_iformula ctxt mono type_enc always_guard_var_in_formula (SOME true) |> close_formula_universally |> bound_tvars type_enc (n > 1) (atomic_types_of T) |> maybe_negate, NONE, isabelle_info generate_info inductiveN helper_rank) end fun lines_of_tags_sym_decl ctxt generate_info mono type_enc n s (j, (s', T_args, T, pred_sym, ary, in_conj)) = let val thy = Proof_Context.theory_of ctxt val level = level_of_type_enc type_enc val ident = tags_sym_formula_prefix ^ s ^ (if n > 1 then "_" ^ string_of_int j else "") val (role, maybe_negate) = honor_conj_sym_role in_conj val (arg_Ts, res_T) = chop_fun ary T val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int) val bounds = bound_names |> map (fn name => ATerm ((name, []), [])) val cst = mk_aterm type_enc (s, s') T_args val eq = maybe_negate oo eq_formula type_enc (atomic_types_of T) [] pred_sym val tag_with = tag_with_type ctxt mono type_enc NONE fun formula c = [Formula ((ident, ""), role, eq (tag_with res_T c) c, NONE, isabelle_info generate_info non_rec_defN helper_rank)] in if pred_sym orelse not (should_encode_type ctxt mono level res_T) then [] else if level = Undercover_Types then let val cover = type_arg_cover thy NONE s ary fun maybe_tag (j, arg_T) = member (op =) cover j ? tag_with arg_T val bounds = bounds |> map2 maybe_tag (0 upto ary - 1 ~~ arg_Ts) in formula (cst bounds) end else formula (cst bounds) end fun result_type_of_decl (_, _, T, _, ary, _) = chop_fun ary T |> snd fun rationalize_decls thy (decls as decl :: (decls' as _ :: _)) = let val T = result_type_of_decl decl |> map_type_tvar (fn (z, _) => TVar (z, \<^sort>\type\)) in if forall (type_generalization thy T o result_type_of_decl) decls' then [decl] else decls end | rationalize_decls _ decls = decls fun lines_of_sym_decls ctxt generate_info mono type_enc (s, decls) = (case type_enc of Native _ => [decl_line_of_sym ctxt type_enc s (hd decls)] | Guards (_, level) => let val thy = Proof_Context.theory_of ctxt val decls = decls |> rationalize_decls thy val n = length decls val decls = decls |> filter (should_encode_type ctxt mono level o result_type_of_decl) in (0 upto length decls - 1, decls) |-> map2 (line_of_guards_sym_decl ctxt generate_info mono type_enc n s) end | Tags (_, level) => if is_type_level_uniform level then [] else let val n = length decls in (0 upto n - 1 ~~ decls) |> maps (lines_of_tags_sym_decl ctxt generate_info mono type_enc n s) end) fun lines_of_sym_decl_table ctxt generate_info mono type_enc mono_Ts sym_decl_tab = let val syms = sym_decl_tab |> Symtab.dest |> sort_by fst val mono_lines = lines_of_mono_types ctxt generate_info mono type_enc mono_Ts val decl_lines = maps (lines_of_sym_decls ctxt generate_info mono type_enc) syms in mono_lines @ decl_lines end fun datatypes_of_sym_table ctxt ctrss (DFG Polymorphic) (type_enc as Native _) uncurried_aliases sym_tab = if is_type_enc_polymorphic type_enc then let val thy = Proof_Context.theory_of ctxt fun do_ctr (s, T) = let val s' = make_fixed_const (SOME type_enc) s val ary = ary_of T fun mk name = SOME (mk_aterm type_enc name (robust_const_type_args thy (s, T)) []) in if T = HOLogic.boolT then (case proxify_const s' of SOME proxy_base => mk (proxy_base |>> prefix const_prefix) | NONE => NONE) else (case Symtab.lookup sym_tab s' of NONE => NONE | SOME ({min_ary, ...} : sym_info) => if ary = min_ary then mk (s', s) else if uncurried_aliases then mk (aliased_uncurried ary (s', s)) else NONE) end fun datatype_of_ctrs (ctrs as (_, T1) :: _) = let val ctrs' = map do_ctr ctrs in (native_atp_type_of_typ type_enc false 0 (body_type T1), map_filter I ctrs', forall is_some ctrs') end in ctrss |> map datatype_of_ctrs |> filter #3 end else [] | datatypes_of_sym_table _ _ _ _ _ _ = [] fun decl_line_of_datatype (ty as AType (((_, s'), _), ty_args), ctrs, exhaust) = let val xs = map (fn AType ((name, _), []) => name) ty_args in Datatype_Decl (datatype_decl_prefix ^ ascii_of s', map (rpair []) xs, ty, ctrs, exhaust) end fun pair_append (xs1, xs2) (ys1, ys2) = (xs1 @ ys1, xs2 @ ys2) fun do_uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab base_s0 types in_conj = let fun do_alias ary = let val thy = Proof_Context.theory_of ctxt val (role, maybe_negate) = honor_conj_sym_role in_conj val base_name = base_s0 |> `(make_fixed_const (SOME type_enc)) val T = (case types of [T] => T | _ => robust_const_type thy base_s0) val T_args = robust_const_type_args thy (base_s0, T) val (base_name as (base_s, _), T_args) = mangle_type_args_in_const type_enc base_name T_args val base_ary = min_ary_of sym_tab0 base_s fun do_const name = IConst (name, T, T_args) val filter_ty_args = filter_type_args_in_iterm thy ctrss type_enc val atp_term_of = atp_term_of_iterm ctxt mono type_enc (SOME true) val name1 as (s1, _) = base_name |> ary - 1 > base_ary ? aliased_uncurried (ary - 1) val name2 as (s2, _) = base_name |> aliased_uncurried ary val (arg_Ts, _) = chop_fun ary T val bound_names = 1 upto ary |> map (`I o make_bound_var o string_of_int) val bounds = bound_names ~~ arg_Ts val (first_bounds, last_bound) = bounds |> map (fn (name, T) => IConst (name, T, [])) |> split_last val tm1 = mk_app_op type_enc (list_app (do_const name1) first_bounds) last_bound |> filter_ty_args val tm2 = list_app (do_const name2) (first_bounds @ [last_bound]) |> filter_ty_args val eq = eq_formula type_enc (atomic_types_of T) (map (apsnd (do_bound_type type_enc)) bounds) false (atp_term_of tm1) (atp_term_of tm2) in ([tm1, tm2], [Formula ((uncurried_alias_eq_prefix ^ s2, ""), role, eq |> maybe_negate, NONE, isabelle_info generate_info non_rec_defN helper_rank)]) |> (if ary - 1 = base_ary orelse Symtab.defined sym_tab s1 then I else pair_append (do_alias (ary - 1))) end in do_alias end fun uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab (s, {min_ary, types, in_conj, ...} : sym_info) = (case unprefix_and_unascii const_prefix s of SOME mangled_s => if String.isSubstring uncurried_alias_sep mangled_s then let val base_s0 = mangled_s |> unmangled_invert_const in do_uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab base_s0 types in_conj min_ary end else ([], []) | NONE => ([], [])) fun uncurried_alias_lines_of_sym_table ctxt generate_info ctrss mono type_enc uncurried_aliases sym_tab0 sym_tab = ([], []) |> uncurried_aliases ? Symtab.fold_rev (pair_append o uncurried_alias_lines_of_sym ctxt generate_info ctrss mono type_enc sym_tab0 sym_tab) sym_tab val implicit_declsN = "Could-be-implicit typings" val explicit_declsN = "Explicit typings" val uncurried_alias_eqsN = "Uncurried aliases" val factsN = "Relevant facts" val subclassesN = "Subclasses" val tconsN = "Type constructors" val helpersN = "Helper facts" val conjsN = "Conjectures" val free_typesN = "Free types" (* TFF allows implicit declarations of types, function symbols, and predicate symbols (with "$i" as the type of individuals), but some provers (e.g., SNARK) require explicit declarations. The situation is similar for THF. *) fun default_type pred_sym = let fun typ 0 0 = if pred_sym then bool_atype else individual_atype | typ 0 tm_ary = AFun (individual_atype, typ 0 (tm_ary - 1)) | typ ty_ary tm_ary = APi (replicate ty_ary tvar_a_name, typ 0 tm_ary) in typ end fun undeclared_in_problem problem = let fun do_sym (name as (s, _)) value = if is_tptp_user_symbol s then Symtab.default (s, (name, value)) else I fun do_class name = apfst (apfst (do_sym name ())) val do_bound_tvars = fold do_class o snd fun do_type (AType ((name, _), tys)) = apfst (apsnd (do_sym name (length tys))) #> fold do_type tys | do_type (AFun (ty1, ty2)) = do_type ty1 #> do_type ty2 | do_type (APi (_, ty)) = do_type ty fun do_term pred_sym (ATerm ((name, tys), tms)) = apsnd (do_sym name (fn _ => default_type pred_sym (length tys) (length tms))) #> fold do_type tys #> fold (do_term false) tms | do_term _ (AAbs (((_, ty), tm), args)) = do_type ty #> do_term false tm #> fold (do_term false) args fun do_formula (ATyQuant (_, xs, phi)) = fold (do_type o fst) xs #> fold (fold do_class o snd) xs #> do_formula phi | do_formula (AQuant (_, xs, phi)) = fold do_type (map_filter snd xs) #> do_formula phi | do_formula (AConn (_, phis)) = fold do_formula phis | do_formula (AAtom tm) = do_term true tm fun do_line (Class_Decl (_, _, cls)) = fold do_class cls | do_line (Type_Decl _) = I | do_line (Sym_Decl (_, _, ty)) = do_type ty | do_line (Datatype_Decl (_, xs, ty, tms, _)) = fold do_bound_tvars xs #> do_type ty #> fold (do_term false) tms | do_line (Class_Memb (_, xs, ty, cl)) = fold do_bound_tvars xs #> do_type ty #> do_class cl | do_line (Formula (_, _, phi, _, _)) = do_formula phi val ((cls, tys), syms) = declared_in_atp_problem problem in ((Symtab.empty, Symtab.empty), Symtab.empty) |>> apfst (fold (fn (s, _) => Symtab.default (s, (("", ""), ()))) cls) |>> apsnd (fold (fn (s, _) => Symtab.default (s, (("", ""), 0))) tys) ||> fold (fn (s, _) => Symtab.default (s, (("", ""), K tvar_a_atype))) syms |> fold (fold do_line o snd) problem end fun declare_undeclared_in_problem heading problem = let val ((cls, tys), syms) = undeclared_in_problem problem val decls = Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *) | (s, (cls, ())) => cons (Class_Decl (class_decl_prefix ^ s, cls, []))) cls [] @ Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *) | (s, (ty, ary)) => cons (Type_Decl (type_decl_prefix ^ s, ty, ary))) tys [] @ Symtab.fold (fn (_, (("", ""), _)) => I (* already declared *) | (s, (sym, ty)) => cons (Sym_Decl (sym_decl_prefix ^ s, sym, ty ()))) syms [] in (heading, decls) :: problem end val all_ctrss_of_datatypes = map (map_filter (try dest_Const) o #ctrs) o Ctr_Sugar.ctr_sugars_of val app_op_and_predicator_threshold = 45 fun generate_atp_problem ctxt generate_info format prem_role type_enc mode lam_trans uncurried_aliases readable_names presimp hyp_ts concl_t facts = let val thy = Proof_Context.theory_of ctxt val type_enc = type_enc |> adjust_type_enc format val completish = (case mode of Sledgehammer_Completish k => k | _ => 0) (* Forcing explicit applications is expensive for polymorphic encodings, because it takes only one existential variable ranging over "'a => 'b" to ruin everything. Hence we do it only if there are few facts (which is normally the case for "metis" and the minimizer). *) val app_op_level = if completish > 0 then Full_App_Op_And_Predicator else if length facts + length hyp_ts >= app_op_and_predicator_threshold then if is_type_enc_polymorphic type_enc then Min_App_Op else Sufficient_App_Op else Sufficient_App_Op_And_Predicator val lam_trans = if lam_trans = keep_lamsN andalso not (is_type_enc_full_higher_order type_enc) then liftingN else lam_trans + val simp_options = {if_simps = not (is_type_enc_fool type_enc)} val (classes, conjs, facts, subclass_pairs, tcon_clauses, lifted) = - translate_formulas ctxt prem_role format type_enc lam_trans presimp hyp_ts concl_t facts + translate_formulas simp_options ctxt prem_role format type_enc lam_trans presimp hyp_ts concl_t facts val (_, sym_tab0) = sym_table_of_facts ctxt type_enc app_op_level conjs facts val mono = conjs @ facts |> mononotonicity_info_of_facts ctxt type_enc completish val ctrss = all_ctrss_of_datatypes ctxt fun firstorderize in_helper = firstorderize_fact thy ctrss type_enc (uncurried_aliases andalso not in_helper) completish sym_tab0 val (conjs, facts) = (conjs, facts) |> apply2 (map (firstorderize false)) val (ho_stuff, sym_tab) = sym_table_of_facts ctxt type_enc Min_App_Op conjs facts val (uncurried_alias_eq_tms, uncurried_alias_eq_lines) = uncurried_alias_lines_of_sym_table ctxt generate_info ctrss mono type_enc uncurried_aliases sym_tab0 sym_tab val (_, sym_tab) = (ho_stuff, sym_tab) |> fold (add_iterm_syms_to_sym_table ctxt Min_App_Op false false) uncurried_alias_eq_tms val helpers = sym_tab |> helper_facts_of_sym_table ctxt format type_enc lam_trans completish |> map (firstorderize true) val all_facts = helpers @ conjs @ facts val mono_Ts = monotonic_types_of_facts ctxt mono type_enc all_facts val datatypes = datatypes_of_sym_table ctxt ctrss format type_enc uncurried_aliases sym_tab val class_decl_lines = decl_lines_of_classes type_enc classes val sym_decl_lines = (conjs, helpers @ facts, uncurried_alias_eq_tms) |> sym_decl_table_of_facts thy type_enc sym_tab |> lines_of_sym_decl_table ctxt generate_info mono type_enc mono_Ts val datatype_decl_lines = map decl_line_of_datatype datatypes val decl_lines = class_decl_lines @ sym_decl_lines @ datatype_decl_lines val num_facts = length facts val freshen = mode <> Exporter andalso mode <> Translator val pos = mode <> Exporter val rank_of = rank_of_fact_num num_facts val fact_lines = map (line_of_fact ctxt generate_info fact_prefix ascii_of I freshen pos mono type_enc rank_of) (0 upto num_facts - 1 ~~ facts) val subclass_lines = maps (lines_of_subclass_pair generate_info type_enc) subclass_pairs val tcon_lines = map (line_of_tcon_clause generate_info type_enc) tcon_clauses val helper_lines = 0 upto length helpers - 1 ~~ helpers |> map (line_of_fact ctxt generate_info helper_prefix I (K "") false true mono type_enc (K default_rank)) val free_type_lines = lines_of_free_types type_enc (facts @ conjs) val conj_lines = map (line_of_conjecture ctxt mono type_enc) conjs (* Reordering these might confuse the proof reconstruction code. *) val problem = [(explicit_declsN, decl_lines), (uncurried_alias_eqsN, uncurried_alias_eq_lines), (factsN, fact_lines), (subclassesN, subclass_lines), (tconsN, tcon_lines), (helpersN, helper_lines), (free_typesN, free_type_lines), (conjsN, conj_lines)] val problem = problem |> (case format of CNF => ensure_cnf_problem | CNF_UEQ => filter_cnf_ueq_problem | FOF => I | _ => declare_undeclared_in_problem implicit_declsN) val (problem, pool) = problem |> nice_atp_problem readable_names format fun add_sym_ary (s, {min_ary, ...} : sym_info) = min_ary > 0 ? Symtab.insert (op =) (s, min_ary) in (problem, Option.map snd pool |> the_default Symtab.empty, lifted, Symtab.fold add_sym_ary sym_tab Symtab.empty) end (* FUDGE *) val conj_weight = 0.0 val hyp_weight = 0.1 val fact_min_weight = 0.2 val fact_max_weight = 1.0 val type_info_default_weight = 0.8 (* Weights are from 0.0 (most important) to 1.0 (least important). *) fun atp_problem_selection_weights problem = let fun add_term_weights weight (ATerm ((s, _), tms)) = is_tptp_user_symbol s ? Symtab.default (s, weight) #> fold (add_term_weights weight) tms | add_term_weights weight (AAbs ((_, tm), args)) = add_term_weights weight tm #> fold (add_term_weights weight) args fun add_line_weights weight (Formula (_, _, phi, _, _)) = formula_fold NONE (K (add_term_weights weight)) phi | add_line_weights _ _ = I fun add_conjectures_weights [] = I | add_conjectures_weights conjs = let val (hyps, conj) = split_last conjs in add_line_weights conj_weight conj #> fold (add_line_weights hyp_weight) hyps end fun add_facts_weights facts = let val num_facts = length facts fun weight_of j = fact_min_weight + (fact_max_weight - fact_min_weight) * Real.fromInt j / Real.fromInt num_facts in map weight_of (0 upto num_facts - 1) ~~ facts |> fold (uncurry add_line_weights) end val get = these o AList.lookup (op =) problem in Symtab.empty |> add_conjectures_weights (get free_typesN @ get conjsN) |> add_facts_weights (get factsN) |> fold (fold (add_line_weights type_info_default_weight) o get) [explicit_declsN, subclassesN, tconsN] |> Symtab.dest |> sort (prod_ord Real.compare string_ord o apply2 swap) end (* Ugly hack: may make innocent victims (collateral damage) *) fun may_be_app s args = String.isPrefix app_op_name s andalso length args = 2 fun may_be_predicator s = member (op =) [predicator_name, prefixed_predicator_name] s fun strip_predicator (tm as ATerm ((s, _), [tm'])) = if may_be_predicator s then tm' else tm | strip_predicator tm = tm fun make_head_roll (ATerm ((s, _), tms)) = if may_be_app s tms then make_head_roll (hd tms) ||> append (tl tms) else (s, tms) | make_head_roll _ = ("", []) fun strip_up_to_predicator (ATyQuant (_, _, phi)) = strip_up_to_predicator phi | strip_up_to_predicator (AQuant (_, _, phi)) = strip_up_to_predicator phi | strip_up_to_predicator (AConn (_, phis)) = maps strip_up_to_predicator phis | strip_up_to_predicator (AAtom tm) = [strip_predicator tm] fun strip_ahorn_etc (ATyQuant (_, _, phi)) = strip_ahorn_etc phi | strip_ahorn_etc (AQuant (_, _, phi)) = strip_ahorn_etc phi | strip_ahorn_etc (AConn (AImplies, [phi1, phi2])) = strip_ahorn_etc phi2 |>> append (strip_up_to_predicator phi1) | strip_ahorn_etc phi = ([], hd (strip_up_to_predicator phi)) fun strip_iff_etc (ATyQuant (_, _, phi)) = strip_iff_etc phi | strip_iff_etc (AQuant (_, _, phi)) = strip_iff_etc phi | strip_iff_etc (AConn (AIff, [phi1, phi2])) = apply2 strip_up_to_predicator (phi1, phi2) | strip_iff_etc _ = ([], []) val max_term_order_weight = 2147483647 fun atp_problem_term_order_info problem = let fun add_edge s s' = Graph.default_node (s, ()) #> Graph.default_node (s', ()) #> Graph.add_edge_acyclic (s, s') fun add_term_deps head (ATerm ((s, _), args)) = if is_tptp_user_symbol head then (if is_tptp_user_symbol s then perhaps (try (add_edge s head)) else I) #> fold (add_term_deps head) args else I | add_term_deps head (AAbs ((_, tm), args)) = add_term_deps head tm #> fold (add_term_deps head) args fun add_intro_deps pred (Formula (_, role, phi, _, info)) = if pred (role, info) then let val (hyps, concl) = strip_ahorn_etc phi in (case make_head_roll concl of (head, args as _ :: _) => fold (add_term_deps head) (hyps @ args) | _ => I) end else I | add_intro_deps _ _ = I fun add_atom_eq_deps (SOME true) (ATerm ((s, _), [lhs as _, rhs])) = if is_tptp_equal s then (case make_head_roll lhs of (head, args as _ :: _) => fold (add_term_deps head) (rhs :: args) | _ => I) else I | add_atom_eq_deps _ _ = I fun add_eq_deps pred (Formula (_, role, phi, _, info)) = if pred (role, info) then (case strip_iff_etc phi of ([lhs], rhs) => (case make_head_roll lhs of (head, args as _ :: _) => fold (add_term_deps head) (rhs @ args) | _ => I) | _ => formula_fold (SOME (role <> Conjecture)) add_atom_eq_deps phi) else I | add_eq_deps _ _ = I fun has_status status (_, info) = extract_isabelle_status info = SOME status fun is_conj (role, _) = (role = Conjecture orelse role = Hypothesis) val graph = Graph.empty |> fold (fold (add_eq_deps (has_status non_rec_defN)) o snd) problem |> fold (fold (add_eq_deps (has_status rec_defN orf has_status simpN orf is_conj)) o snd) problem |> fold (fold (add_intro_deps (has_status inductiveN)) o snd) problem |> fold (fold (add_intro_deps (has_status introN)) o snd) problem fun next_weight w = if w + w <= max_term_order_weight then w + w else w + 1 fun add_weights _ [] = I | add_weights weight syms = fold (AList.update (op =) o rpair weight) syms #> add_weights (next_weight weight) (fold (union (op =) o Graph.immediate_succs graph) syms []) in (* Sorting is not just for aesthetics: It specifies the precedence order for the term ordering (KBO or LPO), from smaller to larger values. *) [] |> add_weights 1 (Graph.minimals graph) |> sort (int_ord o apply2 snd) end end; diff --git a/src/HOL/Tools/Meson/meson.ML b/src/HOL/Tools/Meson/meson.ML --- a/src/HOL/Tools/Meson/meson.ML +++ b/src/HOL/Tools/Meson/meson.ML @@ -1,790 +1,795 @@ (* Title: HOL/Tools/Meson/meson.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen The MESON resolution proof procedure for HOL. When making clauses, avoids using the rewriter -- instead uses RS recursively. *) signature MESON = sig + type simp_options = {if_simps : bool} val trace : bool Config.T val max_clauses : int Config.T val first_order_resolve : Proof.context -> thm -> thm -> thm val size_of_subgoals: thm -> int val has_too_many_clauses: Proof.context -> term -> bool val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context val finish_cnf: thm list -> thm list val presimplified_consts : string list - val presimplify: Proof.context -> thm -> thm - val make_nnf: Proof.context -> thm -> thm + val presimplify: simp_options -> Proof.context -> thm -> thm + val make_nnf: simp_options -> Proof.context -> thm -> thm val choice_theorems : theory -> thm list - val skolemize_with_choice_theorems : Proof.context -> thm list -> thm -> thm - val skolemize : Proof.context -> thm -> thm + val skolemize_with_choice_theorems : simp_options -> Proof.context -> thm list -> thm -> thm + val skolemize : simp_options -> Proof.context -> thm -> thm val cong_extensionalize_thm : Proof.context -> thm -> thm val abs_extensionalize_conv : Proof.context -> conv val abs_extensionalize_thm : Proof.context -> thm -> thm val make_clauses_unsorted: Proof.context -> thm list -> thm list val make_clauses: Proof.context -> thm list -> thm list val make_horns: thm list -> thm list val best_prolog_tac: Proof.context -> (thm -> int) -> thm list -> tactic val depth_prolog_tac: Proof.context -> thm list -> tactic val gocls: thm list -> thm list - val skolemize_prems_tac : Proof.context -> thm list -> int -> tactic + val skolemize_prems_tac : simp_options -> Proof.context -> thm list -> int -> tactic val MESON: tactic -> (thm list -> thm list) -> (thm list -> tactic) -> Proof.context -> int -> tactic val best_meson_tac: (thm -> int) -> Proof.context -> int -> tactic val safe_best_meson_tac: Proof.context -> int -> tactic val depth_meson_tac: Proof.context -> int -> tactic val prolog_step_tac': Proof.context -> thm list -> int -> tactic val iter_deepen_prolog_tac: Proof.context -> thm list -> tactic val iter_deepen_meson_tac: Proof.context -> thm list -> int -> tactic val make_meta_clause: Proof.context -> thm -> thm val make_meta_clauses: Proof.context -> thm list -> thm list val meson_tac: Proof.context -> thm list -> int -> tactic end structure Meson : MESON = struct +type simp_options = {if_simps : bool} + val trace = Attrib.setup_config_bool \<^binding>\meson_trace\ (K false) fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () val max_clauses = Attrib.setup_config_int \<^binding>\meson_max_clauses\ (K 60) (*No known example (on 1-5-2007) needs even thirty*) val iter_deepen_limit = 50; val disj_forward = @{thm disj_forward}; val disj_forward2 = @{thm disj_forward2}; val make_pos_rule = @{thm make_pos_rule}; val make_pos_rule' = @{thm make_pos_rule'}; val make_pos_goal = @{thm make_pos_goal}; val make_neg_rule = @{thm make_neg_rule}; val make_neg_rule' = @{thm make_neg_rule'}; val make_neg_goal = @{thm make_neg_goal}; val conj_forward = @{thm conj_forward}; val all_forward = @{thm all_forward}; val ex_forward = @{thm ex_forward}; val not_conjD = @{thm not_conjD}; val not_disjD = @{thm not_disjD}; val not_notD = @{thm not_notD}; val not_allD = @{thm not_allD}; val not_exD = @{thm not_exD}; val imp_to_disjD = @{thm imp_to_disjD}; val not_impD = @{thm not_impD}; val iff_to_disjD = @{thm iff_to_disjD}; val not_iffD = @{thm not_iffD}; val conj_exD1 = @{thm conj_exD1}; val conj_exD2 = @{thm conj_exD2}; val disj_exD = @{thm disj_exD}; val disj_exD1 = @{thm disj_exD1}; val disj_exD2 = @{thm disj_exD2}; val disj_assoc = @{thm disj_assoc}; val disj_comm = @{thm disj_comm}; val disj_FalseD1 = @{thm disj_FalseD1}; val disj_FalseD2 = @{thm disj_FalseD2}; (**** Operators for forward proof ****) (** First-order Resolution **) (*FIXME: currently does not "rename variables apart"*) fun first_order_resolve ctxt thA thB = (case \<^try>\ let val thy = Proof_Context.theory_of ctxt val tmA = Thm.concl_of thA val Const(\<^const_name>\Pure.imp\,_) $ tmB $ _ = Thm.prop_of thB val tenv = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty) |> snd val insts = Vartab.fold (fn (xi, (_, t)) => cons (xi, Thm.cterm_of ctxt t)) tenv []; in thA RS (infer_instantiate ctxt insts thB) end\ of SOME th => th | NONE => raise THM ("first_order_resolve", 0, [thA, thB])) (* Hack to make it less likely that we lose our precious bound variable names in "rename_bound_vars_RS" below, because of a clash. *) val protect_prefix = "Meson_xyzzy" fun protect_bound_var_names (t $ u) = protect_bound_var_names t $ protect_bound_var_names u | protect_bound_var_names (Abs (s, T, t')) = Abs (protect_prefix ^ s, T, protect_bound_var_names t') | protect_bound_var_names t = t fun fix_bound_var_names old_t new_t = let fun quant_of \<^const_name>\All\ = SOME true | quant_of \<^const_name>\Ball\ = SOME true | quant_of \<^const_name>\Ex\ = SOME false | quant_of \<^const_name>\Bex\ = SOME false | quant_of _ = NONE val flip_quant = Option.map not fun some_eq (SOME x) (SOME y) = x = y | some_eq _ _ = false fun add_names quant (Const (quant_s, _) $ Abs (s, _, t')) = add_names quant t' #> some_eq quant (quant_of quant_s) ? cons s | add_names quant (\<^const>\Not\ $ t) = add_names (flip_quant quant) t | add_names quant (\<^const>\implies\ $ t1 $ t2) = add_names (flip_quant quant) t1 #> add_names quant t2 | add_names quant (t1 $ t2) = fold (add_names quant) [t1, t2] | add_names _ _ = I fun lost_names quant = subtract (op =) (add_names quant new_t []) (add_names quant old_t []) fun aux ((t1 as Const (quant_s, _)) $ (Abs (s, T, t'))) = t1 $ Abs (s |> String.isPrefix protect_prefix s ? perhaps (try (fn _ => hd (lost_names (quant_of quant_s)))), T, aux t') | aux (t1 $ t2) = aux t1 $ aux t2 | aux t = t in aux new_t end (* Forward proof while preserving bound variables names *) fun rename_bound_vars_RS th rl = let val t = Thm.concl_of th val r = Thm.concl_of rl val th' = th RS Thm.rename_boundvars r (protect_bound_var_names r) rl val t' = Thm.concl_of th' in Thm.rename_boundvars t' (fix_bound_var_names t t') th' end (*raises exception if no rules apply*) fun tryres (th, rls) = let fun tryall [] = raise THM("tryres", 0, th::rls) | tryall (rl::rls) = (rename_bound_vars_RS th rl handle THM _ => tryall rls) in tryall rls end; (* Special version of "resolve_tac" that works around an explosion in the unifier. If the goal has the form "?P c", the danger is that resolving it against a property of the form "... c ... c ... c ..." will lead to a huge unification problem, due to the (spurious) choices between projection and imitation. The workaround is to instantiate "?P := (%c. ... c ... c ... c ...)" manually. *) fun quant_resolve_tac ctxt th i st = case (Thm.concl_of st, Thm.prop_of th) of (\<^const>\Trueprop\ $ (Var _ $ (c as Free _)), \<^const>\Trueprop\ $ _) => let val cc = Thm.cterm_of ctxt c val ct = Thm.dest_arg (Thm.cprop_of th) in resolve_tac ctxt [th] i (Thm.instantiate' [] [SOME (Thm.lambda cc ct)] st) end | _ => resolve_tac ctxt [th] i st (*Permits forward proof from rules that discharge assumptions. The supplied proof state st, e.g. from conj_forward, should have the form "[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q" and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*) fun forward_res ctxt nf st = let fun tacf [prem] = quant_resolve_tac ctxt (nf prem) 1 | tacf prems = error (cat_lines ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" :: Thm.string_of_thm ctxt st :: "Premises:" :: map (Thm.string_of_thm ctxt) prems)) in case Seq.pull (ALLGOALS (Misc_Legacy.METAHYPS ctxt tacf) st) of SOME (th, _) => th | NONE => raise THM ("forward_res", 0, [st]) end; (*Are any of the logical connectives in "bs" present in the term?*) fun has_conns bs = let fun has (Const _) = false | has (Const(\<^const_name>\Trueprop\,_) $ p) = has p | has (Const(\<^const_name>\Not\,_) $ p) = has p | has (Const(\<^const_name>\HOL.disj\,_) $ p $ q) = member (op =) bs \<^const_name>\HOL.disj\ orelse has p orelse has q | has (Const(\<^const_name>\HOL.conj\,_) $ p $ q) = member (op =) bs \<^const_name>\HOL.conj\ orelse has p orelse has q | has (Const(\<^const_name>\All\,_) $ Abs(_,_,p)) = member (op =) bs \<^const_name>\All\ orelse has p | has (Const(\<^const_name>\Ex\,_) $ Abs(_,_,p)) = member (op =) bs \<^const_name>\Ex\ orelse has p | has _ = false in has end; (**** Clause handling ****) fun literals (Const(\<^const_name>\Trueprop\,_) $ P) = literals P | literals (Const(\<^const_name>\HOL.disj\,_) $ P $ Q) = literals P @ literals Q | literals (Const(\<^const_name>\Not\,_) $ P) = [(false,P)] | literals P = [(true,P)]; (*number of literals in a term*) val nliterals = length o literals; (*** Tautology Checking ***) fun signed_lits_aux (Const (\<^const_name>\HOL.disj\, _) $ P $ Q) (poslits, neglits) = signed_lits_aux Q (signed_lits_aux P (poslits, neglits)) | signed_lits_aux (Const(\<^const_name>\Not\,_) $ P) (poslits, neglits) = (poslits, P::neglits) | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits); fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (Thm.concl_of th)) ([],[]); (*Literals like X=X are tautologous*) fun taut_poslit (Const(\<^const_name>\HOL.eq\,_) $ t $ u) = t aconv u | taut_poslit (Const(\<^const_name>\True\,_)) = true | taut_poslit _ = false; fun is_taut th = let val (poslits,neglits) = signed_lits th in exists taut_poslit poslits orelse exists (member (op aconv) neglits) (\<^term>\False\ :: poslits) end handle TERM _ => false; (*probably dest_Trueprop on a weird theorem*) (*** To remove trivial negated equality literals from clauses ***) (*They are typically functional reflexivity axioms and are the converses of injectivity equivalences*) val not_refl_disj_D = @{thm not_refl_disj_D}; (*Is either term a Var that does not properly occur in the other term?*) fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u)) | eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u)) | eliminable _ = false; fun refl_clause_aux 0 th = th | refl_clause_aux n th = case HOLogic.dest_Trueprop (Thm.concl_of th) of (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) $ _) => refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*) | (Const (\<^const_name>\HOL.disj\, _) $ (Const(\<^const_name>\Not\,_) $ (Const(\<^const_name>\HOL.eq\,_) $ t $ u)) $ _) => if eliminable(t,u) then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*) else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*) | (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm) | _ => (*not a disjunction*) th; fun notequal_lits_count (Const (\<^const_name>\HOL.disj\, _) $ P $ Q) = notequal_lits_count P + notequal_lits_count Q | notequal_lits_count (Const(\<^const_name>\Not\,_) $ (Const(\<^const_name>\HOL.eq\,_) $ _ $ _)) = 1 | notequal_lits_count _ = 0; (*Simplify a clause by applying reflexivity to its negated equality literals*) fun refl_clause th = let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (Thm.concl_of th)) in zero_var_indexes (refl_clause_aux neqs th) end handle TERM _ => th; (*probably dest_Trueprop on a weird theorem*) (*** Removal of duplicate literals ***) (*Forward proof, passing extra assumptions as theorems to the tactic*) fun forward_res2 ctxt nf hyps st = case Seq.pull (REPEAT (Misc_Legacy.METAHYPS ctxt (fn major::minors => resolve_tac ctxt [nf (minors @ hyps) major] 1) 1) st) of SOME(th,_) => th | NONE => raise THM("forward_res2", 0, [st]); (*Remove duplicates in P|Q by assuming ~P in Q rls (initially []) accumulates assumptions of the form P==>False*) fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc) handle THM _ => tryres(th,rls) handle THM _ => tryres(forward_res2 ctxt (nodups_aux ctxt) rls (th RS disj_forward2), [disj_FalseD1, disj_FalseD2, asm_rl]) handle THM _ => th; (*Remove duplicate literals, if there are any*) fun nodups ctxt th = if has_duplicates (op =) (literals (Thm.prop_of th)) then nodups_aux ctxt [] th else th; (*** The basic CNF transformation ***) fun estimated_num_clauses bound t = let fun sum x y = if x < bound andalso y < bound then x+y else bound fun prod x y = if x < bound andalso y < bound then x*y else bound (*Estimate the number of clauses in order to detect infeasible theorems*) fun signed_nclauses b (Const(\<^const_name>\Trueprop\,_) $ t) = signed_nclauses b t | signed_nclauses b (Const(\<^const_name>\Not\,_) $ t) = signed_nclauses (not b) t | signed_nclauses b (Const(\<^const_name>\HOL.conj\,_) $ t $ u) = if b then sum (signed_nclauses b t) (signed_nclauses b u) else prod (signed_nclauses b t) (signed_nclauses b u) | signed_nclauses b (Const(\<^const_name>\HOL.disj\,_) $ t $ u) = if b then prod (signed_nclauses b t) (signed_nclauses b u) else sum (signed_nclauses b t) (signed_nclauses b u) | signed_nclauses b (Const(\<^const_name>\HOL.implies\,_) $ t $ u) = if b then prod (signed_nclauses (not b) t) (signed_nclauses b u) else sum (signed_nclauses (not b) t) (signed_nclauses b u) | signed_nclauses b (Const(\<^const_name>\HOL.eq\, Type ("fun", [T, _])) $ t $ u) = if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*) if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u)) (prod (signed_nclauses (not b) u) (signed_nclauses b t)) else sum (prod (signed_nclauses b t) (signed_nclauses b u)) (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u)) else 1 | signed_nclauses b (Const(\<^const_name>\Ex\, _) $ Abs (_,_,t)) = signed_nclauses b t | signed_nclauses b (Const(\<^const_name>\All\,_) $ Abs (_,_,t)) = signed_nclauses b t | signed_nclauses _ _ = 1; (* literal *) in signed_nclauses true t end fun has_too_many_clauses ctxt t = let val max_cl = Config.get ctxt max_clauses in estimated_num_clauses (max_cl + 1) t > max_cl end (*Replaces universally quantified variables by FREE variables -- because assumptions may not contain scheme variables. Later, generalize using Variable.export. *) local val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec)))) |> Thm.term_of |> dest_Var; fun name_of (Const (\<^const_name>\All\, _) $ Abs(x, _, _)) = x | name_of _ = Name.uu; in fun freeze_spec th ctxt = let val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (Thm.concl_of th))] ctxt; val spec' = spec |> Thm.instantiate ([], [(spec_var, Thm.cterm_of ctxt' (Free (x, snd spec_var)))]); in (th RS spec', ctxt') end end; fun apply_skolem_theorem ctxt (th, rls) = let fun tryall [] = raise THM ("apply_skolem_theorem", 0, th::rls) | tryall (rl :: rls) = first_order_resolve ctxt th rl handle THM _ => tryall rls in tryall rls end (* Conjunctive normal form, adding clauses from th in front of ths (for foldr). Strips universal quantifiers and breaks up conjunctions. Eliminates existential quantifiers using Skolemization theorems. *) fun cnf old_skolem_ths ctxt (th, ths) = let val ctxt_ref = Unsynchronized.ref ctxt (* FIXME ??? *) fun cnf_aux (th,ths) = if not (can HOLogic.dest_Trueprop (Thm.prop_of th)) then ths (*meta-level: ignore*) else if not (has_conns [\<^const_name>\All\, \<^const_name>\Ex\, \<^const_name>\HOL.conj\] (Thm.prop_of th)) then nodups ctxt th :: ths (*no work to do, terminate*) else case head_of (HOLogic.dest_Trueprop (Thm.concl_of th)) of Const (\<^const_name>\HOL.conj\, _) => (*conjunction*) cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths)) | Const (\<^const_name>\All\, _) => (*universal quantifier*) let val (th', ctxt') = freeze_spec th (! ctxt_ref) in ctxt_ref := ctxt'; cnf_aux (th', ths) end | Const (\<^const_name>\Ex\, _) => (*existential quantifier: Insert Skolem functions*) cnf_aux (apply_skolem_theorem (! ctxt_ref) (th, old_skolem_ths), ths) | Const (\<^const_name>\HOL.disj\, _) => (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using all combinations of converting P, Q to CNF.*) (*There is one assumption, which gets bound to prem and then normalized via cnf_nil. The normal form is given to resolve_tac, instantiate a Boolean variable created by resolution with disj_forward. Since (cnf_nil prem) returns a LIST of theorems, we can backtrack to get all combinations.*) let val tac = Misc_Legacy.METAHYPS ctxt (fn [prem] => resolve_tac ctxt (cnf_nil prem) 1) 1 in Seq.list_of ((tac THEN tac) (th RS disj_forward)) @ ths end | _ => nodups ctxt th :: ths (*no work to do*) and cnf_nil th = cnf_aux (th, []) val cls = if has_too_many_clauses ctxt (Thm.concl_of th) then (trace_msg ctxt (fn () => "cnf is ignoring: " ^ Thm.string_of_thm ctxt th); ths) else cnf_aux (th, ths) in (cls, !ctxt_ref) end fun make_cnf old_skolem_ths th ctxt = cnf old_skolem_ths ctxt (th, []) (*Generalization, removal of redundant equalities, removal of tautologies.*) fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths); (**** Generation of contrapositives ****) fun is_left (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) $ _)) = true | is_left _ = false; (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) fun assoc_right th = if is_left (Thm.prop_of th) then assoc_right (th RS disj_assoc) else th; (*Must check for negative literal first!*) val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; (*For ordinary resolution. *) val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule']; (*Create a goal or support clause, conclusing False*) fun make_goal th = (*Must check for negative literal first!*) make_goal (tryres(th, clause_rules)) handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); fun rigid t = not (is_Var (head_of t)); fun ok4horn (Const (\<^const_name>\Trueprop\,_) $ (Const (\<^const_name>\HOL.disj\, _) $ t $ _)) = rigid t | ok4horn (Const (\<^const_name>\Trueprop\,_) $ t) = rigid t | ok4horn _ = false; (*Create a meta-level Horn clause*) fun make_horn crules th = if ok4horn (Thm.concl_of th) then make_horn crules (tryres(th,crules)) handle THM _ => th else th; (*Generate Horn clauses for all contrapositives of a clause. The input, th, is a HOL disjunction.*) fun add_contras crules th hcs = let fun rots (0,_) = hcs | rots (k,th) = zero_var_indexes (make_horn crules th) :: rots(k-1, assoc_right (th RS disj_comm)) in case nliterals(Thm.prop_of th) of 1 => th::hcs | n => rots(n, assoc_right th) end; (*Use "theorem naming" to label the clauses*) fun name_thms label = let fun name1 th (k, ths) = (k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths) in fn ths => #2 (fold_rev name1 ths (length ths, [])) end; (*Is the given disjunction an all-negative support clause?*) fun is_negative th = forall (not o #1) (literals (Thm.prop_of th)); val neg_clauses = filter is_negative; (***** MESON PROOF PROCEDURE *****) fun rhyps (Const(\<^const_name>\Pure.imp\,_) $ (Const(\<^const_name>\Trueprop\,_) $ A) $ phi, As) = rhyps(phi, A::As) | rhyps (_, As) = As; (** Detecting repeated assumptions in a subgoal **) (*The stringtree detects repeated assumptions.*) fun ins_term t net = Net.insert_term (op aconv) (t, t) net; (*detects repetitions in a list of terms*) fun has_reps [] = false | has_reps [_] = false | has_reps [t,u] = (t aconv u) | has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true; (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) fun TRYING_eq_assume_tac 0 st = Seq.single st | TRYING_eq_assume_tac i st = TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st) handle THM _ => TRYING_eq_assume_tac (i-1) st; fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (Thm.nprems_of st) st; (*Loop checking: FAIL if trying to prove the same thing twice -- if *ANY* subgoal has repeated literals*) fun check_tac st = if exists (fn prem => has_reps (rhyps(prem,[]))) (Thm.prems_of st) then Seq.empty else Seq.single st; (* resolve_from_net_tac actually made it slower... *) fun prolog_step_tac ctxt horns i = (assume_tac ctxt i APPEND resolve_tac ctxt horns i) THEN check_tac THEN TRYALL_eq_assume_tac; (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz; fun size_of_subgoals st = fold_rev addconcl (Thm.prems_of st) 0; (*Negation Normal Form*) val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, not_impD, not_iffD, not_allD, not_exD, not_notD]; fun ok4nnf (Const (\<^const_name>\Trueprop\,_) $ (Const (\<^const_name>\Not\, _) $ t)) = rigid t | ok4nnf (Const (\<^const_name>\Trueprop\,_) $ t) = rigid t | ok4nnf _ = false; fun make_nnf1 ctxt th = if ok4nnf (Thm.concl_of th) then make_nnf1 ctxt (tryres(th, nnf_rls)) handle THM ("tryres", _, _) => forward_res ctxt (make_nnf1 ctxt) (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) handle THM ("tryres", _, _) => th else th (*The simplification removes defined quantifiers and occurrences of True and False. nnf_ss also includes the one-point simprocs, which are needed to avoid the various one-point theorems from generating junk clauses.*) val nnf_simps = @{thms simp_implies_def Ex1_def Ball_def Bex_def if_True if_False if_cancel if_eq_cancel cases_simp} -val nnf_extra_simps = @{thms split_ifs ex_simps all_simps simp_thms} +fun nnf_extra_simps ({if_simps} : simp_options) = + (if if_simps then @{thms split_ifs} else []) @ @{thms ex_simps all_simps simp_thms} (* FIXME: "let_simp" is probably redundant now that we also rewrite with "Let_def [abs_def]". *) -val nnf_ss = +fun nnf_ss simp_options = simpset_of (put_simpset HOL_basic_ss \<^context> - addsimps nnf_extra_simps + addsimps (nnf_extra_simps simp_options) addsimprocs [\<^simproc>\defined_All\, \<^simproc>\defined_Ex\, \<^simproc>\neq\, \<^simproc>\let_simp\]) val presimplified_consts = [\<^const_name>\simp_implies\, \<^const_name>\False\, \<^const_name>\True\, \<^const_name>\Ex1\, \<^const_name>\Ball\, \<^const_name>\Bex\, \<^const_name>\If\, \<^const_name>\Let\] -fun presimplify ctxt = +fun presimplify simp_options ctxt = rewrite_rule ctxt (map safe_mk_meta_eq nnf_simps) - #> simplify (put_simpset nnf_ss ctxt) + #> simplify (put_simpset (nnf_ss simp_options) ctxt) #> rewrite_rule ctxt @{thms Let_def [abs_def]} -fun make_nnf ctxt th = +fun make_nnf simp_options ctxt th = (case Thm.prems_of th of - [] => th |> presimplify ctxt |> make_nnf1 ctxt + [] => th |> presimplify simp_options ctxt |> make_nnf1 ctxt | _ => raise THM ("make_nnf: premises in argument", 0, [th])); fun choice_theorems thy = try (Global_Theory.get_thm thy) "Hilbert_Choice.choice" |> the_list (* Pull existential quantifiers to front. This accomplishes Skolemization for clauses that arise from a subgoal. *) -fun skolemize_with_choice_theorems ctxt choice_ths = +fun skolemize_with_choice_theorems simp_options ctxt choice_ths = let fun aux th = if not (has_conns [\<^const_name>\Ex\] (Thm.prop_of th)) then th else tryres (th, choice_ths @ [conj_exD1, conj_exD2, disj_exD, disj_exD1, disj_exD2]) |> aux handle THM ("tryres", _, _) => tryres (th, [conj_forward, disj_forward, all_forward]) |> forward_res ctxt aux |> aux handle THM ("tryres", _, _) => rename_bound_vars_RS th ex_forward |> forward_res ctxt aux - in aux o make_nnf ctxt end + in aux o make_nnf simp_options ctxt end -fun skolemize ctxt = +fun skolemize simp_options ctxt = let val thy = Proof_Context.theory_of ctxt in - skolemize_with_choice_theorems ctxt (choice_theorems thy) + skolemize_with_choice_theorems simp_options ctxt (choice_theorems thy) end exception NO_F_PATTERN of unit fun get_F_pattern T t u = let fun pat t u = let val ((head1, args1), (head2, args2)) = (t, u) |> apply2 strip_comb in if head1 = head2 then let val pats = map2 pat args1 args2 in case filter (is_some o fst) pats of [(SOME T, _)] => (SOME T, list_comb (head1, map snd pats)) | [] => (NONE, t) | _ => raise NO_F_PATTERN () end else let val T = fastype_of t in if can dest_funT T then (SOME T, Bound 0) else raise NO_F_PATTERN () end end in if T = \<^typ>\bool\ then NONE else case pat t u of (SOME T, p as _ $ _) => SOME (Abs (Name.uu, T, p)) | _ => NONE end handle NO_F_PATTERN () => NONE val ext_cong_neq = @{thm ext_cong_neq} (* Strengthens "f g ~= f h" to "f g ~= f h & (EX x. g x ~= h x)". *) fun cong_extensionalize_thm ctxt th = (case Thm.concl_of th of \<^const>\Trueprop\ $ (\<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ (t as _ $ _) $ (u as _ $ _))) => (case get_F_pattern T t u of SOME p => th RS infer_instantiate ctxt [(("F", 0), Thm.cterm_of ctxt p)] ext_cong_neq | NONE => th) | _ => th) (* Removes the lambdas from an equation of the form "t = (%x1 ... xn. u)". It would be desirable to do this symmetrically but there's at least one existing proof in "Tarski" that relies on the current behavior. *) fun abs_extensionalize_conv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, _) $ _ $ Abs _ => ct |> (Conv.rewr_conv @{thm fun_eq_iff [THEN eq_reflection]} then_conv abs_extensionalize_conv ctxt) | _ $ _ => Conv.comb_conv (abs_extensionalize_conv ctxt) ct | Abs _ => Conv.abs_conv (abs_extensionalize_conv o snd) ctxt ct | _ => Conv.all_conv ct) val abs_extensionalize_thm = Conv.fconv_rule o abs_extensionalize_conv -fun try_skolemize_etc ctxt th = +fun try_skolemize_etc simp_options ctxt th = let val th = th |> cong_extensionalize_thm ctxt in [th] (* Extensionalize lambdas in "th", because that makes sense and that's what Sledgehammer does, but also keep an unextensionalized version of "th" for backward compatibility. *) |> insert Thm.eq_thm_prop (abs_extensionalize_thm ctxt th) - |> map_filter (fn th => th |> try (skolemize ctxt) + |> map_filter (fn th => th |> try (skolemize simp_options ctxt) |> tap (fn NONE => trace_msg ctxt (fn () => "Failed to skolemize " ^ Thm.string_of_thm ctxt th) | _ => ())) end fun add_clauses ctxt th cls = let val (cnfs, ctxt') = ctxt |> Variable.declare_thm th |> make_cnf [] th; in Variable.export ctxt' ctxt cnfs @ cls end; (*Sort clauses by number of literals*) fun fewerlits (th1, th2) = nliterals (Thm.prop_of th1) < nliterals (Thm.prop_of th2) (*Make clauses from a list of theorems, previously Skolemized and put into nnf. The resulting clauses are HOL disjunctions.*) fun make_clauses_unsorted ctxt ths = fold_rev (add_clauses ctxt) ths []; val make_clauses = sort (make_ord fewerlits) oo make_clauses_unsorted; (*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*) fun make_horns ths = name_thms "Horn#" (distinct Thm.eq_thm_prop (fold_rev (add_contras clause_rules) ths [])); (*Could simply use nprems_of, which would count remaining subgoals -- no discrimination as to their size! With BEST_FIRST, fails for problem 41.*) fun best_prolog_tac ctxt sizef horns = BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac ctxt horns 1); fun depth_prolog_tac ctxt horns = DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac ctxt horns 1); (*Return all negative clauses, as possible goal clauses*) fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls)); -fun skolemize_prems_tac ctxt prems = - cut_facts_tac (maps (try_skolemize_etc ctxt) prems) THEN' REPEAT o eresolve_tac ctxt [exE] +fun skolemize_prems_tac simp_options ctxt prems = + cut_facts_tac (maps (try_skolemize_etc simp_options ctxt) prems) THEN' + REPEAT o eresolve_tac ctxt [exE] (*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions. Function mkcl converts theorems to clauses.*) fun MESON preskolem_tac mkcl cltac ctxt i st = SELECT_GOAL (EVERY [Object_Logic.atomize_prems_tac ctxt 1, resolve_tac ctxt @{thms ccontr} 1, preskolem_tac, Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} => - EVERY1 [skolemize_prems_tac ctxt' negs, + EVERY1 [skolemize_prems_tac {if_simps = true} ctxt' negs, Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st handle THM _ => no_tac st; (*probably from make_meta_clause, not first-order*) (** Best-first search versions **) (*ths is a list of additional clauses (HOL disjunctions) to use.*) fun best_meson_tac sizef ctxt = MESON all_tac (make_clauses ctxt) (fn cls => THEN_BEST_FIRST (resolve_tac ctxt (gocls cls) 1) (has_fewer_prems 1, sizef) (prolog_step_tac ctxt (make_horns cls) 1)) ctxt (*First, breaks the goal into independent units*) fun safe_best_meson_tac ctxt = SELECT_GOAL (TRY (safe_tac ctxt) THEN TRYALL (best_meson_tac size_of_subgoals ctxt)); (** Depth-first search version **) fun depth_meson_tac ctxt = MESON all_tac (make_clauses ctxt) (fn cls => EVERY [resolve_tac ctxt (gocls cls) 1, depth_prolog_tac ctxt (make_horns cls)]) ctxt (** Iterative deepening version **) (*This version does only one inference per call; having only one eq_assume_tac speeds it up!*) fun prolog_step_tac' ctxt horns = let val horn0s = (*0 subgoals vs 1 or more*) take_prefix Thm.no_prems horns val nrtac = resolve_from_net_tac ctxt (Tactic.build_net horns) in fn i => eq_assume_tac i ORELSE match_tac ctxt horn0s i ORELSE (*no backtracking if unit MATCHES*) ((assume_tac ctxt i APPEND nrtac i) THEN check_tac) end; fun iter_deepen_prolog_tac ctxt horns = ITER_DEEPEN iter_deepen_limit (has_fewer_prems 1) (prolog_step_tac' ctxt horns); fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON all_tac (make_clauses ctxt) (fn cls => (case (gocls (cls @ ths)) of [] => no_tac (*no goal clauses*) | goes => let val horns = make_horns (cls @ ths) val _ = trace_msg ctxt (fn () => cat_lines ("meson method called:" :: map (Thm.string_of_thm ctxt) (cls @ ths) @ ["clauses:"] @ map (Thm.string_of_thm ctxt) horns)) in THEN_ITER_DEEPEN iter_deepen_limit (resolve_tac ctxt goes 1) (has_fewer_prems 1) (prolog_step_tac' ctxt horns) end)); fun meson_tac ctxt ths = SELECT_GOAL (TRY (safe_tac ctxt) THEN TRYALL (iter_deepen_meson_tac ctxt ths)); (**** Code to support ordinary resolution, rather than Model Elimination ****) (*Convert a list of clauses (disjunctions) to meta-level clauses (==>), with no contrapositives, for ordinary resolution.*) (*Rules to convert the head literal into a negated assumption. If the head literal is already negated, then using notEfalse instead of notEfalse' prevents a double negation.*) val notEfalse = @{lemma "\ P \ P \ False" by (rule notE)}; val notEfalse' = @{lemma "P \ \ P \ False" by (rule notE)}; fun negated_asm_of_head th = th RS notEfalse handle THM _ => th RS notEfalse'; (*Converting one theorem from a disjunction to a meta-level clause*) fun make_meta_clause ctxt th = let val (fth, thaw) = Misc_Legacy.freeze_thaw_robust ctxt th in (zero_var_indexes o Thm.varifyT_global o thaw 0 o negated_asm_of_head o make_horn resolution_clause_rules) fth end; fun make_meta_clauses ctxt ths = name_thms "MClause#" (distinct Thm.eq_thm_prop (map (make_meta_clause ctxt) ths)); end; diff --git a/src/HOL/Tools/Meson/meson_clausify.ML b/src/HOL/Tools/Meson/meson_clausify.ML --- a/src/HOL/Tools/Meson/meson_clausify.ML +++ b/src/HOL/Tools/Meson/meson_clausify.ML @@ -1,389 +1,388 @@ (* Title: HOL/Tools/Meson/meson_clausify.ML Author: Jia Meng, Cambridge University Computer Laboratory and NICTA Author: Jasmin Blanchette, TU Muenchen Transformation of HOL theorems into CNF forms. *) signature MESON_CLAUSIFY = sig val new_skolem_var_prefix : string val new_nonskolem_var_prefix : string val is_zapped_var_name : string -> bool val is_quasi_lambda_free : term -> bool val introduce_combinators_in_cterm : Proof.context -> cterm -> thm val introduce_combinators_in_theorem : Proof.context -> thm -> thm val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool val ss_only : thm list -> Proof.context -> Proof.context - val cnf_axiom : - Proof.context -> bool -> bool -> int -> thm + val cnf_axiom : Meson.simp_options -> Proof.context -> bool -> bool -> int -> thm -> (thm * term) option * thm list end; structure Meson_Clausify : MESON_CLAUSIFY = struct (* the extra "Meson" helps prevent clashes (FIXME) *) val new_skolem_var_prefix = "MesonSK" val new_nonskolem_var_prefix = "MesonV" fun is_zapped_var_name s = exists (fn prefix => String.isPrefix prefix s) [new_skolem_var_prefix, new_nonskolem_var_prefix] (**** Transformation of Elimination Rules into First-Order Formulas****) val cfalse = Thm.cterm_of \<^theory_context>\HOL\ \<^term>\False\; val ctp_false = Thm.cterm_of \<^theory_context>\HOL\ (HOLogic.mk_Trueprop \<^term>\False\); (* Converts an elim-rule into an equivalent theorem that does not have the predicate variable. Leaves other theorems unchanged. We simply instantiate the conclusion variable to False. (Cf. "transform_elim_prop" in "Sledgehammer_Util".) *) fun transform_elim_theorem th = (case Thm.concl_of th of (*conclusion variable*) \<^const>\Trueprop\ $ (Var (v as (_, \<^typ>\bool\))) => Thm.instantiate ([], [(v, cfalse)]) th | Var (v as (_, \<^typ>\prop\)) => Thm.instantiate ([], [(v, ctp_false)]) th | _ => th) (**** SKOLEMIZATION BY INFERENCE (lcp) ****) fun mk_old_skolem_term_wrapper t = let val T = fastype_of t in Const (\<^const_name>\Meson.skolem\, T --> T) $ t end fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') | beta_eta_in_abs_body t = Envir.beta_eta_contract t (*Traverse a theorem, accumulating Skolem function definitions.*) fun old_skolem_defs th = let fun dec_sko (Const (\<^const_name>\Ex\, _) $ (body as Abs (_, T, p))) rhss = (*Existential: declare a Skolem function, then insert into body and continue*) let val args = Misc_Legacy.term_frees body (* Forms a lambda-abstraction over the formal parameters *) val rhs = fold_rev (absfree o dest_Free) args (HOLogic.choice_const T $ beta_eta_in_abs_body body) |> mk_old_skolem_term_wrapper val comb = list_comb (rhs, args) in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end | dec_sko (Const (\<^const_name>\All\,_) $ Abs (a, T, p)) rhss = (*Universal quant: insert a free variable into body and continue*) let val fname = singleton (Name.variant_list (Misc_Legacy.add_term_names (p, []))) a in dec_sko (subst_bound (Free(fname,T), p)) rhss end | dec_sko (\<^const>\conj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q | dec_sko (\<^const>\disj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q | dec_sko (\<^const>\Trueprop\ $ p) rhss = dec_sko p rhss | dec_sko _ rhss = rhss in dec_sko (Thm.prop_of th) [] end; (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) fun is_quasi_lambda_free (Const (\<^const_name>\Meson.skolem\, _) $ _) = true | is_quasi_lambda_free (t1 $ t2) = is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 | is_quasi_lambda_free (Abs _) = false | is_quasi_lambda_free _ = true fun abstract ctxt ct = let val Abs (_, _, body) = Thm.term_of ct val (x, cbody) = Thm.dest_abs NONE ct val (A, cbodyT) = Thm.dest_funT (Thm.ctyp_of_cterm ct) fun makeK () = Thm.instantiate' [SOME A, SOME cbodyT] [SOME cbody] @{thm abs_K} in case body of Const _ => makeK() | Free _ => makeK() | Var _ => makeK() (*though Var isn't expected*) | Bound 0 => Thm.instantiate' [SOME A] [] @{thm abs_I} (*identity: I*) | rator$rand => if Term.is_dependent rator then (*C or S*) if Term.is_dependent rand then (*S*) let val crator = Thm.lambda x (Thm.dest_fun cbody) val crand = Thm.lambda x (Thm.dest_arg cbody) val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) val abs_S' = @{thm abs_S} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_S') in Thm.transitive abs_S' (Conv.binop_conv (abstract ctxt) rhs) end else (*C*) let val crator = Thm.lambda x (Thm.dest_fun cbody) val crand = Thm.dest_arg cbody val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) val abs_C' = @{thm abs_C} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_C') in Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv (abstract ctxt)) rhs) end else if Term.is_dependent rand then (*B or eta*) if rand = Bound 0 then Thm.eta_conversion ct else (*B*) let val crator = Thm.dest_fun cbody val crand = Thm.lambda x (Thm.dest_arg cbody) val (C, B) = Thm.dest_funT (Thm.ctyp_of_cterm crator) val abs_B' = @{thm abs_B} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B') in Thm.transitive abs_B' (Conv.arg_conv (abstract ctxt) rhs) end else makeK () | _ => raise Fail "abstract: Bad term" end; (* Traverse a theorem, remplacing lambda-abstractions with combinators. *) fun introduce_combinators_in_cterm ctxt ct = if is_quasi_lambda_free (Thm.term_of ct) then Thm.reflexive ct else case Thm.term_of ct of Abs _ => let val (cv, cta) = Thm.dest_abs NONE ct val (v, _) = dest_Free (Thm.term_of cv) val u_th = introduce_combinators_in_cterm ctxt cta val cu = Thm.rhs_of u_th val comb_eq = abstract ctxt (Thm.lambda cv cu) in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end | _ $ _ => let val (ct1, ct2) = Thm.dest_comb ct in Thm.combination (introduce_combinators_in_cterm ctxt ct1) (introduce_combinators_in_cterm ctxt ct2) end fun introduce_combinators_in_theorem ctxt th = if is_quasi_lambda_free (Thm.prop_of th) then th else let val th = Drule.eta_contraction_rule th val eqth = introduce_combinators_in_cterm ctxt (Thm.cprop_of th) in Thm.equal_elim eqth th end handle THM (msg, _, _) => (warning ("Error in the combinator translation of " ^ Thm.string_of_thm ctxt th ^ "\nException message: " ^ msg); (* A type variable of sort "{}" will make "abstraction" fail. *) TrueI) (*cterms are used throughout for efficiency*) val cTrueprop = Thm.cterm_of \<^theory_context>\HOL\ HOLogic.Trueprop; (*Given an abstraction over n variables, replace the bound variables by free ones. Return the body, along with the list of free variables.*) fun c_variant_abs_multi (ct0, vars) = let val (cv,ct) = Thm.dest_abs NONE ct0 in c_variant_abs_multi (ct, cv::vars) end handle CTERM _ => (ct0, rev vars); (* Given the definition of a Skolem function, return a theorem to replace an existential formula by a use of that function. Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) fun old_skolem_theorem_of_def ctxt rhs0 = let val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.cterm_of ctxt val rhs' = rhs |> Thm.dest_comb |> snd val (ch, frees) = c_variant_abs_multi (rhs', []) val (hilbert, cabs) = ch |> Thm.dest_comb |>> Thm.term_of val T = case hilbert of Const (_, Type (\<^type_name>\fun\, [_, T])) => T | _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert]) val cex = Thm.cterm_of ctxt (HOLogic.exists_const T) val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) val conc = Drule.list_comb (rhs, frees) |> Drule.beta_conv cabs |> Thm.apply cTrueprop fun tacf [prem] = rewrite_goals_tac ctxt @{thms skolem_def [abs_def]} THEN resolve_tac ctxt [(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]}) RS Global_Theory.get_thm (Proof_Context.theory_of ctxt) "Hilbert_Choice.someI_ex"] 1 in Goal.prove_internal ctxt [ex_tm] conc tacf |> forall_intr_list frees |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) |> Thm.varifyT_global end fun to_definitional_cnf_with_quantifiers ctxt th = let val eqth = CNF.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (Thm.prop_of th)) val eqth = eqth RS @{thm eq_reflection} val eqth = eqth RS @{thm TruepropI} in Thm.equal_elim eqth th end fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s = (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^ "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ string_of_int index_no ^ "_" ^ Name.desymbolize (SOME false) s fun cluster_of_zapped_var_name s = let val get_int = the o Int.fromString o nth (space_explode "_" s) in ((get_int 1, (get_int 2, get_int 3)), String.isPrefix new_skolem_var_prefix s) end fun rename_bound_vars_to_be_zapped ax_no = let fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t = case t of (t1 as Const (s, _)) $ Abs (s', T, t') => if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\ then let val skolem = (pos = (s = \<^const_name>\Ex\)) val (cluster, index_no) = if skolem = cluster_skolem then (cluster, index_no) else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) val s' = zapped_var_name cluster index_no s' in t1 $ Abs (s', T, aux cluster (index_no + 1) pos t') end else t | (t1 as Const (s, _)) $ t2 $ t3 => if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\HOL.implies\ then t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3 else if s = \<^const_name>\HOL.conj\ orelse s = \<^const_name>\HOL.disj\ then t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3 else t | (t1 as Const (s, _)) $ t2 => if s = \<^const_name>\Trueprop\ then t1 $ aux cluster index_no pos t2 else if s = \<^const_name>\Not\ then t1 $ aux cluster index_no (not pos) t2 else t | _ => t in aux ((ax_no, 0), true) 0 true end fun zap pos ct = ct |> (case Thm.term_of ct of Const (s, _) $ Abs (s', _, _) => if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\ then Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos else Conv.all_conv | Const (s, _) $ _ $ _ => if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\implies\ then Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) else if s = \<^const_name>\conj\ orelse s = \<^const_name>\disj\ then Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) else Conv.all_conv | Const (s, _) $ _ => if s = \<^const_name>\Trueprop\ then Conv.arg_conv (zap pos) else if s = \<^const_name>\Not\ then Conv.arg_conv (zap (not pos)) else Conv.all_conv | _ => Conv.all_conv) fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths val cheat_choice = \<^prop>\\x. \y. Q x y \ \f. \x. Q x (f x)\ |> Logic.varify_global |> Skip_Proof.make_thm \<^theory> (* Converts an Isabelle theorem into NNF. *) -fun nnf_axiom choice_ths new_skolem ax_no th ctxt = +fun nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt = let val thy = Proof_Context.theory_of ctxt val th = th |> transform_elim_theorem |> zero_var_indexes |> new_skolem ? forall_intr_vars val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single val th = th |> Conv.fconv_rule (Object_Logic.atomize ctxt) |> Meson.cong_extensionalize_thm ctxt |> Meson.abs_extensionalize_thm ctxt - |> Meson.make_nnf ctxt + |> Meson.make_nnf simp_options ctxt in if new_skolem then let fun skolemize choice_ths = - Meson.skolemize_with_choice_theorems ctxt choice_ths + Meson.skolemize_with_choice_theorems simp_options ctxt choice_ths #> simplify (ss_only @{thms all_simps[symmetric]} ctxt) val no_choice = null choice_ths val pull_out = if no_choice then simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]} ctxt) else skolemize choice_ths val discharger_th = th |> pull_out val discharger_th = discharger_th |> Meson.has_too_many_clauses ctxt (Thm.concl_of discharger_th) ? (to_definitional_cnf_with_quantifiers ctxt #> pull_out) val zapped_th = discharger_th |> Thm.prop_of |> rename_bound_vars_to_be_zapped ax_no |> (if no_choice then Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of else Thm.cterm_of ctxt) |> zap true val fixes = [] |> Term.add_free_names (Thm.prop_of zapped_th) |> filter is_zapped_var_name val ctxt' = ctxt |> Variable.add_fixes_direct fixes val fully_skolemized_t = zapped_th |> singleton (Variable.export ctxt' ctxt) |> Thm.cprop_of |> Thm.dest_equals |> snd |> Thm.term_of in if exists_subterm (fn Var ((s, _), _) => String.isPrefix new_skolem_var_prefix s | _ => false) fully_skolemized_t then let val (fully_skolemized_ct, ctxt) = yield_singleton (Variable.import_terms true) fully_skolemized_t ctxt |>> Thm.cterm_of ctxt in (SOME (discharger_th, fully_skolemized_ct), (Thm.assume fully_skolemized_ct, ctxt)) end else (NONE, (th, ctxt)) end else (NONE, (th |> Meson.has_too_many_clauses ctxt (Thm.concl_of th) ? to_definitional_cnf_with_quantifiers ctxt, ctxt)) end (* Convert a theorem to CNF, with additional premises due to skolemization. *) -fun cnf_axiom ctxt0 new_skolem combinators ax_no th = +fun cnf_axiom simp_options ctxt0 new_skolem combinators ax_no th = let val choice_ths = Meson.choice_theorems (Proof_Context.theory_of ctxt0) val (opt, (nnf_th, ctxt1)) = - nnf_axiom choice_ths new_skolem ax_no th ctxt0 + nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt0 fun clausify th = Meson.make_cnf (if new_skolem orelse null choice_ths then [] else map (old_skolem_theorem_of_def ctxt1) (old_skolem_defs th)) th ctxt1 val (cnf_ths, ctxt2) = clausify nnf_th fun intr_imp ct th = Thm.instantiate ([], [((("i", 0), \<^typ>\nat\), Thm.cterm_of ctxt2 (HOLogic.mk_nat ax_no))]) (zero_var_indexes @{thm skolem_COMBK_D}) RS Thm.implies_intr ct th in (opt |> Option.map (I #>> singleton (Variable.export ctxt2 ctxt0) ##> (Thm.term_of #> HOLogic.dest_Trueprop #> singleton (Variable.export_terms ctxt2 ctxt0))), cnf_ths |> map (combinators ? introduce_combinators_in_theorem ctxt2 #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf |> map (Thm.close_derivation \<^here>)) end handle THM _ => (NONE, []) end; diff --git a/src/HOL/Tools/Meson/meson_tactic.ML b/src/HOL/Tools/Meson/meson_tactic.ML --- a/src/HOL/Tools/Meson/meson_tactic.ML +++ b/src/HOL/Tools/Meson/meson_tactic.ML @@ -1,26 +1,26 @@ (* Title: HOL/Tools/Meson/meson_tactic.ML Author: Jia Meng, Cambridge University Computer Laboratory and NICTA Author: Jasmin Blanchette, TU Muenchen The "meson" proof method for HOL. *) signature MESON_TACTIC = sig val meson_general_tac : Proof.context -> thm list -> int -> tactic end; structure Meson_Tactic : MESON_TACTIC = struct fun meson_general_tac ctxt ths = let val ctxt' = put_claset HOL_cs ctxt - in Meson.meson_tac ctxt' (maps (snd o Meson_Clausify.cnf_axiom ctxt' false true 0) ths) end + in Meson.meson_tac ctxt' (maps (snd o Meson_Clausify.cnf_axiom {if_simps = true} ctxt' false true 0) ths) end val _ = Theory.setup (Method.setup \<^binding>\meson\ (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ctxt ths))) "MESON resolution proof procedure") end; diff --git a/src/HOL/Tools/Metis/metis_tactic.ML b/src/HOL/Tools/Metis/metis_tactic.ML --- a/src/HOL/Tools/Metis/metis_tactic.ML +++ b/src/HOL/Tools/Metis/metis_tactic.ML @@ -1,310 +1,310 @@ (* Title: HOL/Tools/Metis/metis_tactic.ML Author: Kong W. Susanto, Cambridge University Computer Laboratory Author: Lawrence C. Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Copyright Cambridge University 2007 HOL setup for the Metis prover. *) signature METIS_TACTIC = sig val trace : bool Config.T val verbose : bool Config.T val new_skolem : bool Config.T val advisory_simp : bool Config.T val metis_tac_unused : string list -> string -> Proof.context -> thm list -> int -> thm -> thm list * thm Seq.seq val metis_tac : string list -> string -> Proof.context -> thm list -> int -> tactic val metis_method : (string list option * string option) * thm list -> Proof.context -> thm list -> tactic val metis_lam_transs : string list val parse_metis_options : (string list option * string option) parser end structure Metis_Tactic : METIS_TACTIC = struct open ATP_Problem_Generate open ATP_Proof_Reconstruct open Metis_Generate open Metis_Reconstruct val new_skolem = Attrib.setup_config_bool \<^binding>\metis_new_skolem\ (K false) val advisory_simp = Attrib.setup_config_bool \<^binding>\metis_advisory_simp\ (K true) (* Designed to work also with monomorphic instances of polymorphic theorems. *) fun have_common_thm ctxt ths1 ths2 = exists (member (Term.aconv_untyped o apply2 Thm.prop_of) ths1) (map (Meson.make_meta_clause ctxt) ths2) (*Determining which axiom clauses are actually used*) fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th) | used_axioms _ _ = NONE (* Lightweight predicate type information comes in two flavors, "t = t'" and "t => t'", where "t" and "t'" are the same term modulo type tags. In Isabelle, type tags are stripped away, so we are left with "t = t" or "t => t". Type tag idempotence is also handled this way. *) fun reflexive_or_trivial_of_metis ctxt type_enc sym_tab concealed mth = (case hol_clause_of_metis ctxt type_enc sym_tab concealed mth of Const (\<^const_name>\HOL.eq\, _) $ _ $ t => let val ct = Thm.cterm_of ctxt t val cT = Thm.ctyp_of_cterm ct in refl |> Thm.instantiate' [SOME cT] [SOME ct] end | Const (\<^const_name>\disj\, _) $ t1 $ t2 => (if can HOLogic.dest_not t1 then t2 else t1) |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt |> Thm.trivial | _ => raise Fail "expected reflexive or trivial clause") |> Meson.make_meta_clause ctxt fun lam_lifted_of_metis ctxt type_enc sym_tab concealed mth = let val tac = rewrite_goals_tac ctxt @{thms lambda_def [abs_def]} THEN resolve_tac ctxt [refl] 1 val t = hol_clause_of_metis ctxt type_enc sym_tab concealed mth val ct = Thm.cterm_of ctxt (HOLogic.mk_Trueprop t) in Goal.prove_internal ctxt [] ct (K tac) |> Meson.make_meta_clause ctxt end fun add_vars_and_frees (t $ u) = fold (add_vars_and_frees) [t, u] | add_vars_and_frees (Abs (_, _, t)) = add_vars_and_frees t | add_vars_and_frees (t as Var _) = insert (op =) t | add_vars_and_frees (t as Free _) = insert (op =) t | add_vars_and_frees _ = I fun introduce_lam_wrappers ctxt th = if Meson_Clausify.is_quasi_lambda_free (Thm.prop_of th) then th else let fun conv first ctxt ct = if Meson_Clausify.is_quasi_lambda_free (Thm.term_of ct) then Thm.reflexive ct else (case Thm.term_of ct of Abs (_, _, u) => if first then (case add_vars_and_frees u [] of [] => Conv.abs_conv (conv false o snd) ctxt ct |> (fn th => Meson.first_order_resolve ctxt th @{thm Metis.eq_lambdaI}) | v :: _ => Abs (Name.uu, fastype_of v, abstract_over (v, Thm.term_of ct)) $ v |> Thm.cterm_of ctxt |> Conv.comb_conv (conv true ctxt)) else Conv.abs_conv (conv false o snd) ctxt ct | Const (\<^const_name>\Meson.skolem\, _) $ _ => Thm.reflexive ct | _ => Conv.comb_conv (conv true ctxt) ct) val eq_th = conv true ctxt (Thm.cprop_of th) (* We replace the equation's left-hand side with a beta-equivalent term so that "Thm.equal_elim" works below. *) val t0 $ _ $ t2 = Thm.prop_of eq_th val eq_ct = t0 $ Thm.prop_of th $ t2 |> Thm.cterm_of ctxt val eq_th' = Goal.prove_internal ctxt [] eq_ct (K (resolve_tac ctxt [eq_th] 1)) in Thm.equal_elim eq_th' th end fun clause_params ordering = {ordering = ordering, orderLiterals = Metis_Clause.UnsignedLiteralOrder, orderTerms = true} fun active_params ordering = {clause = clause_params ordering, prefactor = #prefactor Metis_Active.default, postfactor = #postfactor Metis_Active.default} val waiting_params = {symbolsWeight = 1.0, variablesWeight = 0.05, literalsWeight = 0.01, models = []} fun resolution_params ordering = {active = active_params ordering, waiting = waiting_params} fun kbo_advisory_simp_ordering ord_info = let fun weight (m, _) = AList.lookup (op =) ord_info (Metis_Name.toString m) |> the_default 1 fun precedence p = (case int_ord (apply2 weight p) of EQUAL => #precedence Metis_KnuthBendixOrder.default p | ord => ord) in {weight = weight, precedence = precedence} end fun metis_call type_enc lam_trans = let val type_enc = (case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases type_enc of [alias] => alias | _ => type_enc) val opts = [] |> type_enc <> partial_typesN ? cons type_enc |> lam_trans <> default_metis_lam_trans ? cons lam_trans in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end exception METIS_UNPROVABLE of unit (* Main function to start Metis proof and reconstruction *) fun FOL_SOLVE unused type_encs lam_trans ctxt cls ths0 = let val thy = Proof_Context.theory_of ctxt val new_skolem = Config.get ctxt new_skolem orelse null (Meson.choice_theorems thy) val do_lams = lam_trans = liftingN ? introduce_lam_wrappers ctxt val th_cls_pairs = map2 (fn j => fn th => (Thm.get_name_hint th, th |> Drule.eta_contraction_rule - |> Meson_Clausify.cnf_axiom ctxt new_skolem (lam_trans = combsN) j + |> Meson_Clausify.cnf_axiom {if_simps = true} ctxt new_skolem (lam_trans = combsN) j ||> map do_lams)) (0 upto length ths0 - 1) ths0 val ths = maps (snd o snd) th_cls_pairs val dischargers = map (fst o snd) th_cls_pairs val cls = cls |> map (Drule.eta_contraction_rule #> do_lams) val _ = trace_msg ctxt (K "FOL_SOLVE: CONJECTURE CLAUSES") val _ = List.app (fn th => trace_msg ctxt (fn () => Thm.string_of_thm ctxt th)) cls val type_enc :: fallback_type_encs = type_encs val _ = trace_msg ctxt (fn () => "type_enc = " ^ type_enc) val type_enc = type_enc_of_string Strict type_enc val (sym_tab, axioms, ord_info, concealed) = generate_metis_problem ctxt type_enc lam_trans cls ths fun get_isa_thm mth Isa_Reflexive_or_Trivial = reflexive_or_trivial_of_metis ctxt type_enc sym_tab concealed mth | get_isa_thm mth Isa_Lambda_Lifted = lam_lifted_of_metis ctxt type_enc sym_tab concealed mth | get_isa_thm _ (Isa_Raw ith) = ith val axioms = axioms |> map (fn (mth, ith) => (mth, get_isa_thm mth ith |> Thm.close_derivation \<^here>)) val _ = trace_msg ctxt (K "ISABELLE CLAUSES") val _ = List.app (fn (_, ith) => trace_msg ctxt (fn () => Thm.string_of_thm ctxt ith)) axioms val _ = trace_msg ctxt (K "METIS CLAUSES") val _ = List.app (fn (mth, _) => trace_msg ctxt (fn () => Metis_Thm.toString mth)) axioms val _ = trace_msg ctxt (K "START METIS PROVE PROCESS") val ordering = if Config.get ctxt advisory_simp then kbo_advisory_simp_ordering (ord_info ()) else Metis_KnuthBendixOrder.default fun fall_back () = (verbose_warning ctxt ("Falling back on " ^ quote (metis_call (hd fallback_type_encs) lam_trans) ^ "..."); FOL_SOLVE unused fallback_type_encs lam_trans ctxt cls ths0) in (case filter (fn t => Thm.prop_of t aconv \<^prop>\False\) cls of false_th :: _ => [false_th RS @{thm FalseE}] | [] => (case Metis_Resolution.loop (Metis_Resolution.new (resolution_params ordering) {axioms = axioms |> map fst, conjecture = []}) of Metis_Resolution.Contradiction mth => let val _ = trace_msg ctxt (fn () => "METIS RECONSTRUCTION START: " ^ Metis_Thm.toString mth) val ctxt' = fold Variable.declare_constraints (map Thm.prop_of cls) ctxt (*add constraints arising from converting goal to clause form*) val proof = Metis_Proof.proof mth val result = fold (replay_one_inference ctxt' type_enc concealed sym_tab) proof axioms val used = map_filter (used_axioms axioms) proof val _ = trace_msg ctxt (K "METIS COMPLETED; clauses actually used:") val _ = List.app (fn th => trace_msg ctxt (fn () => Thm.string_of_thm ctxt th)) used val unused_th_cls_pairs = filter_out (have_common_thm ctxt used o #2 o #2) th_cls_pairs val _ = unused := maps (#2 o #2) unused_th_cls_pairs; val _ = if not (null unused_th_cls_pairs) then verbose_warning ctxt ("Unused theorems: " ^ commas_quote (map #1 unused_th_cls_pairs)) else (); val _ = if not (null cls) andalso not (have_common_thm ctxt used cls) then verbose_warning ctxt "The assumptions are inconsistent" else (); in (case result of (_, ith) :: _ => (trace_msg ctxt (fn () => "Success: " ^ Thm.string_of_thm ctxt ith); [discharge_skolem_premises ctxt dischargers ith]) | _ => (trace_msg ctxt (K "Metis: No result"); [])) end | Metis_Resolution.Satisfiable _ => (trace_msg ctxt (K "Metis: No first-order proof with the supplied lemmas"); raise METIS_UNPROVABLE ())) handle METIS_UNPROVABLE () => if null fallback_type_encs then [] else fall_back () | METIS_RECONSTRUCT (loc, msg) => if null fallback_type_encs then (verbose_warning ctxt ("Failed to replay Metis proof\n" ^ loc ^ ": " ^ msg); []) else fall_back ()) end fun neg_clausify ctxt combinators = single #> Meson.make_clauses_unsorted ctxt #> combinators ? map (Meson_Clausify.introduce_combinators_in_theorem ctxt) #> Meson.finish_cnf fun preskolem_tac ctxt st0 = (if exists (Meson.has_too_many_clauses ctxt) (Logic.prems_of_goal (Thm.prop_of st0) 1) then Simplifier.full_simp_tac (Meson_Clausify.ss_only @{thms not_all not_ex} ctxt) 1 THEN CNF.cnfx_rewrite_tac ctxt 1 else all_tac) st0 fun metis_tac_unused type_encs0 lam_trans ctxt ths i st0 = let val unused = Unsynchronized.ref [] val type_encs = if null type_encs0 then partial_type_encs else type_encs0 val _ = trace_msg ctxt (fn () => "Metis called with theorems\n" ^ cat_lines (map (Thm.string_of_thm ctxt) ths)) val type_encs = type_encs |> maps unalias_type_enc val combs = (lam_trans = combsN) fun tac clause = resolve_tac ctxt (FOL_SOLVE unused type_encs lam_trans ctxt clause ths) 1 val seq = Meson.MESON (preskolem_tac ctxt) (maps (neg_clausify ctxt combs)) tac ctxt i st0 in (!unused, seq) end fun metis_tac type_encs lam_trans ctxt ths i = snd o metis_tac_unused type_encs lam_trans ctxt ths i (* Whenever "X" has schematic type variables, we treat "using X by metis" as "by (metis X)" to prevent "Subgoal.FOCUS" from freezing the type variables. We don't do it for nonschematic facts "X" because this breaks a few proofs (in the rare and subtle case where a proof relied on extensionality not being applied) and brings few benefits. *) val has_tvar = exists_type (exists_subtype (fn TVar _ => true | _ => false)) o Thm.prop_of fun metis_method ((override_type_encs, lam_trans), ths) ctxt facts = let val (schem_facts, nonschem_facts) = List.partition has_tvar facts in HEADGOAL (Method.insert_tac ctxt nonschem_facts THEN' CHANGED_PROP o metis_tac (these override_type_encs) (the_default default_metis_lam_trans lam_trans) ctxt (schem_facts @ ths)) end val metis_lam_transs = [opaque_liftingN, liftingN, combsN] fun set_opt _ x NONE = SOME x | set_opt get x (SOME x0) = error ("Cannot specify both " ^ quote (get x0) ^ " and " ^ quote (get x)) fun consider_opt s = if s = "hide_lams" then error "\"hide_lams\" has been renamed \"opaque_lifting\"" else if member (op =) metis_lam_transs s then apsnd (set_opt I s) else apfst (set_opt hd [s]) val parse_metis_options = Scan.optional (Args.parens (Args.name -- Scan.option (\<^keyword>\,\ |-- Args.name)) >> (fn (s, s') => (NONE, NONE) |> consider_opt s |> (case s' of SOME s' => consider_opt s' | _ => I))) (NONE, NONE) val _ = Theory.setup (Method.setup \<^binding>\metis\ (Scan.lift parse_metis_options -- Attrib.thms >> (METHOD oo metis_method)) "Metis for FOL and HOL problems") end; diff --git a/src/HOL/ex/Meson_Test.thy b/src/HOL/ex/Meson_Test.thy --- a/src/HOL/ex/Meson_Test.thy +++ b/src/HOL/ex/Meson_Test.thy @@ -1,2900 +1,2900 @@ section \Meson test cases\ theory Meson_Test imports Main begin text \ WARNING: there are many potential conflicts between variables used below and constants declared in HOL! \ hide_const (open) implies union inter subset quotient sum text \ Test data for the MESON proof procedure (Excludes the equality problems 51, 52, 56, 58) \ subsection \Interactive examples\ lemma problem_25: "(\x. P x) & (\x. L x --> ~ (M x & R x)) & (\x. P x --> (M x & L x)) & ((\x. P x --> Q x) | (\x. P x & R x)) --> (\x. Q x & P x)" apply (rule ccontr) ML_prf \ val ctxt = \<^context>; val prem25 = Thm.assume \<^cprop>\\ ?thesis\; - val nnf25 = Meson.make_nnf ctxt prem25; - val xsko25 = Meson.skolemize ctxt nnf25; + val nnf25 = Meson.make_nnf {if_simps = true} ctxt prem25; + val xsko25 = Meson.skolemize {if_simps = true} ctxt nnf25; \ apply (tactic \cut_tac xsko25 1 THEN REPEAT (eresolve_tac \<^context> [exE] 1)\) ML_val \ val ctxt = \<^context>; val [_, sko25] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal}))); val clauses25 = Meson.make_clauses ctxt [sko25]; (*7 clauses*) val horns25 = Meson.make_horns clauses25; (*16 Horn clauses*) val go25 :: _ = Meson.gocls clauses25; val ctxt' = fold Thm.declare_hyps (maps Thm.chyps_of (go25 :: horns25)) ctxt; Goal.prove ctxt' [] [] \<^prop>\False\ (fn _ => resolve_tac ctxt' [go25] 1 THEN Meson.depth_prolog_tac ctxt' horns25); \ oops lemma problem_26: "((\x. p x) = (\x. q x)) & (\x. \y. p x & q y --> (r x = s y)) --> ((\x. p x --> r x) = (\x. q x --> s x))" apply (rule ccontr) ML_prf \ val ctxt = \<^context>; val prem26 = Thm.assume \<^cprop>\\ ?thesis\ - val nnf26 = Meson.make_nnf ctxt prem26; - val xsko26 = Meson.skolemize ctxt nnf26; + val nnf26 = Meson.make_nnf {if_simps = true} ctxt prem26; + val xsko26 = Meson.skolemize {if_simps = true} ctxt nnf26; \ apply (tactic \cut_tac xsko26 1 THEN REPEAT (eresolve_tac \<^context> [exE] 1)\) ML_val \ val ctxt = \<^context>; val [_, sko26] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal}))); val clauses26 = Meson.make_clauses ctxt [sko26]; val _ = \<^assert> (length clauses26 = 9); val horns26 = Meson.make_horns clauses26; val _ = \<^assert> (length horns26 = 24); val go26 :: _ = Meson.gocls clauses26; val ctxt' = fold Thm.declare_hyps (maps Thm.chyps_of (go26 :: horns26)) ctxt; Goal.prove ctxt' [] [] \<^prop>\False\ (fn _ => resolve_tac ctxt' [go26] 1 THEN Meson.depth_prolog_tac ctxt' horns26); (*7 ms*) (*Proof is of length 107!!*) \ oops lemma problem_43: \ \NOW PROVED AUTOMATICALLY!!\ (*16 Horn clauses*) "(\x. \y. q x y = (\z. p z x = (p z y::bool))) --> (\x. (\y. q x y = (q y x::bool)))" apply (rule ccontr) ML_prf \ val ctxt = \<^context>; val prem43 = Thm.assume \<^cprop>\\ ?thesis\; - val nnf43 = Meson.make_nnf ctxt prem43; - val xsko43 = Meson.skolemize ctxt nnf43; + val nnf43 = Meson.make_nnf {if_simps = true} ctxt prem43; + val xsko43 = Meson.skolemize {if_simps = true} ctxt nnf43; \ apply (tactic \cut_tac xsko43 1 THEN REPEAT (eresolve_tac \<^context> [exE] 1)\) ML_val \ val ctxt = \<^context>; val [_, sko43] = #prems (#1 (Subgoal.focus ctxt 1 NONE (#goal @{Isar.goal}))); val clauses43 = Meson.make_clauses ctxt [sko43]; val _ = \<^assert> (length clauses43 = 6); val horns43 = Meson.make_horns clauses43; val _ = \<^assert> (length horns43 = 16); val go43 :: _ = Meson.gocls clauses43; val ctxt' = fold Thm.declare_hyps (maps Thm.chyps_of (go43 :: horns43)) ctxt; Goal.prove ctxt' [] [] \<^prop>\False\ (fn _ => resolve_tac ctxt' [go43] 1 THEN Meson.best_prolog_tac ctxt' Meson.size_of_subgoals horns43); (*7ms*) \ oops (* #1 (q x xa ==> ~ q x xa) ==> q xa x #2 (q xa x ==> ~ q xa x) ==> q x xa #3 (~ q x xa ==> q x xa) ==> ~ q xa x #4 (~ q xa x ==> q xa x) ==> ~ q x xa #5 [| ~ q ?U ?V ==> q ?U ?V; ~ p ?W ?U ==> p ?W ?U |] ==> p ?W ?V #6 [| ~ p ?W ?U ==> p ?W ?U; p ?W ?V ==> ~ p ?W ?V |] ==> ~ q ?U ?V #7 [| p ?W ?V ==> ~ p ?W ?V; ~ q ?U ?V ==> q ?U ?V |] ==> ~ p ?W ?U #8 [| ~ q ?U ?V ==> q ?U ?V; ~ p ?W ?V ==> p ?W ?V |] ==> p ?W ?U #9 [| ~ p ?W ?V ==> p ?W ?V; p ?W ?U ==> ~ p ?W ?U |] ==> ~ q ?U ?V #10 [| p ?W ?U ==> ~ p ?W ?U; ~ q ?U ?V ==> q ?U ?V |] ==> ~ p ?W ?V #11 [| p (xb ?U ?V) ?U ==> ~ p (xb ?U ?V) ?U; p (xb ?U ?V) ?V ==> ~ p (xb ?U ?V) ?V |] ==> q ?U ?V #12 [| p (xb ?U ?V) ?V ==> ~ p (xb ?U ?V) ?V; q ?U ?V ==> ~ q ?U ?V |] ==> p (xb ?U ?V) ?U #13 [| q ?U ?V ==> ~ q ?U ?V; p (xb ?U ?V) ?U ==> ~ p (xb ?U ?V) ?U |] ==> p (xb ?U ?V) ?V #14 [| ~ p (xb ?U ?V) ?U ==> p (xb ?U ?V) ?U; ~ p (xb ?U ?V) ?V ==> p (xb ?U ?V) ?V |] ==> q ?U ?V #15 [| ~ p (xb ?U ?V) ?V ==> p (xb ?U ?V) ?V; q ?U ?V ==> ~ q ?U ?V |] ==> ~ p (xb ?U ?V) ?U #16 [| q ?U ?V ==> ~ q ?U ?V; ~ p (xb ?U ?V) ?U ==> p (xb ?U ?V) ?U |] ==> ~ p (xb ?U ?V) ?V And here is the proof! (Unkn is the start state after use of goal clause) [Unkn, Res ([Thm "#14"], false, 1), Res ([Thm "#5"], false, 1), Res ([Thm "#1"], false, 1), Asm 1, Res ([Thm "#13"], false, 1), Asm 2, Asm 1, Res ([Thm "#13"], false, 1), Asm 1, Res ([Thm "#10"], false, 1), Res ([Thm "#16"], false, 1), Asm 2, Asm 1, Res ([Thm "#1"], false, 1), Asm 1, Res ([Thm "#14"], false, 1), Res ([Thm "#5"], false, 1), Res ([Thm "#2"], false, 1), Asm 1, Res ([Thm "#13"], false, 1), Asm 2, Asm 1, Res ([Thm "#8"], false, 1), Res ([Thm "#2"], false, 1), Asm 1, Res ([Thm "#12"], false, 1), Asm 2, Asm 1] : lderiv list *) text \ MORE and MUCH HARDER test data for the MESON proof procedure (courtesy John Harrison). \ (* ========================================================================= *) (* 100 problems selected from the TPTP library *) (* ========================================================================= *) (* * Original timings for John Harrison's MESON_TAC. * Timings below on a 600MHz Pentium III (perch) * Some timings below refer to griffon, which is a dual 2.5GHz Power Mac G5. * * A few variable names have been changed to avoid clashing with constants. * * Changed numeric constants e.g. 0, 1, 2... to num0, num1, num2... * * Here's a list giving typical CPU times, as well as common names and * literature references. * * BOO003-1 34.6 B2 part 1 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob2_part1.ver1.in [ANL] * BOO004-1 36.7 B2 part 2 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob2_part2.ver1 [ANL] * BOO005-1 47.4 B3 part 1 [McCharen, et al., 1976]; B5 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob3_part1.ver1.in [ANL] * BOO006-1 48.4 B3 part 2 [McCharen, et al., 1976]; B6 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob3_part2.ver1 [ANL] * BOO011-1 19.0 B7 [McCharen, et al., 1976]; prob7.ver1 [ANL] * CAT001-3 45.2 C1 [McCharen, et al., 1976]; p1.ver3.in [ANL] * CAT003-3 10.5 C3 [McCharen, et al., 1976]; p3.ver3.in [ANL] * CAT005-1 480.1 C5 [McCharen, et al., 1976]; p5.ver1.in [ANL] * CAT007-1 11.9 C7 [McCharen, et al., 1976]; p7.ver1.in [ANL] * CAT018-1 81.3 p18.ver1.in [ANL] * COL001-2 16.0 C1 [Wos & McCune, 1988] * COL023-1 5.1 [McCune & Wos, 1988] * COL032-1 15.8 [McCune & Wos, 1988] * COL052-2 13.2 bird4.ver2.in [ANL] * COL075-2 116.9 [Jech, 1994] * COM001-1 1.7 shortburst [Wilson & Minker, 1976] * COM002-1 4.4 burstall [Wilson & Minker, 1976] * COM002-2 7.4 * COM003-2 22.1 [Brushi, 1991] * COM004-1 45.1 * GEO003-1 71.7 T3 [McCharen, et al., 1976]; t3.ver1.in [ANL] * GEO017-2 78.8 D4.1 [Quaife, 1989] * GEO027-3 181.5 D10.1 [Quaife, 1989] * GEO058-2 104.0 R4 [Quaife, 1989] * GEO079-1 2.4 GEOMETRY THEOREM [Slagle, 1967] * GRP001-1 47.8 CADE-11 Competition 1 [Overbeek, 1990]; G1 [McCharen, et al., 1976]; THEOREM 1 [Lusk & McCune, 1993]; wos10 [Wilson & Minker, 1976]; xsquared.ver1.in [ANL]; [Robinson, 1963] * GRP008-1 50.4 Problem 4 [Wos, 1965]; wos4 [Wilson & Minker, 1976] * GRP013-1 40.2 Problem 11 [Wos, 1965]; wos11 [Wilson & Minker, 1976] * GRP037-3 43.8 Problem 17 [Wos, 1965]; wos17 [Wilson & Minker, 1976] * GRP031-2 3.2 ls23 [Lawrence & Starkey, 1974]; ls23 [Wilson & Minker, 1976] * GRP034-4 2.5 ls26 [Lawrence & Starkey, 1974]; ls26 [Wilson & Minker, 1976] * GRP047-2 11.7 [Veroff, 1992] * GRP130-1 170.6 Bennett QG8 [TPTP]; QG8 [Slaney, 1993] * GRP156-1 48.7 ax_mono1c [Schulz, 1995] * GRP168-1 159.1 p01a [Schulz, 1995] * HEN003-3 39.9 HP3 [McCharen, et al., 1976] * HEN007-2 125.7 H7 [McCharen, et al., 1976] * HEN008-4 62.0 H8 [McCharen, et al., 1976] * HEN009-5 136.3 H9 [McCharen, et al., 1976]; hp9.ver3.in [ANL] * HEN012-3 48.5 new.ver2.in [ANL] * LCL010-1 370.9 EC-73 [McCune & Wos, 1992]; ec_yq.in [OTTER] * LCL077-2 51.6 morgan.two.ver1.in [ANL] * LCL082-1 14.6 IC-1.1 [Wos, et al., 1990]; IC-65 [McCune & Wos, 1992]; ls2 [SETHEO]; S1 [Pfenning, 1988] * LCL111-1 585.6 CADE-11 Competition 6 [Overbeek, 1990]; mv25.in [OTTER]; MV-57 [McCune & Wos, 1992]; mv.in part 2 [OTTER]; ovb6 [SETHEO]; THEOREM 6 [Lusk & McCune, 1993] * LCL143-1 10.9 Lattice structure theorem 2 [Bonacina, 1991] * LCL182-1 271.6 Problem 2.16 [Whitehead & Russell, 1927] * LCL200-1 12.0 Problem 2.46 [Whitehead & Russell, 1927] * LCL215-1 214.4 Problem 2.62 [Whitehead & Russell, 1927]; Problem 2.63 [Whitehead & Russell, 1927] * LCL230-2 0.2 Pelletier 5 [Pelletier, 1986] * LDA003-1 68.5 Problem 3 [Jech, 1993] * MSC002-1 9.2 DBABHP [Michie, et al., 1972]; DBABHP [Wilson & Minker, 1976] * MSC003-1 3.2 HASPARTS-T1 [Wilson & Minker, 1976] * MSC004-1 9.3 HASPARTS-T2 [Wilson & Minker, 1976] * MSC005-1 1.8 Problem 5.1 [Plaisted, 1982] * MSC006-1 39.0 nonob.lop [SETHEO] * NUM001-1 14.0 Chang-Lee-10a [Chang, 1970]; ls28 [Lawrence & Starkey, 1974]; ls28 [Wilson & Minker, 1976] * NUM021-1 52.3 ls65 [Lawrence & Starkey, 1974]; ls65 [Wilson & Minker, 1976] * NUM024-1 64.6 ls75 [Lawrence & Starkey, 1974]; ls75 [Wilson & Minker, 1976] * NUM180-1 621.2 LIM2.1 [Quaife] * NUM228-1 575.9 TRECDEF4 cor. [Quaife] * PLA002-1 37.4 Problem 5.7 [Plaisted, 1982] * PLA006-1 7.2 [Segre & Elkan, 1994] * PLA017-1 484.8 [Segre & Elkan, 1994] * PLA022-1 19.1 [Segre & Elkan, 1994] * PLA022-2 19.7 [Segre & Elkan, 1994] * PRV001-1 10.3 PV1 [McCharen, et al., 1976] * PRV003-1 3.9 E2 [McCharen, et al., 1976]; v2.lop [SETHEO] * PRV005-1 4.3 E4 [McCharen, et al., 1976]; v4.lop [SETHEO] * PRV006-1 6.0 E5 [McCharen, et al., 1976]; v5.lop [SETHEO] * PRV009-1 2.2 Hoares FIND [Bledsoe, 1977]; Problem 5.5 [Plaisted, 1982] * PUZ012-1 3.5 Boxes-of-fruit [Wos, 1988]; Boxes-of-fruit [Wos, et al., 1992]; boxes.ver1.in [ANL] * PUZ020-1 56.6 knightknave.in [ANL] * PUZ025-1 58.4 Problem 35 [Smullyan, 1978]; tandl35.ver1.in [ANL] * PUZ029-1 5.1 pigs.ver1.in [ANL] * RNG001-3 82.4 EX6-T? [Wilson & Minker, 1976]; ex6.lop [SETHEO]; Example 6a [Fleisig, et al., 1974]; FEX6T1 [SPRFN]; FEX6T2 [SPRFN] * RNG001-5 399.8 Problem 21 [Wos, 1965]; wos21 [Wilson & Minker, 1976] * RNG011-5 8.4 CADE-11 Competition Eq-10 [Overbeek, 1990]; PROBLEM 10 [Zhang, 1993]; THEOREM EQ-10 [Lusk & McCune, 1993] * RNG023-6 9.1 [Stevens, 1987] * RNG028-2 9.3 PROOF III [Anantharaman & Hsiang, 1990] * RNG038-2 16.2 Problem 27 [Wos, 1965]; wos27 [Wilson & Minker, 1976] * RNG040-2 180.5 Problem 29 [Wos, 1965]; wos29 [Wilson & Minker, 1976] * RNG041-1 35.8 Problem 30 [Wos, 1965]; wos30 [Wilson & Minker, 1976] * ROB010-1 205.0 Lemma 3.3 [Winker, 1990]; RA2 [Lusk & Wos, 1992] * ROB013-1 23.6 Lemma 3.5 [Winker, 1990] * ROB016-1 15.2 Corollary 3.7 [Winker, 1990] * ROB021-1 230.4 [McCune, 1992] * SET005-1 192.2 ls108 [Lawrence & Starkey, 1974]; ls108 [Wilson & Minker, 1976] * SET009-1 10.5 ls116 [Lawrence & Starkey, 1974]; ls116 [Wilson & Minker, 1976] * SET025-4 694.7 Lemma 10 [Boyer, et al, 1986] * SET046-5 2.3 p42.in [ANL]; Pelletier 42 [Pelletier, 1986] * SET047-5 3.7 p43.in [ANL]; Pelletier 43 [Pelletier, 1986] * SYN034-1 2.8 QW [Michie, et al., 1972]; QW [Wilson & Minker, 1976] * SYN071-1 1.9 Pelletier 48 [Pelletier, 1986] * SYN349-1 61.7 Ch17N5 [Tammet, 1994] * SYN352-1 5.5 Ch18N4 [Tammet, 1994] * TOP001-2 61.1 Lemma 1a [Wick & McCune, 1989] * TOP002-2 0.4 Lemma 1b [Wick & McCune, 1989] * TOP004-1 181.6 Lemma 1d [Wick & McCune, 1989] * TOP004-2 9.0 Lemma 1d [Wick & McCune, 1989] * TOP005-2 139.8 Lemma 1e [Wick & McCune, 1989] *) abbreviation "EQU001_0_ax equal \ (\X. equal(X::'a,X)) & (\Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & (\Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z))" abbreviation "BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum \ (\X Y. sum(X::'a,Y,add(X::'a,Y))) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) & (\Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) & (\X. sum(additive_identity::'a,X,X)) & (\X. sum(X::'a,additive_identity,X)) & (\X. product(multiplicative_identity::'a,X,X)) & (\X. product(X::'a,multiplicative_identity,X)) & (\Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) & (\Y Z V3 X V1 V2 V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & product(V3::'a,X,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 V3 X V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(V3::'a,X,V4)) & (\Y Z X V3 V1 V2 V4. sum(X::'a,Y,V1) & sum(X::'a,Z,V2) & product(Y::'a,Z,V3) & sum(X::'a,V3,V4) --> product(V1::'a,V2,V4)) & (\Y Z V1 V2 X V3 V4. sum(X::'a,Y,V1) & sum(X::'a,Z,V2) & product(Y::'a,Z,V3) & product(V1::'a,V2,V4) --> sum(X::'a,V3,V4)) & (\Y Z V3 X V1 V2 V4. sum(Y::'a,X,V1) & sum(Z::'a,X,V2) & product(Y::'a,Z,V3) & sum(V3::'a,X,V4) --> product(V1::'a,V2,V4)) & (\Y Z V1 V2 V3 X V4. sum(Y::'a,X,V1) & sum(Z::'a,X,V2) & product(Y::'a,Z,V3) & product(V1::'a,V2,V4) --> sum(V3::'a,X,V4)) & (\X. sum(INVERSE(X),X,multiplicative_identity)) & (\X. sum(X::'a,INVERSE(X),multiplicative_identity)) & (\X. product(INVERSE(X),X,additive_identity)) & (\X. product(X::'a,INVERSE(X),additive_identity)) & (\X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) & (\X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V))" abbreviation "BOO002_0_eq INVERSE multiply add product sum equal \ (\X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) & (\X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) & (\X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & (\X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & (\X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) & (\X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) & (\X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y)))" (*51194 inferences so far. Searching to depth 13. 232.9 secs*) lemma BOO003_1: "EQU001_0_ax equal & BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum & BOO002_0_eq INVERSE multiply add product sum equal & (~product(x::'a,x,x)) --> False" by meson (*51194 inferences so far. Searching to depth 13. 204.6 secs Strange! The previous problem also has 51194 inferences at depth 13. They must be very similar!*) lemma BOO004_1: "EQU001_0_ax equal & BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum & BOO002_0_eq INVERSE multiply add product sum equal & (~sum(x::'a,x,x)) --> False" by meson (*74799 inferences so far. Searching to depth 13. 290.0 secs*) lemma BOO005_1: "EQU001_0_ax equal & BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum & BOO002_0_eq INVERSE multiply add product sum equal & (~sum(x::'a,multiplicative_identity,multiplicative_identity)) --> False" by meson (*74799 inferences so far. Searching to depth 13. 314.6 secs*) lemma BOO006_1: "EQU001_0_ax equal & BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum & BOO002_0_eq INVERSE multiply add product sum equal & (~product(x::'a,additive_identity,additive_identity)) --> False" by meson (*5 inferences so far. Searching to depth 5. 1.3 secs*) lemma BOO011_1: "EQU001_0_ax equal & BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum & BOO002_0_eq INVERSE multiply add product sum equal & (~equal(INVERSE(additive_identity),multiplicative_identity)) --> False" by meson abbreviation "CAT003_0_ax f1 compos codomain domain equal there_exists equivalent \ (\Y X. equivalent(X::'a,Y) --> there_exists(X)) & (\X Y. equivalent(X::'a,Y) --> equal(X::'a,Y)) & (\X Y. there_exists(X) & equal(X::'a,Y) --> equivalent(X::'a,Y)) & (\X. there_exists(domain(X)) --> there_exists(X)) & (\X. there_exists(codomain(X)) --> there_exists(X)) & (\Y X. there_exists(compos(X::'a,Y)) --> there_exists(domain(X))) & (\X Y. there_exists(compos(X::'a,Y)) --> equal(domain(X),codomain(Y))) & (\X Y. there_exists(domain(X)) & equal(domain(X),codomain(Y)) --> there_exists(compos(X::'a,Y))) & (\X Y Z. equal(compos(X::'a,compos(Y::'a,Z)),compos(compos(X::'a,Y),Z))) & (\X. equal(compos(X::'a,domain(X)),X)) & (\X. equal(compos(codomain(X),X),X)) & (\X Y. equivalent(X::'a,Y) --> there_exists(Y)) & (\X Y. there_exists(X) & there_exists(Y) & equal(X::'a,Y) --> equivalent(X::'a,Y)) & (\Y X. there_exists(compos(X::'a,Y)) --> there_exists(codomain(X))) & (\X Y. there_exists(f1(X::'a,Y)) | equal(X::'a,Y)) & (\X Y. equal(X::'a,f1(X::'a,Y)) | equal(Y::'a,f1(X::'a,Y)) | equal(X::'a,Y)) & (\X Y. equal(X::'a,f1(X::'a,Y)) & equal(Y::'a,f1(X::'a,Y)) --> equal(X::'a,Y))" abbreviation "CAT003_0_eq f1 compos codomain domain equivalent there_exists equal \ (\X Y. equal(X::'a,Y) & there_exists(X) --> there_exists(Y)) & (\X Y Z. equal(X::'a,Y) & equivalent(X::'a,Z) --> equivalent(Y::'a,Z)) & (\X Z Y. equal(X::'a,Y) & equivalent(Z::'a,X) --> equivalent(Z::'a,Y)) & (\X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & (\X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) & (\X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) & (\X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & (\A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E)))" (*4007 inferences so far. Searching to depth 9. 13 secs*) lemma CAT001_3: "EQU001_0_ax equal & CAT003_0_ax f1 compos codomain domain equal there_exists equivalent & CAT003_0_eq f1 compos codomain domain equivalent there_exists equal & (there_exists(compos(a::'a,b))) & (\Y X Z. equal(compos(compos(a::'a,b),X),Y) & equal(compos(compos(a::'a,b),Z),Y) --> equal(X::'a,Z)) & (there_exists(compos(b::'a,h))) & (equal(compos(b::'a,h),compos(b::'a,g))) & (~equal(h::'a,g)) --> False" by meson (*245 inferences so far. Searching to depth 7. 1.0 secs*) lemma CAT003_3: "EQU001_0_ax equal & CAT003_0_ax f1 compos codomain domain equal there_exists equivalent & CAT003_0_eq f1 compos codomain domain equivalent there_exists equal & (there_exists(compos(a::'a,b))) & (\Y X Z. equal(compos(X::'a,compos(a::'a,b)),Y) & equal(compos(Z::'a,compos(a::'a,b)),Y) --> equal(X::'a,Z)) & (there_exists(h)) & (equal(compos(h::'a,a),compos(g::'a,a))) & (~equal(g::'a,h)) --> False" by meson abbreviation "CAT001_0_ax equal codomain domain identity_map compos product defined \ (\X Y. defined(X::'a,Y) --> product(X::'a,Y,compos(X::'a,Y))) & (\Z X Y. product(X::'a,Y,Z) --> defined(X::'a,Y)) & (\X Xy Y Z. product(X::'a,Y,Xy) & defined(Xy::'a,Z) --> defined(Y::'a,Z)) & (\Y Xy Z X Yz. product(X::'a,Y,Xy) & product(Y::'a,Z,Yz) & defined(Xy::'a,Z) --> defined(X::'a,Yz)) & (\Xy Y Z X Yz Xyz. product(X::'a,Y,Xy) & product(Xy::'a,Z,Xyz) & product(Y::'a,Z,Yz) --> product(X::'a,Yz,Xyz)) & (\Z Yz X Y. product(Y::'a,Z,Yz) & defined(X::'a,Yz) --> defined(X::'a,Y)) & (\Y X Yz Xy Z. product(Y::'a,Z,Yz) & product(X::'a,Y,Xy) & defined(X::'a,Yz) --> defined(Xy::'a,Z)) & (\Yz X Y Xy Z Xyz. product(Y::'a,Z,Yz) & product(X::'a,Yz,Xyz) & product(X::'a,Y,Xy) --> product(Xy::'a,Z,Xyz)) & (\Y X Z. defined(X::'a,Y) & defined(Y::'a,Z) & identity_map(Y) --> defined(X::'a,Z)) & (\X. identity_map(domain(X))) & (\X. identity_map(codomain(X))) & (\X. defined(X::'a,domain(X))) & (\X. defined(codomain(X),X)) & (\X. product(X::'a,domain(X),X)) & (\X. product(codomain(X),X,X)) & (\X Y. defined(X::'a,Y) & identity_map(X) --> product(X::'a,Y,Y)) & (\Y X. defined(X::'a,Y) & identity_map(Y) --> product(X::'a,Y,X)) & (\X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W))" abbreviation "CAT001_0_eq compos defined identity_map codomain domain product equal \ (\X Y Z W. equal(X::'a,Y) & product(X::'a,Z,W) --> product(Y::'a,Z,W)) & (\X Z Y W. equal(X::'a,Y) & product(Z::'a,X,W) --> product(Z::'a,Y,W)) & (\X Z W Y. equal(X::'a,Y) & product(Z::'a,W,X) --> product(Z::'a,W,Y)) & (\X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & (\X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) & (\X Y. equal(X::'a,Y) & identity_map(X) --> identity_map(Y)) & (\X Y Z. equal(X::'a,Y) & defined(X::'a,Z) --> defined(Y::'a,Z)) & (\X Z Y. equal(X::'a,Y) & defined(Z::'a,X) --> defined(Z::'a,Y)) & (\X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & (\X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z)))" (*54288 inferences so far. Searching to depth 14. 118.0 secs*) lemma CAT005_1: "EQU001_0_ax equal & CAT001_0_ax equal codomain domain identity_map compos product defined & CAT001_0_eq compos defined identity_map codomain domain product equal & (defined(a::'a,d)) & (identity_map(d)) & (~equal(domain(a),d)) --> False" by meson (*1728 inferences so far. Searching to depth 10. 5.8 secs*) lemma CAT007_1: "EQU001_0_ax equal & CAT001_0_ax equal codomain domain identity_map compos product defined & CAT001_0_eq compos defined identity_map codomain domain product equal & (equal(domain(a),codomain(b))) & (~defined(a::'a,b)) --> False" by meson (*82895 inferences so far. Searching to depth 13. 355 secs*) lemma CAT018_1: "EQU001_0_ax equal & CAT001_0_ax equal codomain domain identity_map compos product defined & CAT001_0_eq compos defined identity_map codomain domain product equal & (defined(a::'a,b)) & (defined(b::'a,c)) & (~defined(a::'a,compos(b::'a,c))) --> False" by meson (*1118 inferences so far. Searching to depth 8. 2.3 secs*) lemma COL001_2: "EQU001_0_ax equal & (\X Y Z. equal(apply(apply(apply(s::'a,X),Y),Z),apply(apply(X::'a,Z),apply(Y::'a,Z)))) & (\Y X. equal(apply(apply(k::'a,X),Y),X)) & (\X Y Z. equal(apply(apply(apply(b::'a,X),Y),Z),apply(X::'a,apply(Y::'a,Z)))) & (\X. equal(apply(i::'a,X),X)) & (\A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) & (\X. equal(apply(apply(apply(s::'a,apply(b::'a,X)),i),apply(apply(s::'a,apply(b::'a,X)),i)),apply(x::'a,apply(apply(apply(s::'a,apply(b::'a,X)),i),apply(apply(s::'a,apply(b::'a,X)),i))))) & (\Y. ~equal(Y::'a,apply(combinator::'a,Y))) --> False" by meson (*500 inferences so far. Searching to depth 8. 0.9 secs*) lemma COL023_1: "EQU001_0_ax equal & (\X Y Z. equal(apply(apply(apply(b::'a,X),Y),Z),apply(X::'a,apply(Y::'a,Z)))) & (\X Y Z. equal(apply(apply(apply(n::'a,X),Y),Z),apply(apply(apply(X::'a,Z),Y),Z))) & (\A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) & (\Y. ~equal(Y::'a,apply(combinator::'a,Y))) --> False" by meson (*3018 inferences so far. Searching to depth 10. 4.3 secs*) lemma COL032_1: "EQU001_0_ax equal & (\X. equal(apply(m::'a,X),apply(X::'a,X))) & (\Y X Z. equal(apply(apply(apply(q::'a,X),Y),Z),apply(Y::'a,apply(X::'a,Z)))) & (\A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) & (\G H. equal(G::'a,H) --> equal(f(G),f(H))) & (\Y. ~equal(apply(Y::'a,f(Y)),apply(f(Y),apply(Y::'a,f(Y))))) --> False" by meson (*381878 inferences so far. Searching to depth 13. 670.4 secs*) lemma COL052_2: "EQU001_0_ax equal & (\X Y W. equal(response(compos(X::'a,Y),W),response(X::'a,response(Y::'a,W)))) & (\X Y. agreeable(X) --> equal(response(X::'a,common_bird(Y)),response(Y::'a,common_bird(Y)))) & (\Z X. equal(response(X::'a,Z),response(compatible(X),Z)) --> agreeable(X)) & (\A B. equal(A::'a,B) --> equal(common_bird(A),common_bird(B))) & (\C D. equal(C::'a,D) --> equal(compatible(C),compatible(D))) & (\Q R. equal(Q::'a,R) & agreeable(Q) --> agreeable(R)) & (\A B C. equal(A::'a,B) --> equal(compos(A::'a,C),compos(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(compos(F'::'a,D),compos(F'::'a,E))) & (\G H I'. equal(G::'a,H) --> equal(response(G::'a,I'),response(H::'a,I'))) & (\J L K'. equal(J::'a,K') --> equal(response(L::'a,J),response(L::'a,K'))) & (agreeable(c)) & (~agreeable(a)) & (equal(c::'a,compos(a::'a,b))) --> False" by meson (*13201 inferences so far. Searching to depth 11. 31.9 secs*) lemma COL075_2: "EQU001_0_ax equal & (\Y X. equal(apply(apply(k::'a,X),Y),X)) & (\X Y Z. equal(apply(apply(apply(abstraction::'a,X),Y),Z),apply(apply(X::'a,apply(k::'a,Z)),apply(Y::'a,Z)))) & (\D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) & (\G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) & (\A B. equal(A::'a,B) --> equal(b(A),b(B))) & (\C D. equal(C::'a,D) --> equal(c(C),c(D))) & (\Y. ~equal(apply(apply(Y::'a,b(Y)),c(Y)),apply(b(Y),b(Y)))) --> False" by meson (*33 inferences so far. Searching to depth 7. 0.1 secs*) lemma COM001_1: "(\Goal_state Start_state. follows(Goal_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) & (\Goal_state Intermediate_state Start_state. succeeds(Goal_state::'a,Intermediate_state) & succeeds(Intermediate_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) & (\Start_state Label Goal_state. has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state) --> succeeds(Goal_state::'a,Start_state)) & (\Start_state Condition Goal_state. has(Start_state::'a,ifthen(Condition::'a,Goal_state)) --> succeeds(Goal_state::'a,Start_state)) & (labels(loop::'a,p3)) & (has(p3::'a,ifthen(equal(register_j::'a,n),p4))) & (has(p4::'a,goto(out))) & (follows(p5::'a,p4)) & (follows(p8::'a,p3)) & (has(p8::'a,goto(loop))) & (~succeeds(p3::'a,p3)) --> False" by meson (*533 inferences so far. Searching to depth 13. 0.3 secs*) lemma COM002_1: "(\Goal_state Start_state. follows(Goal_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) & (\Goal_state Intermediate_state Start_state. succeeds(Goal_state::'a,Intermediate_state) & succeeds(Intermediate_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) & (\Start_state Label Goal_state. has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state) --> succeeds(Goal_state::'a,Start_state)) & (\Start_state Condition Goal_state. has(Start_state::'a,ifthen(Condition::'a,Goal_state)) --> succeeds(Goal_state::'a,Start_state)) & (has(p1::'a,assign(register_j::'a,num0))) & (follows(p2::'a,p1)) & (has(p2::'a,assign(register_k::'a,num1))) & (labels(loop::'a,p3)) & (follows(p3::'a,p2)) & (has(p3::'a,ifthen(equal(register_j::'a,n),p4))) & (has(p4::'a,goto(out))) & (follows(p5::'a,p4)) & (follows(p6::'a,p3)) & (has(p6::'a,assign(register_k::'a,mtimes(num2::'a,register_k)))) & (follows(p7::'a,p6)) & (has(p7::'a,assign(register_j::'a,mplus(register_j::'a,num1)))) & (follows(p8::'a,p7)) & (has(p8::'a,goto(loop))) & (~succeeds(p3::'a,p3)) --> False" by meson (*4821 inferences so far. Searching to depth 14. 1.3 secs*) lemma COM002_2: "(\Goal_state Start_state. ~(fails(Goal_state::'a,Start_state) & follows(Goal_state::'a,Start_state))) & (\Goal_state Intermediate_state Start_state. fails(Goal_state::'a,Start_state) --> fails(Goal_state::'a,Intermediate_state) | fails(Intermediate_state::'a,Start_state)) & (\Start_state Label Goal_state. ~(fails(Goal_state::'a,Start_state) & has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state))) & (\Start_state Condition Goal_state. ~(fails(Goal_state::'a,Start_state) & has(Start_state::'a,ifthen(Condition::'a,Goal_state)))) & (has(p1::'a,assign(register_j::'a,num0))) & (follows(p2::'a,p1)) & (has(p2::'a,assign(register_k::'a,num1))) & (labels(loop::'a,p3)) & (follows(p3::'a,p2)) & (has(p3::'a,ifthen(equal(register_j::'a,n),p4))) & (has(p4::'a,goto(out))) & (follows(p5::'a,p4)) & (follows(p6::'a,p3)) & (has(p6::'a,assign(register_k::'a,mtimes(num2::'a,register_k)))) & (follows(p7::'a,p6)) & (has(p7::'a,assign(register_j::'a,mplus(register_j::'a,num1)))) & (follows(p8::'a,p7)) & (has(p8::'a,goto(loop))) & (fails(p3::'a,p3)) --> False" by meson (*98 inferences so far. Searching to depth 10. 1.1 secs*) lemma COM003_2: "(\X Y Z. program_decides(X) & program(Y) --> decides(X::'a,Y,Z)) & (\X. program_decides(X) | program(f2(X))) & (\X. decides(X::'a,f2(X),f1(X)) --> program_decides(X)) & (\X. program_program_decides(X) --> program(X)) & (\X. program_program_decides(X) --> program_decides(X)) & (\X. program(X) & program_decides(X) --> program_program_decides(X)) & (\X. algorithm_program_decides(X) --> algorithm(X)) & (\X. algorithm_program_decides(X) --> program_decides(X)) & (\X. algorithm(X) & program_decides(X) --> algorithm_program_decides(X)) & (\Y X. program_halts2(X::'a,Y) --> program(X)) & (\X Y. program_halts2(X::'a,Y) --> halts2(X::'a,Y)) & (\X Y. program(X) & halts2(X::'a,Y) --> program_halts2(X::'a,Y)) & (\W X Y Z. halts3_outputs(X::'a,Y,Z,W) --> halts3(X::'a,Y,Z)) & (\Y Z X W. halts3_outputs(X::'a,Y,Z,W) --> outputs(X::'a,W)) & (\Y Z X W. halts3(X::'a,Y,Z) & outputs(X::'a,W) --> halts3_outputs(X::'a,Y,Z,W)) & (\Y X. program_not_halts2(X::'a,Y) --> program(X)) & (\X Y. ~(program_not_halts2(X::'a,Y) & halts2(X::'a,Y))) & (\X Y. program(X) --> program_not_halts2(X::'a,Y) | halts2(X::'a,Y)) & (\W X Y. halts2_outputs(X::'a,Y,W) --> halts2(X::'a,Y)) & (\Y X W. halts2_outputs(X::'a,Y,W) --> outputs(X::'a,W)) & (\Y X W. halts2(X::'a,Y) & outputs(X::'a,W) --> halts2_outputs(X::'a,Y,W)) & (\X W Y Z. program_halts2_halts3_outputs(X::'a,Y,Z,W) --> program_halts2(Y::'a,Z)) & (\X Y Z W. program_halts2_halts3_outputs(X::'a,Y,Z,W) --> halts3_outputs(X::'a,Y,Z,W)) & (\X Y Z W. program_halts2(Y::'a,Z) & halts3_outputs(X::'a,Y,Z,W) --> program_halts2_halts3_outputs(X::'a,Y,Z,W)) & (\X W Y Z. program_not_halts2_halts3_outputs(X::'a,Y,Z,W) --> program_not_halts2(Y::'a,Z)) & (\X Y Z W. program_not_halts2_halts3_outputs(X::'a,Y,Z,W) --> halts3_outputs(X::'a,Y,Z,W)) & (\X Y Z W. program_not_halts2(Y::'a,Z) & halts3_outputs(X::'a,Y,Z,W) --> program_not_halts2_halts3_outputs(X::'a,Y,Z,W)) & (\X W Y. program_halts2_halts2_outputs(X::'a,Y,W) --> program_halts2(Y::'a,Y)) & (\X Y W. program_halts2_halts2_outputs(X::'a,Y,W) --> halts2_outputs(X::'a,Y,W)) & (\X Y W. program_halts2(Y::'a,Y) & halts2_outputs(X::'a,Y,W) --> program_halts2_halts2_outputs(X::'a,Y,W)) & (\X W Y. program_not_halts2_halts2_outputs(X::'a,Y,W) --> program_not_halts2(Y::'a,Y)) & (\X Y W. program_not_halts2_halts2_outputs(X::'a,Y,W) --> halts2_outputs(X::'a,Y,W)) & (\X Y W. program_not_halts2(Y::'a,Y) & halts2_outputs(X::'a,Y,W) --> program_not_halts2_halts2_outputs(X::'a,Y,W)) & (\X. algorithm_program_decides(X) --> program_program_decides(c1)) & (\W Y Z. program_program_decides(W) --> program_halts2_halts3_outputs(W::'a,Y,Z,good)) & (\W Y Z. program_program_decides(W) --> program_not_halts2_halts3_outputs(W::'a,Y,Z,bad)) & (\W. program(W) & program_halts2_halts3_outputs(W::'a,f3(W),f3(W),good) & program_not_halts2_halts3_outputs(W::'a,f3(W),f3(W),bad) --> program(c2)) & (\W Y. program(W) & program_halts2_halts3_outputs(W::'a,f3(W),f3(W),good) & program_not_halts2_halts3_outputs(W::'a,f3(W),f3(W),bad) --> program_halts2_halts2_outputs(c2::'a,Y,good)) & (\W Y. program(W) & program_halts2_halts3_outputs(W::'a,f3(W),f3(W),good) & program_not_halts2_halts3_outputs(W::'a,f3(W),f3(W),bad) --> program_not_halts2_halts2_outputs(c2::'a,Y,bad)) & (\V. program(V) & program_halts2_halts2_outputs(V::'a,f4(V),good) & program_not_halts2_halts2_outputs(V::'a,f4(V),bad) --> program(c3)) & (\V Y. program(V) & program_halts2_halts2_outputs(V::'a,f4(V),good) & program_not_halts2_halts2_outputs(V::'a,f4(V),bad) & program_halts2(Y::'a,Y) --> halts2(c3::'a,Y)) & (\V Y. program(V) & program_halts2_halts2_outputs(V::'a,f4(V),good) & program_not_halts2_halts2_outputs(V::'a,f4(V),bad) --> program_not_halts2_halts2_outputs(c3::'a,Y,bad)) & (algorithm_program_decides(c4)) --> False" by meson (*2100398 inferences so far. Searching to depth 12. 1256s (21 mins) on griffon*) lemma COM004_1: "EQU001_0_ax equal & (\C D P Q X Y. failure_node(X::'a,or(C::'a,P)) & failure_node(Y::'a,or(D::'a,Q)) & contradictory(P::'a,Q) & siblings(X::'a,Y) --> failure_node(parent_of(X::'a,Y),or(C::'a,D))) & (\X. contradictory(negate(X),X)) & (\X. contradictory(X::'a,negate(X))) & (\X. siblings(left_child_of(X),right_child_of(X))) & (\D E. equal(D::'a,E) --> equal(left_child_of(D),left_child_of(E))) & (\F' G. equal(F'::'a,G) --> equal(negate(F'),negate(G))) & (\H I' J. equal(H::'a,I') --> equal(or(H::'a,J),or(I'::'a,J))) & (\K' M L. equal(K'::'a,L) --> equal(or(M::'a,K'),or(M::'a,L))) & (\N O' P. equal(N::'a,O') --> equal(parent_of(N::'a,P),parent_of(O'::'a,P))) & (\Q S' R. equal(Q::'a,R) --> equal(parent_of(S'::'a,Q),parent_of(S'::'a,R))) & (\T' U. equal(T'::'a,U) --> equal(right_child_of(T'),right_child_of(U))) & (\V W X. equal(V::'a,W) & contradictory(V::'a,X) --> contradictory(W::'a,X)) & (\Y A1 Z. equal(Y::'a,Z) & contradictory(A1::'a,Y) --> contradictory(A1::'a,Z)) & (\B1 C1 D1. equal(B1::'a,C1) & failure_node(B1::'a,D1) --> failure_node(C1::'a,D1)) & (\E1 G1 F1. equal(E1::'a,F1) & failure_node(G1::'a,E1) --> failure_node(G1::'a,F1)) & (\H1 I1 J1. equal(H1::'a,I1) & siblings(H1::'a,J1) --> siblings(I1::'a,J1)) & (\K1 M1 L1. equal(K1::'a,L1) & siblings(M1::'a,K1) --> siblings(M1::'a,L1)) & (failure_node(n_left::'a,or(EMPTY::'a,atom))) & (failure_node(n_right::'a,or(EMPTY::'a,negate(atom)))) & (equal(n_left::'a,left_child_of(n))) & (equal(n_right::'a,right_child_of(n))) & (\Z. ~failure_node(Z::'a,or(EMPTY::'a,EMPTY))) --> False" oops abbreviation "GEO001_0_ax continuous lower_dimension_point_3 lower_dimension_point_2 lower_dimension_point_1 extension euclid2 euclid1 outer_pasch equidistant equal between \ (\X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) & (\V X Y Z. between(X::'a,Y,V) & between(Y::'a,Z,V) --> between(X::'a,Y,Z)) & (\Y X V Z. between(X::'a,Y,Z) & between(X::'a,Y,V) --> equal(X::'a,Y) | between(X::'a,Z,V) | between(X::'a,V,Z)) & (\Y X. equidistant(X::'a,Y,Y,X)) & (\Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) & (\X Y Z V V2 W. equidistant(X::'a,Y,Z,V) & equidistant(X::'a,Y,V2,W) --> equidistant(Z::'a,V,V2,W)) & (\W X Z V Y. between(X::'a,W,V) & between(Y::'a,V,Z) --> between(X::'a,outer_pasch(W::'a,X,Y,Z,V),Y)) & (\W X Y Z V. between(X::'a,W,V) & between(Y::'a,V,Z) --> between(Z::'a,W,outer_pasch(W::'a,X,Y,Z,V))) & (\W X Y Z V. between(X::'a,V,W) & between(Y::'a,V,Z) --> equal(X::'a,V) | between(X::'a,Z,euclid1(W::'a,X,Y,Z,V))) & (\W X Y Z V. between(X::'a,V,W) & between(Y::'a,V,Z) --> equal(X::'a,V) | between(X::'a,Y,euclid2(W::'a,X,Y,Z,V))) & (\W X Y Z V. between(X::'a,V,W) & between(Y::'a,V,Z) --> equal(X::'a,V) | between(euclid1(W::'a,X,Y,Z,V),W,euclid2(W::'a,X,Y,Z,V))) & (\X1 Y1 X Y Z V Z1 V1. equidistant(X::'a,Y,X1,Y1) & equidistant(Y::'a,Z,Y1,Z1) & equidistant(X::'a,V,X1,V1) & equidistant(Y::'a,V,Y1,V1) & between(X::'a,Y,Z) & between(X1::'a,Y1,Z1) --> equal(X::'a,Y) | equidistant(Z::'a,V,Z1,V1)) & (\X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) & (\X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) & (~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) & (~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) & (~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) & (\Z X Y W V. equidistant(X::'a,W,X,V) & equidistant(Y::'a,W,Y,V) & equidistant(Z::'a,W,Z,V) --> between(X::'a,Y,Z) | between(Y::'a,Z,X) | between(Z::'a,X,Y) | equal(W::'a,V)) & (\X Y Z X1 Z1 V. equidistant(V::'a,X,V,X1) & equidistant(V::'a,Z,V,Z1) & between(V::'a,X,Z) & between(X::'a,Y,Z) --> equidistant(V::'a,Y,Z,continuous(X::'a,Y,Z,X1,Z1,V))) & (\X Y Z X1 V Z1. equidistant(V::'a,X,V,X1) & equidistant(V::'a,Z,V,Z1) & between(V::'a,X,Z) & between(X::'a,Y,Z) --> between(X1::'a,continuous(X::'a,Y,Z,X1,Z1,V),Z1))" abbreviation "GEO001_0_eq continuous extension euclid2 euclid1 outer_pasch equidistant between equal \ (\X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) & (\X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) & (\X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) & (\X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) & (\X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) & (\X Y V1 V2 V3 V4. equal(X::'a,Y) --> equal(outer_pasch(X::'a,V1,V2,V3,V4),outer_pasch(Y::'a,V1,V2,V3,V4))) & (\X V1 Y V2 V3 V4. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,X,V2,V3,V4),outer_pasch(V1::'a,Y,V2,V3,V4))) & (\X V1 V2 Y V3 V4. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,V2,X,V3,V4),outer_pasch(V1::'a,V2,Y,V3,V4))) & (\X V1 V2 V3 Y V4. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,V2,V3,X,V4),outer_pasch(V1::'a,V2,V3,Y,V4))) & (\X V1 V2 V3 V4 Y. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,V2,V3,V4,X),outer_pasch(V1::'a,V2,V3,V4,Y))) & (\A B C D E F'. equal(A::'a,B) --> equal(euclid1(A::'a,C,D,E,F'),euclid1(B::'a,C,D,E,F'))) & (\G I' H J K' L. equal(G::'a,H) --> equal(euclid1(I'::'a,G,J,K',L),euclid1(I'::'a,H,J,K',L))) & (\M O' P N Q R. equal(M::'a,N) --> equal(euclid1(O'::'a,P,M,Q,R),euclid1(O'::'a,P,N,Q,R))) & (\S' U V W T' X. equal(S'::'a,T') --> equal(euclid1(U::'a,V,W,S',X),euclid1(U::'a,V,W,T',X))) & (\Y A1 B1 C1 D1 Z. equal(Y::'a,Z) --> equal(euclid1(A1::'a,B1,C1,D1,Y),euclid1(A1::'a,B1,C1,D1,Z))) & (\E1 F1 G1 H1 I1 J1. equal(E1::'a,F1) --> equal(euclid2(E1::'a,G1,H1,I1,J1),euclid2(F1::'a,G1,H1,I1,J1))) & (\K1 M1 L1 N1 O1 P1. equal(K1::'a,L1) --> equal(euclid2(M1::'a,K1,N1,O1,P1),euclid2(M1::'a,L1,N1,O1,P1))) & (\Q1 S1 T1 R1 U1 V1. equal(Q1::'a,R1) --> equal(euclid2(S1::'a,T1,Q1,U1,V1),euclid2(S1::'a,T1,R1,U1,V1))) & (\W1 Y1 Z1 A2 X1 B2. equal(W1::'a,X1) --> equal(euclid2(Y1::'a,Z1,A2,W1,B2),euclid2(Y1::'a,Z1,A2,X1,B2))) & (\C2 E2 F2 G2 H2 D2. equal(C2::'a,D2) --> equal(euclid2(E2::'a,F2,G2,H2,C2),euclid2(E2::'a,F2,G2,H2,D2))) & (\X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) & (\X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) & (\X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) & (\X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) & (\X Y V1 V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(X::'a,V1,V2,V3,V4,V5),continuous(Y::'a,V1,V2,V3,V4,V5))) & (\X V1 Y V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,X,V2,V3,V4,V5),continuous(V1::'a,Y,V2,V3,V4,V5))) & (\X V1 V2 Y V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,X,V3,V4,V5),continuous(V1::'a,V2,Y,V3,V4,V5))) & (\X V1 V2 V3 Y V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,X,V4,V5),continuous(V1::'a,V2,V3,Y,V4,V5))) & (\X V1 V2 V3 V4 Y V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,X,V5),continuous(V1::'a,V2,V3,V4,Y,V5))) & (\X V1 V2 V3 V4 V5 Y. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,V5,X),continuous(V1::'a,V2,V3,V4,V5,Y)))" (*179 inferences so far. Searching to depth 7. 3.9 secs*) lemma GEO003_1: "EQU001_0_ax equal & GEO001_0_ax continuous lower_dimension_point_3 lower_dimension_point_2 lower_dimension_point_1 extension euclid2 euclid1 outer_pasch equidistant equal between & GEO001_0_eq continuous extension euclid2 euclid1 outer_pasch equidistant between equal & (~between(a::'a,b,b)) --> False" by meson abbreviation "GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3 lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension between equal equidistant \ (\Y X. equidistant(X::'a,Y,Y,X)) & (\X Y Z V V2 W. equidistant(X::'a,Y,Z,V) & equidistant(X::'a,Y,V2,W) --> equidistant(Z::'a,V,V2,W)) & (\Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) & (\X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) & (\X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) & (\X1 Y1 X Y Z V Z1 V1. equidistant(X::'a,Y,X1,Y1) & equidistant(Y::'a,Z,Y1,Z1) & equidistant(X::'a,V,X1,V1) & equidistant(Y::'a,V,Y1,V1) & between(X::'a,Y,Z) & between(X1::'a,Y1,Z1) --> equal(X::'a,Y) | equidistant(Z::'a,V,Z1,V1)) & (\X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) & (\U V W X Y. between(U::'a,V,W) & between(Y::'a,X,W) --> between(V::'a,inner_pasch(U::'a,V,W,X,Y),Y)) & (\V W X Y U. between(U::'a,V,W) & between(Y::'a,X,W) --> between(X::'a,inner_pasch(U::'a,V,W,X,Y),U)) & (~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) & (~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) & (~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) & (\Z X Y W V. equidistant(X::'a,W,X,V) & equidistant(Y::'a,W,Y,V) & equidistant(Z::'a,W,Z,V) --> between(X::'a,Y,Z) | between(Y::'a,Z,X) | between(Z::'a,X,Y) | equal(W::'a,V)) & (\U V W X Y. between(U::'a,W,Y) & between(V::'a,W,X) --> equal(U::'a,W) | between(U::'a,V,euclid1(U::'a,V,W,X,Y))) & (\U V W X Y. between(U::'a,W,Y) & between(V::'a,W,X) --> equal(U::'a,W) | between(U::'a,X,euclid2(U::'a,V,W,X,Y))) & (\U V W X Y. between(U::'a,W,Y) & between(V::'a,W,X) --> equal(U::'a,W) | between(euclid1(U::'a,V,W,X,Y),Y,euclid2(U::'a,V,W,X,Y))) & (\U V V1 W X X1. equidistant(U::'a,V,U,V1) & equidistant(U::'a,X,U,X1) & between(U::'a,V,X) & between(V::'a,W,X) --> between(V1::'a,continuous(U::'a,V,V1,W,X,X1),X1)) & (\U V V1 W X X1. equidistant(U::'a,V,U,V1) & equidistant(U::'a,X,U,X1) & between(U::'a,V,X) & between(V::'a,W,X) --> equidistant(U::'a,W,U,continuous(U::'a,V,V1,W,X,X1))) & (\X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) & (\X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) & (\X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) & (\X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) & (\X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) & (\X Y V1 V2 V3 V4. equal(X::'a,Y) --> equal(inner_pasch(X::'a,V1,V2,V3,V4),inner_pasch(Y::'a,V1,V2,V3,V4))) & (\X V1 Y V2 V3 V4. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,X,V2,V3,V4),inner_pasch(V1::'a,Y,V2,V3,V4))) & (\X V1 V2 Y V3 V4. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,V2,X,V3,V4),inner_pasch(V1::'a,V2,Y,V3,V4))) & (\X V1 V2 V3 Y V4. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,V2,V3,X,V4),inner_pasch(V1::'a,V2,V3,Y,V4))) & (\X V1 V2 V3 V4 Y. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,V2,V3,V4,X),inner_pasch(V1::'a,V2,V3,V4,Y))) & (\A B C D E F'. equal(A::'a,B) --> equal(euclid1(A::'a,C,D,E,F'),euclid1(B::'a,C,D,E,F'))) & (\G I' H J K' L. equal(G::'a,H) --> equal(euclid1(I'::'a,G,J,K',L),euclid1(I'::'a,H,J,K',L))) & (\M O' P N Q R. equal(M::'a,N) --> equal(euclid1(O'::'a,P,M,Q,R),euclid1(O'::'a,P,N,Q,R))) & (\S' U V W T' X. equal(S'::'a,T') --> equal(euclid1(U::'a,V,W,S',X),euclid1(U::'a,V,W,T',X))) & (\Y A1 B1 C1 D1 Z. equal(Y::'a,Z) --> equal(euclid1(A1::'a,B1,C1,D1,Y),euclid1(A1::'a,B1,C1,D1,Z))) & (\E1 F1 G1 H1 I1 J1. equal(E1::'a,F1) --> equal(euclid2(E1::'a,G1,H1,I1,J1),euclid2(F1::'a,G1,H1,I1,J1))) & (\K1 M1 L1 N1 O1 P1. equal(K1::'a,L1) --> equal(euclid2(M1::'a,K1,N1,O1,P1),euclid2(M1::'a,L1,N1,O1,P1))) & (\Q1 S1 T1 R1 U1 V1. equal(Q1::'a,R1) --> equal(euclid2(S1::'a,T1,Q1,U1,V1),euclid2(S1::'a,T1,R1,U1,V1))) & (\W1 Y1 Z1 A2 X1 B2. equal(W1::'a,X1) --> equal(euclid2(Y1::'a,Z1,A2,W1,B2),euclid2(Y1::'a,Z1,A2,X1,B2))) & (\C2 E2 F2 G2 H2 D2. equal(C2::'a,D2) --> equal(euclid2(E2::'a,F2,G2,H2,C2),euclid2(E2::'a,F2,G2,H2,D2))) & (\X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) & (\X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) & (\X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) & (\X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) & (\X Y V1 V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(X::'a,V1,V2,V3,V4,V5),continuous(Y::'a,V1,V2,V3,V4,V5))) & (\X V1 Y V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,X,V2,V3,V4,V5),continuous(V1::'a,Y,V2,V3,V4,V5))) & (\X V1 V2 Y V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,X,V3,V4,V5),continuous(V1::'a,V2,Y,V3,V4,V5))) & (\X V1 V2 V3 Y V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,X,V4,V5),continuous(V1::'a,V2,V3,Y,V4,V5))) & (\X V1 V2 V3 V4 Y V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,X,V5),continuous(V1::'a,V2,V3,V4,Y,V5))) & (\X V1 V2 V3 V4 V5 Y. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,V5,X),continuous(V1::'a,V2,V3,V4,V5,Y)))" (*4272 inferences so far. Searching to depth 10. 29.4 secs*) lemma GEO017_2: "EQU001_0_ax equal & GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3 lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension between equal equidistant & (equidistant(u::'a,v,w,x)) & (~equidistant(u::'a,v,x,w)) --> False" by meson (*382903 inferences so far. Searching to depth 9. 1022s (17 mins) on griffon*) lemma GEO027_3: "EQU001_0_ax equal & GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3 lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension between equal equidistant & (\U V. equal(reflection(U::'a,V),extension(U::'a,V,U,V))) & (\X Y Z. equal(X::'a,Y) --> equal(reflection(X::'a,Z),reflection(Y::'a,Z))) & (\A1 C1 B1. equal(A1::'a,B1) --> equal(reflection(C1::'a,A1),reflection(C1::'a,B1))) & (\U V. equidistant(U::'a,V,U,V)) & (\W X U V. equidistant(U::'a,V,W,X) --> equidistant(W::'a,X,U,V)) & (\V U W X. equidistant(U::'a,V,W,X) --> equidistant(V::'a,U,W,X)) & (\U V X W. equidistant(U::'a,V,W,X) --> equidistant(U::'a,V,X,W)) & (\V U X W. equidistant(U::'a,V,W,X) --> equidistant(V::'a,U,X,W)) & (\W X V U. equidistant(U::'a,V,W,X) --> equidistant(W::'a,X,V,U)) & (\X W U V. equidistant(U::'a,V,W,X) --> equidistant(X::'a,W,U,V)) & (\X W V U. equidistant(U::'a,V,W,X) --> equidistant(X::'a,W,V,U)) & (\W X U V Y Z. equidistant(U::'a,V,W,X) & equidistant(W::'a,X,Y,Z) --> equidistant(U::'a,V,Y,Z)) & (\U V W. equal(V::'a,extension(U::'a,V,W,W))) & (\W X U V Y. equal(Y::'a,extension(U::'a,V,W,X)) --> between(U::'a,V,Y)) & (\U V. between(U::'a,V,reflection(U::'a,V))) & (\U V. equidistant(V::'a,reflection(U::'a,V),U,V)) & (\U V. equal(U::'a,V) --> equal(V::'a,reflection(U::'a,V))) & (\U. equal(U::'a,reflection(U::'a,U))) & (\U V. equal(V::'a,reflection(U::'a,V)) --> equal(U::'a,V)) & (\U V. equidistant(U::'a,U,V,V)) & (\V V1 U W U1 W1. equidistant(U::'a,V,U1,V1) & equidistant(V::'a,W,V1,W1) & between(U::'a,V,W) & between(U1::'a,V1,W1) --> equidistant(U::'a,W,U1,W1)) & (\U V W X. between(U::'a,V,W) & between(U::'a,V,X) & equidistant(V::'a,W,V,X) --> equal(U::'a,V) | equal(W::'a,X)) & (between(u::'a,v,w)) & (~equal(u::'a,v)) & (~equal(w::'a,extension(u::'a,v,v,w))) --> False" oops (*313884 inferences so far. Searching to depth 10. 887 secs: 15 mins.*) lemma GEO058_2: "EQU001_0_ax equal & GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3 lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension between equal equidistant & (\U V. equal(reflection(U::'a,V),extension(U::'a,V,U,V))) & (\X Y Z. equal(X::'a,Y) --> equal(reflection(X::'a,Z),reflection(Y::'a,Z))) & (\A1 C1 B1. equal(A1::'a,B1) --> equal(reflection(C1::'a,A1),reflection(C1::'a,B1))) & (equal(v::'a,reflection(u::'a,v))) & (~equal(u::'a,v)) --> False" oops (*0 inferences so far. Searching to depth 0. 0.2 secs*) lemma GEO079_1: "(\U V W X Y Z. right_angle(U::'a,V,W) & right_angle(X::'a,Y,Z) --> eq(U::'a,V,W,X,Y,Z)) & (\U V W X Y Z. CONGRUENT(U::'a,V,W,X,Y,Z) --> eq(U::'a,V,W,X,Y,Z)) & (\V W U X. trapezoid(U::'a,V,W,X) --> parallel(V::'a,W,U,X)) & (\U V X Y. parallel(U::'a,V,X,Y) --> eq(X::'a,V,U,V,X,Y)) & (trapezoid(a::'a,b,c,d)) & (~eq(a::'a,c,b,c,a,d)) --> False" by meson abbreviation "GRP003_0_ax equal multiply INVERSE identity product \ (\X. product(identity::'a,X,X)) & (\X. product(X::'a,identity,X)) & (\X. product(INVERSE(X),X,identity)) & (\X. product(X::'a,INVERSE(X),identity)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W))" abbreviation "GRP003_0_eq product multiply INVERSE equal \ (\X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) & (\X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) & (\X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) & (\X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y))" (*2032008 inferences so far. Searching to depth 16. 658s (11 mins) on griffon*) lemma GRP001_1: "EQU001_0_ax equal & GRP003_0_ax equal multiply INVERSE identity product & GRP003_0_eq product multiply INVERSE equal & (\X. product(X::'a,X,identity)) & (product(a::'a,b,c)) & (~product(b::'a,a,c)) --> False" oops (*2386 inferences so far. Searching to depth 11. 8.7 secs*) lemma GRP008_1: "EQU001_0_ax equal & GRP003_0_ax equal multiply INVERSE identity product & GRP003_0_eq product multiply INVERSE equal & (\A B. equal(A::'a,B) --> equal(h(A),h(B))) & (\C D. equal(C::'a,D) --> equal(j(C),j(D))) & (\A B. equal(A::'a,B) & q(A) --> q(B)) & (\B A C. q(A) & product(A::'a,B,C) --> product(B::'a,A,C)) & (\A. product(j(A),A,h(A)) | product(A::'a,j(A),h(A)) | q(A)) & (\A. product(j(A),A,h(A)) & product(A::'a,j(A),h(A)) --> q(A)) & (~q(identity)) --> False" by meson (*8625 inferences so far. Searching to depth 11. 20 secs*) lemma GRP013_1: "EQU001_0_ax equal & GRP003_0_ax equal multiply INVERSE identity product & GRP003_0_eq product multiply INVERSE equal & (\A. product(A::'a,A,identity)) & (product(a::'a,b,c)) & (product(INVERSE(a),INVERSE(b),d)) & (\A C B. product(INVERSE(A),INVERSE(B),C) --> product(A::'a,C,B)) & (~product(c::'a,d,identity)) --> False" by meson (*2448 inferences so far. Searching to depth 10. 7.2 secs*) lemma GRP037_3: "EQU001_0_ax equal & GRP003_0_ax equal multiply INVERSE identity product & GRP003_0_eq product multiply INVERSE equal & (\A B C. subgroup_member(A) & subgroup_member(B) & product(A::'a,INVERSE(B),C) --> subgroup_member(C)) & (\A B. equal(A::'a,B) & subgroup_member(A) --> subgroup_member(B)) & (\A. subgroup_member(A) --> product(Gidentity::'a,A,A)) & (\A. subgroup_member(A) --> product(A::'a,Gidentity,A)) & (\A. subgroup_member(A) --> product(A::'a,Ginverse(A),Gidentity)) & (\A. subgroup_member(A) --> product(Ginverse(A),A,Gidentity)) & (\A. subgroup_member(A) --> subgroup_member(Ginverse(A))) & (\A B. equal(A::'a,B) --> equal(Ginverse(A),Ginverse(B))) & (\A C D B. product(A::'a,B,C) & product(A::'a,D,C) --> equal(D::'a,B)) & (\B C D A. product(A::'a,B,C) & product(D::'a,B,C) --> equal(D::'a,A)) & (subgroup_member(a)) & (subgroup_member(Gidentity)) & (~equal(INVERSE(a),Ginverse(a))) --> False" by meson (*163 inferences so far. Searching to depth 11. 0.3 secs*) lemma GRP031_2: "(\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) & (\A. product(A::'a,INVERSE(A),identity)) & (\A. product(A::'a,identity,A)) & (\A. ~product(A::'a,a,identity)) --> False" by meson (*47 inferences so far. Searching to depth 11. 0.2 secs*) lemma GRP034_4: "(\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X. product(identity::'a,X,X)) & (\X. product(X::'a,identity,X)) & (\X. product(X::'a,INVERSE(X),identity)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) & (\B A C. subgroup_member(A) & subgroup_member(B) & product(B::'a,INVERSE(A),C) --> subgroup_member(C)) & (subgroup_member(a)) & (~subgroup_member(INVERSE(a))) --> False" by meson (*3287 inferences so far. Searching to depth 14. 3.5 secs*) lemma GRP047_2: "(\X. product(identity::'a,X,X)) & (\X. product(INVERSE(X),X,identity)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) & (\X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) & (equal(a::'a,b)) & (~equal(multiply(c::'a,a),multiply(c::'a,b))) --> False" by meson (*25559 inferences so far. Searching to depth 19. 16.9 secs*) lemma GRP130_1_002: "(group_element(e_1)) & (group_element(e_2)) & (~equal(e_1::'a,e_2)) & (~equal(e_2::'a,e_1)) & (\X Y. group_element(X) & group_element(Y) --> product(X::'a,Y,e_1) | product(X::'a,Y,e_2)) & (\X Y W Z. product(X::'a,Y,W) & product(X::'a,Y,Z) --> equal(W::'a,Z)) & (\X Y W Z. product(X::'a,W,Y) & product(X::'a,Z,Y) --> equal(W::'a,Z)) & (\Y X W Z. product(W::'a,Y,X) & product(Z::'a,Y,X) --> equal(W::'a,Z)) & (\Z1 Z2 Y X. product(X::'a,Y,Z1) & product(X::'a,Z1,Z2) --> product(Z2::'a,Y,X)) --> False" by meson abbreviation "GRP004_0_ax INVERSE identity multiply equal \ (\X. equal(multiply(identity::'a,X),X)) & (\X. equal(multiply(INVERSE(X),X),identity)) & (\X Y Z. equal(multiply(multiply(X::'a,Y),Z),multiply(X::'a,multiply(Y::'a,Z)))) & (\A B. equal(A::'a,B) --> equal(INVERSE(A),INVERSE(B))) & (\C D E. equal(C::'a,D) --> equal(multiply(C::'a,E),multiply(D::'a,E))) & (\F' H G. equal(F'::'a,G) --> equal(multiply(H::'a,F'),multiply(H::'a,G)))" abbreviation "GRP004_2_ax multiply least_upper_bound greatest_lower_bound equal \ (\Y X. equal(greatest_lower_bound(X::'a,Y),greatest_lower_bound(Y::'a,X))) & (\Y X. equal(least_upper_bound(X::'a,Y),least_upper_bound(Y::'a,X))) & (\X Y Z. equal(greatest_lower_bound(X::'a,greatest_lower_bound(Y::'a,Z)),greatest_lower_bound(greatest_lower_bound(X::'a,Y),Z))) & (\X Y Z. equal(least_upper_bound(X::'a,least_upper_bound(Y::'a,Z)),least_upper_bound(least_upper_bound(X::'a,Y),Z))) & (\X. equal(least_upper_bound(X::'a,X),X)) & (\X. equal(greatest_lower_bound(X::'a,X),X)) & (\Y X. equal(least_upper_bound(X::'a,greatest_lower_bound(X::'a,Y)),X)) & (\Y X. equal(greatest_lower_bound(X::'a,least_upper_bound(X::'a,Y)),X)) & (\Y X Z. equal(multiply(X::'a,least_upper_bound(Y::'a,Z)),least_upper_bound(multiply(X::'a,Y),multiply(X::'a,Z)))) & (\Y X Z. equal(multiply(X::'a,greatest_lower_bound(Y::'a,Z)),greatest_lower_bound(multiply(X::'a,Y),multiply(X::'a,Z)))) & (\Y Z X. equal(multiply(least_upper_bound(Y::'a,Z),X),least_upper_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) & (\Y Z X. equal(multiply(greatest_lower_bound(Y::'a,Z),X),greatest_lower_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) & (\A B C. equal(A::'a,B) --> equal(greatest_lower_bound(A::'a,C),greatest_lower_bound(B::'a,C))) & (\A C B. equal(A::'a,B) --> equal(greatest_lower_bound(C::'a,A),greatest_lower_bound(C::'a,B))) & (\A B C. equal(A::'a,B) --> equal(least_upper_bound(A::'a,C),least_upper_bound(B::'a,C))) & (\A C B. equal(A::'a,B) --> equal(least_upper_bound(C::'a,A),least_upper_bound(C::'a,B))) & (\A B C. equal(A::'a,B) --> equal(multiply(A::'a,C),multiply(B::'a,C))) & (\A C B. equal(A::'a,B) --> equal(multiply(C::'a,A),multiply(C::'a,B)))" (*3468 inferences so far. Searching to depth 10. 9.1 secs*) lemma GRP156_1: "EQU001_0_ax equal & GRP004_0_ax INVERSE identity multiply equal & GRP004_2_ax multiply least_upper_bound greatest_lower_bound equal & (equal(least_upper_bound(a::'a,b),b)) & (~equal(greatest_lower_bound(multiply(a::'a,c),multiply(b::'a,c)),multiply(a::'a,c))) --> False" by meson (*4394 inferences so far. Searching to depth 10. 8.2 secs*) lemma GRP168_1: "EQU001_0_ax equal & GRP004_0_ax INVERSE identity multiply equal & GRP004_2_ax multiply least_upper_bound greatest_lower_bound equal & (equal(least_upper_bound(a::'a,b),b)) & (~equal(least_upper_bound(multiply(INVERSE(c),multiply(a::'a,c)),multiply(INVERSE(c),multiply(b::'a,c))),multiply(INVERSE(c),multiply(b::'a,c)))) --> False" by meson abbreviation "HEN002_0_ax identity Zero Divide equal mless_equal \ (\X Y. mless_equal(X::'a,Y) --> equal(Divide(X::'a,Y),Zero)) & (\X Y. equal(Divide(X::'a,Y),Zero) --> mless_equal(X::'a,Y)) & (\Y X. mless_equal(Divide(X::'a,Y),X)) & (\X Y Z. mless_equal(Divide(Divide(X::'a,Z),Divide(Y::'a,Z)),Divide(Divide(X::'a,Y),Z))) & (\X. mless_equal(Zero::'a,X)) & (\X Y. mless_equal(X::'a,Y) & mless_equal(Y::'a,X) --> equal(X::'a,Y)) & (\X. mless_equal(X::'a,identity))" abbreviation "HEN002_0_eq mless_equal Divide equal \ (\A B C. equal(A::'a,B) --> equal(Divide(A::'a,C),Divide(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(Divide(F'::'a,D),Divide(F'::'a,E))) & (\G H I'. equal(G::'a,H) & mless_equal(G::'a,I') --> mless_equal(H::'a,I')) & (\J L K'. equal(J::'a,K') & mless_equal(L::'a,J) --> mless_equal(L::'a,K'))" (*250258 inferences so far. Searching to depth 16. 406.2 secs*) lemma HEN003_3: "EQU001_0_ax equal & HEN002_0_ax identity Zero Divide equal mless_equal & HEN002_0_eq mless_equal Divide equal & (~equal(Divide(a::'a,a),Zero)) --> False" by meson (*202177 inferences so far. Searching to depth 14. 451 secs*) lemma HEN007_2: "EQU001_0_ax equal & (\X Y. mless_equal(X::'a,Y) --> quotient(X::'a,Y,Zero)) & (\X Y. quotient(X::'a,Y,Zero) --> mless_equal(X::'a,Y)) & (\Y Z X. quotient(X::'a,Y,Z) --> mless_equal(Z::'a,X)) & (\Y X V3 V2 V1 Z V4 V5. quotient(X::'a,Y,V1) & quotient(Y::'a,Z,V2) & quotient(X::'a,Z,V3) & quotient(V3::'a,V2,V4) & quotient(V1::'a,Z,V5) --> mless_equal(V4::'a,V5)) & (\X. mless_equal(Zero::'a,X)) & (\X Y. mless_equal(X::'a,Y) & mless_equal(Y::'a,X) --> equal(X::'a,Y)) & (\X. mless_equal(X::'a,identity)) & (\X Y. quotient(X::'a,Y,Divide(X::'a,Y))) & (\X Y Z W. quotient(X::'a,Y,Z) & quotient(X::'a,Y,W) --> equal(Z::'a,W)) & (\X Y W Z. equal(X::'a,Y) & quotient(X::'a,W,Z) --> quotient(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & quotient(W::'a,X,Z) --> quotient(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & quotient(W::'a,Z,X) --> quotient(W::'a,Z,Y)) & (\X Z Y. equal(X::'a,Y) & mless_equal(Z::'a,X) --> mless_equal(Z::'a,Y)) & (\X Y Z. equal(X::'a,Y) & mless_equal(X::'a,Z) --> mless_equal(Y::'a,Z)) & (\X Y W. equal(X::'a,Y) --> equal(Divide(X::'a,W),Divide(Y::'a,W))) & (\X W Y. equal(X::'a,Y) --> equal(Divide(W::'a,X),Divide(W::'a,Y))) & (\X. quotient(X::'a,identity,Zero)) & (\X. quotient(Zero::'a,X,Zero)) & (\X. quotient(X::'a,X,Zero)) & (\X. quotient(X::'a,Zero,X)) & (\Y X Z. mless_equal(X::'a,Y) & mless_equal(Y::'a,Z) --> mless_equal(X::'a,Z)) & (\W1 X Z W2 Y. quotient(X::'a,Y,W1) & mless_equal(W1::'a,Z) & quotient(X::'a,Z,W2) --> mless_equal(W2::'a,Y)) & (mless_equal(x::'a,y)) & (quotient(z::'a,y,zQy)) & (quotient(z::'a,x,zQx)) & (~mless_equal(zQy::'a,zQx)) --> False" by meson (*60026 inferences so far. Searching to depth 12. 42.2 secs*) lemma HEN008_4: "EQU001_0_ax equal & HEN002_0_ax identity Zero Divide equal mless_equal & HEN002_0_eq mless_equal Divide equal & (\X. equal(Divide(X::'a,identity),Zero)) & (\X. equal(Divide(Zero::'a,X),Zero)) & (\X. equal(Divide(X::'a,X),Zero)) & (equal(Divide(a::'a,Zero),a)) & (\Y X Z. mless_equal(X::'a,Y) & mless_equal(Y::'a,Z) --> mless_equal(X::'a,Z)) & (\X Z Y. mless_equal(Divide(X::'a,Y),Z) --> mless_equal(Divide(X::'a,Z),Y)) & (\Y Z X. mless_equal(X::'a,Y) --> mless_equal(Divide(Z::'a,Y),Divide(Z::'a,X))) & (mless_equal(a::'a,b)) & (~mless_equal(Divide(a::'a,c),Divide(b::'a,c))) --> False" by meson (*3160 inferences so far. Searching to depth 11. 3.5 secs*) lemma HEN009_5: "EQU001_0_ax equal & (\Y X. equal(Divide(Divide(X::'a,Y),X),Zero)) & (\X Y Z. equal(Divide(Divide(Divide(X::'a,Z),Divide(Y::'a,Z)),Divide(Divide(X::'a,Y),Z)),Zero)) & (\X. equal(Divide(Zero::'a,X),Zero)) & (\X Y. equal(Divide(X::'a,Y),Zero) & equal(Divide(Y::'a,X),Zero) --> equal(X::'a,Y)) & (\X. equal(Divide(X::'a,identity),Zero)) & (\A B C. equal(A::'a,B) --> equal(Divide(A::'a,C),Divide(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(Divide(F'::'a,D),Divide(F'::'a,E))) & (\Y X Z. equal(Divide(X::'a,Y),Zero) & equal(Divide(Y::'a,Z),Zero) --> equal(Divide(X::'a,Z),Zero)) & (\X Z Y. equal(Divide(Divide(X::'a,Y),Z),Zero) --> equal(Divide(Divide(X::'a,Z),Y),Zero)) & (\Y Z X. equal(Divide(X::'a,Y),Zero) --> equal(Divide(Divide(Z::'a,Y),Divide(Z::'a,X)),Zero)) & (~equal(Divide(identity::'a,a),Divide(identity::'a,Divide(identity::'a,Divide(identity::'a,a))))) & (equal(Divide(identity::'a,a),b)) & (equal(Divide(identity::'a,b),c)) & (equal(Divide(identity::'a,c),d)) & (~equal(b::'a,d)) --> False" by meson (*970373 inferences so far. Searching to depth 17. 890.0 secs*) lemma HEN012_3: "EQU001_0_ax equal & HEN002_0_ax identity Zero Divide equal mless_equal & HEN002_0_eq mless_equal Divide equal & (~mless_equal(a::'a,a)) --> False" by meson (*1063 inferences so far. Searching to depth 20. 1.0 secs*) lemma LCL010_1: "(\X Y. is_a_theorem(equivalent(X::'a,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) & (\X Z Y. is_a_theorem(equivalent(equivalent(X::'a,Y),equivalent(equivalent(X::'a,Z),equivalent(Z::'a,Y))))) & (~is_a_theorem(equivalent(equivalent(a::'a,b),equivalent(equivalent(c::'a,b),equivalent(a::'a,c))))) --> False" by meson (*2549 inferences so far. Searching to depth 12. 1.4 secs*) lemma LCL077_2: "(\X Y. is_a_theorem(implies(X,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) & (\Y X. is_a_theorem(implies(X,implies(Y,X)))) & (\Y X Z. is_a_theorem(implies(implies(X,implies(Y,Z)),implies(implies(X,Y),implies(X,Z))))) & (\Y X. is_a_theorem(implies(implies(not(X),not(Y)),implies(Y,X)))) & (\X2 X1 X3. is_a_theorem(implies(X1,X2)) & is_a_theorem(implies(X2,X3)) --> is_a_theorem(implies(X1,X3))) & (~is_a_theorem(implies(not(not(a)),a))) --> False" by meson (*2036 inferences so far. Searching to depth 20. 1.5 secs*) lemma LCL082_1: "(\X Y. is_a_theorem(implies(X::'a,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) & (\Y Z U X. is_a_theorem(implies(implies(implies(X::'a,Y),Z),implies(implies(Z::'a,X),implies(U::'a,X))))) & (~is_a_theorem(implies(a::'a,implies(b::'a,a)))) --> False" by meson (*1100 inferences so far. Searching to depth 13. 1.0 secs*) lemma LCL111_1: "(\X Y. is_a_theorem(implies(X,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) & (\Y X. is_a_theorem(implies(X,implies(Y,X)))) & (\Y X Z. is_a_theorem(implies(implies(X,Y),implies(implies(Y,Z),implies(X,Z))))) & (\Y X. is_a_theorem(implies(implies(implies(X,Y),Y),implies(implies(Y,X),X)))) & (\Y X. is_a_theorem(implies(implies(not(X),not(Y)),implies(Y,X)))) & (~is_a_theorem(implies(implies(a,b),implies(implies(c,a),implies(c,b))))) --> False" by meson (*667 inferences so far. Searching to depth 9. 1.4 secs*) lemma LCL143_1: "(\X. equal(X,X)) & (\Y X. equal(X,Y) --> equal(Y,X)) & (\Y X Z. equal(X,Y) & equal(Y,Z) --> equal(X,Z)) & (\X. equal(implies(true,X),X)) & (\Y X Z. equal(implies(implies(X,Y),implies(implies(Y,Z),implies(X,Z))),true)) & (\Y X. equal(implies(implies(X,Y),Y),implies(implies(Y,X),X))) & (\Y X. equal(implies(implies(not(X),not(Y)),implies(Y,X)),true)) & (\A B C. equal(A,B) --> equal(implies(A,C),implies(B,C))) & (\D F' E. equal(D,E) --> equal(implies(F',D),implies(F',E))) & (\G H. equal(G,H) --> equal(not(G),not(H))) & (\X Y. equal(big_V(X,Y),implies(implies(X,Y),Y))) & (\X Y. equal(big_hat(X,Y),not(big_V(not(X),not(Y))))) & (\X Y. ordered(X,Y) --> equal(implies(X,Y),true)) & (\X Y. equal(implies(X,Y),true) --> ordered(X,Y)) & (\A B C. equal(A,B) --> equal(big_V(A,C),big_V(B,C))) & (\D F' E. equal(D,E) --> equal(big_V(F',D),big_V(F',E))) & (\G H I'. equal(G,H) --> equal(big_hat(G,I'),big_hat(H,I'))) & (\J L K'. equal(J,K') --> equal(big_hat(L,J),big_hat(L,K'))) & (\M N O'. equal(M,N) & ordered(M,O') --> ordered(N,O')) & (\P R Q. equal(P,Q) & ordered(R,P) --> ordered(R,Q)) & (ordered(x,y)) & (~ordered(implies(z,x),implies(z,y))) --> False" by meson (*5245 inferences so far. Searching to depth 12. 4.6 secs*) lemma LCL182_1: "(\A. axiom(or(not(or(A,A)),A))) & (\B A. axiom(or(not(A),or(B,A)))) & (\B A. axiom(or(not(or(A,B)),or(B,A)))) & (\B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) & (\A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) & (\X. axiom(X) --> theorem(X)) & (\X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) & (\X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) & (~theorem(or(not(or(not(p),q)),or(not(not(q)),not(p))))) --> False" by meson (*410 inferences so far. Searching to depth 10. 0.3 secs*) lemma LCL200_1: "(\A. axiom(or(not(or(A,A)),A))) & (\B A. axiom(or(not(A),or(B,A)))) & (\B A. axiom(or(not(or(A,B)),or(B,A)))) & (\B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) & (\A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) & (\X. axiom(X) --> theorem(X)) & (\X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) & (\X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) & (~theorem(or(not(not(or(p,q))),not(q)))) --> False" by meson (*5849 inferences so far. Searching to depth 12. 5.6 secs*) lemma LCL215_1: "(\A. axiom(or(not(or(A,A)),A))) & (\B A. axiom(or(not(A),or(B,A)))) & (\B A. axiom(or(not(or(A,B)),or(B,A)))) & (\B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) & (\A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) & (\X. axiom(X) --> theorem(X)) & (\X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) & (\X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) & (~theorem(or(not(or(not(p),q)),or(not(or(p,q)),q)))) --> False" by meson (*0 secs. Not sure that a search even starts!*) lemma LCL230_2: "(q --> p | r) & (~p) & (q) & (~r) --> False" by meson (*119585 inferences so far. Searching to depth 14. 262.4 secs*) lemma LDA003_1: "EQU001_0_ax equal & (\Y X Z. equal(f(X::'a,f(Y::'a,Z)),f(f(X::'a,Y),f(X::'a,Z)))) & (\X Y. left(X::'a,f(X::'a,Y))) & (\Y X Z. left(X::'a,Y) & left(Y::'a,Z) --> left(X::'a,Z)) & (equal(num2::'a,f(num1::'a,num1))) & (equal(num3::'a,f(num2::'a,num1))) & (equal(u::'a,f(num2::'a,num2))) & (\A B C. equal(A::'a,B) --> equal(f(A::'a,C),f(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(f(F'::'a,D),f(F'::'a,E))) & (\G H I'. equal(G::'a,H) & left(G::'a,I') --> left(H::'a,I')) & (\J L K'. equal(J::'a,K') & left(L::'a,J) --> left(L::'a,K')) & (~left(num3::'a,u)) --> False" by meson (*2392 inferences so far. Searching to depth 12. 2.2 secs*) lemma MSC002_1: "(at(something::'a,here,now)) & (\Place Situation. hand_at(Place::'a,Situation) --> hand_at(Place::'a,let_go(Situation))) & (\Place Another_place Situation. hand_at(Place::'a,Situation) --> hand_at(Another_place::'a,go(Another_place::'a,Situation))) & (\Thing Situation. ~held(Thing::'a,let_go(Situation))) & (\Situation Thing. at(Thing::'a,here,Situation) --> red(Thing)) & (\Thing Place Situation. at(Thing::'a,Place,Situation) --> at(Thing::'a,Place,let_go(Situation))) & (\Thing Place Situation. at(Thing::'a,Place,Situation) --> at(Thing::'a,Place,pick_up(Situation))) & (\Thing Place Situation. at(Thing::'a,Place,Situation) --> grabbed(Thing::'a,pick_up(go(Place::'a,let_go(Situation))))) & (\Thing Situation. red(Thing) & put(Thing::'a,there,Situation) --> answer(Situation)) & (\Place Thing Another_place Situation. at(Thing::'a,Place,Situation) & grabbed(Thing::'a,Situation) --> put(Thing::'a,Another_place,go(Another_place::'a,Situation))) & (\Thing Place Another_place Situation. at(Thing::'a,Place,Situation) --> held(Thing::'a,Situation) | at(Thing::'a,Place,go(Another_place::'a,Situation))) & (\One_place Thing Place Situation. hand_at(One_place::'a,Situation) & held(Thing::'a,Situation) --> at(Thing::'a,Place,go(Place::'a,Situation))) & (\Place Thing Situation. hand_at(Place::'a,Situation) & at(Thing::'a,Place,Situation) --> held(Thing::'a,pick_up(Situation))) & (\Situation. ~answer(Situation)) --> False" by meson (*73 inferences so far. Searching to depth 10. 0.2 secs*) lemma MSC003_1: "(\Number_of_small_parts Small_part Big_part Number_of_mid_parts Mid_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) --> in'(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Mid_part) | has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) & (\Big_part Mid_part Number_of_mid_parts Number_of_small_parts Small_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) & has_parts(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Number_of_small_parts,Small_part) --> has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) & (in'(john::'a,boy)) & (\X. in'(X::'a,boy) --> in'(X::'a,human)) & (\X. in'(X::'a,hand) --> has_parts(X::'a,num5,fingers)) & (\X. in'(X::'a,human) --> has_parts(X::'a,num2,arm)) & (\X. in'(X::'a,arm) --> has_parts(X::'a,num1,hand)) & (~has_parts(john::'a,mtimes(num2::'a,num1),hand)) --> False" by meson (*1486 inferences so far. Searching to depth 20. 1.2 secs*) lemma MSC004_1: "(\Number_of_small_parts Small_part Big_part Number_of_mid_parts Mid_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) --> in'(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Mid_part) | has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) & (\Big_part Mid_part Number_of_mid_parts Number_of_small_parts Small_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) & has_parts(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Number_of_small_parts,Small_part) --> has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) & (in'(john::'a,boy)) & (\X. in'(X::'a,boy) --> in'(X::'a,human)) & (\X. in'(X::'a,hand) --> has_parts(X::'a,num5,fingers)) & (\X. in'(X::'a,human) --> has_parts(X::'a,num2,arm)) & (\X. in'(X::'a,arm) --> has_parts(X::'a,num1,hand)) & (~has_parts(john::'a,mtimes(mtimes(num2::'a,num1),num5),fingers)) --> False" by meson (*100 inferences so far. Searching to depth 12. 0.1 secs*) lemma MSC005_1: "(value(truth::'a,truth)) & (value(falsity::'a,falsity)) & (\X Y. value(X::'a,truth) & value(Y::'a,truth) --> value(xor(X::'a,Y),falsity)) & (\X Y. value(X::'a,truth) & value(Y::'a,falsity) --> value(xor(X::'a,Y),truth)) & (\X Y. value(X::'a,falsity) & value(Y::'a,truth) --> value(xor(X::'a,Y),truth)) & (\X Y. value(X::'a,falsity) & value(Y::'a,falsity) --> value(xor(X::'a,Y),falsity)) & (\Value. ~value(xor(xor(xor(xor(truth::'a,falsity),falsity),truth),falsity),Value)) --> False" by meson (*19116 inferences so far. Searching to depth 16. 15.9 secs*) lemma MSC006_1: "(\Y X Z. p(X::'a,Y) & p(Y::'a,Z) --> p(X::'a,Z)) & (\Y X Z. q(X::'a,Y) & q(Y::'a,Z) --> q(X::'a,Z)) & (\Y X. q(X::'a,Y) --> q(Y::'a,X)) & (\X Y. p(X::'a,Y) | q(X::'a,Y)) & (~p(a::'a,b)) & (~q(c::'a,d)) --> False" by meson (*1713 inferences so far. Searching to depth 10. 2.8 secs*) lemma NUM001_1: "(\A. equal(A::'a,A)) & (\B A C. equal(A::'a,B) & equal(B::'a,C) --> equal(A::'a,C)) & (\B A. equal(add(A::'a,B),add(B::'a,A))) & (\A B C. equal(add(A::'a,add(B::'a,C)),add(add(A::'a,B),C))) & (\B A. equal(subtract(add(A::'a,B),B),A)) & (\A B. equal(A::'a,subtract(add(A::'a,B),B))) & (\A C B. equal(add(subtract(A::'a,B),C),subtract(add(A::'a,C),B))) & (\A C B. equal(subtract(add(A::'a,B),C),add(subtract(A::'a,C),B))) & (\A C B D. equal(A::'a,B) & equal(C::'a,add(A::'a,D)) --> equal(C::'a,add(B::'a,D))) & (\A C D B. equal(A::'a,B) & equal(C::'a,add(D::'a,A)) --> equal(C::'a,add(D::'a,B))) & (\A C B D. equal(A::'a,B) & equal(C::'a,subtract(A::'a,D)) --> equal(C::'a,subtract(B::'a,D))) & (\A C D B. equal(A::'a,B) & equal(C::'a,subtract(D::'a,A)) --> equal(C::'a,subtract(D::'a,B))) & (~equal(add(add(a::'a,b),c),add(a::'a,add(b::'a,c)))) --> False" by meson abbreviation "NUM001_0_ax multiply successor num0 add equal \ (\A. equal(add(A::'a,num0),A)) & (\A B. equal(add(A::'a,successor(B)),successor(add(A::'a,B)))) & (\A. equal(multiply(A::'a,num0),num0)) & (\B A. equal(multiply(A::'a,successor(B)),add(multiply(A::'a,B),A))) & (\A B. equal(successor(A),successor(B)) --> equal(A::'a,B)) & (\A B. equal(A::'a,B) --> equal(successor(A),successor(B)))" abbreviation "NUM001_1_ax predecessor_of_1st_minus_2nd successor add equal mless \ (\A C B. mless(A::'a,B) & mless(C::'a,A) --> mless(C::'a,B)) & (\A B C. equal(add(successor(A),B),C) --> mless(B::'a,C)) & (\A B. mless(A::'a,B) --> equal(add(successor(predecessor_of_1st_minus_2nd(B::'a,A)),A),B))" abbreviation "NUM001_2_ax equal mless divides \ (\A B. divides(A::'a,B) --> mless(A::'a,B) | equal(A::'a,B)) & (\A B. mless(A::'a,B) --> divides(A::'a,B)) & (\A B. equal(A::'a,B) --> divides(A::'a,B))" (*20717 inferences so far. Searching to depth 11. 13.7 secs*) lemma NUM021_1: "EQU001_0_ax equal & NUM001_0_ax multiply successor num0 add equal & NUM001_1_ax predecessor_of_1st_minus_2nd successor add equal mless & NUM001_2_ax equal mless divides & (mless(b::'a,c)) & (~mless(b::'a,a)) & (divides(c::'a,a)) & (\A. ~equal(successor(A),num0)) --> False" by meson (*26320 inferences so far. Searching to depth 10. 26.4 secs*) lemma NUM024_1: "EQU001_0_ax equal & NUM001_0_ax multiply successor num0 add equal & NUM001_1_ax predecessor_of_1st_minus_2nd successor add equal mless & (\B A. equal(add(A::'a,B),add(B::'a,A))) & (\B A C. equal(add(A::'a,B),add(C::'a,B)) --> equal(A::'a,C)) & (mless(a::'a,a)) & (\A. ~equal(successor(A),num0)) --> False" by meson abbreviation "SET004_0_ax not_homomorphism2 not_homomorphism1 homomorphism compatible operation cantor diagonalise subset_relation one_to_one choice apply regular function identity_relation single_valued_class compos powerClass sum_class omega inductive successor_relation successor image' rng domain range_of INVERSE flip rot domain_of null_class restrct difference union complement intersection element_relation second first cross_product ordered_pair singleton unordered_pair equal universal_class not_subclass_element member subclass \ (\X U Y. subclass(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) & (\X Y. member(not_subclass_element(X::'a,Y),X) | subclass(X::'a,Y)) & (\X Y. member(not_subclass_element(X::'a,Y),Y) --> subclass(X::'a,Y)) & (\X. subclass(X::'a,universal_class)) & (\X Y. equal(X::'a,Y) --> subclass(X::'a,Y)) & (\Y X. equal(X::'a,Y) --> subclass(Y::'a,X)) & (\X Y. subclass(X::'a,Y) & subclass(Y::'a,X) --> equal(X::'a,Y)) & (\X U Y. member(U::'a,unordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) & (\X Y. member(X::'a,universal_class) --> member(X::'a,unordered_pair(X::'a,Y))) & (\X Y. member(Y::'a,universal_class) --> member(Y::'a,unordered_pair(X::'a,Y))) & (\X Y. member(unordered_pair(X::'a,Y),universal_class)) & (\X. equal(unordered_pair(X::'a,X),singleton(X))) & (\X Y. equal(unordered_pair(singleton(X),unordered_pair(X::'a,singleton(Y))),ordered_pair(X::'a,Y))) & (\V Y U X. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(U::'a,X)) & (\U X V Y. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(V::'a,Y)) & (\U V X Y. member(U::'a,X) & member(V::'a,Y) --> member(ordered_pair(U::'a,V),cross_product(X::'a,Y))) & (\X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> equal(ordered_pair(first(Z),second(Z)),Z)) & (subclass(element_relation::'a,cross_product(universal_class::'a,universal_class))) & (\X Y. member(ordered_pair(X::'a,Y),element_relation) --> member(X::'a,Y)) & (\X Y. member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) & member(X::'a,Y) --> member(ordered_pair(X::'a,Y),element_relation)) & (\Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) & (\X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) & (\Z X Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) & (\Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) & (\Z X. member(Z::'a,universal_class) --> member(Z::'a,complement(X)) | member(Z::'a,X)) & (\X Y. equal(complement(intersection(complement(X),complement(Y))),union(X::'a,Y))) & (\X Y. equal(intersection(complement(intersection(X::'a,Y)),complement(intersection(complement(X),complement(Y)))),difference(X::'a,Y))) & (\Xr X Y. equal(intersection(Xr::'a,cross_product(X::'a,Y)),restrct(Xr::'a,X,Y))) & (\Xr X Y. equal(intersection(cross_product(X::'a,Y),Xr),restrct(Xr::'a,X,Y))) & (\Z X. ~(equal(restrct(X::'a,singleton(Z),universal_class),null_class) & member(Z::'a,domain_of(X)))) & (\Z X. member(Z::'a,universal_class) --> equal(restrct(X::'a,singleton(Z),universal_class),null_class) | member(Z::'a,domain_of(X))) & (\X. subclass(rot(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) & (\V W U X. member(ordered_pair(ordered_pair(U::'a,V),W),rot(X)) --> member(ordered_pair(ordered_pair(V::'a,W),U),X)) & (\U V W X. member(ordered_pair(ordered_pair(V::'a,W),U),X) & member(ordered_pair(ordered_pair(U::'a,V),W),cross_product(cross_product(universal_class::'a,universal_class),universal_class)) --> member(ordered_pair(ordered_pair(U::'a,V),W),rot(X))) & (\X. subclass(flip(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) & (\V U W X. member(ordered_pair(ordered_pair(U::'a,V),W),flip(X)) --> member(ordered_pair(ordered_pair(V::'a,U),W),X)) & (\U V W X. member(ordered_pair(ordered_pair(V::'a,U),W),X) & member(ordered_pair(ordered_pair(U::'a,V),W),cross_product(cross_product(universal_class::'a,universal_class),universal_class)) --> member(ordered_pair(ordered_pair(U::'a,V),W),flip(X))) & (\Y. equal(domain_of(flip(cross_product(Y::'a,universal_class))),INVERSE(Y))) & (\Z. equal(domain_of(INVERSE(Z)),range_of(Z))) & (\Z X Y. equal(first(not_subclass_element(restrct(Z::'a,X,singleton(Y)),null_class)),domain(Z::'a,X,Y))) & (\Z X Y. equal(second(not_subclass_element(restrct(Z::'a,singleton(X),Y),null_class)),rng(Z::'a,X,Y))) & (\Xr X. equal(range_of(restrct(Xr::'a,X,universal_class)),image'(Xr::'a,X))) & (\X. equal(union(X::'a,singleton(X)),successor(X))) & (subclass(successor_relation::'a,cross_product(universal_class::'a,universal_class))) & (\X Y. member(ordered_pair(X::'a,Y),successor_relation) --> equal(successor(X),Y)) & (\X Y. equal(successor(X),Y) & member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) --> member(ordered_pair(X::'a,Y),successor_relation)) & (\X. inductive(X) --> member(null_class::'a,X)) & (\X. inductive(X) --> subclass(image'(successor_relation::'a,X),X)) & (\X. member(null_class::'a,X) & subclass(image'(successor_relation::'a,X),X) --> inductive(X)) & (inductive(omega)) & (\Y. inductive(Y) --> subclass(omega::'a,Y)) & (member(omega::'a,universal_class)) & (\X. equal(domain_of(restrct(element_relation::'a,universal_class,X)),sum_class(X))) & (\X. member(X::'a,universal_class) --> member(sum_class(X),universal_class)) & (\X. equal(complement(image'(element_relation::'a,complement(X))),powerClass(X))) & (\U. member(U::'a,universal_class) --> member(powerClass(U),universal_class)) & (\Yr Xr. subclass(compos(Yr::'a,Xr),cross_product(universal_class::'a,universal_class))) & (\Z Yr Xr Y. member(ordered_pair(Y::'a,Z),compos(Yr::'a,Xr)) --> member(Z::'a,image'(Yr::'a,image'(Xr::'a,singleton(Y))))) & (\Y Z Yr Xr. member(Z::'a,image'(Yr::'a,image'(Xr::'a,singleton(Y)))) & member(ordered_pair(Y::'a,Z),cross_product(universal_class::'a,universal_class)) --> member(ordered_pair(Y::'a,Z),compos(Yr::'a,Xr))) & (\X. single_valued_class(X) --> subclass(compos(X::'a,INVERSE(X)),identity_relation)) & (\X. subclass(compos(X::'a,INVERSE(X)),identity_relation) --> single_valued_class(X)) & (\Xf. function(Xf) --> subclass(Xf::'a,cross_product(universal_class::'a,universal_class))) & (\Xf. function(Xf) --> subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation)) & (\Xf. subclass(Xf::'a,cross_product(universal_class::'a,universal_class)) & subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation) --> function(Xf)) & (\Xf X. function(Xf) & member(X::'a,universal_class) --> member(image'(Xf::'a,X),universal_class)) & (\X. equal(X::'a,null_class) | member(regular(X),X)) & (\X. equal(X::'a,null_class) | equal(intersection(X::'a,regular(X)),null_class)) & (\Xf Y. equal(sum_class(image'(Xf::'a,singleton(Y))),apply(Xf::'a,Y))) & (function(choice)) & (\Y. member(Y::'a,universal_class) --> equal(Y::'a,null_class) | member(apply(choice::'a,Y),Y)) & (\Xf. one_to_one(Xf) --> function(Xf)) & (\Xf. one_to_one(Xf) --> function(INVERSE(Xf))) & (\Xf. function(INVERSE(Xf)) & function(Xf) --> one_to_one(Xf)) & (equal(intersection(cross_product(universal_class::'a,universal_class),intersection(cross_product(universal_class::'a,universal_class),complement(compos(complement(element_relation),INVERSE(element_relation))))),subset_relation)) & (equal(intersection(INVERSE(subset_relation),subset_relation),identity_relation)) & (\Xr. equal(complement(domain_of(intersection(Xr::'a,identity_relation))),diagonalise(Xr))) & (\X. equal(intersection(domain_of(X),diagonalise(compos(INVERSE(element_relation),X))),cantor(X))) & (\Xf. operation(Xf) --> function(Xf)) & (\Xf. operation(Xf) --> equal(cross_product(domain_of(domain_of(Xf)),domain_of(domain_of(Xf))),domain_of(Xf))) & (\Xf. operation(Xf) --> subclass(range_of(Xf),domain_of(domain_of(Xf)))) & (\Xf. function(Xf) & equal(cross_product(domain_of(domain_of(Xf)),domain_of(domain_of(Xf))),domain_of(Xf)) & subclass(range_of(Xf),domain_of(domain_of(Xf))) --> operation(Xf)) & (\Xf1 Xf2 Xh. compatible(Xh::'a,Xf1,Xf2) --> function(Xh)) & (\Xf2 Xf1 Xh. compatible(Xh::'a,Xf1,Xf2) --> equal(domain_of(domain_of(Xf1)),domain_of(Xh))) & (\Xf1 Xh Xf2. compatible(Xh::'a,Xf1,Xf2) --> subclass(range_of(Xh),domain_of(domain_of(Xf2)))) & (\Xh Xh1 Xf1 Xf2. function(Xh) & equal(domain_of(domain_of(Xf1)),domain_of(Xh)) & subclass(range_of(Xh),domain_of(domain_of(Xf2))) --> compatible(Xh1::'a,Xf1,Xf2)) & (\Xh Xf2 Xf1. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf1)) & (\Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf2)) & (\Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> compatible(Xh::'a,Xf1,Xf2)) & (\Xf2 Xh Xf1 X Y. homomorphism(Xh::'a,Xf1,Xf2) & member(ordered_pair(X::'a,Y),domain_of(Xf1)) --> equal(apply(Xf2::'a,ordered_pair(apply(Xh::'a,X),apply(Xh::'a,Y))),apply(Xh::'a,apply(Xf1::'a,ordered_pair(X::'a,Y))))) & (\Xh Xf1 Xf2. operation(Xf1) & operation(Xf2) & compatible(Xh::'a,Xf1,Xf2) --> member(ordered_pair(not_homomorphism1(Xh::'a,Xf1,Xf2),not_homomorphism2(Xh::'a,Xf1,Xf2)),domain_of(Xf1)) | homomorphism(Xh::'a,Xf1,Xf2)) & (\Xh Xf1 Xf2. operation(Xf1) & operation(Xf2) & compatible(Xh::'a,Xf1,Xf2) & equal(apply(Xf2::'a,ordered_pair(apply(Xh::'a,not_homomorphism1(Xh::'a,Xf1,Xf2)),apply(Xh::'a,not_homomorphism2(Xh::'a,Xf1,Xf2)))),apply(Xh::'a,apply(Xf1::'a,ordered_pair(not_homomorphism1(Xh::'a,Xf1,Xf2),not_homomorphism2(Xh::'a,Xf1,Xf2))))) --> homomorphism(Xh::'a,Xf1,Xf2))" abbreviation "SET004_0_eq subclass single_valued_class operation one_to_one member inductive homomorphism function compatible unordered_pair union sum_class successor singleton second rot restrct regular range_of rng powerClass ordered_pair not_subclass_element not_homomorphism2 not_homomorphism1 INVERSE intersection image' flip first domain_of domain difference diagonalise cross_product compos complement cantor apply equal \ (\D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) & (\G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) & (\J K'. equal(J::'a,K') --> equal(cantor(J),cantor(K'))) & (\L M. equal(L::'a,M) --> equal(complement(L),complement(M))) & (\N O' P. equal(N::'a,O') --> equal(compos(N::'a,P),compos(O'::'a,P))) & (\Q S' R. equal(Q::'a,R) --> equal(compos(S'::'a,Q),compos(S'::'a,R))) & (\T' U V. equal(T'::'a,U) --> equal(cross_product(T'::'a,V),cross_product(U::'a,V))) & (\W Y X. equal(W::'a,X) --> equal(cross_product(Y::'a,W),cross_product(Y::'a,X))) & (\Z A1. equal(Z::'a,A1) --> equal(diagonalise(Z),diagonalise(A1))) & (\B1 C1 D1. equal(B1::'a,C1) --> equal(difference(B1::'a,D1),difference(C1::'a,D1))) & (\E1 G1 F1. equal(E1::'a,F1) --> equal(difference(G1::'a,E1),difference(G1::'a,F1))) & (\H1 I1 J1 K1. equal(H1::'a,I1) --> equal(domain(H1::'a,J1,K1),domain(I1::'a,J1,K1))) & (\L1 N1 M1 O1. equal(L1::'a,M1) --> equal(domain(N1::'a,L1,O1),domain(N1::'a,M1,O1))) & (\P1 R1 S1 Q1. equal(P1::'a,Q1) --> equal(domain(R1::'a,S1,P1),domain(R1::'a,S1,Q1))) & (\T1 U1. equal(T1::'a,U1) --> equal(domain_of(T1),domain_of(U1))) & (\V1 W1. equal(V1::'a,W1) --> equal(first(V1),first(W1))) & (\X1 Y1. equal(X1::'a,Y1) --> equal(flip(X1),flip(Y1))) & (\Z1 A2 B2. equal(Z1::'a,A2) --> equal(image'(Z1::'a,B2),image'(A2::'a,B2))) & (\C2 E2 D2. equal(C2::'a,D2) --> equal(image'(E2::'a,C2),image'(E2::'a,D2))) & (\F2 G2 H2. equal(F2::'a,G2) --> equal(intersection(F2::'a,H2),intersection(G2::'a,H2))) & (\I2 K2 J2. equal(I2::'a,J2) --> equal(intersection(K2::'a,I2),intersection(K2::'a,J2))) & (\L2 M2. equal(L2::'a,M2) --> equal(INVERSE(L2),INVERSE(M2))) & (\N2 O2 P2 Q2. equal(N2::'a,O2) --> equal(not_homomorphism1(N2::'a,P2,Q2),not_homomorphism1(O2::'a,P2,Q2))) & (\R2 T2 S2 U2. equal(R2::'a,S2) --> equal(not_homomorphism1(T2::'a,R2,U2),not_homomorphism1(T2::'a,S2,U2))) & (\V2 X2 Y2 W2. equal(V2::'a,W2) --> equal(not_homomorphism1(X2::'a,Y2,V2),not_homomorphism1(X2::'a,Y2,W2))) & (\Z2 A3 B3 C3. equal(Z2::'a,A3) --> equal(not_homomorphism2(Z2::'a,B3,C3),not_homomorphism2(A3::'a,B3,C3))) & (\D3 F3 E3 G3. equal(D3::'a,E3) --> equal(not_homomorphism2(F3::'a,D3,G3),not_homomorphism2(F3::'a,E3,G3))) & (\H3 J3 K3 I3. equal(H3::'a,I3) --> equal(not_homomorphism2(J3::'a,K3,H3),not_homomorphism2(J3::'a,K3,I3))) & (\L3 M3 N3. equal(L3::'a,M3) --> equal(not_subclass_element(L3::'a,N3),not_subclass_element(M3::'a,N3))) & (\O3 Q3 P3. equal(O3::'a,P3) --> equal(not_subclass_element(Q3::'a,O3),not_subclass_element(Q3::'a,P3))) & (\R3 S3 T3. equal(R3::'a,S3) --> equal(ordered_pair(R3::'a,T3),ordered_pair(S3::'a,T3))) & (\U3 W3 V3. equal(U3::'a,V3) --> equal(ordered_pair(W3::'a,U3),ordered_pair(W3::'a,V3))) & (\X3 Y3. equal(X3::'a,Y3) --> equal(powerClass(X3),powerClass(Y3))) & (\Z3 A4 B4 C4. equal(Z3::'a,A4) --> equal(rng(Z3::'a,B4,C4),rng(A4::'a,B4,C4))) & (\D4 F4 E4 G4. equal(D4::'a,E4) --> equal(rng(F4::'a,D4,G4),rng(F4::'a,E4,G4))) & (\H4 J4 K4 I4. equal(H4::'a,I4) --> equal(rng(J4::'a,K4,H4),rng(J4::'a,K4,I4))) & (\L4 M4. equal(L4::'a,M4) --> equal(range_of(L4),range_of(M4))) & (\N4 O4. equal(N4::'a,O4) --> equal(regular(N4),regular(O4))) & (\P4 Q4 R4 S4. equal(P4::'a,Q4) --> equal(restrct(P4::'a,R4,S4),restrct(Q4::'a,R4,S4))) & (\T4 V4 U4 W4. equal(T4::'a,U4) --> equal(restrct(V4::'a,T4,W4),restrct(V4::'a,U4,W4))) & (\X4 Z4 A5 Y4. equal(X4::'a,Y4) --> equal(restrct(Z4::'a,A5,X4),restrct(Z4::'a,A5,Y4))) & (\B5 C5. equal(B5::'a,C5) --> equal(rot(B5),rot(C5))) & (\D5 E5. equal(D5::'a,E5) --> equal(second(D5),second(E5))) & (\F5 G5. equal(F5::'a,G5) --> equal(singleton(F5),singleton(G5))) & (\H5 I5. equal(H5::'a,I5) --> equal(successor(H5),successor(I5))) & (\J5 K5. equal(J5::'a,K5) --> equal(sum_class(J5),sum_class(K5))) & (\L5 M5 N5. equal(L5::'a,M5) --> equal(union(L5::'a,N5),union(M5::'a,N5))) & (\O5 Q5 P5. equal(O5::'a,P5) --> equal(union(Q5::'a,O5),union(Q5::'a,P5))) & (\R5 S5 T5. equal(R5::'a,S5) --> equal(unordered_pair(R5::'a,T5),unordered_pair(S5::'a,T5))) & (\U5 W5 V5. equal(U5::'a,V5) --> equal(unordered_pair(W5::'a,U5),unordered_pair(W5::'a,V5))) & (\X5 Y5 Z5 A6. equal(X5::'a,Y5) & compatible(X5::'a,Z5,A6) --> compatible(Y5::'a,Z5,A6)) & (\B6 D6 C6 E6. equal(B6::'a,C6) & compatible(D6::'a,B6,E6) --> compatible(D6::'a,C6,E6)) & (\F6 H6 I6 G6. equal(F6::'a,G6) & compatible(H6::'a,I6,F6) --> compatible(H6::'a,I6,G6)) & (\J6 K6. equal(J6::'a,K6) & function(J6) --> function(K6)) & (\L6 M6 N6 O6. equal(L6::'a,M6) & homomorphism(L6::'a,N6,O6) --> homomorphism(M6::'a,N6,O6)) & (\P6 R6 Q6 S6. equal(P6::'a,Q6) & homomorphism(R6::'a,P6,S6) --> homomorphism(R6::'a,Q6,S6)) & (\T6 V6 W6 U6. equal(T6::'a,U6) & homomorphism(V6::'a,W6,T6) --> homomorphism(V6::'a,W6,U6)) & (\X6 Y6. equal(X6::'a,Y6) & inductive(X6) --> inductive(Y6)) & (\Z6 A7 B7. equal(Z6::'a,A7) & member(Z6::'a,B7) --> member(A7::'a,B7)) & (\C7 E7 D7. equal(C7::'a,D7) & member(E7::'a,C7) --> member(E7::'a,D7)) & (\F7 G7. equal(F7::'a,G7) & one_to_one(F7) --> one_to_one(G7)) & (\H7 I7. equal(H7::'a,I7) & operation(H7) --> operation(I7)) & (\J7 K7. equal(J7::'a,K7) & single_valued_class(J7) --> single_valued_class(K7)) & (\L7 M7 N7. equal(L7::'a,M7) & subclass(L7::'a,N7) --> subclass(M7::'a,N7)) & (\O7 Q7 P7. equal(O7::'a,P7) & subclass(Q7::'a,O7) --> subclass(Q7::'a,P7))" abbreviation "SET004_1_ax range_of function maps apply application_function singleton_relation element_relation complement intersection single_valued3 singleton image' domain single_valued2 second single_valued1 identity_relation INVERSE not_subclass_element first domain_of domain_relation composition_function compos equal ordered_pair member universal_class cross_product compose_class subclass \ (\X. subclass(compose_class(X),cross_product(universal_class::'a,universal_class))) & (\X Y Z. member(ordered_pair(Y::'a,Z),compose_class(X)) --> equal(compos(X::'a,Y),Z)) & (\Y Z X. member(ordered_pair(Y::'a,Z),cross_product(universal_class::'a,universal_class)) & equal(compos(X::'a,Y),Z) --> member(ordered_pair(Y::'a,Z),compose_class(X))) & (subclass(composition_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) & (\X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),composition_function) --> equal(compos(X::'a,Y),Z)) & (\X Y. member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) --> member(ordered_pair(X::'a,ordered_pair(Y::'a,compos(X::'a,Y))),composition_function)) & (subclass(domain_relation::'a,cross_product(universal_class::'a,universal_class))) & (\X Y. member(ordered_pair(X::'a,Y),domain_relation) --> equal(domain_of(X),Y)) & (\X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,domain_of(X)),domain_relation)) & (\X. equal(first(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued1(X))) & (\X. equal(second(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued2(X))) & (\X. equal(domain(X::'a,image'(INVERSE(X),singleton(single_valued1(X))),single_valued2(X)),single_valued3(X))) & (equal(intersection(complement(compos(element_relation::'a,complement(identity_relation))),element_relation),singleton_relation)) & (subclass(application_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) & (\Z Y X. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> member(Y::'a,domain_of(X))) & (\X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> equal(apply(X::'a,Y),Z)) & (\Z X Y. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class))) & member(Y::'a,domain_of(X)) --> member(ordered_pair(X::'a,ordered_pair(Y::'a,apply(X::'a,Y))),application_function)) & (\X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) & (\Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) & (\X Xf Y. maps(Xf::'a,X,Y) --> subclass(range_of(Xf),Y)) & (\Xf Y. function(Xf) & subclass(range_of(Xf),Y) --> maps(Xf::'a,domain_of(Xf),Y))" abbreviation "SET004_1_eq maps single_valued3 single_valued2 single_valued1 compose_class equal \ (\L M. equal(L::'a,M) --> equal(compose_class(L),compose_class(M))) & (\N2 O2. equal(N2::'a,O2) --> equal(single_valued1(N2),single_valued1(O2))) & (\P2 Q2. equal(P2::'a,Q2) --> equal(single_valued2(P2),single_valued2(Q2))) & (\R2 S2. equal(R2::'a,S2) --> equal(single_valued3(R2),single_valued3(S2))) & (\X2 Y2 Z2 A3. equal(X2::'a,Y2) & maps(X2::'a,Z2,A3) --> maps(Y2::'a,Z2,A3)) & (\B3 D3 C3 E3. equal(B3::'a,C3) & maps(D3::'a,B3,E3) --> maps(D3::'a,C3,E3)) & (\F3 H3 I3 G3. equal(F3::'a,G3) & maps(H3::'a,I3,F3) --> maps(H3::'a,I3,G3))" abbreviation "NUM004_0_ax integer_of omega ordinal_multiply add_relation ordinal_add recursion apply range_of union_of_range_map function recursion_equation_functions rest_relation rest_of limit_ordinals kind_1_ordinals successor_relation image' universal_class sum_class element_relation ordinal_numbers section not_well_ordering ordered_pair least member well_ordering singleton domain_of segment null_class intersection asymmetric compos transitive cross_product connected identity_relation complement restrct subclass irreflexive symmetrization_of INVERSE union equal \ (\X. equal(union(X::'a,INVERSE(X)),symmetrization_of(X))) & (\X Y. irreflexive(X::'a,Y) --> subclass(restrct(X::'a,Y,Y),complement(identity_relation))) & (\X Y. subclass(restrct(X::'a,Y,Y),complement(identity_relation)) --> irreflexive(X::'a,Y)) & (\Y X. connected(X::'a,Y) --> subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X)))) & (\X Y. subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X))) --> connected(X::'a,Y)) & (\Xr Y. transitive(Xr::'a,Y) --> subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y))) & (\Xr Y. subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y)) --> transitive(Xr::'a,Y)) & (\Xr Y. asymmetric(Xr::'a,Y) --> equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class)) & (\Xr Y. equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class) --> asymmetric(Xr::'a,Y)) & (\Xr Y Z. equal(segment(Xr::'a,Y,Z),domain_of(restrct(Xr::'a,Y,singleton(Z))))) & (\X Y. well_ordering(X::'a,Y) --> connected(X::'a,Y)) & (\Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(U::'a,null_class) | member(least(Xr::'a,U),U)) & (\Y V Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) & member(V::'a,U) --> member(least(Xr::'a,U),U)) & (\Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(segment(Xr::'a,U,least(Xr::'a,U)),null_class)) & (\Y V U Xr. ~(well_ordering(Xr::'a,Y) & subclass(U::'a,Y) & member(V::'a,U) & member(ordered_pair(V::'a,least(Xr::'a,U)),Xr))) & (\Xr Y. connected(Xr::'a,Y) & equal(not_well_ordering(Xr::'a,Y),null_class) --> well_ordering(Xr::'a,Y)) & (\Xr Y. connected(Xr::'a,Y) --> subclass(not_well_ordering(Xr::'a,Y),Y) | well_ordering(Xr::'a,Y)) & (\V Xr Y. member(V::'a,not_well_ordering(Xr::'a,Y)) & equal(segment(Xr::'a,not_well_ordering(Xr::'a,Y),V),null_class) & connected(Xr::'a,Y) --> well_ordering(Xr::'a,Y)) & (\Xr Y Z. section(Xr::'a,Y,Z) --> subclass(Y::'a,Z)) & (\Xr Z Y. section(Xr::'a,Y,Z) --> subclass(domain_of(restrct(Xr::'a,Z,Y)),Y)) & (\Xr Y Z. subclass(Y::'a,Z) & subclass(domain_of(restrct(Xr::'a,Z,Y)),Y) --> section(Xr::'a,Y,Z)) & (\X. member(X::'a,ordinal_numbers) --> well_ordering(element_relation::'a,X)) & (\X. member(X::'a,ordinal_numbers) --> subclass(sum_class(X),X)) & (\X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) & member(X::'a,universal_class) --> member(X::'a,ordinal_numbers)) & (\X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) --> member(X::'a,ordinal_numbers) | equal(X::'a,ordinal_numbers)) & (equal(union(singleton(null_class),image'(successor_relation::'a,ordinal_numbers)),kind_1_ordinals)) & (equal(intersection(complement(kind_1_ordinals),ordinal_numbers),limit_ordinals)) & (\X. subclass(rest_of(X),cross_product(universal_class::'a,universal_class))) & (\V U X. member(ordered_pair(U::'a,V),rest_of(X)) --> member(U::'a,domain_of(X))) & (\X U V. member(ordered_pair(U::'a,V),rest_of(X)) --> equal(restrct(X::'a,U,universal_class),V)) & (\U V X. member(U::'a,domain_of(X)) & equal(restrct(X::'a,U,universal_class),V) --> member(ordered_pair(U::'a,V),rest_of(X))) & (subclass(rest_relation::'a,cross_product(universal_class::'a,universal_class))) & (\X Y. member(ordered_pair(X::'a,Y),rest_relation) --> equal(rest_of(X),Y)) & (\X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,rest_of(X)),rest_relation)) & (\X Z. member(X::'a,recursion_equation_functions(Z)) --> function(Z)) & (\Z X. member(X::'a,recursion_equation_functions(Z)) --> function(X)) & (\Z X. member(X::'a,recursion_equation_functions(Z)) --> member(domain_of(X),ordinal_numbers)) & (\Z X. member(X::'a,recursion_equation_functions(Z)) --> equal(compos(Z::'a,rest_of(X)),X)) & (\X Z. function(Z) & function(X) & member(domain_of(X),ordinal_numbers) & equal(compos(Z::'a,rest_of(X)),X) --> member(X::'a,recursion_equation_functions(Z))) & (subclass(union_of_range_map::'a,cross_product(universal_class::'a,universal_class))) & (\X Y. member(ordered_pair(X::'a,Y),union_of_range_map) --> equal(sum_class(range_of(X)),Y)) & (\X Y. member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) & equal(sum_class(range_of(X)),Y) --> member(ordered_pair(X::'a,Y),union_of_range_map)) & (\X Y. equal(apply(recursion(X::'a,successor_relation,union_of_range_map),Y),ordinal_add(X::'a,Y))) & (\X Y. equal(recursion(null_class::'a,apply(add_relation::'a,X),union_of_range_map),ordinal_multiply(X::'a,Y))) & (\X. member(X::'a,omega) --> equal(integer_of(X),X)) & (\X. member(X::'a,omega) | equal(integer_of(X),null_class))" abbreviation "NUM004_0_eq well_ordering transitive section irreflexive connected asymmetric symmetrization_of segment rest_of recursion_equation_functions recursion ordinal_multiply ordinal_add not_well_ordering least integer_of equal \ (\D E. equal(D::'a,E) --> equal(integer_of(D),integer_of(E))) & (\F' G H. equal(F'::'a,G) --> equal(least(F'::'a,H),least(G::'a,H))) & (\I' K' J. equal(I'::'a,J) --> equal(least(K'::'a,I'),least(K'::'a,J))) & (\L M N. equal(L::'a,M) --> equal(not_well_ordering(L::'a,N),not_well_ordering(M::'a,N))) & (\O' Q P. equal(O'::'a,P) --> equal(not_well_ordering(Q::'a,O'),not_well_ordering(Q::'a,P))) & (\R S' T'. equal(R::'a,S') --> equal(ordinal_add(R::'a,T'),ordinal_add(S'::'a,T'))) & (\U W V. equal(U::'a,V) --> equal(ordinal_add(W::'a,U),ordinal_add(W::'a,V))) & (\X Y Z. equal(X::'a,Y) --> equal(ordinal_multiply(X::'a,Z),ordinal_multiply(Y::'a,Z))) & (\A1 C1 B1. equal(A1::'a,B1) --> equal(ordinal_multiply(C1::'a,A1),ordinal_multiply(C1::'a,B1))) & (\F1 G1 H1 I1. equal(F1::'a,G1) --> equal(recursion(F1::'a,H1,I1),recursion(G1::'a,H1,I1))) & (\J1 L1 K1 M1. equal(J1::'a,K1) --> equal(recursion(L1::'a,J1,M1),recursion(L1::'a,K1,M1))) & (\N1 P1 Q1 O1. equal(N1::'a,O1) --> equal(recursion(P1::'a,Q1,N1),recursion(P1::'a,Q1,O1))) & (\R1 S1. equal(R1::'a,S1) --> equal(recursion_equation_functions(R1),recursion_equation_functions(S1))) & (\T1 U1. equal(T1::'a,U1) --> equal(rest_of(T1),rest_of(U1))) & (\V1 W1 X1 Y1. equal(V1::'a,W1) --> equal(segment(V1::'a,X1,Y1),segment(W1::'a,X1,Y1))) & (\Z1 B2 A2 C2. equal(Z1::'a,A2) --> equal(segment(B2::'a,Z1,C2),segment(B2::'a,A2,C2))) & (\D2 F2 G2 E2. equal(D2::'a,E2) --> equal(segment(F2::'a,G2,D2),segment(F2::'a,G2,E2))) & (\H2 I2. equal(H2::'a,I2) --> equal(symmetrization_of(H2),symmetrization_of(I2))) & (\J2 K2 L2. equal(J2::'a,K2) & asymmetric(J2::'a,L2) --> asymmetric(K2::'a,L2)) & (\M2 O2 N2. equal(M2::'a,N2) & asymmetric(O2::'a,M2) --> asymmetric(O2::'a,N2)) & (\P2 Q2 R2. equal(P2::'a,Q2) & connected(P2::'a,R2) --> connected(Q2::'a,R2)) & (\S2 U2 T2. equal(S2::'a,T2) & connected(U2::'a,S2) --> connected(U2::'a,T2)) & (\V2 W2 X2. equal(V2::'a,W2) & irreflexive(V2::'a,X2) --> irreflexive(W2::'a,X2)) & (\Y2 A3 Z2. equal(Y2::'a,Z2) & irreflexive(A3::'a,Y2) --> irreflexive(A3::'a,Z2)) & (\B3 C3 D3 E3. equal(B3::'a,C3) & section(B3::'a,D3,E3) --> section(C3::'a,D3,E3)) & (\F3 H3 G3 I3. equal(F3::'a,G3) & section(H3::'a,F3,I3) --> section(H3::'a,G3,I3)) & (\J3 L3 M3 K3. equal(J3::'a,K3) & section(L3::'a,M3,J3) --> section(L3::'a,M3,K3)) & (\N3 O3 P3. equal(N3::'a,O3) & transitive(N3::'a,P3) --> transitive(O3::'a,P3)) & (\Q3 S3 R3. equal(Q3::'a,R3) & transitive(S3::'a,Q3) --> transitive(S3::'a,R3)) & (\T3 U3 V3. equal(T3::'a,U3) & well_ordering(T3::'a,V3) --> well_ordering(U3::'a,V3)) & (\W3 Y3 X3. equal(W3::'a,X3) & well_ordering(Y3::'a,W3) --> well_ordering(Y3::'a,X3))" (*1345 inferences so far. Searching to depth 7. 23.3 secs. BIG*) lemma NUM180_1: "EQU001_0_ax equal & SET004_0_ax not_homomorphism2 not_homomorphism1 homomorphism compatible operation cantor diagonalise subset_relation one_to_one choice apply regular function identity_relation single_valued_class compos powerClass sum_class omega inductive successor_relation successor image' rng domain range_of INVERSE flip rot domain_of null_class restrct difference union complement intersection element_relation second first cross_product ordered_pair singleton unordered_pair equal universal_class not_subclass_element member subclass & SET004_0_eq subclass single_valued_class operation one_to_one member inductive homomorphism function compatible unordered_pair union sum_class successor singleton second rot restrct regular range_of rng powerClass ordered_pair not_subclass_element not_homomorphism2 not_homomorphism1 INVERSE intersection image' flip first domain_of domain difference diagonalise cross_product compos complement cantor apply equal & SET004_1_ax range_of function maps apply application_function singleton_relation element_relation complement intersection single_valued3 singleton image' domain single_valued2 second single_valued1 identity_relation INVERSE not_subclass_element first domain_of domain_relation composition_function compos equal ordered_pair member universal_class cross_product compose_class subclass & SET004_1_eq maps single_valued3 single_valued2 single_valued1 compose_class equal & NUM004_0_ax integer_of omega ordinal_multiply add_relation ordinal_add recursion apply range_of union_of_range_map function recursion_equation_functions rest_relation rest_of limit_ordinals kind_1_ordinals successor_relation image' universal_class sum_class element_relation ordinal_numbers section not_well_ordering ordered_pair least member well_ordering singleton domain_of segment null_class intersection asymmetric compos transitive cross_product connected identity_relation complement restrct subclass irreflexive symmetrization_of INVERSE union equal & NUM004_0_eq well_ordering transitive section irreflexive connected asymmetric symmetrization_of segment rest_of recursion_equation_functions recursion ordinal_multiply ordinal_add not_well_ordering least integer_of equal & (~subclass(limit_ordinals::'a,ordinal_numbers)) --> False" by meson (*0 inferences so far. Searching to depth 0. 16.8 secs. BIG*) lemma NUM228_1: "EQU001_0_ax equal & SET004_0_ax not_homomorphism2 not_homomorphism1 homomorphism compatible operation cantor diagonalise subset_relation one_to_one choice apply regular function identity_relation single_valued_class compos powerClass sum_class omega inductive successor_relation successor image' rng domain range_of INVERSE flip rot domain_of null_class restrct difference union complement intersection element_relation second first cross_product ordered_pair singleton unordered_pair equal universal_class not_subclass_element member subclass & SET004_0_eq subclass single_valued_class operation one_to_one member inductive homomorphism function compatible unordered_pair union sum_class successor singleton second rot restrct regular range_of rng powerClass ordered_pair not_subclass_element not_homomorphism2 not_homomorphism1 INVERSE intersection image' flip first domain_of domain difference diagonalise cross_product compos complement cantor apply equal & SET004_1_ax range_of function maps apply application_function singleton_relation element_relation complement intersection single_valued3 singleton image' domain single_valued2 second single_valued1 identity_relation INVERSE not_subclass_element first domain_of domain_relation composition_function compos equal ordered_pair member universal_class cross_product compose_class subclass & SET004_1_eq maps single_valued3 single_valued2 single_valued1 compose_class equal & NUM004_0_ax integer_of omega ordinal_multiply add_relation ordinal_add recursion apply range_of union_of_range_map function recursion_equation_functions rest_relation rest_of limit_ordinals kind_1_ordinals successor_relation image' universal_class sum_class element_relation ordinal_numbers section not_well_ordering ordered_pair least member well_ordering singleton domain_of segment null_class intersection asymmetric compos transitive cross_product connected identity_relation complement restrct subclass irreflexive symmetrization_of INVERSE union equal & NUM004_0_eq well_ordering transitive section irreflexive connected asymmetric symmetrization_of segment rest_of recursion_equation_functions recursion ordinal_multiply ordinal_add not_well_ordering least integer_of equal & (~function(z)) & (~equal(recursion_equation_functions(z),null_class)) --> False" by meson (*4868 inferences so far. Searching to depth 12. 4.3 secs*) lemma PLA002_1: "(\Situation1 Situation2. warm(Situation1) | cold(Situation2)) & (\Situation. at(a::'a,Situation) --> at(b::'a,walk(b::'a,Situation))) & (\Situation. at(a::'a,Situation) --> at(b::'a,drive(b::'a,Situation))) & (\Situation. at(b::'a,Situation) --> at(a::'a,walk(a::'a,Situation))) & (\Situation. at(b::'a,Situation) --> at(a::'a,drive(a::'a,Situation))) & (\Situation. cold(Situation) & at(b::'a,Situation) --> at(c::'a,skate(c::'a,Situation))) & (\Situation. cold(Situation) & at(c::'a,Situation) --> at(b::'a,skate(b::'a,Situation))) & (\Situation. warm(Situation) & at(b::'a,Situation) --> at(d::'a,climb(d::'a,Situation))) & (\Situation. warm(Situation) & at(d::'a,Situation) --> at(b::'a,climb(b::'a,Situation))) & (\Situation. at(c::'a,Situation) --> at(d::'a,go(d::'a,Situation))) & (\Situation. at(d::'a,Situation) --> at(c::'a,go(c::'a,Situation))) & (\Situation. at(c::'a,Situation) --> at(e::'a,go(e::'a,Situation))) & (\Situation. at(e::'a,Situation) --> at(c::'a,go(c::'a,Situation))) & (\Situation. at(d::'a,Situation) --> at(f::'a,go(f::'a,Situation))) & (\Situation. at(f::'a,Situation) --> at(d::'a,go(d::'a,Situation))) & (at(f::'a,s0)) & (\S'. ~at(a::'a,S')) --> False" by meson abbreviation "PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds \ (\X Y State. holds(X::'a,State) & holds(Y::'a,State) --> holds(and'(X::'a,Y),State)) & (\State X. holds(EMPTY::'a,State) & holds(clear(X),State) & differ(X::'a,table) --> holds(holding(X),do(pickup(X),State))) & (\Y X State. holds(on(X::'a,Y),State) & holds(clear(X),State) & holds(EMPTY::'a,State) --> holds(clear(Y),do(pickup(X),State))) & (\Y State X Z. holds(on(X::'a,Y),State) & differ(X::'a,Z) --> holds(on(X::'a,Y),do(pickup(Z),State))) & (\State X Z. holds(clear(X),State) & differ(X::'a,Z) --> holds(clear(X),do(pickup(Z),State))) & (\X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(EMPTY::'a,do(putdown(X::'a,Y),State))) & (\X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(on(X::'a,Y),do(putdown(X::'a,Y),State))) & (\X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(clear(X),do(putdown(X::'a,Y),State))) & (\Z W X Y State. holds(on(X::'a,Y),State) --> holds(on(X::'a,Y),do(putdown(Z::'a,W),State))) & (\X State Z Y. holds(clear(Z),State) & differ(Z::'a,Y) --> holds(clear(Z),do(putdown(X::'a,Y),State)))" abbreviation "PLA001_1_ax EMPTY clear s0 on holds table d c b a differ \ (\Y X. differ(Y::'a,X) --> differ(X::'a,Y)) & (differ(a::'a,b)) & (differ(a::'a,c)) & (differ(a::'a,d)) & (differ(a::'a,table)) & (differ(b::'a,c)) & (differ(b::'a,d)) & (differ(b::'a,table)) & (differ(c::'a,d)) & (differ(c::'a,table)) & (differ(d::'a,table)) & (holds(on(a::'a,table),s0)) & (holds(on(b::'a,table),s0)) & (holds(on(c::'a,d),s0)) & (holds(on(d::'a,table),s0)) & (holds(clear(a),s0)) & (holds(clear(b),s0)) & (holds(clear(c),s0)) & (holds(EMPTY::'a,s0)) & (\State. holds(clear(table),State))" (*190 inferences so far. Searching to depth 10. 0.6 secs*) lemma PLA006_1: "PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds & PLA001_1_ax EMPTY clear s0 on holds table d c b a differ & (\State. ~holds(on(c::'a,table),State)) --> False" by meson (*190 inferences so far. Searching to depth 10. 0.5 secs*) lemma PLA017_1: "PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds & PLA001_1_ax EMPTY clear s0 on holds table d c b a differ & (\State. ~holds(on(a::'a,c),State)) --> False" by meson (*13732 inferences so far. Searching to depth 16. 9.8 secs*) lemma PLA022_1: "PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds & PLA001_1_ax EMPTY clear s0 on holds table d c b a differ & (\State. ~holds(and'(on(c::'a,d),on(a::'a,c)),State)) --> False" by meson (*217 inferences so far. Searching to depth 13. 0.7 secs*) lemma PLA022_2: "PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds & PLA001_1_ax EMPTY clear s0 on holds table d c b a differ & (\State. ~holds(and'(on(a::'a,c),on(c::'a,d)),State)) --> False" by meson (*948 inferences so far. Searching to depth 18. 1.1 secs*) lemma PRV001_1: "(\X Y Z. q1(X::'a,Y,Z) & mless_or_equal(X::'a,Y) --> q2(X::'a,Y,Z)) & (\X Y Z. q1(X::'a,Y,Z) --> mless_or_equal(X::'a,Y) | q3(X::'a,Y,Z)) & (\Z X Y. q2(X::'a,Y,Z) --> q4(X::'a,Y,Y)) & (\Z Y X. q3(X::'a,Y,Z) --> q4(X::'a,Y,X)) & (\X. mless_or_equal(X::'a,X)) & (\X Y. mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,X) --> equal(X::'a,Y)) & (\Y X Z. mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,Z) --> mless_or_equal(X::'a,Z)) & (\Y X. mless_or_equal(X::'a,Y) | mless_or_equal(Y::'a,X)) & (\X Y. equal(X::'a,Y) --> mless_or_equal(X::'a,Y)) & (\X Y Z. equal(X::'a,Y) & mless_or_equal(X::'a,Z) --> mless_or_equal(Y::'a,Z)) & (\X Z Y. equal(X::'a,Y) & mless_or_equal(Z::'a,X) --> mless_or_equal(Z::'a,Y)) & (q1(a::'a,b,c)) & (\W. ~(q4(a::'a,b,W) & mless_or_equal(a::'a,W) & mless_or_equal(b::'a,W) & mless_or_equal(W::'a,a))) & (\W. ~(q4(a::'a,b,W) & mless_or_equal(a::'a,W) & mless_or_equal(b::'a,W) & mless_or_equal(W::'a,b))) --> False" by meson (*PRV is now called SWV (software verification) *) abbreviation "SWV001_1_ax mless_THAN successor predecessor equal \ (\X. equal(predecessor(successor(X)),X)) & (\X. equal(successor(predecessor(X)),X)) & (\X Y. equal(predecessor(X),predecessor(Y)) --> equal(X::'a,Y)) & (\X Y. equal(successor(X),successor(Y)) --> equal(X::'a,Y)) & (\X. mless_THAN(predecessor(X),X)) & (\X. mless_THAN(X::'a,successor(X))) & (\X Y Z. mless_THAN(X::'a,Y) & mless_THAN(Y::'a,Z) --> mless_THAN(X::'a,Z)) & (\X Y. mless_THAN(X::'a,Y) | mless_THAN(Y::'a,X) | equal(X::'a,Y)) & (\X. ~mless_THAN(X::'a,X)) & (\Y X. ~(mless_THAN(X::'a,Y) & mless_THAN(Y::'a,X))) & (\Y X Z. equal(X::'a,Y) & mless_THAN(X::'a,Z) --> mless_THAN(Y::'a,Z)) & (\Y Z X. equal(X::'a,Y) & mless_THAN(Z::'a,X) --> mless_THAN(Z::'a,Y))" abbreviation "SWV001_0_eq a successor predecessor equal \ (\X Y. equal(X::'a,Y) --> equal(predecessor(X),predecessor(Y))) & (\X Y. equal(X::'a,Y) --> equal(successor(X),successor(Y))) & (\X Y. equal(X::'a,Y) --> equal(a(X),a(Y)))" (*21 inferences so far. Searching to depth 5. 0.4 secs*) lemma PRV003_1: "EQU001_0_ax equal & SWV001_1_ax mless_THAN successor predecessor equal & SWV001_0_eq a successor predecessor equal & (~mless_THAN(n::'a,j)) & (mless_THAN(k::'a,j)) & (~mless_THAN(k::'a,i)) & (mless_THAN(i::'a,n)) & (mless_THAN(a(j),a(k))) & (\X. mless_THAN(X::'a,j) & mless_THAN(a(X),a(k)) --> mless_THAN(X::'a,i)) & (\X. mless_THAN(One::'a,i) & mless_THAN(a(X),a(predecessor(i))) --> mless_THAN(X::'a,i) | mless_THAN(n::'a,X)) & (\X. ~(mless_THAN(One::'a,X) & mless_THAN(X::'a,i) & mless_THAN(a(X),a(predecessor(X))))) & (mless_THAN(j::'a,i)) --> False" by meson (*584 inferences so far. Searching to depth 7. 1.1 secs*) lemma PRV005_1: "EQU001_0_ax equal & SWV001_1_ax mless_THAN successor predecessor equal & SWV001_0_eq a successor predecessor equal & (~mless_THAN(n::'a,k)) & (~mless_THAN(k::'a,l)) & (~mless_THAN(k::'a,i)) & (mless_THAN(l::'a,n)) & (mless_THAN(One::'a,l)) & (mless_THAN(a(k),a(predecessor(l)))) & (\X. mless_THAN(X::'a,successor(n)) & mless_THAN(a(X),a(k)) --> mless_THAN(X::'a,l)) & (\X. mless_THAN(One::'a,l) & mless_THAN(a(X),a(predecessor(l))) --> mless_THAN(X::'a,l) | mless_THAN(n::'a,X)) & (\X. ~(mless_THAN(One::'a,X) & mless_THAN(X::'a,l) & mless_THAN(a(X),a(predecessor(X))))) --> False" by meson (*2343 inferences so far. Searching to depth 8. 3.5 secs*) lemma PRV006_1: "EQU001_0_ax equal & SWV001_1_ax mless_THAN successor predecessor equal & SWV001_0_eq a successor predecessor equal & (~mless_THAN(n::'a,m)) & (mless_THAN(i::'a,m)) & (mless_THAN(i::'a,n)) & (~mless_THAN(i::'a,One)) & (mless_THAN(a(i),a(m))) & (\X. mless_THAN(X::'a,successor(n)) & mless_THAN(a(X),a(m)) --> mless_THAN(X::'a,i)) & (\X. mless_THAN(One::'a,i) & mless_THAN(a(X),a(predecessor(i))) --> mless_THAN(X::'a,i) | mless_THAN(n::'a,X)) & (\X. ~(mless_THAN(One::'a,X) & mless_THAN(X::'a,i) & mless_THAN(a(X),a(predecessor(X))))) --> False" by meson (*86 inferences so far. Searching to depth 14. 0.1 secs*) lemma PRV009_1: "(\Y X. mless_or_equal(X::'a,Y) | mless(Y::'a,X)) & (mless(j::'a,i)) & (mless_or_equal(m::'a,p)) & (mless_or_equal(p::'a,q)) & (mless_or_equal(q::'a,n)) & (\X Y. mless_or_equal(m::'a,X) & mless(X::'a,i) & mless(j::'a,Y) & mless_or_equal(Y::'a,n) --> mless_or_equal(a(X),a(Y))) & (\X Y. mless_or_equal(m::'a,X) & mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,j) --> mless_or_equal(a(X),a(Y))) & (\X Y. mless_or_equal(i::'a,X) & mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,n) --> mless_or_equal(a(X),a(Y))) & (~mless_or_equal(a(p),a(q))) --> False" by meson (*222 inferences so far. Searching to depth 8. 0.4 secs*) lemma PUZ012_1: "(\X. equal_fruits(X::'a,X)) & (\X. equal_boxes(X::'a,X)) & (\X Y. ~(label(X::'a,Y) & contains(X::'a,Y))) & (\X. contains(boxa::'a,X) | contains(boxb::'a,X) | contains(boxc::'a,X)) & (\X. contains(X::'a,apples) | contains(X::'a,bananas) | contains(X::'a,oranges)) & (\X Y Z. contains(X::'a,Y) & contains(X::'a,Z) --> equal_fruits(Y::'a,Z)) & (\Y X Z. contains(X::'a,Y) & contains(Z::'a,Y) --> equal_boxes(X::'a,Z)) & (~equal_boxes(boxa::'a,boxb)) & (~equal_boxes(boxb::'a,boxc)) & (~equal_boxes(boxa::'a,boxc)) & (~equal_fruits(apples::'a,bananas)) & (~equal_fruits(bananas::'a,oranges)) & (~equal_fruits(apples::'a,oranges)) & (label(boxa::'a,apples)) & (label(boxb::'a,oranges)) & (label(boxc::'a,bananas)) & (contains(boxb::'a,apples)) & (~(contains(boxa::'a,bananas) & contains(boxc::'a,oranges))) --> False" by meson (*35 inferences so far. Searching to depth 5. 3.2 secs*) lemma PUZ020_1: "EQU001_0_ax equal & (\A B. equal(A::'a,B) --> equal(statement_by(A),statement_by(B))) & (\X. person(X) --> knight(X) | knave(X)) & (\X. ~(person(X) & knight(X) & knave(X))) & (\X Y. says(X::'a,Y) & a_truth(Y) --> a_truth(Y)) & (\X Y. ~(says(X::'a,Y) & equal(X::'a,Y))) & (\Y X. says(X::'a,Y) --> equal(Y::'a,statement_by(X))) & (\X Y. ~(person(X) & equal(X::'a,statement_by(Y)))) & (\X. person(X) & a_truth(statement_by(X)) --> knight(X)) & (\X. person(X) --> a_truth(statement_by(X)) | knave(X)) & (\X Y. equal(X::'a,Y) & knight(X) --> knight(Y)) & (\X Y. equal(X::'a,Y) & knave(X) --> knave(Y)) & (\X Y. equal(X::'a,Y) & person(X) --> person(Y)) & (\X Y Z. equal(X::'a,Y) & says(X::'a,Z) --> says(Y::'a,Z)) & (\X Z Y. equal(X::'a,Y) & says(Z::'a,X) --> says(Z::'a,Y)) & (\X Y. equal(X::'a,Y) & a_truth(X) --> a_truth(Y)) & (\X Y. knight(X) & says(X::'a,Y) --> a_truth(Y)) & (\X Y. ~(knave(X) & says(X::'a,Y) & a_truth(Y))) & (person(husband)) & (person(wife)) & (~equal(husband::'a,wife)) & (says(husband::'a,statement_by(husband))) & (a_truth(statement_by(husband)) & knight(husband) --> knight(wife)) & (knight(husband) --> a_truth(statement_by(husband))) & (a_truth(statement_by(husband)) | knight(wife)) & (knight(wife) --> a_truth(statement_by(husband))) & (~knight(husband)) --> False" by meson (*121806 inferences so far. Searching to depth 17. 63.0 secs*) lemma PUZ025_1: "(\X. a_truth(truthteller(X)) | a_truth(liar(X))) & (\X. ~(a_truth(truthteller(X)) & a_truth(liar(X)))) & (\Truthteller Statement. a_truth(truthteller(Truthteller)) & a_truth(says(Truthteller::'a,Statement)) --> a_truth(Statement)) & (\Liar Statement. ~(a_truth(liar(Liar)) & a_truth(says(Liar::'a,Statement)) & a_truth(Statement))) & (\Statement Truthteller. a_truth(Statement) & a_truth(says(Truthteller::'a,Statement)) --> a_truth(truthteller(Truthteller))) & (\Statement Liar. a_truth(says(Liar::'a,Statement)) --> a_truth(Statement) | a_truth(liar(Liar))) & (\Z X Y. people(X::'a,Y,Z) & a_truth(liar(X)) & a_truth(liar(Y)) --> a_truth(equal_type(X::'a,Y))) & (\Z X Y. people(X::'a,Y,Z) & a_truth(truthteller(X)) & a_truth(truthteller(Y)) --> a_truth(equal_type(X::'a,Y))) & (\X Y. a_truth(equal_type(X::'a,Y)) & a_truth(truthteller(X)) --> a_truth(truthteller(Y))) & (\X Y. a_truth(equal_type(X::'a,Y)) & a_truth(liar(X)) --> a_truth(liar(Y))) & (\X Y. a_truth(truthteller(X)) --> a_truth(equal_type(X::'a,Y)) | a_truth(liar(Y))) & (\X Y. a_truth(liar(X)) --> a_truth(equal_type(X::'a,Y)) | a_truth(truthteller(Y))) & (\Y X. a_truth(equal_type(X::'a,Y)) --> a_truth(equal_type(Y::'a,X))) & (\X Y. ask_1_if_2(X::'a,Y) & a_truth(truthteller(X)) & a_truth(Y) --> answer(yes)) & (\X Y. ask_1_if_2(X::'a,Y) & a_truth(truthteller(X)) --> a_truth(Y) | answer(no)) & (\X Y. ask_1_if_2(X::'a,Y) & a_truth(liar(X)) & a_truth(Y) --> answer(no)) & (\X Y. ask_1_if_2(X::'a,Y) & a_truth(liar(X)) --> a_truth(Y) | answer(yes)) & (people(b::'a,c,a)) & (people(a::'a,b,a)) & (people(a::'a,c,b)) & (people(c::'a,b,a)) & (a_truth(says(a::'a,equal_type(b::'a,c)))) & (ask_1_if_2(c::'a,equal_type(a::'a,b))) & (\Answer. ~answer(Answer)) --> False" by meson (*621 inferences so far. Searching to depth 18. 0.2 secs*) lemma PUZ029_1: "(\X. dances_on_tightropes(X) | eats_pennybuns(X) | old(X)) & (\X. pig(X) & liable_to_giddiness(X) --> treated_with_respect(X)) & (\X. wise(X) & balloonist(X) --> has_umbrella(X)) & (\X. ~(looks_ridiculous(X) & eats_pennybuns(X) & eats_lunch_in_public(X))) & (\X. balloonist(X) & young(X) --> liable_to_giddiness(X)) & (\X. fat(X) & looks_ridiculous(X) --> dances_on_tightropes(X) | eats_lunch_in_public(X)) & (\X. ~(liable_to_giddiness(X) & wise(X) & dances_on_tightropes(X))) & (\X. pig(X) & has_umbrella(X) --> looks_ridiculous(X)) & (\X. treated_with_respect(X) --> dances_on_tightropes(X) | fat(X)) & (\X. young(X) | old(X)) & (\X. ~(young(X) & old(X))) & (wise(piggy)) & (young(piggy)) & (pig(piggy)) & (balloonist(piggy)) --> False" by meson abbreviation "RNG001_0_ax equal additive_inverse add multiply product additive_identity sum \ (\X. sum(additive_identity::'a,X,X)) & (\X. sum(X::'a,additive_identity,X)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y. sum(X::'a,Y,add(X::'a,Y))) & (\X. sum(additive_inverse(X),X,additive_identity)) & (\X. sum(X::'a,additive_inverse(X),additive_identity)) & (\Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) & (\Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) & (\Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) & (\Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) & (\Y Z V3 X V1 V2 V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & product(V3::'a,X,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 V3 X V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(V3::'a,X,V4)) & (\X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) & (\X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V))" abbreviation "RNG001_0_eq product multiply sum add additive_inverse equal \ (\X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) & (\X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & (\X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & (\X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) & (\X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) & (\X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) & (\X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y))" (*93620 inferences so far. Searching to depth 24. 65.9 secs*) lemma RNG001_3: "(\X. sum(additive_identity::'a,X,X)) & (\X. sum(additive_inverse(X),X,additive_identity)) & (\Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) & (\Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) & (~product(a::'a,additive_identity,additive_identity)) --> False" by meson abbreviation "RNG_other_ax multiply add equal product additive_identity additive_inverse sum \ (\X. sum(X::'a,additive_inverse(X),additive_identity)) & (\Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) & (\Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) & (\Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) & (\Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) & (\Y Z V3 X V1 V2 V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & product(V3::'a,X,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 V3 X V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(V3::'a,X,V4)) & (\X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) & (\X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) & (\X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) & (\X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & (\X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) & (\X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) & (\X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) & (\X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) & (\X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y))" (****************SLOW 76385914 inferences so far. Searching to depth 18 No proof after 5 1/2 hours! (griffon) ****************) lemma RNG001_5: "EQU001_0_ax equal & (\X. sum(additive_identity::'a,X,X)) & (\X. sum(X::'a,additive_identity,X)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y. sum(X::'a,Y,add(X::'a,Y))) & (\X. sum(additive_inverse(X),X,additive_identity)) & RNG_other_ax multiply add equal product additive_identity additive_inverse sum & (~product(a::'a,additive_identity,additive_identity)) --> False" oops (*0 inferences so far. Searching to depth 0. 0.5 secs*) lemma RNG011_5: "EQU001_0_ax equal & (\A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) & (\G H. equal(G::'a,H) --> equal(additive_inverse(G),additive_inverse(H))) & (\I' J K'. equal(I'::'a,J) --> equal(multiply(I'::'a,K'),multiply(J::'a,K'))) & (\L N M. equal(L::'a,M) --> equal(multiply(N::'a,L),multiply(N::'a,M))) & (\A B C D. equal(A::'a,B) --> equal(associator(A::'a,C,D),associator(B::'a,C,D))) & (\E G F' H. equal(E::'a,F') --> equal(associator(G::'a,E,H),associator(G::'a,F',H))) & (\I' K' L J. equal(I'::'a,J) --> equal(associator(K'::'a,L,I'),associator(K'::'a,L,J))) & (\M N O'. equal(M::'a,N) --> equal(commutator(M::'a,O'),commutator(N::'a,O'))) & (\P R Q. equal(P::'a,Q) --> equal(commutator(R::'a,P),commutator(R::'a,Q))) & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) & (\X. equal(add(X::'a,additive_identity),X)) & (\X. equal(add(additive_identity::'a,X),X)) & (\X. equal(add(X::'a,additive_inverse(X)),additive_identity)) & (\X. equal(add(additive_inverse(X),X),additive_identity)) & (equal(additive_inverse(additive_identity),additive_identity)) & (\X Y. equal(add(X::'a,add(additive_inverse(X),Y)),Y)) & (\X Y. equal(additive_inverse(add(X::'a,Y)),add(additive_inverse(X),additive_inverse(Y)))) & (\X. equal(additive_inverse(additive_inverse(X)),X)) & (\X. equal(multiply(X::'a,additive_identity),additive_identity)) & (\X. equal(multiply(additive_identity::'a,X),additive_identity)) & (\X Y. equal(multiply(additive_inverse(X),additive_inverse(Y)),multiply(X::'a,Y))) & (\X Y. equal(multiply(X::'a,additive_inverse(Y)),additive_inverse(multiply(X::'a,Y)))) & (\X Y. equal(multiply(additive_inverse(X),Y),additive_inverse(multiply(X::'a,Y)))) & (\Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) & (\X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) & (\X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) & (\X Y Z. equal(associator(X::'a,Y,Z),add(multiply(multiply(X::'a,Y),Z),additive_inverse(multiply(X::'a,multiply(Y::'a,Z)))))) & (\X Y. equal(commutator(X::'a,Y),add(multiply(Y::'a,X),additive_inverse(multiply(X::'a,Y))))) & (\X Y. equal(multiply(multiply(associator(X::'a,X,Y),X),associator(X::'a,X,Y)),additive_identity)) & (~equal(multiply(multiply(associator(a::'a,a,b),a),associator(a::'a,a,b)),additive_identity)) --> False" by meson (*202 inferences so far. Searching to depth 8. 0.6 secs*) lemma RNG023_6: "EQU001_0_ax equal & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(X::'a,add(Y::'a,Z)),add(add(X::'a,Y),Z))) & (\X. equal(add(additive_identity::'a,X),X)) & (\X. equal(add(X::'a,additive_identity),X)) & (\X. equal(multiply(additive_identity::'a,X),additive_identity)) & (\X. equal(multiply(X::'a,additive_identity),additive_identity)) & (\X. equal(add(additive_inverse(X),X),additive_identity)) & (\X. equal(add(X::'a,additive_inverse(X)),additive_identity)) & (\Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) & (\X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) & (\X. equal(additive_inverse(additive_inverse(X)),X)) & (\X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) & (\X Y. equal(multiply(multiply(X::'a,X),Y),multiply(X::'a,multiply(X::'a,Y)))) & (\X Y Z. equal(associator(X::'a,Y,Z),add(multiply(multiply(X::'a,Y),Z),additive_inverse(multiply(X::'a,multiply(Y::'a,Z)))))) & (\X Y. equal(commutator(X::'a,Y),add(multiply(Y::'a,X),additive_inverse(multiply(X::'a,Y))))) & (\D E F'. equal(D::'a,E) --> equal(add(D::'a,F'),add(E::'a,F'))) & (\G I' H. equal(G::'a,H) --> equal(add(I'::'a,G),add(I'::'a,H))) & (\J K'. equal(J::'a,K') --> equal(additive_inverse(J),additive_inverse(K'))) & (\L M N O'. equal(L::'a,M) --> equal(associator(L::'a,N,O'),associator(M::'a,N,O'))) & (\P R Q S'. equal(P::'a,Q) --> equal(associator(R::'a,P,S'),associator(R::'a,Q,S'))) & (\T' V W U. equal(T'::'a,U) --> equal(associator(V::'a,W,T'),associator(V::'a,W,U))) & (\X Y Z. equal(X::'a,Y) --> equal(commutator(X::'a,Z),commutator(Y::'a,Z))) & (\A1 C1 B1. equal(A1::'a,B1) --> equal(commutator(C1::'a,A1),commutator(C1::'a,B1))) & (\D1 E1 F1. equal(D1::'a,E1) --> equal(multiply(D1::'a,F1),multiply(E1::'a,F1))) & (\G1 I1 H1. equal(G1::'a,H1) --> equal(multiply(I1::'a,G1),multiply(I1::'a,H1))) & (~equal(associator(x::'a,x,y),additive_identity)) --> False" by meson (*0 inferences so far. Searching to depth 0. 0.6 secs*) lemma RNG028_2: "EQU001_0_ax equal & (\X. equal(add(additive_identity::'a,X),X)) & (\X. equal(multiply(additive_identity::'a,X),additive_identity)) & (\X. equal(multiply(X::'a,additive_identity),additive_identity)) & (\X. equal(add(additive_inverse(X),X),additive_identity)) & (\X Y. equal(additive_inverse(add(X::'a,Y)),add(additive_inverse(X),additive_inverse(Y)))) & (\X. equal(additive_inverse(additive_inverse(X)),X)) & (\Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) & (\X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) & (\X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) & (\X Y. equal(multiply(multiply(X::'a,X),Y),multiply(X::'a,multiply(X::'a,Y)))) & (\X Y. equal(multiply(additive_inverse(X),Y),additive_inverse(multiply(X::'a,Y)))) & (\X Y. equal(multiply(X::'a,additive_inverse(Y)),additive_inverse(multiply(X::'a,Y)))) & (equal(additive_inverse(additive_identity),additive_identity)) & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(X::'a,add(Y::'a,Z)),add(add(X::'a,Y),Z))) & (\Z X Y. equal(add(X::'a,Z),add(Y::'a,Z)) --> equal(X::'a,Y)) & (\Z X Y. equal(add(Z::'a,X),add(Z::'a,Y)) --> equal(X::'a,Y)) & (\D E F'. equal(D::'a,E) --> equal(add(D::'a,F'),add(E::'a,F'))) & (\G I' H. equal(G::'a,H) --> equal(add(I'::'a,G),add(I'::'a,H))) & (\J K'. equal(J::'a,K') --> equal(additive_inverse(J),additive_inverse(K'))) & (\D1 E1 F1. equal(D1::'a,E1) --> equal(multiply(D1::'a,F1),multiply(E1::'a,F1))) & (\G1 I1 H1. equal(G1::'a,H1) --> equal(multiply(I1::'a,G1),multiply(I1::'a,H1))) & (\X Y Z. equal(associator(X::'a,Y,Z),add(multiply(multiply(X::'a,Y),Z),additive_inverse(multiply(X::'a,multiply(Y::'a,Z)))))) & (\L M N O'. equal(L::'a,M) --> equal(associator(L::'a,N,O'),associator(M::'a,N,O'))) & (\P R Q S'. equal(P::'a,Q) --> equal(associator(R::'a,P,S'),associator(R::'a,Q,S'))) & (\T' V W U. equal(T'::'a,U) --> equal(associator(V::'a,W,T'),associator(V::'a,W,U))) & (\X Y. ~equal(multiply(multiply(Y::'a,X),Y),multiply(Y::'a,multiply(X::'a,Y)))) & (\X Y Z. ~equal(associator(Y::'a,X,Z),additive_inverse(associator(X::'a,Y,Z)))) & (\X Y Z. ~equal(associator(Z::'a,Y,X),additive_inverse(associator(X::'a,Y,Z)))) & (~equal(multiply(multiply(cx::'a,multiply(cy::'a,cx)),cz),multiply(cx::'a,multiply(cy::'a,multiply(cx::'a,cz))))) --> False" by meson (*209 inferences so far. Searching to depth 9. 1.2 secs*) lemma RNG038_2: "(\X. sum(X::'a,additive_identity,X)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y. sum(X::'a,Y,add(X::'a,Y))) & RNG_other_ax multiply add equal product additive_identity additive_inverse sum & (\X. product(additive_identity::'a,X,additive_identity)) & (\X. product(X::'a,additive_identity,additive_identity)) & (\X Y. equal(X::'a,additive_identity) --> product(X::'a,h(X::'a,Y),Y)) & (product(a::'a,b,additive_identity)) & (~equal(a::'a,additive_identity)) & (~equal(b::'a,additive_identity)) --> False" by meson (*2660 inferences so far. Searching to depth 10. 7.0 secs*) lemma RNG040_2: "EQU001_0_ax equal & RNG001_0_eq product multiply sum add additive_inverse equal & (\X. sum(additive_identity::'a,X,X)) & (\X. sum(X::'a,additive_identity,X)) & (\X Y. product(X::'a,Y,multiply(X::'a,Y))) & (\X Y. sum(X::'a,Y,add(X::'a,Y))) & (\X. sum(additive_inverse(X),X,additive_identity)) & (\X. sum(X::'a,additive_inverse(X),additive_identity)) & (\Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) & (\Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) & (\Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) & (\Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) & (\Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) & (\Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) & (\Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) & (\X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) & (\X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) & (\A. product(A::'a,multiplicative_identity,A)) & (\A. product(multiplicative_identity::'a,A,A)) & (\A. product(A::'a,h(A),multiplicative_identity) | equal(A::'a,additive_identity)) & (\A. product(h(A),A,multiplicative_identity) | equal(A::'a,additive_identity)) & (\B A C. product(A::'a,B,C) --> product(B::'a,A,C)) & (\A B. equal(A::'a,B) --> equal(h(A),h(B))) & (sum(b::'a,c,d)) & (product(d::'a,a,additive_identity)) & (product(b::'a,a,l)) & (product(c::'a,a,n)) & (~sum(l::'a,n,additive_identity)) --> False" by meson (*8991 inferences so far. Searching to depth 9. 22.2 secs*) lemma RNG041_1: "EQU001_0_ax equal & RNG001_0_ax equal additive_inverse add multiply product additive_identity sum & RNG001_0_eq product multiply sum add additive_inverse equal & (\A B. equal(A::'a,B) --> equal(h(A),h(B))) & (\A. product(additive_identity::'a,A,additive_identity)) & (\A. product(A::'a,additive_identity,additive_identity)) & (\A. product(A::'a,multiplicative_identity,A)) & (\A. product(multiplicative_identity::'a,A,A)) & (\A. product(A::'a,h(A),multiplicative_identity) | equal(A::'a,additive_identity)) & (\A. product(h(A),A,multiplicative_identity) | equal(A::'a,additive_identity)) & (product(a::'a,b,additive_identity)) & (~equal(a::'a,additive_identity)) & (~equal(b::'a,additive_identity)) --> False" by meson (*101319 inferences so far. Searching to depth 14. 76.0 secs*) lemma ROB010_1: "EQU001_0_ax equal & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) & (\Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) & (\A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) & (\G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & (equal(negate(add(a::'a,negate(b))),c)) & (~equal(negate(add(c::'a,negate(add(b::'a,a)))),a)) --> False" by meson (*6933 inferences so far. Searching to depth 12. 5.1 secs*) lemma ROB013_1: "EQU001_0_ax equal & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) & (\Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) & (\A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) & (\G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & (equal(negate(add(a::'a,b)),c)) & (~equal(negate(add(c::'a,negate(add(negate(b),a)))),a)) --> False" by meson (*6614 inferences so far. Searching to depth 11. 20.4 secs*) lemma ROB016_1: "EQU001_0_ax equal & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) & (\Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) & (\A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) & (\G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & (\J K' L. equal(J::'a,K') --> equal(multiply(J::'a,L),multiply(K'::'a,L))) & (\M O' N. equal(M::'a,N) --> equal(multiply(O'::'a,M),multiply(O'::'a,N))) & (\P Q. equal(P::'a,Q) --> equal(successor(P),successor(Q))) & (\R S'. equal(R::'a,S') & positive_integer(R) --> positive_integer(S')) & (\X. equal(multiply(One::'a,X),X)) & (\V X. positive_integer(X) --> equal(multiply(successor(V),X),add(X::'a,multiply(V::'a,X)))) & (positive_integer(One)) & (\X. positive_integer(X) --> positive_integer(successor(X))) & (equal(negate(add(d::'a,e)),negate(e))) & (positive_integer(k)) & (\Vk X Y. equal(negate(add(negate(Y),negate(add(X::'a,negate(Y))))),X) & positive_integer(Vk) --> equal(negate(add(Y::'a,multiply(Vk::'a,add(X::'a,negate(add(X::'a,negate(Y))))))),negate(Y))) & (~equal(negate(add(e::'a,multiply(k::'a,add(d::'a,negate(add(d::'a,negate(e))))))),negate(e))) --> False" by meson (*14077 inferences so far. Searching to depth 11. 32.8 secs*) lemma ROB021_1: "EQU001_0_ax equal & (\Y X. equal(add(X::'a,Y),add(Y::'a,X))) & (\X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) & (\Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) & (\A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) & (\G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & (\X Y. equal(negate(X),negate(Y)) --> equal(X::'a,Y)) & (~equal(add(negate(add(a::'a,negate(b))),negate(add(negate(a),negate(b)))),b)) --> False" by meson (*35532 inferences so far. Searching to depth 19. 54.3 secs*) lemma SET005_1: "(\Subset Element Superset. member(Element::'a,Subset) & subset(Subset::'a,Superset) --> member(Element::'a,Superset)) & (\Superset Subset. subset(Subset::'a,Superset) | member(member_of_1_not_of_2(Subset::'a,Superset),Subset)) & (\Subset Superset. member(member_of_1_not_of_2(Subset::'a,Superset),Superset) --> subset(Subset::'a,Superset)) & (\Subset Superset. equal_sets(Subset::'a,Superset) --> subset(Subset::'a,Superset)) & (\Subset Superset. equal_sets(Superset::'a,Subset) --> subset(Subset::'a,Superset)) & (\Set2 Set1. subset(Set1::'a,Set2) & subset(Set2::'a,Set1) --> equal_sets(Set2::'a,Set1)) & (\Set2 Intersection Element Set1. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Intersection) --> member(Element::'a,Set1)) & (\Set1 Intersection Element Set2. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Intersection) --> member(Element::'a,Set2)) & (\Set2 Set1 Element Intersection. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Set2) & member(Element::'a,Set1) --> member(Element::'a,Intersection)) & (\Set2 Intersection Set1. member(h(Set1::'a,Set2,Intersection),Intersection) | intersection(Set1::'a,Set2,Intersection) | member(h(Set1::'a,Set2,Intersection),Set1)) & (\Set1 Intersection Set2. member(h(Set1::'a,Set2,Intersection),Intersection) | intersection(Set1::'a,Set2,Intersection) | member(h(Set1::'a,Set2,Intersection),Set2)) & (\Set1 Set2 Intersection. member(h(Set1::'a,Set2,Intersection),Intersection) & member(h(Set1::'a,Set2,Intersection),Set2) & member(h(Set1::'a,Set2,Intersection),Set1) --> intersection(Set1::'a,Set2,Intersection)) & (intersection(a::'a,b,aIb)) & (intersection(b::'a,c,bIc)) & (intersection(a::'a,bIc,aIbIc)) & (~intersection(aIb::'a,c,aIbIc)) --> False" by meson (*6450 inferences so far. Searching to depth 14. 4.2 secs*) lemma SET009_1: "(\Subset Element Superset. member(Element::'a,Subset) & ssubset(Subset::'a,Superset) --> member(Element::'a,Superset)) & (\Superset Subset. ssubset(Subset::'a,Superset) | member(member_of_1_not_of_2(Subset::'a,Superset),Subset)) & (\Subset Superset. member(member_of_1_not_of_2(Subset::'a,Superset),Superset) --> ssubset(Subset::'a,Superset)) & (\Subset Superset. equal_sets(Subset::'a,Superset) --> ssubset(Subset::'a,Superset)) & (\Subset Superset. equal_sets(Superset::'a,Subset) --> ssubset(Subset::'a,Superset)) & (\Set2 Set1. ssubset(Set1::'a,Set2) & ssubset(Set2::'a,Set1) --> equal_sets(Set2::'a,Set1)) & (\Set2 Difference Element Set1. difference(Set1::'a,Set2,Difference) & member(Element::'a,Difference) --> member(Element::'a,Set1)) & (\Element A_set Set1 Set2. ~(member(Element::'a,Set1) & member(Element::'a,Set2) & difference(A_set::'a,Set1,Set2))) & (\Set1 Difference Element Set2. member(Element::'a,Set1) & difference(Set1::'a,Set2,Difference) --> member(Element::'a,Difference) | member(Element::'a,Set2)) & (\Set1 Set2 Difference. difference(Set1::'a,Set2,Difference) | member(k(Set1::'a,Set2,Difference),Set1) | member(k(Set1::'a,Set2,Difference),Difference)) & (\Set1 Set2 Difference. member(k(Set1::'a,Set2,Difference),Set2) --> member(k(Set1::'a,Set2,Difference),Difference) | difference(Set1::'a,Set2,Difference)) & (\Set1 Set2 Difference. member(k(Set1::'a,Set2,Difference),Difference) & member(k(Set1::'a,Set2,Difference),Set1) --> member(k(Set1::'a,Set2,Difference),Set2) | difference(Set1::'a,Set2,Difference)) & (ssubset(d::'a,a)) & (difference(b::'a,a,bDa)) & (difference(b::'a,d,bDd)) & (~ssubset(bDa::'a,bDd)) --> False" by meson (*34726 inferences so far. Searching to depth 6. 2420 secs: 40 mins! BIG*) lemma SET025_4: "EQU001_0_ax equal & (\Y X. member(X::'a,Y) --> little_set(X)) & (\X Y. little_set(f1(X::'a,Y)) | equal(X::'a,Y)) & (\X Y. member(f1(X::'a,Y),X) | member(f1(X::'a,Y),Y) | equal(X::'a,Y)) & (\X Y. member(f1(X::'a,Y),X) & member(f1(X::'a,Y),Y) --> equal(X::'a,Y)) & (\X U Y. member(U::'a,non_ordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) & (\Y U X. little_set(U) & equal(U::'a,X) --> member(U::'a,non_ordered_pair(X::'a,Y))) & (\X U Y. little_set(U) & equal(U::'a,Y) --> member(U::'a,non_ordered_pair(X::'a,Y))) & (\X Y. little_set(non_ordered_pair(X::'a,Y))) & (\X. equal(singleton_set(X),non_ordered_pair(X::'a,X))) & (\X Y. equal(ordered_pair(X::'a,Y),non_ordered_pair(singleton_set(X),non_ordered_pair(X::'a,Y)))) & (\X. ordered_pair_predicate(X) --> little_set(f2(X))) & (\X. ordered_pair_predicate(X) --> little_set(f3(X))) & (\X. ordered_pair_predicate(X) --> equal(X::'a,ordered_pair(f2(X),f3(X)))) & (\X Y Z. little_set(Y) & little_set(Z) & equal(X::'a,ordered_pair(Y::'a,Z)) --> ordered_pair_predicate(X)) & (\Z X. member(Z::'a,first(X)) --> little_set(f4(Z::'a,X))) & (\Z X. member(Z::'a,first(X)) --> little_set(f5(Z::'a,X))) & (\Z X. member(Z::'a,first(X)) --> equal(X::'a,ordered_pair(f4(Z::'a,X),f5(Z::'a,X)))) & (\Z X. member(Z::'a,first(X)) --> member(Z::'a,f4(Z::'a,X))) & (\X V Z U. little_set(U) & little_set(V) & equal(X::'a,ordered_pair(U::'a,V)) & member(Z::'a,U) --> member(Z::'a,first(X))) & (\Z X. member(Z::'a,second(X)) --> little_set(f6(Z::'a,X))) & (\Z X. member(Z::'a,second(X)) --> little_set(f7(Z::'a,X))) & (\Z X. member(Z::'a,second(X)) --> equal(X::'a,ordered_pair(f6(Z::'a,X),f7(Z::'a,X)))) & (\Z X. member(Z::'a,second(X)) --> member(Z::'a,f7(Z::'a,X))) & (\X U Z V. little_set(U) & little_set(V) & equal(X::'a,ordered_pair(U::'a,V)) & member(Z::'a,V) --> member(Z::'a,second(X))) & (\Z. member(Z::'a,estin) --> ordered_pair_predicate(Z)) & (\Z. member(Z::'a,estin) --> member(first(Z),second(Z))) & (\Z. little_set(Z) & ordered_pair_predicate(Z) & member(first(Z),second(Z)) --> member(Z::'a,estin)) & (\Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) & (\X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) & (\X Z Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) & (\Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) & (\Z X. little_set(Z) --> member(Z::'a,complement(X)) | member(Z::'a,X)) & (\X Y. equal(union(X::'a,Y),complement(intersection(complement(X),complement(Y))))) & (\Z X. member(Z::'a,domain_of(X)) --> ordered_pair_predicate(f8(Z::'a,X))) & (\Z X. member(Z::'a,domain_of(X)) --> member(f8(Z::'a,X),X)) & (\Z X. member(Z::'a,domain_of(X)) --> equal(Z::'a,first(f8(Z::'a,X)))) & (\X Z Xp. little_set(Z) & ordered_pair_predicate(Xp) & member(Xp::'a,X) & equal(Z::'a,first(Xp)) --> member(Z::'a,domain_of(X))) & (\X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> ordered_pair_predicate(Z)) & (\Y Z X. member(Z::'a,cross_product(X::'a,Y)) --> member(first(Z),X)) & (\X Z Y. member(Z::'a,cross_product(X::'a,Y)) --> member(second(Z),Y)) & (\X Z Y. little_set(Z) & ordered_pair_predicate(Z) & member(first(Z),X) & member(second(Z),Y) --> member(Z::'a,cross_product(X::'a,Y))) & (\X Z. member(Z::'a,inv1 X) --> ordered_pair_predicate(Z)) & (\Z X. member(Z::'a,inv1 X) --> member(ordered_pair(second(Z),first(Z)),X)) & (\Z X. little_set(Z) & ordered_pair_predicate(Z) & member(ordered_pair(second(Z),first(Z)),X) --> member(Z::'a,inv1 X)) & (\Z X. member(Z::'a,rot_right(X)) --> little_set(f9(Z::'a,X))) & (\Z X. member(Z::'a,rot_right(X)) --> little_set(f10(Z::'a,X))) & (\Z X. member(Z::'a,rot_right(X)) --> little_set(f11(Z::'a,X))) & (\Z X. member(Z::'a,rot_right(X)) --> equal(Z::'a,ordered_pair(f9(Z::'a,X),ordered_pair(f10(Z::'a,X),f11(Z::'a,X))))) & (\Z X. member(Z::'a,rot_right(X)) --> member(ordered_pair(f10(Z::'a,X),ordered_pair(f11(Z::'a,X),f9(Z::'a,X))),X)) & (\Z V W U X. little_set(Z) & little_set(U) & little_set(V) & little_set(W) & equal(Z::'a,ordered_pair(U::'a,ordered_pair(V::'a,W))) & member(ordered_pair(V::'a,ordered_pair(W::'a,U)),X) --> member(Z::'a,rot_right(X))) & (\Z X. member(Z::'a,flip_range_of(X)) --> little_set(f12(Z::'a,X))) & (\Z X. member(Z::'a,flip_range_of(X)) --> little_set(f13(Z::'a,X))) & (\Z X. member(Z::'a,flip_range_of(X)) --> little_set(f14(Z::'a,X))) & (\Z X. member(Z::'a,flip_range_of(X)) --> equal(Z::'a,ordered_pair(f12(Z::'a,X),ordered_pair(f13(Z::'a,X),f14(Z::'a,X))))) & (\Z X. member(Z::'a,flip_range_of(X)) --> member(ordered_pair(f12(Z::'a,X),ordered_pair(f14(Z::'a,X),f13(Z::'a,X))),X)) & (\Z U W V X. little_set(Z) & little_set(U) & little_set(V) & little_set(W) & equal(Z::'a,ordered_pair(U::'a,ordered_pair(V::'a,W))) & member(ordered_pair(U::'a,ordered_pair(W::'a,V)),X) --> member(Z::'a,flip_range_of(X))) & (\X. equal(successor(X),union(X::'a,singleton_set(X)))) & (\Z. ~member(Z::'a,empty_set)) & (\Z. little_set(Z) --> member(Z::'a,universal_set)) & (little_set(infinity)) & (member(empty_set::'a,infinity)) & (\X. member(X::'a,infinity) --> member(successor(X),infinity)) & (\Z X. member(Z::'a,sigma(X)) --> member(f16(Z::'a,X),X)) & (\Z X. member(Z::'a,sigma(X)) --> member(Z::'a,f16(Z::'a,X))) & (\X Z Y. member(Y::'a,X) & member(Z::'a,Y) --> member(Z::'a,sigma(X))) & (\U. little_set(U) --> little_set(sigma(U))) & (\X U Y. ssubset(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) & (\Y X. ssubset(X::'a,Y) | member(f17(X::'a,Y),X)) & (\X Y. member(f17(X::'a,Y),Y) --> ssubset(X::'a,Y)) & (\X Y. proper_subset(X::'a,Y) --> ssubset(X::'a,Y)) & (\X Y. ~(proper_subset(X::'a,Y) & equal(X::'a,Y))) & (\X Y. ssubset(X::'a,Y) --> proper_subset(X::'a,Y) | equal(X::'a,Y)) & (\Z X. member(Z::'a,powerset(X)) --> ssubset(Z::'a,X)) & (\Z X. little_set(Z) & ssubset(Z::'a,X) --> member(Z::'a,powerset(X))) & (\U. little_set(U) --> little_set(powerset(U))) & (\Z X. relation(Z) & member(X::'a,Z) --> ordered_pair_predicate(X)) & (\Z. relation(Z) | member(f18(Z),Z)) & (\Z. ordered_pair_predicate(f18(Z)) --> relation(Z)) & (\U X V W. single_valued_set(X) & little_set(U) & little_set(V) & little_set(W) & member(ordered_pair(U::'a,V),X) & member(ordered_pair(U::'a,W),X) --> equal(V::'a,W)) & (\X. single_valued_set(X) | little_set(f19(X))) & (\X. single_valued_set(X) | little_set(f20(X))) & (\X. single_valued_set(X) | little_set(f21(X))) & (\X. single_valued_set(X) | member(ordered_pair(f19(X),f20(X)),X)) & (\X. single_valued_set(X) | member(ordered_pair(f19(X),f21(X)),X)) & (\X. equal(f20(X),f21(X)) --> single_valued_set(X)) & (\Xf. function(Xf) --> relation(Xf)) & (\Xf. function(Xf) --> single_valued_set(Xf)) & (\Xf. relation(Xf) & single_valued_set(Xf) --> function(Xf)) & (\Z X Xf. member(Z::'a,image'(X::'a,Xf)) --> ordered_pair_predicate(f22(Z::'a,X,Xf))) & (\Z X Xf. member(Z::'a,image'(X::'a,Xf)) --> member(f22(Z::'a,X,Xf),Xf)) & (\Z Xf X. member(Z::'a,image'(X::'a,Xf)) --> member(first(f22(Z::'a,X,Xf)),X)) & (\X Xf Z. member(Z::'a,image'(X::'a,Xf)) --> equal(second(f22(Z::'a,X,Xf)),Z)) & (\Xf X Y Z. little_set(Z) & ordered_pair_predicate(Y) & member(Y::'a,Xf) & member(first(Y),X) & equal(second(Y),Z) --> member(Z::'a,image'(X::'a,Xf))) & (\X Xf. little_set(X) & function(Xf) --> little_set(image'(X::'a,Xf))) & (\X U Y. ~(disjoint(X::'a,Y) & member(U::'a,X) & member(U::'a,Y))) & (\Y X. disjoint(X::'a,Y) | member(f23(X::'a,Y),X)) & (\X Y. disjoint(X::'a,Y) | member(f23(X::'a,Y),Y)) & (\X. equal(X::'a,empty_set) | member(f24(X),X)) & (\X. equal(X::'a,empty_set) | disjoint(f24(X),X)) & (function(f25)) & (\X. little_set(X) --> equal(X::'a,empty_set) | member(f26(X),X)) & (\X. little_set(X) --> equal(X::'a,empty_set) | member(ordered_pair(X::'a,f26(X)),f25)) & (\Z X. member(Z::'a,range_of(X)) --> ordered_pair_predicate(f27(Z::'a,X))) & (\Z X. member(Z::'a,range_of(X)) --> member(f27(Z::'a,X),X)) & (\Z X. member(Z::'a,range_of(X)) --> equal(Z::'a,second(f27(Z::'a,X)))) & (\X Z Xp. little_set(Z) & ordered_pair_predicate(Xp) & member(Xp::'a,X) & equal(Z::'a,second(Xp)) --> member(Z::'a,range_of(X))) & (\Z. member(Z::'a,identity_relation) --> ordered_pair_predicate(Z)) & (\Z. member(Z::'a,identity_relation) --> equal(first(Z),second(Z))) & (\Z. little_set(Z) & ordered_pair_predicate(Z) & equal(first(Z),second(Z)) --> member(Z::'a,identity_relation)) & (\X Y. equal(restrct(X::'a,Y),intersection(X::'a,cross_product(Y::'a,universal_set)))) & (\Xf. one_to_one_function(Xf) --> function(Xf)) & (\Xf. one_to_one_function(Xf) --> function(inv1 Xf)) & (\Xf. function(Xf) & function(inv1 Xf) --> one_to_one_function(Xf)) & (\Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> ordered_pair_predicate(f28(Z::'a,Xf,Y))) & (\Z Y Xf. member(Z::'a,apply(Xf::'a,Y)) --> member(f28(Z::'a,Xf,Y),Xf)) & (\Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> equal(first(f28(Z::'a,Xf,Y)),Y)) & (\Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> member(Z::'a,second(f28(Z::'a,Xf,Y)))) & (\Xf Y Z W. ordered_pair_predicate(W) & member(W::'a,Xf) & equal(first(W),Y) & member(Z::'a,second(W)) --> member(Z::'a,apply(Xf::'a,Y))) & (\Xf X Y. equal(apply_to_two_arguments(Xf::'a,X,Y),apply(Xf::'a,ordered_pair(X::'a,Y)))) & (\X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) & (\Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) & (\X Xf Y. maps(Xf::'a,X,Y) --> ssubset(range_of(Xf),Y)) & (\X Xf Y. function(Xf) & equal(domain_of(Xf),X) & ssubset(range_of(Xf),Y) --> maps(Xf::'a,X,Y)) & (\Xf Xs. closed(Xs::'a,Xf) --> little_set(Xs)) & (\Xs Xf. closed(Xs::'a,Xf) --> little_set(Xf)) & (\Xf Xs. closed(Xs::'a,Xf) --> maps(Xf::'a,cross_product(Xs::'a,Xs),Xs)) & (\Xf Xs. little_set(Xs) & little_set(Xf) & maps(Xf::'a,cross_product(Xs::'a,Xs),Xs) --> closed(Xs::'a,Xf)) & (\Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f29(Z::'a,Xf,Xg))) & (\Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f30(Z::'a,Xf,Xg))) & (\Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f31(Z::'a,Xf,Xg))) & (\Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> equal(Z::'a,ordered_pair(f29(Z::'a,Xf,Xg),f30(Z::'a,Xf,Xg)))) & (\Z Xg Xf. member(Z::'a,composition(Xf::'a,Xg)) --> member(ordered_pair(f29(Z::'a,Xf,Xg),f31(Z::'a,Xf,Xg)),Xf)) & (\Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> member(ordered_pair(f31(Z::'a,Xf,Xg),f30(Z::'a,Xf,Xg)),Xg)) & (\Z X Xf W Y Xg. little_set(Z) & little_set(X) & little_set(Y) & little_set(W) & equal(Z::'a,ordered_pair(X::'a,Y)) & member(ordered_pair(X::'a,W),Xf) & member(ordered_pair(W::'a,Y),Xg) --> member(Z::'a,composition(Xf::'a,Xg))) & (\Xh Xs2 Xf2 Xs1 Xf1. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> closed(Xs1::'a,Xf1)) & (\Xh Xs1 Xf1 Xs2 Xf2. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> closed(Xs2::'a,Xf2)) & (\Xf1 Xf2 Xh Xs1 Xs2. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> maps(Xh::'a,Xs1,Xs2)) & (\Xs2 Xs1 Xf1 Xf2 X Xh Y. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) & member(X::'a,Xs1) & member(Y::'a,Xs1) --> equal(apply(Xh::'a,apply_to_two_arguments(Xf1::'a,X,Y)),apply_to_two_arguments(Xf2::'a,apply(Xh::'a,X),apply(Xh::'a,Y)))) & (\Xh Xf1 Xs2 Xf2 Xs1. closed(Xs1::'a,Xf1) & closed(Xs2::'a,Xf2) & maps(Xh::'a,Xs1,Xs2) --> homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) | member(f32(Xh::'a,Xs1,Xf1,Xs2,Xf2),Xs1)) & (\Xh Xf1 Xs2 Xf2 Xs1. closed(Xs1::'a,Xf1) & closed(Xs2::'a,Xf2) & maps(Xh::'a,Xs1,Xs2) --> homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) | member(f33(Xh::'a,Xs1,Xf1,Xs2,Xf2),Xs1)) & (\Xh Xs1 Xf1 Xs2 Xf2. closed(Xs1::'a,Xf1) & closed(Xs2::'a,Xf2) & maps(Xh::'a,Xs1,Xs2) & equal(apply(Xh::'a,apply_to_two_arguments(Xf1::'a,f32(Xh::'a,Xs1,Xf1,Xs2,Xf2),f33(Xh::'a,Xs1,Xf1,Xs2,Xf2))),apply_to_two_arguments(Xf2::'a,apply(Xh::'a,f32(Xh::'a,Xs1,Xf1,Xs2,Xf2)),apply(Xh::'a,f33(Xh::'a,Xs1,Xf1,Xs2,Xf2)))) --> homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2)) & (\A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E))) & (\A2 B2. equal(A2::'a,B2) --> equal(f2(A2),f2(B2))) & (\G4 H4. equal(G4::'a,H4) --> equal(f3(G4),f3(H4))) & (\O7 P7 Q7. equal(O7::'a,P7) --> equal(f4(O7::'a,Q7),f4(P7::'a,Q7))) & (\R7 T7 S7. equal(R7::'a,S7) --> equal(f4(T7::'a,R7),f4(T7::'a,S7))) & (\U7 V7 W7. equal(U7::'a,V7) --> equal(f5(U7::'a,W7),f5(V7::'a,W7))) & (\X7 Z7 Y7. equal(X7::'a,Y7) --> equal(f5(Z7::'a,X7),f5(Z7::'a,Y7))) & (\A8 B8 C8. equal(A8::'a,B8) --> equal(f6(A8::'a,C8),f6(B8::'a,C8))) & (\D8 F8 E8. equal(D8::'a,E8) --> equal(f6(F8::'a,D8),f6(F8::'a,E8))) & (\G8 H8 I8. equal(G8::'a,H8) --> equal(f7(G8::'a,I8),f7(H8::'a,I8))) & (\J8 L8 K8. equal(J8::'a,K8) --> equal(f7(L8::'a,J8),f7(L8::'a,K8))) & (\M8 N8 O8. equal(M8::'a,N8) --> equal(f8(M8::'a,O8),f8(N8::'a,O8))) & (\P8 R8 Q8. equal(P8::'a,Q8) --> equal(f8(R8::'a,P8),f8(R8::'a,Q8))) & (\S8 T8 U8. equal(S8::'a,T8) --> equal(f9(S8::'a,U8),f9(T8::'a,U8))) & (\V8 X8 W8. equal(V8::'a,W8) --> equal(f9(X8::'a,V8),f9(X8::'a,W8))) & (\G H I'. equal(G::'a,H) --> equal(f10(G::'a,I'),f10(H::'a,I'))) & (\J L K'. equal(J::'a,K') --> equal(f10(L::'a,J),f10(L::'a,K'))) & (\M N O'. equal(M::'a,N) --> equal(f11(M::'a,O'),f11(N::'a,O'))) & (\P R Q. equal(P::'a,Q) --> equal(f11(R::'a,P),f11(R::'a,Q))) & (\S' T' U. equal(S'::'a,T') --> equal(f12(S'::'a,U),f12(T'::'a,U))) & (\V X W. equal(V::'a,W) --> equal(f12(X::'a,V),f12(X::'a,W))) & (\Y Z A1. equal(Y::'a,Z) --> equal(f13(Y::'a,A1),f13(Z::'a,A1))) & (\B1 D1 C1. equal(B1::'a,C1) --> equal(f13(D1::'a,B1),f13(D1::'a,C1))) & (\E1 F1 G1. equal(E1::'a,F1) --> equal(f14(E1::'a,G1),f14(F1::'a,G1))) & (\H1 J1 I1. equal(H1::'a,I1) --> equal(f14(J1::'a,H1),f14(J1::'a,I1))) & (\K1 L1 M1. equal(K1::'a,L1) --> equal(f16(K1::'a,M1),f16(L1::'a,M1))) & (\N1 P1 O1. equal(N1::'a,O1) --> equal(f16(P1::'a,N1),f16(P1::'a,O1))) & (\Q1 R1 S1. equal(Q1::'a,R1) --> equal(f17(Q1::'a,S1),f17(R1::'a,S1))) & (\T1 V1 U1. equal(T1::'a,U1) --> equal(f17(V1::'a,T1),f17(V1::'a,U1))) & (\W1 X1. equal(W1::'a,X1) --> equal(f18(W1),f18(X1))) & (\Y1 Z1. equal(Y1::'a,Z1) --> equal(f19(Y1),f19(Z1))) & (\C2 D2. equal(C2::'a,D2) --> equal(f20(C2),f20(D2))) & (\E2 F2. equal(E2::'a,F2) --> equal(f21(E2),f21(F2))) & (\G2 H2 I2 J2. equal(G2::'a,H2) --> equal(f22(G2::'a,I2,J2),f22(H2::'a,I2,J2))) & (\K2 M2 L2 N2. equal(K2::'a,L2) --> equal(f22(M2::'a,K2,N2),f22(M2::'a,L2,N2))) & (\O2 Q2 R2 P2. equal(O2::'a,P2) --> equal(f22(Q2::'a,R2,O2),f22(Q2::'a,R2,P2))) & (\S2 T2 U2. equal(S2::'a,T2) --> equal(f23(S2::'a,U2),f23(T2::'a,U2))) & (\V2 X2 W2. equal(V2::'a,W2) --> equal(f23(X2::'a,V2),f23(X2::'a,W2))) & (\Y2 Z2. equal(Y2::'a,Z2) --> equal(f24(Y2),f24(Z2))) & (\A3 B3. equal(A3::'a,B3) --> equal(f26(A3),f26(B3))) & (\C3 D3 E3. equal(C3::'a,D3) --> equal(f27(C3::'a,E3),f27(D3::'a,E3))) & (\F3 H3 G3. equal(F3::'a,G3) --> equal(f27(H3::'a,F3),f27(H3::'a,G3))) & (\I3 J3 K3 L3. equal(I3::'a,J3) --> equal(f28(I3::'a,K3,L3),f28(J3::'a,K3,L3))) & (\M3 O3 N3 P3. equal(M3::'a,N3) --> equal(f28(O3::'a,M3,P3),f28(O3::'a,N3,P3))) & (\Q3 S3 T3 R3. equal(Q3::'a,R3) --> equal(f28(S3::'a,T3,Q3),f28(S3::'a,T3,R3))) & (\U3 V3 W3 X3. equal(U3::'a,V3) --> equal(f29(U3::'a,W3,X3),f29(V3::'a,W3,X3))) & (\Y3 A4 Z3 B4. equal(Y3::'a,Z3) --> equal(f29(A4::'a,Y3,B4),f29(A4::'a,Z3,B4))) & (\C4 E4 F4 D4. equal(C4::'a,D4) --> equal(f29(E4::'a,F4,C4),f29(E4::'a,F4,D4))) & (\I4 J4 K4 L4. equal(I4::'a,J4) --> equal(f30(I4::'a,K4,L4),f30(J4::'a,K4,L4))) & (\M4 O4 N4 P4. equal(M4::'a,N4) --> equal(f30(O4::'a,M4,P4),f30(O4::'a,N4,P4))) & (\Q4 S4 T4 R4. equal(Q4::'a,R4) --> equal(f30(S4::'a,T4,Q4),f30(S4::'a,T4,R4))) & (\U4 V4 W4 X4. equal(U4::'a,V4) --> equal(f31(U4::'a,W4,X4),f31(V4::'a,W4,X4))) & (\Y4 A5 Z4 B5. equal(Y4::'a,Z4) --> equal(f31(A5::'a,Y4,B5),f31(A5::'a,Z4,B5))) & (\C5 E5 F5 D5. equal(C5::'a,D5) --> equal(f31(E5::'a,F5,C5),f31(E5::'a,F5,D5))) & (\G5 H5 I5 J5 K5 L5. equal(G5::'a,H5) --> equal(f32(G5::'a,I5,J5,K5,L5),f32(H5::'a,I5,J5,K5,L5))) & (\M5 O5 N5 P5 Q5 R5. equal(M5::'a,N5) --> equal(f32(O5::'a,M5,P5,Q5,R5),f32(O5::'a,N5,P5,Q5,R5))) & (\S5 U5 V5 T5 W5 X5. equal(S5::'a,T5) --> equal(f32(U5::'a,V5,S5,W5,X5),f32(U5::'a,V5,T5,W5,X5))) & (\Y5 A6 B6 C6 Z5 D6. equal(Y5::'a,Z5) --> equal(f32(A6::'a,B6,C6,Y5,D6),f32(A6::'a,B6,C6,Z5,D6))) & (\E6 G6 H6 I6 J6 F6. equal(E6::'a,F6) --> equal(f32(G6::'a,H6,I6,J6,E6),f32(G6::'a,H6,I6,J6,F6))) & (\K6 L6 M6 N6 O6 P6. equal(K6::'a,L6) --> equal(f33(K6::'a,M6,N6,O6,P6),f33(L6::'a,M6,N6,O6,P6))) & (\Q6 S6 R6 T6 U6 V6. equal(Q6::'a,R6) --> equal(f33(S6::'a,Q6,T6,U6,V6),f33(S6::'a,R6,T6,U6,V6))) & (\W6 Y6 Z6 X6 A7 B7. equal(W6::'a,X6) --> equal(f33(Y6::'a,Z6,W6,A7,B7),f33(Y6::'a,Z6,X6,A7,B7))) & (\C7 E7 F7 G7 D7 H7. equal(C7::'a,D7) --> equal(f33(E7::'a,F7,G7,C7,H7),f33(E7::'a,F7,G7,D7,H7))) & (\I7 K7 L7 M7 N7 J7. equal(I7::'a,J7) --> equal(f33(K7::'a,L7,M7,N7,I7),f33(K7::'a,L7,M7,N7,J7))) & (\A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) & (\D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) & (\G H I' J. equal(G::'a,H) --> equal(apply_to_two_arguments(G::'a,I',J),apply_to_two_arguments(H::'a,I',J))) & (\K' M L N. equal(K'::'a,L) --> equal(apply_to_two_arguments(M::'a,K',N),apply_to_two_arguments(M::'a,L,N))) & (\O' Q R P. equal(O'::'a,P) --> equal(apply_to_two_arguments(Q::'a,R,O'),apply_to_two_arguments(Q::'a,R,P))) & (\S' T'. equal(S'::'a,T') --> equal(complement(S'),complement(T'))) & (\U V W. equal(U::'a,V) --> equal(composition(U::'a,W),composition(V::'a,W))) & (\X Z Y. equal(X::'a,Y) --> equal(composition(Z::'a,X),composition(Z::'a,Y))) & (\A1 B1. equal(A1::'a,B1) --> equal(inv1 A1,inv1 B1)) & (\C1 D1 E1. equal(C1::'a,D1) --> equal(cross_product(C1::'a,E1),cross_product(D1::'a,E1))) & (\F1 H1 G1. equal(F1::'a,G1) --> equal(cross_product(H1::'a,F1),cross_product(H1::'a,G1))) & (\I1 J1. equal(I1::'a,J1) --> equal(domain_of(I1),domain_of(J1))) & (\I10 J10. equal(I10::'a,J10) --> equal(first(I10),first(J10))) & (\Q10 R10. equal(Q10::'a,R10) --> equal(flip_range_of(Q10),flip_range_of(R10))) & (\S10 T10 U10. equal(S10::'a,T10) --> equal(image'(S10::'a,U10),image'(T10::'a,U10))) & (\V10 X10 W10. equal(V10::'a,W10) --> equal(image'(X10::'a,V10),image'(X10::'a,W10))) & (\Y10 Z10 A11. equal(Y10::'a,Z10) --> equal(intersection(Y10::'a,A11),intersection(Z10::'a,A11))) & (\B11 D11 C11. equal(B11::'a,C11) --> equal(intersection(D11::'a,B11),intersection(D11::'a,C11))) & (\E11 F11 G11. equal(E11::'a,F11) --> equal(non_ordered_pair(E11::'a,G11),non_ordered_pair(F11::'a,G11))) & (\H11 J11 I11. equal(H11::'a,I11) --> equal(non_ordered_pair(J11::'a,H11),non_ordered_pair(J11::'a,I11))) & (\K11 L11 M11. equal(K11::'a,L11) --> equal(ordered_pair(K11::'a,M11),ordered_pair(L11::'a,M11))) & (\N11 P11 O11. equal(N11::'a,O11) --> equal(ordered_pair(P11::'a,N11),ordered_pair(P11::'a,O11))) & (\Q11 R11. equal(Q11::'a,R11) --> equal(powerset(Q11),powerset(R11))) & (\S11 T11. equal(S11::'a,T11) --> equal(range_of(S11),range_of(T11))) & (\U11 V11 W11. equal(U11::'a,V11) --> equal(restrct(U11::'a,W11),restrct(V11::'a,W11))) & (\X11 Z11 Y11. equal(X11::'a,Y11) --> equal(restrct(Z11::'a,X11),restrct(Z11::'a,Y11))) & (\A12 B12. equal(A12::'a,B12) --> equal(rot_right(A12),rot_right(B12))) & (\C12 D12. equal(C12::'a,D12) --> equal(second(C12),second(D12))) & (\K12 L12. equal(K12::'a,L12) --> equal(sigma(K12),sigma(L12))) & (\M12 N12. equal(M12::'a,N12) --> equal(singleton_set(M12),singleton_set(N12))) & (\O12 P12. equal(O12::'a,P12) --> equal(successor(O12),successor(P12))) & (\Q12 R12 S12. equal(Q12::'a,R12) --> equal(union(Q12::'a,S12),union(R12::'a,S12))) & (\T12 V12 U12. equal(T12::'a,U12) --> equal(union(V12::'a,T12),union(V12::'a,U12))) & (\W12 X12 Y12. equal(W12::'a,X12) & closed(W12::'a,Y12) --> closed(X12::'a,Y12)) & (\Z12 B13 A13. equal(Z12::'a,A13) & closed(B13::'a,Z12) --> closed(B13::'a,A13)) & (\C13 D13 E13. equal(C13::'a,D13) & disjoint(C13::'a,E13) --> disjoint(D13::'a,E13)) & (\F13 H13 G13. equal(F13::'a,G13) & disjoint(H13::'a,F13) --> disjoint(H13::'a,G13)) & (\I13 J13. equal(I13::'a,J13) & function(I13) --> function(J13)) & (\K13 L13 M13 N13 O13 P13. equal(K13::'a,L13) & homomorphism(K13::'a,M13,N13,O13,P13) --> homomorphism(L13::'a,M13,N13,O13,P13)) & (\Q13 S13 R13 T13 U13 V13. equal(Q13::'a,R13) & homomorphism(S13::'a,Q13,T13,U13,V13) --> homomorphism(S13::'a,R13,T13,U13,V13)) & (\W13 Y13 Z13 X13 A14 B14. equal(W13::'a,X13) & homomorphism(Y13::'a,Z13,W13,A14,B14) --> homomorphism(Y13::'a,Z13,X13,A14,B14)) & (\C14 E14 F14 G14 D14 H14. equal(C14::'a,D14) & homomorphism(E14::'a,F14,G14,C14,H14) --> homomorphism(E14::'a,F14,G14,D14,H14)) & (\I14 K14 L14 M14 N14 J14. equal(I14::'a,J14) & homomorphism(K14::'a,L14,M14,N14,I14) --> homomorphism(K14::'a,L14,M14,N14,J14)) & (\O14 P14. equal(O14::'a,P14) & little_set(O14) --> little_set(P14)) & (\Q14 R14 S14 T14. equal(Q14::'a,R14) & maps(Q14::'a,S14,T14) --> maps(R14::'a,S14,T14)) & (\U14 W14 V14 X14. equal(U14::'a,V14) & maps(W14::'a,U14,X14) --> maps(W14::'a,V14,X14)) & (\Y14 A15 B15 Z14. equal(Y14::'a,Z14) & maps(A15::'a,B15,Y14) --> maps(A15::'a,B15,Z14)) & (\C15 D15 E15. equal(C15::'a,D15) & member(C15::'a,E15) --> member(D15::'a,E15)) & (\F15 H15 G15. equal(F15::'a,G15) & member(H15::'a,F15) --> member(H15::'a,G15)) & (\I15 J15. equal(I15::'a,J15) & one_to_one_function(I15) --> one_to_one_function(J15)) & (\K15 L15. equal(K15::'a,L15) & ordered_pair_predicate(K15) --> ordered_pair_predicate(L15)) & (\M15 N15 O15. equal(M15::'a,N15) & proper_subset(M15::'a,O15) --> proper_subset(N15::'a,O15)) & (\P15 R15 Q15. equal(P15::'a,Q15) & proper_subset(R15::'a,P15) --> proper_subset(R15::'a,Q15)) & (\S15 T15. equal(S15::'a,T15) & relation(S15) --> relation(T15)) & (\U15 V15. equal(U15::'a,V15) & single_valued_set(U15) --> single_valued_set(V15)) & (\W15 X15 Y15. equal(W15::'a,X15) & ssubset(W15::'a,Y15) --> ssubset(X15::'a,Y15)) & (\Z15 B16 A16. equal(Z15::'a,A16) & ssubset(B16::'a,Z15) --> ssubset(B16::'a,A16)) & (~little_set(ordered_pair(a::'a,b))) --> False" oops (*13 inferences so far. Searching to depth 8. 0 secs*) lemma SET046_5: "(\Y X. ~(element(X::'a,a) & element(X::'a,Y) & element(Y::'a,X))) & (\X. element(X::'a,f(X)) | element(X::'a,a)) & (\X. element(f(X),X) | element(X::'a,a)) --> False" by meson (*33 inferences so far. Searching to depth 9. 0.2 secs*) lemma SET047_5: "(\X Z Y. set_equal(X::'a,Y) & element(Z::'a,X) --> element(Z::'a,Y)) & (\Y Z X. set_equal(X::'a,Y) & element(Z::'a,Y) --> element(Z::'a,X)) & (\X Y. element(f(X::'a,Y),X) | element(f(X::'a,Y),Y) | set_equal(X::'a,Y)) & (\X Y. element(f(X::'a,Y),Y) & element(f(X::'a,Y),X) --> set_equal(X::'a,Y)) & (set_equal(a::'a,b) | set_equal(b::'a,a)) & (~(set_equal(b::'a,a) & set_equal(a::'a,b))) --> False" by meson (*311 inferences so far. Searching to depth 12. 0.1 secs*) lemma SYN034_1: "(\A. p(A::'a,a) | p(A::'a,f(A))) & (\A. p(A::'a,a) | p(f(A),A)) & (\A B. ~(p(A::'a,B) & p(B::'a,A) & p(B::'a,a))) --> False" by meson (*30 inferences so far. Searching to depth 6. 0.2 secs*) lemma SYN071_1: "EQU001_0_ax equal & (equal(a::'a,b) | equal(c::'a,d)) & (equal(a::'a,c) | equal(b::'a,d)) & (~equal(a::'a,d)) & (~equal(b::'a,c)) --> False" by meson (*1897410 inferences so far. Searching to depth 48 206s, nearly 4 mins on griffon.*) lemma SYN349_1: "(\X Y. f(w(X),g(X::'a,Y)) --> f(X::'a,g(X::'a,Y))) & (\X Y. f(X::'a,g(X::'a,Y)) --> f(w(X),g(X::'a,Y))) & (\Y X. f(X::'a,g(X::'a,Y)) & f(Y::'a,g(X::'a,Y)) --> f(g(X::'a,Y),Y) | f(g(X::'a,Y),w(X))) & (\Y X. f(g(X::'a,Y),Y) & f(Y::'a,g(X::'a,Y)) --> f(X::'a,g(X::'a,Y)) | f(g(X::'a,Y),w(X))) & (\Y X. f(X::'a,g(X::'a,Y)) | f(g(X::'a,Y),Y) | f(Y::'a,g(X::'a,Y)) | f(g(X::'a,Y),w(X))) & (\Y X. f(X::'a,g(X::'a,Y)) & f(g(X::'a,Y),Y) --> f(Y::'a,g(X::'a,Y)) | f(g(X::'a,Y),w(X))) & (\Y X. f(X::'a,g(X::'a,Y)) & f(g(X::'a,Y),w(X)) --> f(g(X::'a,Y),Y) | f(Y::'a,g(X::'a,Y))) & (\Y X. f(g(X::'a,Y),Y) & f(g(X::'a,Y),w(X)) --> f(X::'a,g(X::'a,Y)) | f(Y::'a,g(X::'a,Y))) & (\Y X. f(Y::'a,g(X::'a,Y)) & f(g(X::'a,Y),w(X)) --> f(X::'a,g(X::'a,Y)) | f(g(X::'a,Y),Y)) & (\Y X. ~(f(X::'a,g(X::'a,Y)) & f(g(X::'a,Y),Y) & f(Y::'a,g(X::'a,Y)) & f(g(X::'a,Y),w(X)))) --> False" oops (*398 inferences so far. Searching to depth 12. 0.4 secs*) lemma SYN352_1: "(f(a::'a,b)) & (\X Y. f(X::'a,Y) --> f(b::'a,z(X::'a,Y)) | f(Y::'a,z(X::'a,Y))) & (\X Y. f(X::'a,Y) | f(z(X::'a,Y),z(X::'a,Y))) & (\X Y. f(b::'a,z(X::'a,Y)) | f(X::'a,z(X::'a,Y)) | f(z(X::'a,Y),z(X::'a,Y))) & (\X Y. f(b::'a,z(X::'a,Y)) & f(X::'a,z(X::'a,Y)) --> f(z(X::'a,Y),z(X::'a,Y))) & (\X Y. ~(f(X::'a,Y) & f(X::'a,z(X::'a,Y)) & f(Y::'a,z(X::'a,Y)))) & (\X Y. f(X::'a,Y) --> f(X::'a,z(X::'a,Y)) | f(Y::'a,z(X::'a,Y))) --> False" by meson (*5336 inferences so far. Searching to depth 15. 5.3 secs*) lemma TOP001_2: "(\Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) & (\U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) & (\U Uu1 Vf. element_of_set(U::'a,Uu1) & element_of_collection(Uu1::'a,Vf) --> element_of_set(U::'a,union_of_members(Vf))) & (\Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) & (\Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) & (\U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) & (\X. subset_sets(X::'a,X)) & (\X U Y. subset_sets(X::'a,Y) & element_of_set(U::'a,X) --> element_of_set(U::'a,Y)) & (\X Y. equal_sets(X::'a,Y) --> subset_sets(X::'a,Y)) & (\Y X. subset_sets(X::'a,Y) | element_of_set(in_1st_set(X::'a,Y),X)) & (\X Y. element_of_set(in_1st_set(X::'a,Y),Y) --> subset_sets(X::'a,Y)) & (basis(cx::'a,f)) & (~subset_sets(union_of_members(top_of_basis(f)),cx)) --> False" by meson (*0 inferences so far. Searching to depth 0. 0 secs*) lemma TOP002_2: "(\Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & (\X. ~element_of_set(X::'a,empty_set)) & (~element_of_collection(empty_set::'a,top_of_basis(f))) --> False" by meson (*0 inferences so far. Searching to depth 0. 6.5 secs. BIG*) lemma TOP004_1: "(\Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) & (\U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) & (\U Uu1 Vf. element_of_set(U::'a,Uu1) & element_of_collection(Uu1::'a,Vf) --> element_of_set(U::'a,union_of_members(Vf))) & (\Vf U Va. element_of_set(U::'a,intersection_of_members(Vf)) & element_of_collection(Va::'a,Vf) --> element_of_set(U::'a,Va)) & (\U Vf. element_of_set(U::'a,intersection_of_members(Vf)) | element_of_collection(f2(Vf::'a,U),Vf)) & (\Vf U. element_of_set(U::'a,f2(Vf::'a,U)) --> element_of_set(U::'a,intersection_of_members(Vf))) & (\Vt X. topological_space(X::'a,Vt) --> equal_sets(union_of_members(Vt),X)) & (\X Vt. topological_space(X::'a,Vt) --> element_of_collection(empty_set::'a,Vt)) & (\X Vt. topological_space(X::'a,Vt) --> element_of_collection(X::'a,Vt)) & (\X Y Z Vt. topological_space(X::'a,Vt) & element_of_collection(Y::'a,Vt) & element_of_collection(Z::'a,Vt) --> element_of_collection(intersection_of_sets(Y::'a,Z),Vt)) & (\X Vf Vt. topological_space(X::'a,Vt) & subset_collections(Vf::'a,Vt) --> element_of_collection(union_of_members(Vf),Vt)) & (\X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) --> topological_space(X::'a,Vt) | element_of_collection(f3(X::'a,Vt),Vt) | subset_collections(f5(X::'a,Vt),Vt)) & (\X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(union_of_members(f5(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt) | element_of_collection(f3(X::'a,Vt),Vt)) & (\X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) --> topological_space(X::'a,Vt) | element_of_collection(f4(X::'a,Vt),Vt) | subset_collections(f5(X::'a,Vt),Vt)) & (\X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(union_of_members(f5(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt) | element_of_collection(f4(X::'a,Vt),Vt)) & (\X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(intersection_of_sets(f3(X::'a,Vt),f4(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt) | subset_collections(f5(X::'a,Vt),Vt)) & (\X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(intersection_of_sets(f3(X::'a,Vt),f4(X::'a,Vt)),Vt) & element_of_collection(union_of_members(f5(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt)) & (\U X Vt. open(U::'a,X,Vt) --> topological_space(X::'a,Vt)) & (\X U Vt. open(U::'a,X,Vt) --> element_of_collection(U::'a,Vt)) & (\X U Vt. topological_space(X::'a,Vt) & element_of_collection(U::'a,Vt) --> open(U::'a,X,Vt)) & (\U X Vt. closed(U::'a,X,Vt) --> topological_space(X::'a,Vt)) & (\U X Vt. closed(U::'a,X,Vt) --> open(relative_complement_sets(U::'a,X),X,Vt)) & (\U X Vt. topological_space(X::'a,Vt) & open(relative_complement_sets(U::'a,X),X,Vt) --> closed(U::'a,X,Vt)) & (\Vs X Vt. finer(Vt::'a,Vs,X) --> topological_space(X::'a,Vt)) & (\Vt X Vs. finer(Vt::'a,Vs,X) --> topological_space(X::'a,Vs)) & (\X Vs Vt. finer(Vt::'a,Vs,X) --> subset_collections(Vs::'a,Vt)) & (\X Vs Vt. topological_space(X::'a,Vt) & topological_space(X::'a,Vs) & subset_collections(Vs::'a,Vt) --> finer(Vt::'a,Vs,X)) & (\Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) & (\X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_set(Y::'a,f6(X::'a,Vf,Y,Vb1,Vb2))) & (\X Y Vb1 Vb2 Vf. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_collection(f6(X::'a,Vf,Y,Vb1,Vb2),Vf)) & (\X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> subset_sets(f6(X::'a,Vf,Y,Vb1,Vb2),intersection_of_sets(Vb1::'a,Vb2))) & (\Vf X. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_set(f7(X::'a,Vf),X)) & (\X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_collection(f8(X::'a,Vf),Vf)) & (\X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_collection(f9(X::'a,Vf),Vf)) & (\X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_set(f7(X::'a,Vf),intersection_of_sets(f8(X::'a,Vf),f9(X::'a,Vf)))) & (\Uu9 X Vf. equal_sets(union_of_members(Vf),X) & element_of_set(f7(X::'a,Vf),Uu9) & element_of_collection(Uu9::'a,Vf) & subset_sets(Uu9::'a,intersection_of_sets(f8(X::'a,Vf),f9(X::'a,Vf))) --> basis(X::'a,Vf)) & (\Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) & (\U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) & (\Vf X U. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> subset_sets(f10(Vf::'a,U,X),U)) & (\Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & (\Vf Uu11 U. element_of_set(f11(Vf::'a,U),Uu11) & element_of_collection(Uu11::'a,Vf) & subset_sets(Uu11::'a,U) --> element_of_collection(U::'a,top_of_basis(Vf))) & (\U Y X Vt. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> topological_space(X::'a,Vt)) & (\U Vt Y X. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> subset_sets(Y::'a,X)) & (\X Y U Vt. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> element_of_collection(f12(X::'a,Vt,Y,U),Vt)) & (\X Vt Y U. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> equal_sets(U::'a,intersection_of_sets(Y::'a,f12(X::'a,Vt,Y,U)))) & (\X Vt U Y Uu12. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_collection(Uu12::'a,Vt) & equal_sets(U::'a,intersection_of_sets(Y::'a,Uu12)) --> element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y))) & (\U Y X Vt. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) & (\U Vt Y X. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> subset_sets(Y::'a,X)) & (\Y X Vt U. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> element_of_set(U::'a,f13(Y::'a,X,Vt,U))) & (\X Vt U Y. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> subset_sets(f13(Y::'a,X,Vt,U),Y)) & (\Y U X Vt. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> open(f13(Y::'a,X,Vt,U),X,Vt)) & (\U Y Uu13 X Vt. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_set(U::'a,Uu13) & subset_sets(Uu13::'a,Y) & open(Uu13::'a,X,Vt) --> element_of_set(U::'a,interior(Y::'a,X,Vt))) & (\U Y X Vt. element_of_set(U::'a,closure(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) & (\U Vt Y X. element_of_set(U::'a,closure(Y::'a,X,Vt)) --> subset_sets(Y::'a,X)) & (\Y X Vt U V. element_of_set(U::'a,closure(Y::'a,X,Vt)) & subset_sets(Y::'a,V) & closed(V::'a,X,Vt) --> element_of_set(U::'a,V)) & (\Y X Vt U. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) --> element_of_set(U::'a,closure(Y::'a,X,Vt)) | subset_sets(Y::'a,f14(Y::'a,X,Vt,U))) & (\Y U X Vt. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) --> element_of_set(U::'a,closure(Y::'a,X,Vt)) | closed(f14(Y::'a,X,Vt,U),X,Vt)) & (\Y X Vt U. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_set(U::'a,f14(Y::'a,X,Vt,U)) --> element_of_set(U::'a,closure(Y::'a,X,Vt))) & (\U Y X Vt. neighborhood(U::'a,Y,X,Vt) --> topological_space(X::'a,Vt)) & (\Y U X Vt. neighborhood(U::'a,Y,X,Vt) --> open(U::'a,X,Vt)) & (\X Vt Y U. neighborhood(U::'a,Y,X,Vt) --> element_of_set(Y::'a,U)) & (\X Vt Y U. topological_space(X::'a,Vt) & open(U::'a,X,Vt) & element_of_set(Y::'a,U) --> neighborhood(U::'a,Y,X,Vt)) & (\Z Y X Vt. limit_point(Z::'a,Y,X,Vt) --> topological_space(X::'a,Vt)) & (\Z Vt Y X. limit_point(Z::'a,Y,X,Vt) --> subset_sets(Y::'a,X)) & (\Z X Vt U Y. limit_point(Z::'a,Y,X,Vt) & neighborhood(U::'a,Z,X,Vt) --> element_of_set(f15(Z::'a,Y,X,Vt,U),intersection_of_sets(U::'a,Y))) & (\Y X Vt U Z. ~(limit_point(Z::'a,Y,X,Vt) & neighborhood(U::'a,Z,X,Vt) & eq_p(f15(Z::'a,Y,X,Vt,U),Z))) & (\Y Z X Vt. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) --> limit_point(Z::'a,Y,X,Vt) | neighborhood(f16(Z::'a,Y,X,Vt),Z,X,Vt)) & (\X Vt Y Uu16 Z. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_set(Uu16::'a,intersection_of_sets(f16(Z::'a,Y,X,Vt),Y)) --> limit_point(Z::'a,Y,X,Vt) | eq_p(Uu16::'a,Z)) & (\U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) & (\U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> element_of_set(U::'a,closure(Y::'a,X,Vt))) & (\U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> element_of_set(U::'a,closure(relative_complement_sets(Y::'a,X),X,Vt))) & (\U Y X Vt. topological_space(X::'a,Vt) & element_of_set(U::'a,closure(Y::'a,X,Vt)) & element_of_set(U::'a,closure(relative_complement_sets(Y::'a,X),X,Vt)) --> element_of_set(U::'a,boundary(Y::'a,X,Vt))) & (\X Vt. hausdorff(X::'a,Vt) --> topological_space(X::'a,Vt)) & (\X_2 X_1 X Vt. hausdorff(X::'a,Vt) & element_of_set(X_1::'a,X) & element_of_set(X_2::'a,X) --> eq_p(X_1::'a,X_2) | neighborhood(f17(X::'a,Vt,X_1,X_2),X_1,X,Vt)) & (\X_1 X_2 X Vt. hausdorff(X::'a,Vt) & element_of_set(X_1::'a,X) & element_of_set(X_2::'a,X) --> eq_p(X_1::'a,X_2) | neighborhood(f18(X::'a,Vt,X_1,X_2),X_2,X,Vt)) & (\X Vt X_1 X_2. hausdorff(X::'a,Vt) & element_of_set(X_1::'a,X) & element_of_set(X_2::'a,X) --> eq_p(X_1::'a,X_2) | disjoint_s(f17(X::'a,Vt,X_1,X_2),f18(X::'a,Vt,X_1,X_2))) & (\Vt X. topological_space(X::'a,Vt) --> hausdorff(X::'a,Vt) | element_of_set(f19(X::'a,Vt),X)) & (\Vt X. topological_space(X::'a,Vt) --> hausdorff(X::'a,Vt) | element_of_set(f20(X::'a,Vt),X)) & (\X Vt. topological_space(X::'a,Vt) & eq_p(f19(X::'a,Vt),f20(X::'a,Vt)) --> hausdorff(X::'a,Vt)) & (\X Vt Uu19 Uu20. topological_space(X::'a,Vt) & neighborhood(Uu19::'a,f19(X::'a,Vt),X,Vt) & neighborhood(Uu20::'a,f20(X::'a,Vt),X,Vt) & disjoint_s(Uu19::'a,Uu20) --> hausdorff(X::'a,Vt)) & (\Va1 Va2 X Vt. separation(Va1::'a,Va2,X,Vt) --> topological_space(X::'a,Vt)) & (\Va2 X Vt Va1. ~(separation(Va1::'a,Va2,X,Vt) & equal_sets(Va1::'a,empty_set))) & (\Va1 X Vt Va2. ~(separation(Va1::'a,Va2,X,Vt) & equal_sets(Va2::'a,empty_set))) & (\Va2 X Va1 Vt. separation(Va1::'a,Va2,X,Vt) --> element_of_collection(Va1::'a,Vt)) & (\Va1 X Va2 Vt. separation(Va1::'a,Va2,X,Vt) --> element_of_collection(Va2::'a,Vt)) & (\Vt Va1 Va2 X. separation(Va1::'a,Va2,X,Vt) --> equal_sets(union_of_sets(Va1::'a,Va2),X)) & (\X Vt Va1 Va2. separation(Va1::'a,Va2,X,Vt) --> disjoint_s(Va1::'a,Va2)) & (\Vt X Va1 Va2. topological_space(X::'a,Vt) & element_of_collection(Va1::'a,Vt) & element_of_collection(Va2::'a,Vt) & equal_sets(union_of_sets(Va1::'a,Va2),X) & disjoint_s(Va1::'a,Va2) --> separation(Va1::'a,Va2,X,Vt) | equal_sets(Va1::'a,empty_set) | equal_sets(Va2::'a,empty_set)) & (\X Vt. connected_space(X::'a,Vt) --> topological_space(X::'a,Vt)) & (\Va1 Va2 X Vt. ~(connected_space(X::'a,Vt) & separation(Va1::'a,Va2,X,Vt))) & (\X Vt. topological_space(X::'a,Vt) --> connected_space(X::'a,Vt) | separation(f21(X::'a,Vt),f22(X::'a,Vt),X,Vt)) & (\Va X Vt. connected_set(Va::'a,X,Vt) --> topological_space(X::'a,Vt)) & (\Vt Va X. connected_set(Va::'a,X,Vt) --> subset_sets(Va::'a,X)) & (\X Vt Va. connected_set(Va::'a,X,Vt) --> connected_space(Va::'a,subspace_topology(X::'a,Vt,Va))) & (\X Vt Va. topological_space(X::'a,Vt) & subset_sets(Va::'a,X) & connected_space(Va::'a,subspace_topology(X::'a,Vt,Va)) --> connected_set(Va::'a,X,Vt)) & (\Vf X Vt. open_covering(Vf::'a,X,Vt) --> topological_space(X::'a,Vt)) & (\X Vf Vt. open_covering(Vf::'a,X,Vt) --> subset_collections(Vf::'a,Vt)) & (\Vt Vf X. open_covering(Vf::'a,X,Vt) --> equal_sets(union_of_members(Vf),X)) & (\Vt Vf X. topological_space(X::'a,Vt) & subset_collections(Vf::'a,Vt) & equal_sets(union_of_members(Vf),X) --> open_covering(Vf::'a,X,Vt)) & (\X Vt. compact_space(X::'a,Vt) --> topological_space(X::'a,Vt)) & (\X Vt Vf1. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> finite'(f23(X::'a,Vt,Vf1))) & (\X Vt Vf1. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> subset_collections(f23(X::'a,Vt,Vf1),Vf1)) & (\Vf1 X Vt. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> open_covering(f23(X::'a,Vt,Vf1),X,Vt)) & (\X Vt. topological_space(X::'a,Vt) --> compact_space(X::'a,Vt) | open_covering(f24(X::'a,Vt),X,Vt)) & (\Uu24 X Vt. topological_space(X::'a,Vt) & finite'(Uu24) & subset_collections(Uu24::'a,f24(X::'a,Vt)) & open_covering(Uu24::'a,X,Vt) --> compact_space(X::'a,Vt)) & (\Va X Vt. compact_set(Va::'a,X,Vt) --> topological_space(X::'a,Vt)) & (\Vt Va X. compact_set(Va::'a,X,Vt) --> subset_sets(Va::'a,X)) & (\X Vt Va. compact_set(Va::'a,X,Vt) --> compact_space(Va::'a,subspace_topology(X::'a,Vt,Va))) & (\X Vt Va. topological_space(X::'a,Vt) & subset_sets(Va::'a,X) & compact_space(Va::'a,subspace_topology(X::'a,Vt,Va)) --> compact_set(Va::'a,X,Vt)) & (basis(cx::'a,f)) & (\U. element_of_collection(U::'a,top_of_basis(f))) & (\V. element_of_collection(V::'a,top_of_basis(f))) & (\U V. ~element_of_collection(intersection_of_sets(U::'a,V),top_of_basis(f))) --> False" by meson (*0 inferences so far. Searching to depth 0. 0.8 secs*) lemma TOP004_2: "(\U Uu1 Vf. element_of_set(U::'a,Uu1) & element_of_collection(Uu1::'a,Vf) --> element_of_set(U::'a,union_of_members(Vf))) & (\Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) & (\X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_set(Y::'a,f6(X::'a,Vf,Y,Vb1,Vb2))) & (\X Y Vb1 Vb2 Vf. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_collection(f6(X::'a,Vf,Y,Vb1,Vb2),Vf)) & (\X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> subset_sets(f6(X::'a,Vf,Y,Vb1,Vb2),intersection_of_sets(Vb1::'a,Vb2))) & (\Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) & (\U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) & (\Vf X U. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> subset_sets(f10(Vf::'a,U,X),U)) & (\Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & (\Vf Uu11 U. element_of_set(f11(Vf::'a,U),Uu11) & element_of_collection(Uu11::'a,Vf) & subset_sets(Uu11::'a,U) --> element_of_collection(U::'a,top_of_basis(Vf))) & (\Y X Z. subset_sets(X::'a,Y) & subset_sets(Y::'a,Z) --> subset_sets(X::'a,Z)) & (\Y Z X. element_of_set(Z::'a,intersection_of_sets(X::'a,Y)) --> element_of_set(Z::'a,X)) & (\X Z Y. element_of_set(Z::'a,intersection_of_sets(X::'a,Y)) --> element_of_set(Z::'a,Y)) & (\X Z Y. element_of_set(Z::'a,X) & element_of_set(Z::'a,Y) --> element_of_set(Z::'a,intersection_of_sets(X::'a,Y))) & (\X U Y V. subset_sets(X::'a,Y) & subset_sets(U::'a,V) --> subset_sets(intersection_of_sets(X::'a,U),intersection_of_sets(Y::'a,V))) & (\X Z Y. equal_sets(X::'a,Y) & element_of_set(Z::'a,X) --> element_of_set(Z::'a,Y)) & (\Y X. equal_sets(intersection_of_sets(X::'a,Y),intersection_of_sets(Y::'a,X))) & (basis(cx::'a,f)) & (\U. element_of_collection(U::'a,top_of_basis(f))) & (\V. element_of_collection(V::'a,top_of_basis(f))) & (\U V. ~element_of_collection(intersection_of_sets(U::'a,V),top_of_basis(f))) --> False" by meson (*53777 inferences so far. Searching to depth 20. 68.7 secs*) lemma TOP005_2: "(\Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) & (\U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) & (\Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) & (\U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) & (\Vf X U. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> subset_sets(f10(Vf::'a,U,X),U)) & (\Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & (\Vf Uu11 U. element_of_set(f11(Vf::'a,U),Uu11) & element_of_collection(Uu11::'a,Vf) & subset_sets(Uu11::'a,U) --> element_of_collection(U::'a,top_of_basis(Vf))) & (\X U Y. element_of_set(U::'a,X) --> subset_sets(X::'a,Y) | element_of_set(U::'a,Y)) & (\Y X Z. subset_sets(X::'a,Y) & element_of_collection(Y::'a,Z) --> subset_sets(X::'a,union_of_members(Z))) & (\X U Y. subset_collections(X::'a,Y) & element_of_collection(U::'a,X) --> element_of_collection(U::'a,Y)) & (subset_collections(g::'a,top_of_basis(f))) & (~element_of_collection(union_of_members(g),top_of_basis(f))) --> False" by meson end