diff --git a/src/HOL/SMT.thy b/src/HOL/SMT.thy --- a/src/HOL/SMT.thy +++ b/src/HOL/SMT.thy @@ -1,898 +1,898 @@ (* Title: HOL/SMT.thy Author: Sascha Boehme, TU Muenchen Author: Jasmin Blanchette, VU Amsterdam *) section \Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2\ theory SMT imports Divides Numeral_Simprocs keywords "smt_status" :: diag begin subsection \A skolemization tactic and proof method\ lemma choices: "\Q. \x. \y ya. Q x y ya \ \f fa. \x. Q x (f x) (fa x)" "\Q. \x. \y ya yb. Q x y ya yb \ \f fa fb. \x. Q x (f x) (fa x) (fb x)" "\Q. \x. \y ya yb yc. Q x y ya yb yc \ \f fa fb fc. \x. Q x (f x) (fa x) (fb x) (fc x)" "\Q. \x. \y ya yb yc yd. Q x y ya yb yc yd \ \f fa fb fc fd. \x. Q x (f x) (fa x) (fb x) (fc x) (fd x)" "\Q. \x. \y ya yb yc yd ye. Q x y ya yb yc yd ye \ \f fa fb fc fd fe. \x. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x)" "\Q. \x. \y ya yb yc yd ye yf. Q x y ya yb yc yd ye yf \ \f fa fb fc fd fe ff. \x. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x)" "\Q. \x. \y ya yb yc yd ye yf yg. Q x y ya yb yc yd ye yf yg \ \f fa fb fc fd fe ff fg. \x. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x) (fg x)" by metis+ lemma bchoices: "\Q. \x \ S. \y ya. Q x y ya \ \f fa. \x \ S. Q x (f x) (fa x)" "\Q. \x \ S. \y ya yb. Q x y ya yb \ \f fa fb. \x \ S. Q x (f x) (fa x) (fb x)" "\Q. \x \ S. \y ya yb yc. Q x y ya yb yc \ \f fa fb fc. \x \ S. Q x (f x) (fa x) (fb x) (fc x)" "\Q. \x \ S. \y ya yb yc yd. Q x y ya yb yc yd \ \f fa fb fc fd. \x \ S. Q x (f x) (fa x) (fb x) (fc x) (fd x)" "\Q. \x \ S. \y ya yb yc yd ye. Q x y ya yb yc yd ye \ \f fa fb fc fd fe. \x \ S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x)" "\Q. \x \ S. \y ya yb yc yd ye yf. Q x y ya yb yc yd ye yf \ \f fa fb fc fd fe ff. \x \ S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x)" "\Q. \x \ S. \y ya yb yc yd ye yf yg. Q x y ya yb yc yd ye yf yg \ \f fa fb fc fd fe ff fg. \x \ S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x) (fg x)" by metis+ ML \ fun moura_tac ctxt = Atomize_Elim.atomize_elim_tac ctxt THEN' SELECT_GOAL (Clasimp.auto_tac (ctxt addSIs @{thms choice choices bchoice bchoices}) THEN ALLGOALS (Metis_Tactic.metis_tac (take 1 ATP_Proof_Reconstruct.partial_type_encs) ATP_Proof_Reconstruct.default_metis_lam_trans ctxt [] ORELSE' blast_tac ctxt)) \ method_setup moura = \ Scan.succeed (SIMPLE_METHOD' o moura_tac) \ "solve skolemization goals, especially those arising from Z3 proofs" hide_fact (open) choices bchoices subsection \Triggers for quantifier instantiation\ text \ Some SMT solvers support patterns as a quantifier instantiation heuristics. Patterns may either be positive terms (tagged by "pat") triggering quantifier instantiations -- when the solver finds a term matching a positive pattern, it instantiates the corresponding quantifier accordingly -- or negative terms (tagged by "nopat") inhibiting quantifier instantiations. A list of patterns of the same kind is called a multipattern, and all patterns in a multipattern are considered conjunctively for quantifier instantiation. A list of multipatterns is called a trigger, and their multipatterns act disjunctively during quantifier instantiation. Each multipattern should mention at least all quantified variables of the preceding quantifier block. \ typedecl 'a symb_list consts Symb_Nil :: "'a symb_list" Symb_Cons :: "'a \ 'a symb_list \ 'a symb_list" typedecl pattern consts pat :: "'a \ pattern" nopat :: "'a \ pattern" definition trigger :: "pattern symb_list symb_list \ bool \ bool" where "trigger _ P = P" subsection \Higher-order encoding\ text \ Application is made explicit for constants occurring with varying numbers of arguments. This is achieved by the introduction of the following constant. \ definition fun_app :: "'a \ 'a" where "fun_app f = f" text \ Some solvers support a theory of arrays which can be used to encode higher-order functions. The following set of lemmas specifies the properties of such (extensional) arrays. \ lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other fun_upd_upd fun_app_def subsection \Normalization\ lemma case_bool_if[abs_def]: "case_bool x y P = (if P then x else y)" by simp lemmas Ex1_def_raw = Ex1_def[abs_def] lemmas Ball_def_raw = Ball_def[abs_def] lemmas Bex_def_raw = Bex_def[abs_def] lemmas abs_if_raw = abs_if[abs_def] lemmas min_def_raw = min_def[abs_def] lemmas max_def_raw = max_def[abs_def] lemma nat_zero_as_int: "0 = nat 0" by simp lemma nat_one_as_int: "1 = nat 1" by simp lemma nat_numeral_as_int: "numeral = (\i. nat (numeral i))" by simp lemma nat_less_as_int: "(<) = (\a b. int a < int b)" by simp lemma nat_leq_as_int: "(\) = (\a b. int a \ int b)" by simp lemma Suc_as_int: "Suc = (\a. nat (int a + 1))" by (rule ext) simp lemma nat_plus_as_int: "(+) = (\a b. nat (int a + int b))" by (rule ext)+ simp lemma nat_minus_as_int: "(-) = (\a b. nat (int a - int b))" by (rule ext)+ simp lemma nat_times_as_int: "(*) = (\a b. nat (int a * int b))" by (simp add: nat_mult_distrib) lemma nat_div_as_int: "(div) = (\a b. nat (int a div int b))" by (simp add: nat_div_distrib) lemma nat_mod_as_int: "(mod) = (\a b. nat (int a mod int b))" by (simp add: nat_mod_distrib) lemma int_Suc: "int (Suc n) = int n + 1" by simp lemma int_plus: "int (n + m) = int n + int m" by (rule of_nat_add) lemma int_minus: "int (n - m) = int (nat (int n - int m))" by auto lemma nat_int_comparison: fixes a b :: nat shows "(a = b) = (int a = int b)" and "(a < b) = (int a < int b)" and "(a \ b) = (int a \ int b)" by simp_all lemma int_ops: fixes a b :: nat shows "int 0 = 0" and "int 1 = 1" and "int (numeral n) = numeral n" and "int (Suc a) = int a + 1" and "int (a + b) = int a + int b" and "int (a - b) = (if int a < int b then 0 else int a - int b)" and "int (a * b) = int a * int b" and "int (a div b) = int a div int b" and "int (a mod b) = int a mod int b" by (auto intro: zdiv_int zmod_int) lemma int_if: fixes a b :: nat shows "int (if P then a else b) = (if P then int a else int b)" by simp subsection \Integer division and modulo for Z3\ text \ The following Z3-inspired definitions are overspecified for the case where \l = 0\. This Schönheitsfehler is corrected in the \div_as_z3div\ and \mod_as_z3mod\ theorems. \ definition z3div :: "int \ int \ int" where "z3div k l = (if l \ 0 then k div l else - (k div - l))" definition z3mod :: "int \ int \ int" where "z3mod k l = k mod (if l \ 0 then l else - l)" lemma div_as_z3div: "\k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))" by (simp add: z3div_def) lemma mod_as_z3mod: "\k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))" by (simp add: z3mod_def) subsection \Extra theorems for veriT reconstruction\ lemma verit_sko_forall: \(\x. P x) \ P (SOME x. \P x)\ using someI[of \\x. \P x\] by auto lemma verit_sko_forall': \P (SOME x. \P x) = A \ (\x. P x) = A\ by (subst verit_sko_forall) lemma verit_sko_forall'': \B = A \ (SOME x. P x) = A \ (SOME x. P x) = B\ by auto lemma verit_sko_forall_indirect: \x = (SOME x. \P x) \ (\x. P x) \ P x\ using someI[of \\x. \P x\] by auto lemma verit_sko_forall_indirect2: \x = (SOME x. \P x) \ (\x :: 'a. (P x = P' x)) \(\x. P' x) \ P x\ using someI[of \\x. \P x\] by auto lemma verit_sko_ex: \(\x. P x) \ P (SOME x. P x)\ using someI[of \\x. P x\] by auto lemma verit_sko_ex': \P (SOME x. P x) = A \ (\x. P x) = A\ by (subst verit_sko_ex) lemma verit_sko_ex_indirect: \x = (SOME x. P x) \ (\x. P x) \ P x\ using someI[of \\x. P x\] by auto lemma verit_sko_ex_indirect2: \x = (SOME x. P x) \ (\x. P x = P' x) \ (\x. P' x) \ P x\ using someI[of \\x. P x\] by auto lemma verit_Pure_trans: \P \ Q \ Q \ P\ by auto lemma verit_if_cong: assumes \b \ c\ and \c \ x \ u\ and \\ c \ y \ v\ shows \(if b then x else y) \ (if c then u else v)\ using assms if_cong[of b c x u] by auto lemma verit_if_weak_cong': \b \ c \ (if b then x else y) \ (if c then x else y)\ by auto lemma verit_or_neg: \(A \ B) \ B \ \A\ \(\A \ B) \ B \ A\ by auto lemma verit_subst_bool: \P \ f True \ f P\ by auto lemma verit_and_pos: \(a \ \(b \ c) \ A) \ \(a \ b \ c) \ A\ \(a \ b \ A) \ \(a \ b) \ A\ \(a \ A) \ \a \ A\ \(\a \ A) \ a \ A\ by blast+ lemma verit_or_pos: \A \ A' \ (c \ A) \ (\c \ A')\ \A \ A' \ (\c \ A) \ (c \ A')\ by blast+ lemma verit_la_generic: \(a::int) \ x \ a = x \ a \ x\ by linarith lemma verit_bfun_elim: \(if b then P True else P False) = P b\ \(\b. P' b) = (P' False \ P' True)\ \(\b. P' b) = (P' False \ P' True)\ by (cases b) (auto simp: all_bool_eq ex_bool_eq) lemma verit_eq_true_simplify: \(P = True) \ P\ by auto lemma verit_and_neg: \(a \ \b \ A) \ \(a \ b) \ A\ \(a \ A) \ \a \ A\ \(\a \ A) \ a \ A\ by blast+ lemma verit_forall_inst: \A \ B \ \A \ B\ \\A \ B \ A \ B\ \A \ B \ \B \ A\ \A \ \B \ B \ A\ \A \ B \ \A \ B\ \\A \ B \ A \ B\ by blast+ lemma verit_eq_transitive: \A = B \ B = C \ A = C\ \A = B \ C = B \ A = C\ \B = A \ B = C \ A = C\ \B = A \ C = B \ A = C\ by auto lemma verit_bool_simplify: \\(P \ Q) \ P \ \Q\ \\(P \ Q) \ \P \ \Q\ \\(P \ Q) \ \P \ \Q\ \(P \ (Q \ R)) \ ((P \ Q) \ R)\ \((P \ Q) \ Q) \ P \ Q\ \(Q \ (P \ Q)) \ (P \ Q)\ \ \This rule was inverted\ \P \ (P \ Q) \ P \ Q\ \(P \ Q) \ P \ P \ Q\ (* \\Additional rules:\ * \((P \ Q) \ P) \ P\ * \((P \ Q) \ Q) \ P \ Q\ * \(P \ Q) \ P\ *) unfolding not_imp imp_conjL by auto text \We need the last equation for \<^term>\\(\a b. \P a b)\\ lemma verit_connective_def: \ \the definition of XOR is missing as the operator is not generated by Isabelle\ \(A = B) \ ((A \ B) \ (B \ A))\ \(If A B C) = ((A \ B) \ (\A \ C))\ \(\x. P x) \ \(\x. \P x)\ \\(\x. P x) \ (\x. \P x)\ by auto lemma verit_ite_simplify: \(If True B C) = B\ \(If False B C) = C\ \(If A' B B) = B\ \(If (\A') B C) = (If A' C B)\ \(If c (If c A B) C) = (If c A C)\ \(If c C (If c A B)) = (If c C B)\ \(If A' True False) = A'\ \(If A' False True) \ \A'\ \(If A' True B') \ A'\B'\ \(If A' B' False) \ A'\B'\ \(If A' False B') \ \A'\B'\ \(If A' B' True) \ \A'\B'\ \x \ True \ x\ \x \ False \ x\ for B C :: 'a and A' B' C' :: bool by auto lemma verit_and_simplify1: \True \ b \ b\ \b \ True \ b\ \False \ b \ False\ \b \ False \ False\ \(c \ \c) \ False\ \(\c \ c) \ False\ \\\a = a\ by auto lemmas verit_and_simplify = conj_ac de_Morgan_conj disj_not1 lemma verit_or_simplify_1: \False \ b \ b\ \b \ False \ b\ \b \ \b\ \\b \ b\ by auto lemmas verit_or_simplify = disj_ac lemma verit_not_simplify: \\ \b \ b\ \\True \ False\ \\False \ True\ by auto lemma verit_implies_simplify: \(\a \ \b) \ (b \ a)\ \(False \ a) \ True\ \(a \ True) \ True\ \(True \ a) \ a\ \(a \ False) \ \a\ \(a \ a) \ True\ \(\a \ a) \ a\ \(a \ \a) \ \a\ \((a \ b) \ b) \ a \ b\ by auto lemma verit_equiv_simplify: \((\a) = (\b)) \ (a = b)\ \(a = a) \ True\ \(a = (\a)) \ False\ \((\a) = a) \ False\ \(True = a) \ a\ \(a = True) \ a\ \(False = a) \ \a\ \(a = False) \ \a\ \\\a \ a\ \(\ False) = True\ for a b :: bool by auto lemmas verit_eq_simplify = semiring_char_0_class.eq_numeral_simps eq_refl zero_neq_one num.simps neg_equal_zero equal_neg_zero one_neq_zero neg_equal_iff_equal lemma verit_minus_simplify: \(a :: 'a :: cancel_comm_monoid_add) - a = 0\ \(a :: 'a :: cancel_comm_monoid_add) - 0 = a\ \0 - (b :: 'b :: {group_add}) = -b\ \- (- (b :: 'b :: group_add)) = b\ by auto lemma verit_sum_simplify: \(a :: 'a :: cancel_comm_monoid_add) + 0 = a\ by auto lemmas verit_prod_simplify = (* already included: mult_zero_class.mult_zero_right mult_zero_class.mult_zero_left *) mult_1 mult_1_right lemma verit_comp_simplify1: \(a :: 'a ::order) < a \ False\ \a \ a\ \\(b' \ a') \ (a' :: 'b :: linorder) < b'\ by auto lemmas verit_comp_simplify = verit_comp_simplify1 le_numeral_simps le_num_simps less_numeral_simps less_num_simps zero_less_one zero_le_one less_neg_numeral_simps lemma verit_la_disequality: \(a :: 'a ::linorder) = b \ \a \ b \ \b \ a\ by auto context begin text \For the reconstruction, we need to keep the order of the arguments.\ named_theorems smt_arith_multiplication \Theorems to reconstruct arithmetic theorems.\ named_theorems smt_arith_combine \Theorems to reconstruct arithmetic theorems.\ named_theorems smt_arith_simplify \Theorems to combine theorems in the LA procedure\ lemmas [smt_arith_simplify] = div_add dvd_numeral_simp divmod_steps less_num_simps le_num_simps if_True if_False divmod_cancel dvd_mult dvd_mult2 less_irrefl prod.case numeral_plus_one divmod_step_def order.refl le_zero_eq le_numeral_simps less_numeral_simps mult.right_neutral simp_thms divides_aux_eq mult_nonneg_nonneg dvd_imp_mod_0 dvd_add zero_less_one mod_mult_self4 numeral_mod_numeral divmod_trivial prod.sel mult.left_neutral div_pos_pos_trivial arith_simps div_add div_mult_self1 add_le_cancel_left add_le_same_cancel2 not_one_le_zero le_numeral_simps add_le_same_cancel1 zero_neq_one zero_le_one le_num_simps add_Suc mod_div_trivial nat.distinct mult_minus_right add.inverse_inverse distrib_left_numeral mult_num_simps numeral_times_numeral add_num_simps divmod_steps rel_simps if_True if_False numeral_div_numeral divmod_cancel prod.case add_num_simps one_plus_numeral fst_conv arith_simps sub_num_simps dbl_inc_simps dbl_simps mult_1 add_le_cancel_right left_diff_distrib_numeral add_uminus_conv_diff zero_neq_one zero_le_one One_nat_def add_Suc mod_div_trivial nat.distinct of_int_1 numerals numeral_One of_int_numeral add_uminus_conv_diff zle_diff1_eq add_less_same_cancel2 minus_add_distrib add_uminus_conv_diff mult.left_neutral semiring_class.distrib_right add_diff_cancel_left' add_diff_eq ring_distribs mult_minus_left minus_diff_eq lemma [smt_arith_simplify]: \\ (a' :: 'a :: linorder) < b' \ b' \ a'\ \\ (a' :: 'a :: linorder) \ b' \ b' < a'\ \(c::int) mod Numeral1 = 0\ \(a::nat) mod Numeral1 = 0\ \(c::int) div Numeral1 = c\ \a div Numeral1 = a\ \(c::int) mod 1 = 0\ \a mod 1 = 0\ \(c::int) div 1 = c\ \a div 1 = a\ \\(a' \ b') \ a' = b'\ by auto lemma div_mod_decomp: "A = (A div n) * n + (A mod n)" for A :: nat by auto lemma div_less_mono: fixes A B :: nat assumes "A < B" "0 < n" and mod: "A mod n = 0""B mod n = 0" shows "(A div n) < (B div n)" proof - show ?thesis using assms(1) apply (subst (asm) div_mod_decomp[of "A" n]) apply (subst (asm) div_mod_decomp[of "B" n]) unfolding mod by (use assms(2,3) in \auto simp: ac_simps\) qed lemma verit_le_mono_div: fixes A B :: nat assumes "A < B" "0 < n" shows "(A div n) + (if B mod n = 0 then 1 else 0) \ (B div n)" by (auto simp: ac_simps Suc_leI assms less_mult_imp_div_less div_le_mono less_imp_le_nat) lemmas [smt_arith_multiplication] = verit_le_mono_div[THEN mult_le_mono1, unfolded add_mult_distrib] div_le_mono[THEN mult_le_mono2, unfolded add_mult_distrib] lemma div_mod_decomp_int: "A = (A div n) * n + (A mod n)" for A :: int by auto lemma zdiv_mono_strict: fixes A B :: int assumes "A < B" "0 < n" and mod: "A mod n = 0""B mod n = 0" shows "(A div n) < (B div n)" proof - show ?thesis using assms(1) apply (subst (asm) div_mod_decomp_int[of A n]) apply (subst (asm) div_mod_decomp_int[of B n]) unfolding mod by (use assms(2,3) in \auto simp: ac_simps\) qed lemma verit_le_mono_div_int: fixes A B :: int assumes "A < B" "0 < n" shows "(A div n) + (if B mod n = 0 then 1 else 0) \ (B div n)" proof - have f2: "B div n = A div n \ A div n < B div n" by (metis (no_types) assms less_le zdiv_mono1) have "B div n \ A div n \ B mod n \ 0" using assms(1) by (metis Groups.add_ac(2) assms(2) eucl_rel_int eucl_rel_int_iff group_cancel.rule0 le_add_same_cancel2 not_le) then have "B mod n = 0 \ A div n + (if B mod n = 0 then 1 else 0) \ B div n" using f2 by (auto dest: zless_imp_add1_zle) then show ?thesis using assms zdiv_mono1 by auto qed lemma verit_less_mono_div_int2: fixes A B :: int assumes "A \ B" "0 < -n" shows "(A div n) \ (B div n)" using assms(1) assms(2) zdiv_mono1_neg by auto lemmas [smt_arith_multiplication] = verit_le_mono_div_int[THEN mult_left_mono, unfolded int_distrib] zdiv_mono1[THEN mult_left_mono, unfolded int_distrib] lemmas [smt_arith_multiplication] = arg_cong[of _ _ \\a :: nat. a div n * p\ for n p :: nat, THEN sym] arg_cong[of _ _ \\a :: int. a div n * p\ for n p :: int, THEN sym] lemma [smt_arith_combine]: "a < b \ c < d \ a + c + 2 \ b + d" "a < b \ c \ d \ a + c + 1 \ b + d" "a \ b \ c < d \ a + c + 1 \ b + d" for a b c :: int by auto lemma [smt_arith_combine]: "a < b \ c < d \ a + c + 2 \ b + d" "a < b \ c \ d \ a + c + 1 \ b + d" "a \ b \ c < d \ a + c + 1 \ b + d" for a b c :: nat by auto lemmas [smt_arith_combine] = add_strict_mono add_less_le_mono add_mono add_le_less_mono lemma [smt_arith_combine]: \m < n \ c = d \ m + c < n + d\ \m \ n \ c = d \ m + c \ n + d\ \c = d \ m < n \ m + c < n + d\ \c = d \ m \ n \ m + c \ n + d\ for m :: \'a :: ordered_cancel_ab_semigroup_add\ by (auto intro: ordered_cancel_ab_semigroup_add_class.add_strict_right_mono ordered_ab_semigroup_add_class.add_right_mono) lemma verit_negate_coefficient: \a \ (b :: 'a :: {ordered_ab_group_add}) \ -a \ -b\ \a < b \ -a > -b\ \a = b \ -a = -b\ by auto end lemma verit_ite_intro: \(if a then P (if a then a' else b') else Q) \ (if a then P a' else Q)\ \(if a then P' else Q' (if a then a' else b')) \ (if a then P' else Q' b')\ \A = f (if a then R else S) \ (if a then A = f R else A = f S)\ by auto lemma verit_ite_if_cong: fixes x y :: bool assumes "b=c" and "c \ True \ x = u" and "c \ False \ y = v" shows "(if b then x else y) \ (if c then u else v)" proof - have H: "(if b then x else y) = (if c then u else v)" using assms by (auto split: if_splits) show "(if b then x else y) \ (if c then u else v)" by (subst H) auto qed subsection \Setup\ ML_file \Tools/SMT/smt_util.ML\ ML_file \Tools/SMT/smt_failure.ML\ ML_file \Tools/SMT/smt_config.ML\ ML_file \Tools/SMT/smt_builtin.ML\ ML_file \Tools/SMT/smt_datatypes.ML\ ML_file \Tools/SMT/smt_normalize.ML\ ML_file \Tools/SMT/smt_translate.ML\ ML_file \Tools/SMT/smtlib.ML\ ML_file \Tools/SMT/smtlib_interface.ML\ ML_file \Tools/SMT/smtlib_proof.ML\ ML_file \Tools/SMT/smtlib_isar.ML\ ML_file \Tools/SMT/z3_proof.ML\ ML_file \Tools/SMT/z3_isar.ML\ ML_file \Tools/SMT/smt_solver.ML\ ML_file \Tools/SMT/cvc_interface.ML\ ML_file \Tools/SMT/lethe_proof.ML\ ML_file \Tools/SMT/lethe_isar.ML\ ML_file \Tools/SMT/lethe_proof_parse.ML\ ML_file \Tools/SMT/cvc_proof_parse.ML\ -ML_file \Tools/SMT/verit_proof.ML\ ML_file \Tools/SMT/conj_disj_perm.ML\ ML_file \Tools/SMT/smt_replay_methods.ML\ ML_file \Tools/SMT/smt_replay.ML\ ML_file \Tools/SMT/smt_replay_arith.ML\ ML_file \Tools/SMT/z3_interface.ML\ ML_file \Tools/SMT/z3_replay_rules.ML\ ML_file \Tools/SMT/z3_replay_methods.ML\ ML_file \Tools/SMT/z3_replay.ML\ ML_file \Tools/SMT/lethe_replay_methods.ML\ ML_file \Tools/SMT/verit_replay_methods.ML\ +ML_file \Tools/SMT/verit_strategies.ML\ ML_file \Tools/SMT/verit_replay.ML\ ML_file \Tools/SMT/smt_systems.ML\ subsection \Configuration\ text \ The current configuration can be printed by the command \smt_status\, which shows the values of most options. \ subsection \General configuration options\ text \ The option \smt_solver\ can be used to change the target SMT solver. The possible values can be obtained from the \smt_status\ command. \ declare [[smt_solver = z3]] text \ Since SMT solvers are potentially nonterminating, there is a timeout (given in seconds) to restrict their runtime. \ declare [[smt_timeout = 0]] text \ SMT solvers apply randomized heuristics. In case a problem is not solvable by an SMT solver, changing the following option might help. \ declare [[smt_random_seed = 1]] text \ In general, the binding to SMT solvers runs as an oracle, i.e, the SMT solvers are fully trusted without additional checks. The following option can cause the SMT solver to run in proof-producing mode, giving a checkable certificate. This is currently implemented only for veriT and Z3. \ declare [[smt_oracle = false]] text \ Each SMT solver provides several command-line options to tweak its behaviour. They can be passed to the solver by setting the following options. \ declare [[cvc4_options = ""]] declare [[cvc5_options = ""]] declare [[verit_options = ""]] declare [[z3_options = ""]] text \ The SMT method provides an inference mechanism to detect simple triggers in quantified formulas, which might increase the number of problems solvable by SMT solvers (note: triggers guide quantifier instantiations in the SMT solver). To turn it on, set the following option. \ declare [[smt_infer_triggers = false]] text \ Enable the following option to use built-in support for datatypes, codatatypes, and records in CVC4 and cvc5. Currently, this is implemented only in oracle mode. \ declare [[cvc_extensions = false]] text \ Enable the following option to use built-in support for div/mod, datatypes, and records in Z3. Currently, this is implemented only in oracle mode. \ declare [[z3_extensions = false]] subsection \Certificates\ text \ By setting the option \smt_certificates\ to the name of a file, all following applications of an SMT solver a cached in that file. Any further application of the same SMT solver (using the very same configuration) re-uses the cached certificate instead of invoking the solver. An empty string disables caching certificates. The filename should be given as an explicit path. It is good practice to use the name of the current theory (with ending \.certs\ instead of \.thy\) as the certificates file. Certificate files should be used at most once in a certain theory context, to avoid race conditions with other concurrent accesses. \ declare [[smt_certificates = ""]] text \ The option \smt_read_only_certificates\ controls whether only stored certificates should be used or invocation of an SMT solver is allowed. When set to \true\, no SMT solver will ever be invoked and only the existing certificates found in the configured cache are used; when set to \false\ and there is no cached certificate for some proposition, then the configured SMT solver is invoked. \ declare [[smt_read_only_certificates = false]] subsection \Tracing\ text \ The SMT method, when applied, traces important information. To make it entirely silent, set the following option to \false\. \ declare [[smt_verbose = true]] text \ For tracing the generated problem file given to the SMT solver as well as the returned result of the solver, the option \smt_trace\ should be set to \true\. \ declare [[smt_trace = false]] subsection \Schematic rules for Z3 proof reconstruction\ text \ Several prof rules of Z3 are not very well documented. There are two lemma groups which can turn failing Z3 proof reconstruction attempts into succeeding ones: the facts in \z3_rule\ are tried prior to any implemented reconstruction procedure for all uncertain Z3 proof rules; the facts in \z3_simp\ are only fed to invocations of the simplifier when reconstructing theory-specific proof steps. \ lemmas [z3_rule] = refl eq_commute conj_commute disj_commute simp_thms nnf_simps ring_distribs field_simps times_divide_eq_right times_divide_eq_left if_True if_False not_not NO_MATCH_def lemma [z3_rule]: "(P \ Q) = (\ (\ P \ \ Q))" "(P \ Q) = (\ (\ Q \ \ P))" "(\ P \ Q) = (\ (P \ \ Q))" "(\ P \ Q) = (\ (\ Q \ P))" "(P \ \ Q) = (\ (\ P \ Q))" "(P \ \ Q) = (\ (Q \ \ P))" "(\ P \ \ Q) = (\ (P \ Q))" "(\ P \ \ Q) = (\ (Q \ P))" by auto lemma [z3_rule]: "(P \ Q) = (Q \ \ P)" "(\ P \ Q) = (P \ Q)" "(\ P \ Q) = (Q \ P)" "(True \ P) = P" "(P \ True) = True" "(False \ P) = True" "(P \ P) = True" "(\ (A \ \ B)) \ (A \ B)" by auto lemma [z3_rule]: "((P = Q) \ R) = (R \ (Q = (\ P)))" by auto lemma [z3_rule]: "(\ True) = False" "(\ False) = True" "(x = x) = True" "(P = True) = P" "(True = P) = P" "(P = False) = (\ P)" "(False = P) = (\ P)" "((\ P) = P) = False" "(P = (\ P)) = False" "((\ P) = (\ Q)) = (P = Q)" "\ (P = (\ Q)) = (P = Q)" "\ ((\ P) = Q) = (P = Q)" "(P \ Q) = (Q = (\ P))" "(P = Q) = ((\ P \ Q) \ (P \ \ Q))" "(P \ Q) = ((\ P \ \ Q) \ (P \ Q))" by auto lemma [z3_rule]: "(if P then P else \ P) = True" "(if \ P then \ P else P) = True" "(if P then True else False) = P" "(if P then False else True) = (\ P)" "(if P then Q else True) = ((\ P) \ Q)" "(if P then Q else True) = (Q \ (\ P))" "(if P then Q else \ Q) = (P = Q)" "(if P then Q else \ Q) = (Q = P)" "(if P then \ Q else Q) = (P = (\ Q))" "(if P then \ Q else Q) = ((\ Q) = P)" "(if \ P then x else y) = (if P then y else x)" "(if P then (if Q then x else y) else x) = (if P \ (\ Q) then y else x)" "(if P then (if Q then x else y) else x) = (if (\ Q) \ P then y else x)" "(if P then (if Q then x else y) else y) = (if P \ Q then x else y)" "(if P then (if Q then x else y) else y) = (if Q \ P then x else y)" "(if P then x else if P then y else z) = (if P then x else z)" "(if P then x else if Q then x else y) = (if P \ Q then x else y)" "(if P then x else if Q then x else y) = (if Q \ P then x else y)" "(if P then x = y else x = z) = (x = (if P then y else z))" "(if P then x = y else y = z) = (y = (if P then x else z))" "(if P then x = y else z = y) = (y = (if P then x else z))" by auto lemma [z3_rule]: "0 + (x::int) = x" "x + 0 = x" "x + x = 2 * x" "0 * x = 0" "1 * x = x" "x + y = y + x" by auto lemma [z3_rule]: (* for def-axiom *) "P = Q \ P \ Q" "P = Q \ \ P \ \ Q" "(\ P) = Q \ \ P \ Q" "(\ P) = Q \ P \ \ Q" "P = (\ Q) \ \ P \ Q" "P = (\ Q) \ P \ \ Q" "P \ Q \ P \ \ Q" "P \ Q \ \ P \ Q" "P \ (\ Q) \ P \ Q" "(\ P) \ Q \ P \ Q" "P \ Q \ P \ (\ Q)" "P \ Q \ (\ P) \ Q" "P \ \ Q \ P \ Q" "\ P \ Q \ P \ Q" "P \ y = (if P then x else y)" "P \ (if P then x else y) = y" "\ P \ x = (if P then x else y)" "\ P \ (if P then x else y) = x" "P \ R \ \ (if P then Q else R)" "\ P \ Q \ \ (if P then Q else R)" "\ (if P then Q else R) \ \ P \ Q" "\ (if P then Q else R) \ P \ R" "(if P then Q else R) \ \ P \ \ Q" "(if P then Q else R) \ P \ \ R" "(if P then \ Q else R) \ \ P \ Q" "(if P then Q else \ R) \ P \ R" by auto hide_type (open) symb_list pattern hide_const (open) Symb_Nil Symb_Cons trigger pat nopat fun_app z3div z3mod end diff --git a/src/HOL/SMT_Examples/SMT_Examples_Verit.certs b/src/HOL/SMT_Examples/SMT_Examples_Verit.certs --- a/src/HOL/SMT_Examples/SMT_Examples_Verit.certs +++ b/src/HOL/SMT_Examples/SMT_Examples_Verit.certs @@ -1,8868 +1,12431 @@ ae54fcb9dfe8ac3652092131f8427bebbd05402c 9 0 unsat (assume axiom0 (! (not true) :named @p_1)) (step t2 (cl (! (= @p_1 false) :named @p_2)) :rule not_simplify) (step t3 (cl (! (not @p_2) :named @p_4) (! (not @p_1) :named @p_3) false) :rule equiv_pos2) (step t4 (cl (not @p_3) true) :rule not_not) (step t5 (cl @p_4 true false) :rule th_resolution :premises (t4 t3)) (step t6 (cl false) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl (not false)) :rule false) (step t8 (cl) :rule resolution :premises (t6 t7)) ef08efbe2a4fd690de84a4f2f024c061b9c72554 12 0 unsat (assume axiom0 (! (not (! (or p$ (not p$)) :named @p_1)) :named @p_2)) (step t2 (cl (= @p_1 true)) :rule or_simplify) (step t3 (cl (= @p_2 (! (not true) :named @p_3))) :rule cong :premises (t2)) (step t4 (cl (= @p_3 false)) :rule not_simplify) (step t5 (cl (! (= @p_2 false) :named @p_4)) :rule trans :premises (t3 t4)) (step t6 (cl (! (not @p_4) :named @p_6) (! (not @p_2) :named @p_5) false) :rule equiv_pos2) (step t7 (cl (not @p_5) @p_1) :rule not_not) (step t8 (cl @p_6 @p_1 false) :rule th_resolution :premises (t7 t6)) (step t9 (cl false) :rule th_resolution :premises (axiom0 t5 t8)) (step t10 (cl (not false)) :rule false) (step t11 (cl) :rule resolution :premises (t9 t10)) b753d4b1e8a27c132b8339803a4e586242a9eed9 17 0 unsat (assume axiom0 (! (not (! (= (! (and p$ true) :named @p_1) p$) :named @p_3)) :named @p_5)) (step t2 (cl (= @p_1 (! (and p$) :named @p_2))) :rule and_simplify) (step t3 (cl (= @p_2 p$)) :rule and_simplify) (step t4 (cl @p_3) :rule trans :premises (t2 t3)) (step t5 (cl (= @p_3 (! (= p$ p$) :named @p_4))) :rule cong :premises (t4)) (step t6 (cl (= @p_4 true)) :rule equiv_simplify) (step t7 (cl (= @p_3 true)) :rule trans :premises (t5 t6)) (step t8 (cl (= @p_5 (! (not true) :named @p_6))) :rule cong :premises (t7)) (step t9 (cl (= @p_6 false)) :rule not_simplify) (step t10 (cl (! (= @p_5 false) :named @p_7)) :rule trans :premises (t8 t9)) (step t11 (cl (! (not @p_7) :named @p_9) (! (not @p_5) :named @p_8) false) :rule equiv_pos2) (step t12 (cl (not @p_8) @p_3) :rule not_not) (step t13 (cl @p_9 @p_3 false) :rule th_resolution :premises (t12 t11)) (step t14 (cl false) :rule th_resolution :premises (axiom0 t10 t13)) (step t15 (cl (not false)) :rule false) (step t16 (cl) :rule resolution :premises (t14 t15)) 3aef0472082a0952dd5c0083a53b60410a65fb15 15 0 unsat (assume axiom0 (! (not (! (=> (! (and (! (or p$ q$) :named @p_8) (! (not p$) :named @p_9)) :named @p_2) q$) :named @p_6)) :named @p_1)) (step t2 (cl (! (= @p_1 (! (and @p_2 (! (not q$) :named @p_10)) :named @p_4)) :named @p_3)) :rule bool_simplify) (step t3 (cl (! (not @p_3) :named @p_7) (! (not @p_1) :named @p_5) @p_4) :rule equiv_pos2) (step t4 (cl (not @p_5) @p_6) :rule not_not) (step t5 (cl @p_7 @p_6 @p_4) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_4) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl (! (= @p_4 (! (and @p_8 @p_9 @p_10) :named @p_12)) :named @p_11)) :rule ac_simp) (step t8 (cl (not @p_11) (not @p_4) @p_12) :rule equiv_pos2) (step t9 (cl @p_12) :rule th_resolution :premises (t6 t7 t8)) (step t10 (cl @p_8) :rule and :premises (t9)) (step t11 (cl p$ q$) :rule or :premises (t10)) (step t12 (cl @p_9) :rule and :premises (t9)) (step t13 (cl @p_10) :rule and :premises (t9)) (step t14 (cl) :rule resolution :premises (t11 t12 t13)) a4d4a57bc38d3365a195600a7f29533d2f85c08f 12 0 unsat (assume axiom0 (! (not (! (=> (! (or (and a$ b$) (and c$ d$)) :named @p_1) @p_1) :named @p_6)) :named @p_2)) (step t2 (cl (! (= @p_2 (! (and @p_1 (not @p_1)) :named @p_4)) :named @p_3)) :rule bool_simplify) (step t3 (cl (! (not @p_3) :named @p_7) (! (not @p_2) :named @p_5) @p_4) :rule equiv_pos2) (step t4 (cl (not @p_5) @p_6) :rule not_not) (step t5 (cl @p_7 @p_6 @p_4) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_4) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl (! (= @p_4 false) :named @p_8)) :rule and_simplify) (step t8 (cl (not @p_8) (not @p_4) false) :rule equiv_pos2) (step t9 (cl false) :rule th_resolution :premises (t6 t7 t8)) (step t10 (cl (not false)) :rule false) (step t11 (cl) :rule resolution :premises (t9 t10)) f8a0f590e48624e657155c162af04ac8e33a73bd 12 0 unsat (assume axiom0 (! (not (! (=> (! (or (and p1$ p2$) p3$) :named @p_2) (! (or (! (=> p1$ (or (and p3$ p2$) (and p1$ p3$))) :named @p_10) p1$) :named @p_3)) :named @p_7)) :named @p_1)) (step t2 (cl (! (= @p_1 (! (and @p_2 (! (not @p_3) :named @p_9)) :named @p_5)) :named @p_4)) :rule bool_simplify) (step t3 (cl (! (not @p_4) :named @p_8) (! (not @p_1) :named @p_6) @p_5) :rule equiv_pos2) (step t4 (cl (not @p_6) @p_7) :rule not_not) (step t5 (cl @p_8 @p_7 @p_5) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_5) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl @p_9) :rule and :premises (t6)) (step t8 (cl (not @p_10)) :rule not_or :premises (t7)) (step t9 (cl p1$) :rule not_implies1 :premises (t8)) (step t10 (cl (not p1$)) :rule not_or :premises (t7)) (step t11 (cl) :rule resolution :premises (t10 t9)) 2815beadf2b8b0fa1acb410b1f6d788eddf2d2da 29 0 unsat (assume axiom0 (! (not (! (= (! (= (! (= (! (= (! (= (! (= (! (= (! (= (! (= p$ p$) :named @p_1) p$) :named @p_2) p$) :named @p_4) p$) :named @p_5) p$) :named @p_6) p$) :named @p_7) p$) :named @p_8) p$) :named @p_9) p$) :named @p_10)) :named @p_11)) (step t2 (cl (= @p_1 true)) :rule equiv_simplify) (step t3 (cl (= @p_2 (! (= true p$) :named @p_3))) :rule cong :premises (t2)) (step t4 (cl (= @p_3 p$)) :rule equiv_simplify) (step t5 (cl @p_4) :rule trans :premises (t3 t4)) (step t6 (cl (= @p_4 @p_1)) :rule cong :premises (t5)) (step t7 (cl (= @p_4 true)) :rule trans :premises (t6 t2)) (step t8 (cl (= @p_5 @p_3)) :rule cong :premises (t7)) (step t9 (cl @p_6) :rule trans :premises (t8 t4)) (step t10 (cl (= @p_6 @p_1)) :rule cong :premises (t9)) (step t11 (cl (= @p_6 true)) :rule trans :premises (t10 t2)) (step t12 (cl (= @p_7 @p_3)) :rule cong :premises (t11)) (step t13 (cl @p_8) :rule trans :premises (t12 t4)) (step t14 (cl (= @p_8 @p_1)) :rule cong :premises (t13)) (step t15 (cl (= @p_8 true)) :rule trans :premises (t14 t2)) (step t16 (cl (= @p_9 @p_3)) :rule cong :premises (t15)) (step t17 (cl @p_10) :rule trans :premises (t16 t4)) (step t18 (cl (= @p_10 @p_1)) :rule cong :premises (t17)) (step t19 (cl (= @p_10 true)) :rule trans :premises (t18 t2)) (step t20 (cl (= @p_11 (! (not true) :named @p_12))) :rule cong :premises (t19)) (step t21 (cl (= @p_12 false)) :rule not_simplify) (step t22 (cl (! (= @p_11 false) :named @p_13)) :rule trans :premises (t20 t21)) (step t23 (cl (! (not @p_13) :named @p_15) (! (not @p_11) :named @p_14) false) :rule equiv_pos2) (step t24 (cl (not @p_14) @p_10) :rule not_not) (step t25 (cl @p_15 @p_10 false) :rule th_resolution :premises (t24 t23)) (step t26 (cl false) :rule th_resolution :premises (axiom0 t22 t25)) (step t27 (cl (not false)) :rule false) (step t28 (cl) :rule resolution :premises (t26 t27)) 324d7169fc854a8e46b44966c6d6829b24a059d5 59 0 unsat (assume axiom0 (! (or a$ (or b$ (or c$ d$))) :named @p_1)) (assume axiom2 (! (or (! (not (! (or a$ (! (and c$ (! (not c$) :named @p_40)) :named @p_4)) :named @p_5)) :named @p_8) b$) :named @p_9)) (assume axiom3 (! (or (! (not (! (and b$ (! (or x$ (not x$)) :named @p_13)) :named @p_14)) :named @p_17) c$) :named @p_18)) (assume axiom4 (! (or (! (not (! (or d$ false) :named @p_22)) :named @p_24) c$) :named @p_25)) (assume axiom5 (! (not (! (or c$ (! (and (! (not p$) :named @p_34) (! (or p$ (! (and q$ (not q$)) :named @p_29)) :named @p_30)) :named @p_33)) :named @p_36)) :named @p_39)) (step t6 (cl (! (= @p_1 (! (or a$ b$ c$ d$) :named @p_3)) :named @p_2)) :rule ac_simp) (step t7 (cl (not @p_2) (not @p_1) @p_3) :rule equiv_pos2) (step t8 (cl @p_3) :rule th_resolution :premises (axiom0 t6 t7)) (step t9 (cl (= @p_4 false)) :rule and_simplify) (step t10 (cl (= @p_5 (! (or a$ false) :named @p_6))) :rule cong :premises (t9)) (step t11 (cl (= @p_6 (! (or a$) :named @p_7))) :rule or_simplify) (step t12 (cl (= @p_7 a$)) :rule or_simplify) (step t13 (cl (= @p_5 a$)) :rule trans :premises (t10 t11 t12)) (step t14 (cl (= @p_8 (! (not a$) :named @p_10))) :rule cong :premises (t13)) (step t15 (cl (! (= @p_9 (! (or @p_10 b$) :named @p_12)) :named @p_11)) :rule cong :premises (t14)) (step t16 (cl (not @p_11) (not @p_9) @p_12) :rule equiv_pos2) (step t17 (cl @p_12) :rule th_resolution :premises (axiom2 t15 t16)) (step t18 (cl (= @p_13 true)) :rule or_simplify) (step t19 (cl (= @p_14 (! (and b$ true) :named @p_15))) :rule cong :premises (t18)) (step t20 (cl (= @p_15 (! (and b$) :named @p_16))) :rule and_simplify) (step t21 (cl (= @p_16 b$)) :rule and_simplify) (step t22 (cl (= @p_14 b$)) :rule trans :premises (t19 t20 t21)) (step t23 (cl (= @p_17 (! (not b$) :named @p_19))) :rule cong :premises (t22)) (step t24 (cl (! (= @p_18 (! (or @p_19 c$) :named @p_21)) :named @p_20)) :rule cong :premises (t23)) (step t25 (cl (not @p_20) (not @p_18) @p_21) :rule equiv_pos2) (step t26 (cl @p_21) :rule th_resolution :premises (axiom3 t24 t25)) (step t27 (cl (= @p_22 (! (or d$) :named @p_23))) :rule or_simplify) (step t28 (cl (= @p_23 d$)) :rule or_simplify) (step t29 (cl (= @p_22 d$)) :rule trans :premises (t27 t28)) (step t30 (cl (= @p_24 (! (not d$) :named @p_26))) :rule cong :premises (t29)) (step t31 (cl (! (= @p_25 (! (or @p_26 c$) :named @p_28)) :named @p_27)) :rule cong :premises (t30)) (step t32 (cl (not @p_27) (not @p_25) @p_28) :rule equiv_pos2) (step t33 (cl @p_28) :rule th_resolution :premises (axiom4 t31 t32)) (step t34 (cl (= @p_29 false)) :rule and_simplify) (step t35 (cl (= @p_30 (! (or p$ false) :named @p_31))) :rule cong :premises (t34)) (step t36 (cl (= @p_31 (! (or p$) :named @p_32))) :rule or_simplify) (step t37 (cl (= @p_32 p$)) :rule or_simplify) (step t38 (cl (= @p_30 p$)) :rule trans :premises (t35 t36 t37)) (step t39 (cl (= @p_33 (! (and @p_34 p$) :named @p_35))) :rule cong :premises (t38)) (step t40 (cl (= @p_35 false)) :rule and_simplify) (step t41 (cl (= @p_33 false)) :rule trans :premises (t39 t40)) (step t42 (cl (= @p_36 (! (or c$ false) :named @p_37))) :rule cong :premises (t41)) (step t43 (cl (= @p_37 (! (or c$) :named @p_38))) :rule or_simplify) (step t44 (cl (= @p_38 c$)) :rule or_simplify) (step t45 (cl (= @p_36 c$)) :rule trans :premises (t42 t43 t44)) (step t46 (cl (! (= @p_39 @p_40) :named @p_41)) :rule cong :premises (t45)) (step t47 (cl (! (not @p_41) :named @p_43) (! (not @p_39) :named @p_42) @p_40) :rule equiv_pos2) (step t48 (cl (not @p_42) @p_36) :rule not_not) (step t49 (cl @p_43 @p_36 @p_40) :rule th_resolution :premises (t48 t47)) (step t50 (cl @p_40) :rule th_resolution :premises (axiom5 t46 t49)) (step t51 (cl a$ b$ c$ d$) :rule or :premises (t8)) (step t52 (cl @p_10 b$) :rule or :premises (t17)) (step t53 (cl @p_19 c$) :rule or :premises (t26)) (step t54 (cl @p_26 c$) :rule or :premises (t33)) (step t55 (cl @p_19) :rule resolution :premises (t53 t50)) (step t56 (cl @p_26) :rule resolution :premises (t54 t50)) (step t57 (cl a$) :rule resolution :premises (t51 t55 t50 t56)) (step t58 (cl) :rule resolution :premises (t52 t55 t57)) e724b7ec8c17b26e1986911ee66a23de60e69502 38 0 unsat (assume axiom0 (! (forall ((?v0 A$) (?v1 A$)) (! (= (! (symm_f$ ?v0 ?v1) :named @p_2) (! (symm_f$ ?v1 ?v0) :named @p_6)) :named @p_8)) :named @p_1)) (assume axiom1 (! (not (! (and (! (= a$ a$) :named @p_19) (! (= (symm_f$ a$ b$) (symm_f$ b$ a$)) :named @p_21)) :named @p_20)) :named @p_24)) (anchor :step t3 :args ((:= (?v0 A$) veriT_vr0) (:= (?v1 A$) veriT_vr1))) (step t3.t1 (cl (! (= ?v0 veriT_vr0) :named @p_5)) :rule refl) (step t3.t2 (cl (! (= ?v1 veriT_vr1) :named @p_4)) :rule refl) (step t3.t3 (cl (= @p_2 (! (symm_f$ veriT_vr0 veriT_vr1) :named @p_3))) :rule cong :premises (t3.t1 t3.t2)) (step t3.t4 (cl @p_4) :rule refl) (step t3.t5 (cl @p_5) :rule refl) (step t3.t6 (cl (= @p_6 (! (symm_f$ veriT_vr1 veriT_vr0) :named @p_7))) :rule cong :premises (t3.t4 t3.t5)) (step t3.t7 (cl (= @p_8 (! (= @p_3 @p_7) :named @p_9))) :rule cong :premises (t3.t3 t3.t6)) (step t3 (cl (! (= @p_1 (! (forall ((veriT_vr0 A$) (veriT_vr1 A$)) @p_9) :named @p_11)) :named @p_10)) :rule bind) (step t4 (cl (not @p_10) (not @p_1) @p_11) :rule equiv_pos2) (step t5 (cl @p_11) :rule th_resolution :premises (axiom0 t3 t4)) (anchor :step t6 :args ((:= (veriT_vr0 A$) veriT_vr2) (:= (veriT_vr1 A$) veriT_vr3))) (step t6.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_14)) :rule refl) (step t6.t2 (cl (! (= veriT_vr1 veriT_vr3) :named @p_13)) :rule refl) (step t6.t3 (cl (= @p_3 (! (symm_f$ veriT_vr2 veriT_vr3) :named @p_12))) :rule cong :premises (t6.t1 t6.t2)) (step t6.t4 (cl @p_13) :rule refl) (step t6.t5 (cl @p_14) :rule refl) (step t6.t6 (cl (= @p_7 (! (symm_f$ veriT_vr3 veriT_vr2) :named @p_15))) :rule cong :premises (t6.t4 t6.t5)) (step t6.t7 (cl (= @p_9 (! (= @p_12 @p_15) :named @p_16))) :rule cong :premises (t6.t3 t6.t6)) (step t6 (cl (! (= @p_11 (! (forall ((veriT_vr2 A$) (veriT_vr3 A$)) @p_16) :named @p_18)) :named @p_17)) :rule bind) (step t7 (cl (not @p_17) (not @p_11) @p_18) :rule equiv_pos2) (step t8 (cl @p_18) :rule th_resolution :premises (t5 t6 t7)) (step t9 (cl (= @p_19 true)) :rule eq_simplify) (step t10 (cl (= @p_20 (! (and true @p_21) :named @p_22))) :rule cong :premises (t9)) (step t11 (cl (= @p_22 (! (and @p_21) :named @p_23))) :rule and_simplify) (step t12 (cl (= @p_23 @p_21)) :rule and_simplify) (step t13 (cl (= @p_20 @p_21)) :rule trans :premises (t10 t11 t12)) (step t14 (cl (! (= @p_24 (! (not @p_21) :named @p_26)) :named @p_25)) :rule cong :premises (t13)) (step t15 (cl (! (not @p_25) :named @p_28) (! (not @p_24) :named @p_27) @p_26) :rule equiv_pos2) (step t16 (cl (not @p_27) @p_20) :rule not_not) (step t17 (cl @p_28 @p_20 @p_26) :rule th_resolution :premises (t16 t15)) (step t18 (cl @p_26) :rule th_resolution :premises (axiom1 t14 t17)) (step t19 (cl (or (! (not @p_18) :named @p_29) @p_21)) :rule forall_inst :args ((:= veriT_vr2 a$) (:= veriT_vr3 b$))) (step t20 (cl @p_29 @p_21) :rule or :premises (t19)) (step t21 (cl) :rule resolution :premises (t20 t8 t18)) 2b0e7825d38340699ae4fceb3006d5eb0b5a0628 333 0 unsat (assume axiom0 (not x0$)) (assume axiom1 (not x30$)) (assume axiom2 (not x29$)) (assume axiom3 (not x59$)) (assume axiom4 (! (or x1$ (or x31$ x0$)) :named @p_57)) (assume axiom6 (! (or x3$ (or x33$ x2$)) :named @p_60)) (assume axiom7 (! (or x4$ (or x34$ x3$)) :named @p_63)) (assume axiom8 (or x35$ x4$)) (assume axiom9 (! (or x5$ (or x36$ x30$)) :named @p_66)) (assume axiom11 (! (or x7$ (or x38$ (or x6$ x32$))) :named @p_69)) (assume axiom13 (! (or x9$ (or x40$ (or x8$ x34$))) :named @p_72)) (assume axiom16 (! (or x11$ (or x43$ (or x10$ x37$))) :named @p_75)) (assume axiom18 (! (or x13$ (or x45$ (or x12$ x39$))) :named @p_78)) (assume axiom20 (! (or x47$ (or x14$ x41$)) :named @p_81)) (assume axiom21 (! (or x15$ (or x48$ x42$)) :named @p_84)) (assume axiom23 (! (or x17$ (or x50$ (or x16$ x44$))) :named @p_87)) (assume axiom25 (! (or x19$ (or x52$ (or x18$ x46$))) :named @p_90)) (assume axiom28 (! (or x21$ (or x55$ (or x20$ x49$))) :named @p_93)) (assume axiom30 (! (or x23$ (or x57$ (or x22$ x51$))) :named @p_96)) (assume axiom32 (! (or x59$ (or x24$ x53$)) :named @p_99)) (assume axiom33 (or x25$ x54$)) (assume axiom35 (! (or x27$ (or x26$ x56$)) :named @p_102)) (assume axiom37 (! (or x29$ (or x28$ x58$)) :named @p_105)) (assume axiom41 (or (! (not x2$) :named @p_1) (! (not x32$) :named @p_2))) (assume axiom42 (or @p_1 (! (not x1$) :named @p_3))) (assume axiom43 (or @p_2 @p_3)) (assume axiom47 (or (! (not x4$) :named @p_4) (! (not x34$) :named @p_5))) (assume axiom48 (or @p_4 (! (not x3$) :named @p_6))) (assume axiom49 (or @p_5 @p_6)) (assume axiom54 (or (! (not x6$) :named @p_7) (! (not x37$) :named @p_8))) (assume axiom55 (or @p_7 (! (not x5$) :named @p_9))) (assume axiom56 (or @p_7 (! (not x31$) :named @p_10))) (assume axiom57 (or @p_8 @p_9)) (assume axiom58 (or @p_8 @p_10)) (assume axiom59 (or @p_9 @p_10)) (assume axiom63 (or (! (not x38$) :named @p_11) @p_7)) (assume axiom64 (or @p_11 @p_2)) (assume axiom66 (or (! (not x8$) :named @p_12) (! (not x39$) :named @p_13))) (assume axiom67 (or @p_12 (! (not x7$) :named @p_14))) (assume axiom68 (or @p_12 (! (not x33$) :named @p_15))) (assume axiom69 (or @p_13 @p_14)) (assume axiom70 (or @p_13 @p_15)) (assume axiom71 (or @p_14 @p_15)) (assume axiom78 (or (! (not x41$) :named @p_16) (! (not x9$) :named @p_17))) (assume axiom79 (or @p_16 (! (not x35$) :named @p_18))) (assume axiom80 (or @p_17 @p_18)) (assume axiom81 (or (! (not x10$) :named @p_19) (! (not x42$) :named @p_20))) (assume axiom82 (or @p_19 (! (not x36$) :named @p_21))) (assume axiom83 (or @p_20 @p_21)) (assume axiom90 (or (! (not x12$) :named @p_22) (! (not x44$) :named @p_23))) (assume axiom91 (or @p_22 (! (not x11$) :named @p_24))) (assume axiom92 (or @p_22 @p_11)) (assume axiom93 (or @p_23 @p_24)) (assume axiom94 (or @p_23 @p_11)) (assume axiom95 (or @p_24 @p_11)) (assume axiom99 (or (! (not x45$) :named @p_25) @p_22)) (assume axiom100 (or @p_25 @p_13)) (assume axiom102 (or (! (not x14$) :named @p_26) (! (not x46$) :named @p_27))) (assume axiom103 (or @p_26 (! (not x13$) :named @p_28))) (assume axiom104 (or @p_26 (! (not x40$) :named @p_29))) (assume axiom105 (or @p_27 @p_28)) (assume axiom106 (or @p_27 @p_29)) (assume axiom107 (or @p_28 @p_29)) (assume axiom113 (or (! (not x48$) :named @p_41) @p_20)) (assume axiom114 (or (! (not x16$) :named @p_30) (! (not x49$) :named @p_31))) (assume axiom115 (or @p_30 (! (not x15$) :named @p_32))) (assume axiom116 (or @p_30 (! (not x43$) :named @p_33))) (assume axiom117 (or @p_31 @p_32)) (assume axiom118 (or @p_31 @p_33)) (assume axiom119 (or @p_32 @p_33)) (assume axiom126 (or (! (not x18$) :named @p_34) (! (not x51$) :named @p_35))) (assume axiom127 (or @p_34 (! (not x17$) :named @p_36))) (assume axiom128 (or @p_34 @p_25)) (assume axiom129 (or @p_35 @p_36)) (assume axiom130 (or @p_35 @p_25)) (assume axiom131 (or @p_36 @p_25)) (assume axiom134 (or (! (not x19$) :named @p_37) @p_27)) (assume axiom138 (or (! (not x53$) :named @p_38) @p_37)) (assume axiom139 (or @p_38 (! (not x47$) :named @p_39))) (assume axiom140 (or @p_37 @p_39)) (assume axiom141 (or (! (not x20$) :named @p_40) (! (not x54$) :named @p_42))) (assume axiom142 (or @p_40 @p_41)) (assume axiom143 (or @p_42 @p_41)) (assume axiom150 (or (! (not x22$) :named @p_43) (! (not x56$) :named @p_44))) (assume axiom151 (or @p_43 (! (not x21$) :named @p_45))) (assume axiom152 (or @p_43 (! (not x50$) :named @p_46))) (assume axiom153 (or @p_44 @p_45)) (assume axiom154 (or @p_44 @p_46)) (assume axiom155 (or @p_45 @p_46)) (assume axiom162 (or (! (not x24$) :named @p_47) (! (not x58$) :named @p_48))) (assume axiom163 (or @p_47 (! (not x23$) :named @p_49))) (assume axiom164 (or @p_47 (! (not x52$) :named @p_50))) (assume axiom165 (or @p_48 @p_49)) (assume axiom166 (or @p_48 @p_50)) (assume axiom167 (or @p_49 @p_50)) (assume axiom172 (or (! (not x26$) :named @p_51) (! (not x25$) :named @p_52))) (assume axiom173 (or @p_51 (! (not x55$) :named @p_53))) (assume axiom174 (or @p_52 @p_53)) (assume axiom178 (or (! (not x28$) :named @p_54) (! (not x27$) :named @p_55))) (assume axiom179 (or @p_54 (! (not x57$) :named @p_56))) (assume axiom180 (or @p_55 @p_56)) (step t102 (cl (! (= @p_57 (! (or x1$ x31$ x0$) :named @p_59)) :named @p_58)) :rule ac_simp) (step t103 (cl (not @p_58) (not @p_57) @p_59) :rule equiv_pos2) (step t104 (cl @p_59) :rule th_resolution :premises (axiom4 t102 t103)) (step t105 (cl (! (= @p_60 (! (or x3$ x33$ x2$) :named @p_62)) :named @p_61)) :rule ac_simp) (step t106 (cl (not @p_61) (not @p_60) @p_62) :rule equiv_pos2) (step t107 (cl @p_62) :rule th_resolution :premises (axiom6 t105 t106)) (step t108 (cl (! (= @p_63 (! (or x4$ x34$ x3$) :named @p_65)) :named @p_64)) :rule ac_simp) (step t109 (cl (not @p_64) (not @p_63) @p_65) :rule equiv_pos2) (step t110 (cl @p_65) :rule th_resolution :premises (axiom7 t108 t109)) (step t111 (cl (! (= @p_66 (! (or x5$ x36$ x30$) :named @p_68)) :named @p_67)) :rule ac_simp) (step t112 (cl (not @p_67) (not @p_66) @p_68) :rule equiv_pos2) (step t113 (cl @p_68) :rule th_resolution :premises (axiom9 t111 t112)) (step t114 (cl (! (= @p_69 (! (or x7$ x38$ x6$ x32$) :named @p_71)) :named @p_70)) :rule ac_simp) (step t115 (cl (not @p_70) (not @p_69) @p_71) :rule equiv_pos2) (step t116 (cl @p_71) :rule th_resolution :premises (axiom11 t114 t115)) (step t117 (cl (! (= @p_72 (! (or x9$ x40$ x8$ x34$) :named @p_74)) :named @p_73)) :rule ac_simp) (step t118 (cl (not @p_73) (not @p_72) @p_74) :rule equiv_pos2) (step t119 (cl @p_74) :rule th_resolution :premises (axiom13 t117 t118)) (step t120 (cl (! (= @p_75 (! (or x11$ x43$ x10$ x37$) :named @p_77)) :named @p_76)) :rule ac_simp) (step t121 (cl (not @p_76) (not @p_75) @p_77) :rule equiv_pos2) (step t122 (cl @p_77) :rule th_resolution :premises (axiom16 t120 t121)) (step t123 (cl (! (= @p_78 (! (or x13$ x45$ x12$ x39$) :named @p_80)) :named @p_79)) :rule ac_simp) (step t124 (cl (not @p_79) (not @p_78) @p_80) :rule equiv_pos2) (step t125 (cl @p_80) :rule th_resolution :premises (axiom18 t123 t124)) (step t126 (cl (! (= @p_81 (! (or x47$ x14$ x41$) :named @p_83)) :named @p_82)) :rule ac_simp) (step t127 (cl (not @p_82) (not @p_81) @p_83) :rule equiv_pos2) (step t128 (cl @p_83) :rule th_resolution :premises (axiom20 t126 t127)) (step t129 (cl (! (= @p_84 (! (or x15$ x48$ x42$) :named @p_86)) :named @p_85)) :rule ac_simp) (step t130 (cl (not @p_85) (not @p_84) @p_86) :rule equiv_pos2) (step t131 (cl @p_86) :rule th_resolution :premises (axiom21 t129 t130)) (step t132 (cl (! (= @p_87 (! (or x17$ x50$ x16$ x44$) :named @p_89)) :named @p_88)) :rule ac_simp) (step t133 (cl (not @p_88) (not @p_87) @p_89) :rule equiv_pos2) (step t134 (cl @p_89) :rule th_resolution :premises (axiom23 t132 t133)) (step t135 (cl (! (= @p_90 (! (or x19$ x52$ x18$ x46$) :named @p_92)) :named @p_91)) :rule ac_simp) (step t136 (cl (not @p_91) (not @p_90) @p_92) :rule equiv_pos2) (step t137 (cl @p_92) :rule th_resolution :premises (axiom25 t135 t136)) (step t138 (cl (! (= @p_93 (! (or x21$ x55$ x20$ x49$) :named @p_95)) :named @p_94)) :rule ac_simp) (step t139 (cl (not @p_94) (not @p_93) @p_95) :rule equiv_pos2) (step t140 (cl @p_95) :rule th_resolution :premises (axiom28 t138 t139)) (step t141 (cl (! (= @p_96 (! (or x23$ x57$ x22$ x51$) :named @p_98)) :named @p_97)) :rule ac_simp) (step t142 (cl (not @p_97) (not @p_96) @p_98) :rule equiv_pos2) (step t143 (cl @p_98) :rule th_resolution :premises (axiom30 t141 t142)) (step t144 (cl (! (= @p_99 (! (or x59$ x24$ x53$) :named @p_101)) :named @p_100)) :rule ac_simp) (step t145 (cl (not @p_100) (not @p_99) @p_101) :rule equiv_pos2) (step t146 (cl @p_101) :rule th_resolution :premises (axiom32 t144 t145)) (step t147 (cl (! (= @p_102 (! (or x27$ x26$ x56$) :named @p_104)) :named @p_103)) :rule ac_simp) (step t148 (cl (not @p_103) (not @p_102) @p_104) :rule equiv_pos2) (step t149 (cl @p_104) :rule th_resolution :premises (axiom35 t147 t148)) (step t150 (cl (! (= @p_105 (! (or x29$ x28$ x58$) :named @p_107)) :named @p_106)) :rule ac_simp) (step t151 (cl (not @p_106) (not @p_105) @p_107) :rule equiv_pos2) (step t152 (cl @p_107) :rule th_resolution :premises (axiom37 t150 t151)) (step t153 (cl x1$ x31$ x0$) :rule or :premises (t104)) (step t154 (cl x1$ x31$) :rule resolution :premises (t153 axiom0)) (step t155 (cl x3$ x33$ x2$) :rule or :premises (t107)) (step t156 (cl x4$ x34$ x3$) :rule or :premises (t110)) (step t157 (cl x35$ x4$) :rule or :premises (axiom8)) (step t158 (cl x5$ x36$ x30$) :rule or :premises (t113)) (step t159 (cl x5$ x36$) :rule resolution :premises (t158 axiom1)) (step t160 (cl x7$ x38$ x6$ x32$) :rule or :premises (t116)) (step t161 (cl x9$ x40$ x8$ x34$) :rule or :premises (t119)) (step t162 (cl x11$ x43$ x10$ x37$) :rule or :premises (t122)) (step t163 (cl x13$ x45$ x12$ x39$) :rule or :premises (t125)) (step t164 (cl x47$ x14$ x41$) :rule or :premises (t128)) (step t165 (cl x15$ x48$ x42$) :rule or :premises (t131)) (step t166 (cl x17$ x50$ x16$ x44$) :rule or :premises (t134)) (step t167 (cl x19$ x52$ x18$ x46$) :rule or :premises (t137)) (step t168 (cl x21$ x55$ x20$ x49$) :rule or :premises (t140)) (step t169 (cl x23$ x57$ x22$ x51$) :rule or :premises (t143)) (step t170 (cl x59$ x24$ x53$) :rule or :premises (t146)) (step t171 (cl x24$ x53$) :rule resolution :premises (t170 axiom3)) (step t172 (cl x25$ x54$) :rule or :premises (axiom33)) (step t173 (cl x27$ x26$ x56$) :rule or :premises (t149)) (step t174 (cl x29$ x28$ x58$) :rule or :premises (t152)) (step t175 (cl x28$ x58$) :rule resolution :premises (t174 axiom2)) (step t176 (cl @p_1 @p_2) :rule or :premises (axiom41)) (step t177 (cl @p_1 @p_3) :rule or :premises (axiom42)) (step t178 (cl @p_2 @p_3) :rule or :premises (axiom43)) (step t179 (cl @p_4 @p_5) :rule or :premises (axiom47)) (step t180 (cl @p_4 @p_6) :rule or :premises (axiom48)) (step t181 (cl @p_5 @p_6) :rule or :premises (axiom49)) (step t182 (cl @p_7 @p_8) :rule or :premises (axiom54)) (step t183 (cl @p_7 @p_9) :rule or :premises (axiom55)) (step t184 (cl @p_7 @p_10) :rule or :premises (axiom56)) (step t185 (cl @p_8 @p_9) :rule or :premises (axiom57)) (step t186 (cl @p_8 @p_10) :rule or :premises (axiom58)) (step t187 (cl @p_9 @p_10) :rule or :premises (axiom59)) (step t188 (cl @p_11 @p_7) :rule or :premises (axiom63)) (step t189 (cl @p_11 @p_2) :rule or :premises (axiom64)) (step t190 (cl @p_12 @p_13) :rule or :premises (axiom66)) (step t191 (cl @p_12 @p_14) :rule or :premises (axiom67)) (step t192 (cl @p_12 @p_15) :rule or :premises (axiom68)) (step t193 (cl @p_13 @p_14) :rule or :premises (axiom69)) (step t194 (cl @p_13 @p_15) :rule or :premises (axiom70)) (step t195 (cl @p_14 @p_15) :rule or :premises (axiom71)) (step t196 (cl @p_16 @p_17) :rule or :premises (axiom78)) (step t197 (cl @p_16 @p_18) :rule or :premises (axiom79)) (step t198 (cl @p_17 @p_18) :rule or :premises (axiom80)) (step t199 (cl @p_19 @p_20) :rule or :premises (axiom81)) (step t200 (cl @p_19 @p_21) :rule or :premises (axiom82)) (step t201 (cl @p_20 @p_21) :rule or :premises (axiom83)) (step t202 (cl @p_22 @p_23) :rule or :premises (axiom90)) (step t203 (cl @p_22 @p_24) :rule or :premises (axiom91)) (step t204 (cl @p_22 @p_11) :rule or :premises (axiom92)) (step t205 (cl @p_23 @p_24) :rule or :premises (axiom93)) (step t206 (cl @p_23 @p_11) :rule or :premises (axiom94)) (step t207 (cl @p_24 @p_11) :rule or :premises (axiom95)) (step t208 (cl @p_25 @p_22) :rule or :premises (axiom99)) (step t209 (cl @p_25 @p_13) :rule or :premises (axiom100)) (step t210 (cl @p_26 @p_27) :rule or :premises (axiom102)) (step t211 (cl @p_26 @p_28) :rule or :premises (axiom103)) (step t212 (cl @p_26 @p_29) :rule or :premises (axiom104)) (step t213 (cl @p_27 @p_28) :rule or :premises (axiom105)) (step t214 (cl @p_27 @p_29) :rule or :premises (axiom106)) (step t215 (cl @p_28 @p_29) :rule or :premises (axiom107)) (step t216 (cl @p_41 @p_20) :rule or :premises (axiom113)) (step t217 (cl @p_30 @p_31) :rule or :premises (axiom114)) (step t218 (cl @p_30 @p_32) :rule or :premises (axiom115)) (step t219 (cl @p_30 @p_33) :rule or :premises (axiom116)) (step t220 (cl @p_31 @p_32) :rule or :premises (axiom117)) (step t221 (cl @p_31 @p_33) :rule or :premises (axiom118)) (step t222 (cl @p_32 @p_33) :rule or :premises (axiom119)) (step t223 (cl @p_34 @p_35) :rule or :premises (axiom126)) (step t224 (cl @p_34 @p_36) :rule or :premises (axiom127)) (step t225 (cl @p_34 @p_25) :rule or :premises (axiom128)) (step t226 (cl @p_35 @p_36) :rule or :premises (axiom129)) (step t227 (cl @p_35 @p_25) :rule or :premises (axiom130)) (step t228 (cl @p_36 @p_25) :rule or :premises (axiom131)) (step t229 (cl @p_37 @p_27) :rule or :premises (axiom134)) (step t230 (cl @p_38 @p_37) :rule or :premises (axiom138)) (step t231 (cl @p_38 @p_39) :rule or :premises (axiom139)) (step t232 (cl @p_37 @p_39) :rule or :premises (axiom140)) (step t233 (cl @p_40 @p_42) :rule or :premises (axiom141)) (step t234 (cl @p_40 @p_41) :rule or :premises (axiom142)) (step t235 (cl @p_42 @p_41) :rule or :premises (axiom143)) (step t236 (cl @p_43 @p_44) :rule or :premises (axiom150)) (step t237 (cl @p_43 @p_45) :rule or :premises (axiom151)) (step t238 (cl @p_43 @p_46) :rule or :premises (axiom152)) (step t239 (cl @p_44 @p_45) :rule or :premises (axiom153)) (step t240 (cl @p_44 @p_46) :rule or :premises (axiom154)) (step t241 (cl @p_45 @p_46) :rule or :premises (axiom155)) (step t242 (cl @p_47 @p_48) :rule or :premises (axiom162)) (step t243 (cl @p_47 @p_49) :rule or :premises (axiom163)) (step t244 (cl @p_47 @p_50) :rule or :premises (axiom164)) (step t245 (cl @p_48 @p_49) :rule or :premises (axiom165)) (step t246 (cl @p_48 @p_50) :rule or :premises (axiom166)) (step t247 (cl @p_49 @p_50) :rule or :premises (axiom167)) (step t248 (cl @p_51 @p_52) :rule or :premises (axiom172)) (step t249 (cl @p_51 @p_53) :rule or :premises (axiom173)) (step t250 (cl @p_52 @p_53) :rule or :premises (axiom174)) (step t251 (cl @p_54 @p_55) :rule or :premises (axiom178)) (step t252 (cl @p_54 @p_56) :rule or :premises (axiom179)) (step t253 (cl @p_55 @p_56) :rule or :premises (axiom180)) (step t254 (cl x48$ x47$ @p_27) :rule resolution :premises (t222 t165 t162 t201 t200 t207 t159 t160 t187 t186 t184 t154 t177 t176 t155 t192 t191 t161 t180 t179 t157 t197 t196 t164 t214 t210)) (step t255 (cl x47$ x45$ x12$ x38$ @p_9) :rule resolution :premises (t196 t161 t164 t181 t215 t211 t155 t163 t195 t193 t191 t160 t178 t177 t154 t187 t183)) (step t256 (cl x47$ x45$ x17$ x50$ @p_32) :rule resolution :premises (t196 t161 t164 t181 t215 t211 t155 t163 t195 t193 t191 t160 t178 t177 t154 t186 t182 t162 t200 t159 t255 t206 t205 t202 t166 t222 t218)) (step t257 (cl x6$ x32$ x45$ x2$ x47$ @p_17) :rule resolution :premises (t204 t160 t163 t195 t194 t155 t180 t211 t157 t164 t198 t196)) (step t258 (cl x10$ x37$ x17$ x50$ @p_11) :rule resolution :premises (t219 t162 t166 t207 t206)) (step t259 (cl x50$ x47$ x48$ x19$ x52$) :rule resolution :premises (t219 t162 t166 t203 t202 t255 t258 t185 t159 t201 t199 t165 t256 t225 t224 t167 t254)) (step t260 (cl x47$ x38$ @p_9) :rule resolution :premises (t212 t161 t164 t198 t197 t157 t181 t180 t155 t195 t191 t160 t178 t177 t154 t187 t183)) (step t261 (cl x50$ x16$ x9$ x19$ x52$ x47$ @p_6) :rule resolution :premises (t202 t163 t166 t190 t225 t224 t161 t167 t212 t211 t210 t164 t197 t157 t181 t180)) (step t262 (cl x38$ x3$ x43$ x10$ x50$ x16$ x19$ x52$ @p_26) :rule resolution :premises (t176 t160 t155 t182 t194 t193 t162 t163 t205 t202 t166 t225 t224 t167 t211 t210)) (step t263 (cl x45$ x12$ x9$ x34$ @p_15) :rule resolution :premises (t215 t163 t161 t194 t192)) (step t264 (cl x45$ x12$ x9$ x34$ x38$ x3$) :rule resolution :premises (t215 t163 t161 t193 t191 t160 t184 t154 t177 t176 t155 t263)) (step t265 (cl x38$ x9$ x34$ x19$ x52$ x43$ x10$ x50$ x16$ @p_22) :rule resolution :premises (t191 t160 t161 t178 t214 t154 t167 t186 t182 t224 t162 t166 t203 t202)) (step t266 (cl x38$ x43$ x10$ x9$ x50$ x16$ x19$ x52$ x47$) :rule resolution :premises (t176 t155 t160 t182 t192 t191 t162 t161 t205 t214 t166 t167 t228 t225 t264 t265 t179 t157 t197 t164 t262 t261)) (step t267 (cl x38$ x43$ x50$ x16$ x19$ x52$ x47$) :rule resolution :premises (t180 t262 t157 t164 t198 t196 t266 t200 t159 t260)) (step t268 (cl x47$ x19$ x52$ @p_44) :rule resolution :premises (t179 t161 t157 t190 t197 t196 t163 t164 t214 t213 t210 t167 t228 t224 t166 t206 t204 t267 t221 t217 t168 t250 t172 t235 t234 t259 t240 t239)) (step t269 (cl x56$ x23$ x11$ @p_11 x12$ x19$ x52$ x47$ @p_5) :rule resolution :premises (t248 t172 t173 t235 t253 t165 t169 t222 t238 t162 t258 t201 t200 t159 t187 t186 t154 t177 t155 t194 t163 t225 t224 t223 t167 t211 t210 t164 t197 t157 t181 t179)) (step t270 (cl x9$ x34$ x45$ x12$ @p_26) :rule resolution :premises (t190 t161 t163 t212 t211)) (step t271 (cl x44$ x23$ x56$ x11$ @p_10) :rule resolution :premises (t226 t166 t169 t241 t237 t253 t168 t173 t250 t248 t172 t235 t234 t165 t222 t221 t219 t162 t201 t200 t159 t187 t186)) (step t272 (cl x47$ x2$ x9$ x34$ x45$ x12$) :rule resolution :premises (t157 t197 t180 t164 t155 t270 t263)) (step t273 (cl x2$ x9$ x34$ x47$ x19$ x52$ x12$) :rule resolution :premises (t180 t155 t157 t192 t197 t161 t164 t214 t210 t167 t225 t272)) (step t274 (cl x19$ x52$ x47$ x23$ x56$ @p_11) :rule resolution :premises (t167 t210 t225 t164 t257 t196 t273 t177 t154 t271 t269 t207 t206 t204 t189 t188)) (step t275 (cl x38$ x12$ x19$ x52$ x47$ @p_17) :rule resolution :premises (t184 t154 t160 t177 t176 t155 t194 t193 t163 t225 t167 t180 t211 t210 t157 t164 t198 t196)) (step t276 (cl x38$ x45$ x12$ x9$ @p_26) :rule resolution :premises (t184 t154 t160 t177 t176 t155 t194 t193 t181 t163 t270 t211)) (step t277 (cl x45$ x12$ x9$ x38$ x47$) :rule resolution :premises (t264 t180 t179 t157 t197 t164 t276)) (step t278 (cl x38$ x47$ x19$ x52$ x12$) :rule resolution :premises (t184 t154 t160 t191 t177 t176 t161 t273 t179 t157 t197 t164 t214 t210 t167 t225 t277 t275)) (step t279 (cl x42$ x44$ x11$ x10$ x19$ x52$ x38$ x32$ x9$ x2$ @p_4) :rule resolution :premises (t250 t172 t168 t235 t234 t241 t165 t166 t222 t221 t219 t224 t162 t167 t182 t214 t160 t161 t195 t192 t155 t180 t179)) (step t280 (cl x42$ x9$ x19$ x52$ x47$ x11$ x10$ x38$ x32$ x40$ x34$ x4$) :rule resolution :premises (t250 t172 t168 t235 t234 t241 t165 t261 t222 t221 t219 t162 t182 t160 t191 t161 t156)) (step t281 (cl x42$ x44$ x11$ x10$ x38$ x32$ x2$ x9$ x19$ x52$ x47$) :rule resolution :premises (t250 t172 t168 t235 t234 t241 t165 t166 t222 t221 t219 t162 t182 t160 t195 t155 t181 t280 t224 t167 t212 t210 t164 t197 t157 t279)) (step t282 (cl @p_48) :rule resolution :premises (t250 t172 t168 t235 t234 t241 t165 t166 t224 t222 t221 t219 t167 t162 t210 t182 t164 t257 t196 t281 t178 t177 t154 t271 t208 t203 t202 t278 t201 t200 t159 t260 t274 t268 t231 t230 t171 t246 t245 t242)) (step t283 (cl x28$) :rule resolution :premises (t175 t282)) (step t284 (cl @p_55) :rule resolution :premises (t251 t283)) (step t285 (cl @p_56) :rule resolution :premises (t252 t283)) (step t286 (cl x42$ x47$ x19$ x52$) :rule resolution :premises (t168 t220 t241 t165 t259 t235 t233 t172 t249 t248 t173 t268 t284)) (step t287 (cl x42$ @p_33) :rule resolution :premises (t239 t168 t173 t250 t248 t172 t235 t234 t165 t222 t221 t284)) (step t288 (cl x47$ @p_17 x43$ x10$ @p_22) :rule resolution :premises (t257 t178 t177 t154 t186 t182 t162 t208 t203)) (step t289 (cl x38$ x3$ @p_13) :rule resolution :premises (t184 t154 t160 t177 t176 t155 t194 t193)) (step t290 (cl x17$ x44$ @p_43) :rule resolution :premises (t233 t168 t172 t217 t249 t248 t166 t173 t238 t237 t236 t284)) (step t291 (cl x18$ x46$ @p_49) :rule resolution :premises (t230 t167 t171 t247 t243)) (step t292 (cl x39$ x47$ @p_17 x38$ x41$) :rule resolution :premises (t220 t165 t168 t235 t233 t172 t249 t248 t173 t237 t236 t169 t291 t227 t225 t163 t288 t287 t201 t200 t159 t211 t210 t260 t164 t284 t285)) (step t293 (cl x23$ @p_25) :rule resolution :premises (t185 t159 t162 t201 t199 t165 t221 t220 t168 t235 t233 t172 t249 t248 t205 t173 t290 t237 t236 t169 t228 t227 t284 t285)) (step t294 (cl x23$ @p_34) :rule resolution :premises (t185 t159 t162 t201 t199 t165 t221 t220 t168 t235 t233 t172 t249 t248 t205 t173 t290 t237 t236 t169 t224 t223 t284 t285)) (step t295 (cl x38$ x3$ @p_8) :rule resolution :premises (t195 t160 t155 t178 t177 t154 t186 t182)) (step t296 (cl x38$ x3$ @p_9) :rule resolution :premises (t195 t160 t155 t178 t177 t154 t187 t183)) (step t297 (cl x38$ x3$ @p_17 x41$) :rule resolution :premises (t287 t162 t201 t200 t159 t296 t295 t203 t163 t213 t167 t294 t293 t244 t243 t171 t232 t231 t292 t289)) (step t298 (cl x12$ x39$ @p_39) :rule resolution :premises (t213 t167 t163 t294 t293 t244 t243 t171 t232 t231)) (step t299 (cl x6$ x32$ @p_17 x12$ x3$ @p_20) :rule resolution :premises (t291 t293 t225 t257 t254 t298 t194 t155 t177 t154 t187 t159 t216 t201)) (step t300 (cl x46$ @p_25) :rule resolution :premises (t291 t293 t225)) (step t301 (cl x11$ x43$ @p_10) :rule resolution :premises (t200 t159 t162 t187 t186)) (step t302 (cl @p_17) :rule resolution :premises (t300 t210 t257 t164 t298 t194 t155 t177 t154 t301 t287 t299 t207 t204 t189 t188 t297 t180 t157 t198 t196)) (step t303 (cl @p_25) :rule resolution :premises (t216 t286 t254 t229 t247 t298 t300 t293 t209 t208)) (step t304 (cl x12$ x47$ x2$) :rule resolution :premises (t211 t163 t164 t194 t197 t155 t157 t181 t179 t272 t303 t302)) (step t305 (cl x12$ x47$) :rule resolution :premises (t201 t287 t159 t301 t187 t154 t177 t304 t207 t277 t302 t303)) (step t306 (cl x47$) :rule resolution :premises (t212 t164 t161 t197 t191 t157 t160 t181 t180 t178 t295 t182 t154 t162 t301 t287 t201 t200 t159 t260 t204 t203 t305 t302)) (step t307 (cl @p_38) :rule resolution :premises (t231 t306)) (step t308 (cl @p_37) :rule resolution :premises (t232 t306)) (step t309 (cl x24$) :rule resolution :premises (t171 t307)) (step t310 (cl @p_49) :rule resolution :premises (t243 t309)) (step t311 (cl @p_50) :rule resolution :premises (t244 t309)) (step t312 (cl @p_34) :rule resolution :premises (t294 t310)) (step t313 (cl x46$) :rule resolution :premises (t167 t312 t308 t311)) (step t314 (cl @p_29) :rule resolution :premises (t214 t313)) (step t315 (cl x11$ x43$ @p_7) :rule resolution :premises (t200 t159 t162 t183 t182)) (step t316 (cl x38$ x11$ x43$) :rule resolution :premises (t181 t155 t161 t195 t191 t160 t315 t178 t177 t154 t301 t302 t314)) (step t317 (cl @p_22) :rule resolution :premises (t191 t160 t161 t178 t181 t154 t296 t187 t183 t159 t201 t287 t316 t204 t203 t302 t314)) (step t318 (cl x39$) :rule resolution :premises (t298 t317 t306)) (step t319 (cl @p_12) :rule resolution :premises (t190 t318)) (step t320 (cl @p_15) :rule resolution :premises (t194 t318)) (step t321 (cl x34$) :rule resolution :premises (t161 t319 t302 t314)) (step t322 (cl @p_6) :rule resolution :premises (t181 t321)) (step t323 (cl x2$) :rule resolution :premises (t155 t322 t320)) (step t324 (cl x38$) :rule resolution :premises (t289 t322 t318)) (step t325 (cl @p_3) :rule resolution :premises (t177 t323)) (step t326 (cl @p_24) :rule resolution :premises (t207 t324)) (step t327 (cl x31$) :rule resolution :premises (t154 t325)) (step t328 (cl @p_9) :rule resolution :premises (t187 t327)) (step t329 (cl x43$) :rule resolution :premises (t301 t327 t326)) (step t330 (cl x36$) :rule resolution :premises (t159 t328)) (step t331 (cl x42$) :rule resolution :premises (t287 t329)) (step t332 (cl) :rule resolution :premises (t201 t330 t331)) 5895c6070af96c275f3f46cd6f9c0ddc1803c656 64 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr2 Int)) (not (! (=> (! (p$ veriT_vr2) :named @p_20) (! (forall ((veriT_vr3 Int)) (! (or @p_20 (! (p$ veriT_vr3) :named @p_24)) :named @p_25)) :named @p_21)) :named @p_26))) :named @p_33)) (define-fun veriT_sk1 () Int (! (choice ((veriT_vr3 Int)) (not (or (p$ @p_33) @p_24))) :named @p_37)) (assume axiom0 (! (not (! (forall ((?v0 Int)) (! (=> (! (p$ ?v0) :named @p_1) (! (forall ((?v1 Int)) (! (or @p_1 (! (p$ ?v1) :named @p_8)) :named @p_10)) :named @p_4)) :named @p_12)) :named @p_2)) :named @p_14)) (anchor :step t2 :args ((:= (?v0 Int) veriT_vr0))) (step t2.t1 (cl (! (= ?v0 veriT_vr0) :named @p_6)) :rule refl) (step t2.t2 (cl (! (= @p_1 (! (p$ veriT_vr0) :named @p_3)) :named @p_7)) :rule cong :premises (t2.t1)) (anchor :step t2.t3 :args ((:= (?v1 Int) veriT_vr1))) (step t2.t3.t1 (cl @p_6) :rule refl) (step t2.t3.t2 (cl @p_7) :rule cong :premises (t2.t3.t1)) (step t2.t3.t3 (cl (= ?v1 veriT_vr1)) :rule refl) (step t2.t3.t4 (cl (= @p_8 (! (p$ veriT_vr1) :named @p_9))) :rule cong :premises (t2.t3.t3)) (step t2.t3.t5 (cl (= @p_10 (! (or @p_3 @p_9) :named @p_11))) :rule cong :premises (t2.t3.t2 t2.t3.t4)) (step t2.t3 (cl (= @p_4 (! (forall ((veriT_vr1 Int)) @p_11) :named @p_5))) :rule bind) (step t2.t4 (cl (= @p_12 (! (=> @p_3 @p_5) :named @p_13))) :rule cong :premises (t2.t2 t2.t3)) (step t2 (cl (= @p_2 (! (forall ((veriT_vr0 Int)) @p_13) :named @p_15))) :rule bind) (step t3 (cl (! (= @p_14 (! (not @p_15) :named @p_17)) :named @p_16)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_16) :named @p_19) (! (not @p_14) :named @p_18) @p_17) :rule equiv_pos2) (step t5 (cl (not @p_18) @p_2) :rule not_not) (step t6 (cl @p_19 @p_2 @p_17) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_17) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((:= (veriT_vr0 Int) veriT_vr2))) (step t8.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_22)) :rule refl) (step t8.t2 (cl (! (= @p_3 @p_20) :named @p_23)) :rule cong :premises (t8.t1)) (anchor :step t8.t3 :args ((:= (veriT_vr1 Int) veriT_vr3))) (step t8.t3.t1 (cl @p_22) :rule refl) (step t8.t3.t2 (cl @p_23) :rule cong :premises (t8.t3.t1)) (step t8.t3.t3 (cl (= veriT_vr1 veriT_vr3)) :rule refl) (step t8.t3.t4 (cl (= @p_9 @p_24)) :rule cong :premises (t8.t3.t3)) (step t8.t3.t5 (cl (= @p_11 @p_25)) :rule cong :premises (t8.t3.t2 t8.t3.t4)) (step t8.t3 (cl (= @p_5 @p_21)) :rule bind) (step t8.t4 (cl (= @p_13 @p_26)) :rule cong :premises (t8.t2 t8.t3)) (step t8 (cl (= @p_15 (! (forall ((veriT_vr2 Int)) @p_26) :named @p_27))) :rule bind) (step t9 (cl (! (= @p_17 (! (not @p_27) :named @p_29)) :named @p_28)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_28) :named @p_31) (! (not @p_17) :named @p_30) @p_29) :rule equiv_pos2) (step t11 (cl (not @p_30) @p_15) :rule not_not) (step t12 (cl @p_31 @p_15 @p_29) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_29) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((:= (veriT_vr2 Int) veriT_sk0))) (step t14.t1 (cl (! (= veriT_vr2 veriT_sk0) :named @p_35)) :rule refl) (step t14.t2 (cl (! (= @p_20 (! (p$ veriT_sk0) :named @p_32)) :named @p_36)) :rule cong :premises (t14.t1)) (anchor :step t14.t3 :args ((:= (veriT_vr3 Int) veriT_sk1))) (step t14.t3.t1 (cl @p_35) :rule refl) (step t14.t3.t2 (cl @p_36) :rule cong :premises (t14.t3.t1)) (step t14.t3.t3 (cl (= veriT_vr3 veriT_sk1)) :rule refl) (step t14.t3.t4 (cl (= @p_24 (! (p$ veriT_sk1) :named @p_38))) :rule cong :premises (t14.t3.t3)) (step t14.t3.t5 (cl (= @p_25 (! (or @p_32 @p_38) :named @p_34))) :rule cong :premises (t14.t3.t2 t14.t3.t4)) (step t14.t3 (cl (= @p_21 @p_34)) :rule sko_forall) (step t14.t4 (cl (= @p_26 (! (=> @p_32 @p_34) :named @p_39))) :rule cong :premises (t14.t2 t14.t3)) (step t14 (cl (= @p_27 @p_39)) :rule sko_forall) (step t15 (cl (! (= @p_29 (! (not @p_39) :named @p_41)) :named @p_40)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_40) :named @p_43) (! (not @p_29) :named @p_42) @p_41) :rule equiv_pos2) (step t17 (cl (not @p_42) @p_27) :rule not_not) (step t18 (cl @p_43 @p_27 @p_41) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_41) :rule th_resolution :premises (t13 t15 t18)) (step t20 (cl (! (= @p_41 (! (and @p_32 (! (not @p_34) :named @p_48)) :named @p_45)) :named @p_44)) :rule bool_simplify) (step t21 (cl (! (not @p_44) :named @p_47) (! (not @p_41) :named @p_46) @p_45) :rule equiv_pos2) (step t22 (cl (not @p_46) @p_39) :rule not_not) (step t23 (cl @p_47 @p_39 @p_45) :rule th_resolution :premises (t22 t21)) (step t24 (cl @p_45) :rule th_resolution :premises (t19 t20 t23)) (step t25 (cl @p_32) :rule and :premises (t24)) (step t26 (cl @p_48) :rule and :premises (t24)) (step t27 (cl (not @p_32)) :rule not_or :premises (t26)) (step t28 (cl) :rule resolution :premises (t27 t25)) 92a1094a80c0dcb33184d755df79398cf322af19 155 0 unsat (define-fun veriT_sk0 () A$ (! (choice ((veriT_vr3 A$)) (! (ite x$ (! (p$ true veriT_vr3) :named @p_48) (! (p$ false veriT_vr3) :named @p_50)) :named @p_51)) :named @p_62)) (assume axiom0 (forall ((?v0 Bool) (?v1 A$)) (= (p$ ?v0 ?v1) ?v0))) (assume axiom1 (not (= (exists ((?v0 A$)) (p$ x$ ?v0)) (p$ x$ c$)))) (step t3 (cl (! (forall ((?v1 A$)) (! (and (! (= (! (p$ false ?v1) :named @p_2) false) :named @p_4) (! (= (! (p$ true ?v1) :named @p_7) true) :named @p_9)) :named @p_11)) :named @p_1)) :rule bfun_elim :premises (axiom0)) (step t4 (cl (! (not (! (= (! (exists ((?v0 A$)) (! (ite x$ (! (p$ true ?v0) :named @p_27) (! (p$ false ?v0) :named @p_30)) :named @p_32)) :named @p_26) (! (ite x$ (! (p$ true c$) :named @p_91) (! (p$ false c$) :named @p_93)) :named @p_36)) :named @p_34)) :named @p_37)) :rule bfun_elim :premises (axiom1)) (anchor :step t5 :args ((:= (?v1 A$) veriT_vr0))) (step t5.t1 (cl (! (= ?v1 veriT_vr0) :named @p_6)) :rule refl) (step t5.t2 (cl (= @p_2 (! (p$ false veriT_vr0) :named @p_3))) :rule cong :premises (t5.t1)) (step t5.t3 (cl (= @p_4 (! (= @p_3 false) :named @p_5))) :rule cong :premises (t5.t2)) (step t5.t4 (cl @p_6) :rule refl) (step t5.t5 (cl (= @p_7 (! (p$ true veriT_vr0) :named @p_8))) :rule cong :premises (t5.t4)) (step t5.t6 (cl (= @p_9 (! (= @p_8 true) :named @p_10))) :rule cong :premises (t5.t5)) (step t5.t7 (cl (= @p_11 (! (and @p_5 @p_10) :named @p_12))) :rule cong :premises (t5.t3 t5.t6)) (step t5 (cl (! (= @p_1 (! (forall ((veriT_vr0 A$)) @p_12) :named @p_14)) :named @p_13)) :rule bind) (step t6 (cl (not @p_13) (not @p_1) @p_14) :rule equiv_pos2) (step t7 (cl @p_14) :rule th_resolution :premises (t3 t5 t6)) (anchor :step t8 :args ((veriT_vr0 A$))) (step t8.t1 (cl (= @p_5 (! (not @p_3) :named @p_15))) :rule equiv_simplify) (step t8.t2 (cl (= @p_10 @p_8)) :rule equiv_simplify) (step t8.t3 (cl (= @p_12 (! (and @p_15 @p_8) :named @p_16))) :rule cong :premises (t8.t1 t8.t2)) (step t8 (cl (! (= @p_14 (! (forall ((veriT_vr0 A$)) @p_16) :named @p_18)) :named @p_17)) :rule bind) (step t9 (cl (not @p_17) (not @p_14) @p_18) :rule equiv_pos2) (step t10 (cl @p_18) :rule th_resolution :premises (t7 t8 t9)) (anchor :step t11 :args ((:= (veriT_vr0 A$) veriT_vr1))) (step t11.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_21)) :rule refl) (step t11.t2 (cl (= @p_3 (! (p$ false veriT_vr1) :named @p_19))) :rule cong :premises (t11.t1)) (step t11.t3 (cl (= @p_15 (! (not @p_19) :named @p_20))) :rule cong :premises (t11.t2)) (step t11.t4 (cl @p_21) :rule refl) (step t11.t5 (cl (= @p_8 (! (p$ true veriT_vr1) :named @p_22))) :rule cong :premises (t11.t4)) (step t11.t6 (cl (= @p_16 (! (and @p_20 @p_22) :named @p_23))) :rule cong :premises (t11.t3 t11.t5)) (step t11 (cl (! (= @p_18 (! (forall ((veriT_vr1 A$)) @p_23) :named @p_25)) :named @p_24)) :rule bind) (step t12 (cl (not @p_24) (not @p_18) @p_25) :rule equiv_pos2) (step t13 (cl @p_25) :rule th_resolution :premises (t10 t11 t12)) (anchor :step t14 :args ((:= (?v0 A$) veriT_vr2))) (step t14.t1 (cl (! (= ?v0 veriT_vr2) :named @p_29)) :rule refl) (step t14.t2 (cl (= @p_27 (! (p$ true veriT_vr2) :named @p_28))) :rule cong :premises (t14.t1)) (step t14.t3 (cl @p_29) :rule refl) (step t14.t4 (cl (= @p_30 (! (p$ false veriT_vr2) :named @p_31))) :rule cong :premises (t14.t3)) (step t14.t5 (cl (= @p_32 (! (ite x$ @p_28 @p_31) :named @p_33))) :rule cong :premises (t14.t2 t14.t4)) (step t14 (cl (= @p_26 (! (exists ((veriT_vr2 A$)) @p_33) :named @p_35))) :rule bind) (step t15 (cl (= @p_34 (! (= @p_35 @p_36) :named @p_38))) :rule cong :premises (t14)) (step t16 (cl (! (= @p_37 (! (not @p_38) :named @p_40)) :named @p_39)) :rule cong :premises (t15)) (step t17 (cl (! (not @p_39) :named @p_42) (! (not @p_37) :named @p_41) @p_40) :rule equiv_pos2) (step t18 (cl (not @p_41) @p_34) :rule not_not) (step t19 (cl @p_42 @p_34 @p_40) :rule th_resolution :premises (t18 t17)) (step t20 (cl @p_40) :rule th_resolution :premises (t4 t16 t19)) (step t21 (cl (= @p_38 (! (and (! (=> @p_35 @p_36) :named @p_52) (! (=> @p_36 @p_35) :named @p_54)) :named @p_43))) :rule connective_def) (step t22 (cl (! (= @p_40 (! (not @p_43) :named @p_45)) :named @p_44)) :rule cong :premises (t21)) (step t23 (cl (! (not @p_44) :named @p_47) (! (not @p_40) :named @p_46) @p_45) :rule equiv_pos2) (step t24 (cl (not @p_46) @p_38) :rule not_not) (step t25 (cl @p_47 @p_38 @p_45) :rule th_resolution :premises (t24 t23)) (step t26 (cl @p_45) :rule th_resolution :premises (t20 t22 t25)) (anchor :step t27 :args ((:= (veriT_vr2 A$) veriT_vr3))) (step t27.t1 (cl (! (= veriT_vr2 veriT_vr3) :named @p_49)) :rule refl) (step t27.t2 (cl (= @p_28 @p_48)) :rule cong :premises (t27.t1)) (step t27.t3 (cl @p_49) :rule refl) (step t27.t4 (cl (= @p_31 @p_50)) :rule cong :premises (t27.t3)) (step t27.t5 (cl (= @p_33 @p_51)) :rule cong :premises (t27.t2 t27.t4)) (step t27 (cl (= @p_35 (! (exists ((veriT_vr3 A$)) @p_51) :named @p_53))) :rule bind) (step t28 (cl (= @p_52 (! (=> @p_53 @p_36) :named @p_55))) :rule cong :premises (t27)) (step t29 (cl (= @p_54 (! (=> @p_36 @p_53) :named @p_56))) :rule cong :premises (t27)) (step t30 (cl (= @p_43 (! (and @p_55 @p_56) :named @p_57))) :rule cong :premises (t28 t29)) (step t31 (cl (! (= @p_45 (! (not @p_57) :named @p_59)) :named @p_58)) :rule cong :premises (t30)) (step t32 (cl (! (not @p_58) :named @p_61) (! (not @p_45) :named @p_60) @p_59) :rule equiv_pos2) (step t33 (cl (not @p_60) @p_43) :rule not_not) (step t34 (cl @p_61 @p_43 @p_59) :rule th_resolution :premises (t33 t32)) (step t35 (cl @p_59) :rule th_resolution :premises (t26 t31 t34)) (anchor :step t36 :args ((:= (veriT_vr3 A$) veriT_sk0))) (step t36.t1 (cl (! (= veriT_vr3 veriT_sk0) :named @p_64)) :rule refl) (step t36.t2 (cl (= @p_48 (! (p$ true veriT_sk0) :named @p_63))) :rule cong :premises (t36.t1)) (step t36.t3 (cl @p_64) :rule refl) (step t36.t4 (cl (= @p_50 (! (p$ false veriT_sk0) :named @p_65))) :rule cong :premises (t36.t3)) (step t36.t5 (cl (= @p_51 (! (ite x$ @p_63 @p_65) :named @p_66))) :rule cong :premises (t36.t2 t36.t4)) (step t36 (cl (= @p_53 @p_66)) :rule sko_ex) (step t37 (cl (= @p_55 (! (=> @p_66 @p_36) :named @p_67))) :rule cong :premises (t36)) (step t38 (cl (= @p_57 (! (and @p_67 @p_56) :named @p_68))) :rule cong :premises (t37)) (step t39 (cl (! (= @p_59 (! (not @p_68) :named @p_70)) :named @p_69)) :rule cong :premises (t38)) (step t40 (cl (! (not @p_69) :named @p_72) (! (not @p_59) :named @p_71) @p_70) :rule equiv_pos2) (step t41 (cl (not @p_71) @p_57) :rule not_not) (step t42 (cl @p_72 @p_57 @p_70) :rule th_resolution :premises (t41 t40)) (step t43 (cl @p_70) :rule th_resolution :premises (t35 t39 t42)) (anchor :step t44 :args ((:= (veriT_vr3 A$) veriT_vr4))) (step t44.t1 (cl (! (= veriT_vr3 veriT_vr4) :named @p_74)) :rule refl) (step t44.t2 (cl (= @p_48 (! (p$ true veriT_vr4) :named @p_73))) :rule cong :premises (t44.t1)) (step t44.t3 (cl @p_74) :rule refl) (step t44.t4 (cl (= @p_50 (! (p$ false veriT_vr4) :named @p_75))) :rule cong :premises (t44.t3)) (step t44.t5 (cl (= @p_51 (! (ite x$ @p_73 @p_75) :named @p_76))) :rule cong :premises (t44.t2 t44.t4)) (step t44 (cl (= @p_53 (! (exists ((veriT_vr4 A$)) @p_76) :named @p_77))) :rule bind) (step t45 (cl (= @p_56 (! (=> @p_36 @p_77) :named @p_78))) :rule cong :premises (t44)) (step t46 (cl (= @p_68 (! (and @p_67 @p_78) :named @p_79))) :rule cong :premises (t45)) (step t47 (cl (! (= @p_70 (! (not @p_79) :named @p_81)) :named @p_80)) :rule cong :premises (t46)) (step t48 (cl (! (not @p_80) :named @p_83) (! (not @p_70) :named @p_82) @p_81) :rule equiv_pos2) (step t49 (cl (not @p_82) @p_68) :rule not_not) (step t50 (cl @p_83 @p_68 @p_81) :rule th_resolution :premises (t49 t48)) (step t51 (cl @p_81) :rule th_resolution :premises (t43 t47 t50)) (step t52 (cl (= @p_77 (! (not (! (forall ((veriT_vr4 A$)) (not @p_76)) :named @p_95)) :named @p_84))) :rule connective_def) (step t53 (cl (= @p_78 (! (=> @p_36 @p_84) :named @p_85))) :rule cong :premises (t52)) (step t54 (cl (= @p_79 (! (and @p_67 @p_85) :named @p_86))) :rule cong :premises (t53)) (step t55 (cl (! (= @p_81 (! (not @p_86) :named @p_88)) :named @p_87)) :rule cong :premises (t54)) (step t56 (cl (! (not @p_87) :named @p_90) (! (not @p_81) :named @p_89) @p_88) :rule equiv_pos2) (step t57 (cl (not @p_89) @p_79) :rule not_not) (step t58 (cl @p_90 @p_79 @p_88) :rule th_resolution :premises (t57 t56)) (step t59 (cl @p_88) :rule th_resolution :premises (t51 t55 t58)) (step t60 (cl (not @p_66) x$ @p_65) :rule ite_pos1) (step t61 (cl @p_67 @p_66) :rule implies_neg1) (step t62 (cl @p_36 (not x$) (! (not @p_91) :named @p_109)) :rule ite_neg2) (step t63 (cl @p_67 (! (not @p_36) :named @p_92)) :rule implies_neg2) (step t64 (cl @p_92 x$ @p_93) :rule ite_pos1) (step t65 (cl @p_85 @p_36) :rule implies_neg1) (step t66 (cl @p_85 (! (not @p_84) :named @p_94)) :rule implies_neg2) (step t67 (cl (not @p_94) @p_95) :rule not_not) (step t68 (cl @p_85 @p_95) :rule th_resolution :premises (t67 t66)) (step t69 (cl (not @p_67) (! (not @p_85) :named @p_110)) :rule not_and :premises (t59)) (step t70 (cl (or (! (not @p_25) :named @p_96) (! (forall ((veriT_vr1 A$)) @p_20) :named @p_97))) :rule qnt_cnf) (step t71 (cl (or @p_96 (! (forall ((veriT_vr1 A$)) @p_22) :named @p_106))) :rule qnt_cnf) (step t72 (cl @p_96 @p_97) :rule or :premises (t70)) (step t73 (cl (or (! (not @p_97) :named @p_98) (! (not @p_65) :named @p_99))) :rule forall_inst :args ((:= veriT_vr1 veriT_sk0))) (step t74 (cl @p_98 @p_99) :rule or :premises (t73)) (step t75 (cl (! (or @p_96 @p_99) :named @p_101) (! (not @p_96) :named @p_100)) :rule or_neg) (step t76 (cl (not @p_100) @p_25) :rule not_not) (step t77 (cl @p_101 @p_25) :rule th_resolution :premises (t76 t75)) (step t78 (cl @p_101 (! (not @p_99) :named @p_102)) :rule or_neg) (step t79 (cl (not @p_102) @p_65) :rule not_not) (step t80 (cl @p_101 @p_65) :rule th_resolution :premises (t79 t78)) (step t81 (cl @p_101) :rule th_resolution :premises (t72 t74 t77 t80)) (step t82 (cl @p_96 @p_99) :rule or :premises (t81)) (step t83 (cl @p_99) :rule resolution :premises (t82 t13)) (step t84 (cl (or @p_98 (! (not @p_93) :named @p_103))) :rule forall_inst :args ((:= veriT_vr1 c$))) (step t85 (cl @p_98 @p_103) :rule or :premises (t84)) (step t86 (cl (! (or @p_96 @p_103) :named @p_104) @p_100) :rule or_neg) (step t87 (cl @p_104 @p_25) :rule th_resolution :premises (t76 t86)) (step t88 (cl @p_104 (! (not @p_103) :named @p_105)) :rule or_neg) (step t89 (cl (not @p_105) @p_93) :rule not_not) (step t90 (cl @p_104 @p_93) :rule th_resolution :premises (t89 t88)) (step t91 (cl @p_104) :rule th_resolution :premises (t72 t85 t87 t90)) (step t92 (cl @p_96 @p_103) :rule or :premises (t91)) (step t93 (cl @p_103) :rule resolution :premises (t92 t13)) (step t94 (cl x$) :rule resolution :premises (t69 t65 t61 t64 t60 t93 t83)) (step t95 (cl @p_96 @p_106) :rule or :premises (t71)) (step t96 (cl (or (! (not @p_106) :named @p_107) @p_91)) :rule forall_inst :args ((:= veriT_vr1 c$))) (step t97 (cl @p_107 @p_91) :rule or :premises (t96)) (step t98 (cl (! (or @p_96 @p_91) :named @p_108) @p_100) :rule or_neg) (step t99 (cl @p_108 @p_25) :rule th_resolution :premises (t76 t98)) (step t100 (cl @p_108 @p_109) :rule or_neg) (step t101 (cl @p_108) :rule th_resolution :premises (t95 t97 t99 t100)) (step t102 (cl @p_96 @p_91) :rule or :premises (t101)) (step t103 (cl @p_91) :rule resolution :premises (t102 t13)) (step t104 (cl @p_36) :rule resolution :premises (t62 t103 t94)) (step t105 (cl @p_67) :rule resolution :premises (t63 t104)) (step t106 (cl @p_110) :rule resolution :premises (t69 t105)) (step t107 (cl @p_95) :rule resolution :premises (t68 t106)) (step t108 (cl (or @p_84 @p_92)) :rule forall_inst :args ((:= veriT_vr4 c$))) (step t109 (cl @p_84 @p_92) :rule or :premises (t108)) (step t110 (cl) :rule resolution :premises (t109 t104 t107)) 811ced85456c84c67b4bd2d5cedbf22b20ed06ce 143 0 unsat (define-fun veriT_sk2 () A$ (! (choice ((veriT_vr9 A$)) (! (ite x$ (! (p$ true veriT_vr9) :named @p_48) (! (p$ false veriT_vr9) :named @p_50)) :named @p_51)) :named @p_62)) (assume axiom0 (forall ((?v0 Bool) (?v1 A$)) (= (p$ ?v0 ?v1) ?v0))) (assume axiom2 (not (= (exists ((?v0 A$)) (p$ x$ ?v0)) (p$ x$ c$)))) (step t3 (cl (! (forall ((?v1 A$)) (! (and (! (= (! (p$ false ?v1) :named @p_2) false) :named @p_4) (! (= (! (p$ true ?v1) :named @p_7) true) :named @p_9)) :named @p_11)) :named @p_1)) :rule bfun_elim :premises (axiom0)) (step t4 (cl (! (not (! (= (! (exists ((?v0 A$)) (! (ite x$ (! (p$ true ?v0) :named @p_27) (! (p$ false ?v0) :named @p_30)) :named @p_32)) :named @p_26) (! (ite x$ (! (p$ true c$) :named @p_91) (p$ false c$)) :named @p_36)) :named @p_34)) :named @p_37)) :rule bfun_elim :premises (axiom2)) (anchor :step t5 :args ((:= (?v1 A$) veriT_vr0))) (step t5.t1 (cl (! (= ?v1 veriT_vr0) :named @p_6)) :rule refl) (step t5.t2 (cl (= @p_2 (! (p$ false veriT_vr0) :named @p_3))) :rule cong :premises (t5.t1)) (step t5.t3 (cl (= @p_4 (! (= @p_3 false) :named @p_5))) :rule cong :premises (t5.t2)) (step t5.t4 (cl @p_6) :rule refl) (step t5.t5 (cl (= @p_7 (! (p$ true veriT_vr0) :named @p_8))) :rule cong :premises (t5.t4)) (step t5.t6 (cl (= @p_9 (! (= @p_8 true) :named @p_10))) :rule cong :premises (t5.t5)) (step t5.t7 (cl (= @p_11 (! (and @p_5 @p_10) :named @p_12))) :rule cong :premises (t5.t3 t5.t6)) (step t5 (cl (! (= @p_1 (! (forall ((veriT_vr0 A$)) @p_12) :named @p_14)) :named @p_13)) :rule bind) (step t6 (cl (not @p_13) (not @p_1) @p_14) :rule equiv_pos2) (step t7 (cl @p_14) :rule th_resolution :premises (t3 t5 t6)) (anchor :step t8 :args ((veriT_vr0 A$))) (step t8.t1 (cl (= @p_5 (! (not @p_3) :named @p_15))) :rule equiv_simplify) (step t8.t2 (cl (= @p_10 @p_8)) :rule equiv_simplify) (step t8.t3 (cl (= @p_12 (! (and @p_15 @p_8) :named @p_16))) :rule cong :premises (t8.t1 t8.t2)) (step t8 (cl (! (= @p_14 (! (forall ((veriT_vr0 A$)) @p_16) :named @p_18)) :named @p_17)) :rule bind) (step t9 (cl (not @p_17) (not @p_14) @p_18) :rule equiv_pos2) (step t10 (cl @p_18) :rule th_resolution :premises (t7 t8 t9)) (anchor :step t11 :args ((:= (veriT_vr0 A$) veriT_vr1))) (step t11.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_21)) :rule refl) (step t11.t2 (cl (= @p_3 (! (p$ false veriT_vr1) :named @p_19))) :rule cong :premises (t11.t1)) (step t11.t3 (cl (= @p_15 (! (not @p_19) :named @p_20))) :rule cong :premises (t11.t2)) (step t11.t4 (cl @p_21) :rule refl) (step t11.t5 (cl (= @p_8 (! (p$ true veriT_vr1) :named @p_22))) :rule cong :premises (t11.t4)) (step t11.t6 (cl (= @p_16 (! (and @p_20 @p_22) :named @p_23))) :rule cong :premises (t11.t3 t11.t5)) (step t11 (cl (! (= @p_18 (! (forall ((veriT_vr1 A$)) @p_23) :named @p_25)) :named @p_24)) :rule bind) (step t12 (cl (not @p_24) (not @p_18) @p_25) :rule equiv_pos2) (step t13 (cl @p_25) :rule th_resolution :premises (t10 t11 t12)) (anchor :step t14 :args ((:= (?v0 A$) veriT_vr8))) (step t14.t1 (cl (! (= ?v0 veriT_vr8) :named @p_29)) :rule refl) (step t14.t2 (cl (= @p_27 (! (p$ true veriT_vr8) :named @p_28))) :rule cong :premises (t14.t1)) (step t14.t3 (cl @p_29) :rule refl) (step t14.t4 (cl (= @p_30 (! (p$ false veriT_vr8) :named @p_31))) :rule cong :premises (t14.t3)) (step t14.t5 (cl (= @p_32 (! (ite x$ @p_28 @p_31) :named @p_33))) :rule cong :premises (t14.t2 t14.t4)) (step t14 (cl (= @p_26 (! (exists ((veriT_vr8 A$)) @p_33) :named @p_35))) :rule bind) (step t15 (cl (= @p_34 (! (= @p_35 @p_36) :named @p_38))) :rule cong :premises (t14)) (step t16 (cl (! (= @p_37 (! (not @p_38) :named @p_40)) :named @p_39)) :rule cong :premises (t15)) (step t17 (cl (! (not @p_39) :named @p_42) (! (not @p_37) :named @p_41) @p_40) :rule equiv_pos2) (step t18 (cl (not @p_41) @p_34) :rule not_not) (step t19 (cl @p_42 @p_34 @p_40) :rule th_resolution :premises (t18 t17)) (step t20 (cl @p_40) :rule th_resolution :premises (t4 t16 t19)) (step t21 (cl (= @p_38 (! (and (! (=> @p_35 @p_36) :named @p_52) (! (=> @p_36 @p_35) :named @p_54)) :named @p_43))) :rule connective_def) (step t22 (cl (! (= @p_40 (! (not @p_43) :named @p_45)) :named @p_44)) :rule cong :premises (t21)) (step t23 (cl (! (not @p_44) :named @p_47) (! (not @p_40) :named @p_46) @p_45) :rule equiv_pos2) (step t24 (cl (not @p_46) @p_38) :rule not_not) (step t25 (cl @p_47 @p_38 @p_45) :rule th_resolution :premises (t24 t23)) (step t26 (cl @p_45) :rule th_resolution :premises (t20 t22 t25)) (anchor :step t27 :args ((:= (veriT_vr8 A$) veriT_vr9))) (step t27.t1 (cl (! (= veriT_vr8 veriT_vr9) :named @p_49)) :rule refl) (step t27.t2 (cl (= @p_28 @p_48)) :rule cong :premises (t27.t1)) (step t27.t3 (cl @p_49) :rule refl) (step t27.t4 (cl (= @p_31 @p_50)) :rule cong :premises (t27.t3)) (step t27.t5 (cl (= @p_33 @p_51)) :rule cong :premises (t27.t2 t27.t4)) (step t27 (cl (= @p_35 (! (exists ((veriT_vr9 A$)) @p_51) :named @p_53))) :rule bind) (step t28 (cl (= @p_52 (! (=> @p_53 @p_36) :named @p_55))) :rule cong :premises (t27)) (step t29 (cl (= @p_54 (! (=> @p_36 @p_53) :named @p_56))) :rule cong :premises (t27)) (step t30 (cl (= @p_43 (! (and @p_55 @p_56) :named @p_57))) :rule cong :premises (t28 t29)) (step t31 (cl (! (= @p_45 (! (not @p_57) :named @p_59)) :named @p_58)) :rule cong :premises (t30)) (step t32 (cl (! (not @p_58) :named @p_61) (! (not @p_45) :named @p_60) @p_59) :rule equiv_pos2) (step t33 (cl (not @p_60) @p_43) :rule not_not) (step t34 (cl @p_61 @p_43 @p_59) :rule th_resolution :premises (t33 t32)) (step t35 (cl @p_59) :rule th_resolution :premises (t26 t31 t34)) (anchor :step t36 :args ((:= (veriT_vr9 A$) veriT_sk2))) (step t36.t1 (cl (! (= veriT_vr9 veriT_sk2) :named @p_64)) :rule refl) (step t36.t2 (cl (= @p_48 (! (p$ true veriT_sk2) :named @p_63))) :rule cong :premises (t36.t1)) (step t36.t3 (cl @p_64) :rule refl) (step t36.t4 (cl (= @p_50 (! (p$ false veriT_sk2) :named @p_65))) :rule cong :premises (t36.t3)) (step t36.t5 (cl (= @p_51 (! (ite x$ @p_63 @p_65) :named @p_66))) :rule cong :premises (t36.t2 t36.t4)) (step t36 (cl (= @p_53 @p_66)) :rule sko_ex) (step t37 (cl (= @p_55 (! (=> @p_66 @p_36) :named @p_67))) :rule cong :premises (t36)) (step t38 (cl (= @p_57 (! (and @p_67 @p_56) :named @p_68))) :rule cong :premises (t37)) (step t39 (cl (! (= @p_59 (! (not @p_68) :named @p_70)) :named @p_69)) :rule cong :premises (t38)) (step t40 (cl (! (not @p_69) :named @p_72) (! (not @p_59) :named @p_71) @p_70) :rule equiv_pos2) (step t41 (cl (not @p_71) @p_57) :rule not_not) (step t42 (cl @p_72 @p_57 @p_70) :rule th_resolution :premises (t41 t40)) (step t43 (cl @p_70) :rule th_resolution :premises (t35 t39 t42)) (anchor :step t44 :args ((:= (veriT_vr9 A$) veriT_vr10))) (step t44.t1 (cl (! (= veriT_vr9 veriT_vr10) :named @p_74)) :rule refl) (step t44.t2 (cl (= @p_48 (! (p$ true veriT_vr10) :named @p_73))) :rule cong :premises (t44.t1)) (step t44.t3 (cl @p_74) :rule refl) (step t44.t4 (cl (= @p_50 (! (p$ false veriT_vr10) :named @p_75))) :rule cong :premises (t44.t3)) (step t44.t5 (cl (= @p_51 (! (ite x$ @p_73 @p_75) :named @p_76))) :rule cong :premises (t44.t2 t44.t4)) (step t44 (cl (= @p_53 (! (exists ((veriT_vr10 A$)) @p_76) :named @p_77))) :rule bind) (step t45 (cl (= @p_56 (! (=> @p_36 @p_77) :named @p_78))) :rule cong :premises (t44)) (step t46 (cl (= @p_68 (! (and @p_67 @p_78) :named @p_79))) :rule cong :premises (t45)) (step t47 (cl (! (= @p_70 (! (not @p_79) :named @p_81)) :named @p_80)) :rule cong :premises (t46)) (step t48 (cl (! (not @p_80) :named @p_83) (! (not @p_70) :named @p_82) @p_81) :rule equiv_pos2) (step t49 (cl (not @p_82) @p_68) :rule not_not) (step t50 (cl @p_83 @p_68 @p_81) :rule th_resolution :premises (t49 t48)) (step t51 (cl @p_81) :rule th_resolution :premises (t43 t47 t50)) (step t52 (cl (= @p_77 (! (not (! (forall ((veriT_vr10 A$)) (not @p_76)) :named @p_93)) :named @p_84))) :rule connective_def) (step t53 (cl (= @p_78 (! (=> @p_36 @p_84) :named @p_85))) :rule cong :premises (t52)) (step t54 (cl (= @p_79 (! (and @p_67 @p_85) :named @p_86))) :rule cong :premises (t53)) (step t55 (cl (! (= @p_81 (! (not @p_86) :named @p_88)) :named @p_87)) :rule cong :premises (t54)) (step t56 (cl (! (not @p_87) :named @p_90) (! (not @p_81) :named @p_89) @p_88) :rule equiv_pos2) (step t57 (cl (not @p_89) @p_79) :rule not_not) (step t58 (cl @p_90 @p_79 @p_88) :rule th_resolution :premises (t57 t56)) (step t59 (cl @p_88) :rule th_resolution :premises (t51 t55 t58)) (step t60 (cl (not @p_66) x$ @p_65) :rule ite_pos1) (step t61 (cl @p_67 @p_66) :rule implies_neg1) (step t62 (cl @p_36 (not x$) (! (not @p_91) :named @p_104)) :rule ite_neg2) (step t63 (cl @p_67 (! (not @p_36) :named @p_106)) :rule implies_neg2) (step t64 (cl @p_85 @p_36) :rule implies_neg1) (step t65 (cl @p_85 (! (not @p_84) :named @p_92)) :rule implies_neg2) (step t66 (cl (not @p_92) @p_93) :rule not_not) (step t67 (cl @p_85 @p_93) :rule th_resolution :premises (t66 t65)) (step t68 (cl (not @p_67) (! (not @p_85) :named @p_105)) :rule not_and :premises (t59)) (step t69 (cl (or (! (not @p_25) :named @p_94) (! (forall ((veriT_vr1 A$)) @p_20) :named @p_95))) :rule qnt_cnf) (step t70 (cl (or @p_94 (! (forall ((veriT_vr1 A$)) @p_22) :named @p_101))) :rule qnt_cnf) (step t71 (cl @p_94 @p_95) :rule or :premises (t69)) (step t72 (cl (or (! (not @p_95) :named @p_96) (! (not @p_65) :named @p_97))) :rule forall_inst :args ((:= veriT_vr1 veriT_sk2))) (step t73 (cl @p_96 @p_97) :rule or :premises (t72)) (step t74 (cl (! (or @p_94 @p_97) :named @p_99) (! (not @p_94) :named @p_98)) :rule or_neg) (step t75 (cl (not @p_98) @p_25) :rule not_not) (step t76 (cl @p_99 @p_25) :rule th_resolution :premises (t75 t74)) (step t77 (cl @p_99 (! (not @p_97) :named @p_100)) :rule or_neg) (step t78 (cl (not @p_100) @p_65) :rule not_not) (step t79 (cl @p_99 @p_65) :rule th_resolution :premises (t78 t77)) (step t80 (cl @p_99) :rule th_resolution :premises (t71 t73 t76 t79)) (step t81 (cl @p_94 @p_97) :rule or :premises (t80)) (step t82 (cl @p_97) :rule resolution :premises (t81 t13)) (step t83 (cl @p_94 @p_101) :rule or :premises (t70)) (step t84 (cl (or (! (not @p_101) :named @p_102) @p_91)) :rule forall_inst :args ((:= veriT_vr1 c$))) (step t85 (cl @p_102 @p_91) :rule or :premises (t84)) (step t86 (cl (! (or @p_94 @p_91) :named @p_103) @p_98) :rule or_neg) (step t87 (cl @p_103 @p_25) :rule th_resolution :premises (t75 t86)) (step t88 (cl @p_103 @p_104) :rule or_neg) (step t89 (cl @p_103) :rule th_resolution :premises (t83 t85 t87 t88)) (step t90 (cl @p_94 @p_91) :rule or :premises (t89)) (step t91 (cl @p_91) :rule resolution :premises (t90 t13)) (step t92 (cl @p_67) :rule resolution :premises (t62 t60 t63 t61 t91 t82)) (step t93 (cl @p_105) :rule resolution :premises (t68 t92)) (step t94 (cl @p_36) :rule resolution :premises (t64 t93)) (step t95 (cl @p_93) :rule resolution :premises (t67 t93)) (step t96 (cl (or @p_84 @p_106)) :rule forall_inst :args ((:= veriT_vr10 c$))) (step t97 (cl @p_84 @p_106) :rule or :premises (t96)) (step t98 (cl) :rule resolution :premises (t97 t94 t95)) fa33723879319b6b43d6be3aebc3d4bf112ced6c 57 0 unsat (assume axiom0 (! (ite (! (p$ x$) :named @p_2) (! (not (! (exists ((?v0 A$)) (! (p$ ?v0) :named @p_1)) :named @p_3)) :named @p_5) (! (forall ((?v0 A$)) (! (not @p_1) :named @p_10)) :named @p_7)) :named @p_12)) (assume axiom1 (! (not (! (=> @p_2 (! (p$ y$) :named @p_35)) :named @p_39)) :named @p_34)) (anchor :step t3 :args ((:= (?v0 A$) veriT_vr0))) (step t3.t1 (cl (! (= ?v0 veriT_vr0) :named @p_8)) :rule refl) (step t3.t2 (cl (! (= @p_1 (! (p$ veriT_vr0) :named @p_4)) :named @p_9)) :rule cong :premises (t3.t1)) (step t3 (cl (= @p_3 (! (exists ((veriT_vr0 A$)) @p_4) :named @p_6))) :rule bind) (step t4 (cl (= @p_5 (! (not @p_6) :named @p_13))) :rule cong :premises (t3)) (anchor :step t5 :args ((:= (?v0 A$) veriT_vr0))) (step t5.t1 (cl @p_8) :rule refl) (step t5.t2 (cl @p_9) :rule cong :premises (t5.t1)) (step t5.t3 (cl (= @p_10 (! (not @p_4) :named @p_11))) :rule cong :premises (t5.t2)) (step t5 (cl (= @p_7 (! (forall ((veriT_vr0 A$)) @p_11) :named @p_14))) :rule bind) (step t6 (cl (! (= @p_12 (! (ite @p_2 @p_13 @p_14) :named @p_16)) :named @p_15)) :rule cong :premises (t4 t5)) (step t7 (cl (not @p_15) (not @p_12) @p_16) :rule equiv_pos2) (step t8 (cl @p_16) :rule th_resolution :premises (axiom0 t6 t7)) (anchor :step t9 :args ((:= (veriT_vr0 A$) veriT_vr1))) (step t9.t1 (cl (= veriT_vr0 veriT_vr1)) :rule refl) (step t9.t2 (cl (= @p_4 (! (p$ veriT_vr1) :named @p_17))) :rule cong :premises (t9.t1)) (step t9.t3 (cl (= @p_11 (! (not @p_17) :named @p_18))) :rule cong :premises (t9.t2)) (step t9 (cl (= @p_14 (! (forall ((veriT_vr1 A$)) @p_18) :named @p_19))) :rule bind) (step t10 (cl (! (= @p_16 (! (ite @p_2 @p_13 @p_19) :named @p_21)) :named @p_20)) :rule cong :premises (t9)) (step t11 (cl (not @p_20) (not @p_16) @p_21) :rule equiv_pos2) (step t12 (cl @p_21) :rule th_resolution :premises (t8 t10 t11)) (anchor :step t13 :args ((:= (veriT_vr0 A$) veriT_vr2))) (step t13.t1 (cl (= veriT_vr0 veriT_vr2)) :rule refl) (step t13.t2 (cl (= @p_4 (! (p$ veriT_vr2) :named @p_22))) :rule cong :premises (t13.t1)) (step t13 (cl (= @p_6 (! (exists ((veriT_vr2 A$)) @p_22) :named @p_23))) :rule bind) (step t14 (cl (= @p_13 (! (not @p_23) :named @p_26))) :rule cong :premises (t13)) (anchor :step t15 :args ((:= (veriT_vr1 A$) veriT_vr3))) (step t15.t1 (cl (= veriT_vr1 veriT_vr3)) :rule refl) (step t15.t2 (cl (= @p_17 (! (p$ veriT_vr3) :named @p_24))) :rule cong :premises (t15.t1)) (step t15.t3 (cl (= @p_18 (! (not @p_24) :named @p_25))) :rule cong :premises (t15.t2)) (step t15 (cl (= @p_19 (! (forall ((veriT_vr3 A$)) @p_25) :named @p_27))) :rule bind) (step t16 (cl (! (= @p_21 (! (ite @p_2 @p_26 @p_27) :named @p_29)) :named @p_28)) :rule cong :premises (t14 t15)) (step t17 (cl (not @p_28) (not @p_21) @p_29) :rule equiv_pos2) (step t18 (cl @p_29) :rule th_resolution :premises (t12 t16 t17)) (step t19 (cl (= @p_23 (! (not (! (forall ((veriT_vr2 A$)) (not @p_22)) :named @p_41)) :named @p_30))) :rule connective_def) (step t20 (cl (= @p_26 (! (not @p_30) :named @p_31))) :rule cong :premises (t19)) (step t21 (cl (! (= @p_29 (! (ite @p_2 @p_31 @p_27) :named @p_33)) :named @p_32)) :rule cong :premises (t20)) (step t22 (cl (not @p_32) (not @p_29) @p_33) :rule equiv_pos2) (step t23 (cl @p_33) :rule th_resolution :premises (t18 t21 t22)) (step t24 (cl (! (= @p_34 (! (and @p_2 (not @p_35)) :named @p_37)) :named @p_36)) :rule bool_simplify) (step t25 (cl (! (not @p_36) :named @p_40) (! (not @p_34) :named @p_38) @p_37) :rule equiv_pos2) (step t26 (cl (not @p_38) @p_39) :rule not_not) (step t27 (cl @p_40 @p_39 @p_37) :rule th_resolution :premises (t26 t25)) (step t28 (cl @p_37) :rule th_resolution :premises (axiom1 t24 t27)) (step t29 (cl (= @p_31 @p_41)) :rule not_simplify) (step t30 (cl (! (= @p_33 (! (ite @p_2 @p_41 @p_27) :named @p_43)) :named @p_42)) :rule cong :premises (t29)) (step t31 (cl (not @p_42) (not @p_33) @p_43) :rule equiv_pos2) (step t32 (cl @p_43) :rule th_resolution :premises (t23 t30 t31)) (step t33 (cl (! (not @p_2) :named @p_44) @p_41) :rule ite2 :premises (t32)) (step t34 (cl @p_2) :rule and :premises (t28)) (step t35 (cl @p_41) :rule resolution :premises (t33 t34)) (step t36 (cl (or @p_30 @p_44)) :rule forall_inst :args ((:= veriT_vr2 x$))) (step t37 (cl @p_30 @p_44) :rule or :premises (t36)) (step t38 (cl) :rule resolution :premises (t37 t34 t35)) 3a011b0429bfc23b05e60ee65e41832adbfa1a5a 12 0 unsat (assume axiom0 (! (not (! (= 3 3) :named @p_1)) :named @p_2)) (step t2 (cl (= @p_1 true)) :rule eq_simplify) (step t3 (cl (= @p_2 (! (not true) :named @p_3))) :rule cong :premises (t2)) (step t4 (cl (= @p_3 false)) :rule not_simplify) (step t5 (cl (! (= @p_2 false) :named @p_4)) :rule trans :premises (t3 t4)) (step t6 (cl (! (not @p_4) :named @p_6) (! (not @p_2) :named @p_5) false) :rule equiv_pos2) (step t7 (cl (not @p_5) @p_1) :rule not_not) (step t8 (cl @p_6 @p_1 false) :rule th_resolution :premises (t7 t6)) (step t9 (cl false) :rule th_resolution :premises (axiom0 t5 t8)) (step t10 (cl (not false)) :rule false) (step t11 (cl) :rule resolution :premises (t9 t10)) ed9c8e3326f1e52cfa128cd451ceb8afc62fa2fc 12 0 unsat (assume axiom0 (! (not (! (= 3.0 3.0) :named @p_1)) :named @p_2)) (step t2 (cl (= @p_1 true)) :rule eq_simplify) (step t3 (cl (= @p_2 (! (not true) :named @p_3))) :rule cong :premises (t2)) (step t4 (cl (= @p_3 false)) :rule not_simplify) (step t5 (cl (! (= @p_2 false) :named @p_4)) :rule trans :premises (t3 t4)) (step t6 (cl (! (not @p_4) :named @p_6) (! (not @p_2) :named @p_5) false) :rule equiv_pos2) (step t7 (cl (not @p_5) @p_1) :rule not_not) (step t8 (cl @p_6 @p_1 false) :rule th_resolution :premises (t7 t6)) (step t9 (cl false) :rule th_resolution :premises (axiom0 t5 t8)) (step t10 (cl (not false)) :rule false) (step t11 (cl) :rule resolution :premises (t9 t10)) 1702d82a39b3f9713bf57e315ab761c4a81dbe59 9 0 unsat (assume axiom0 (not (! (= (! (+ x$ (+ y$ z$)) :named @p_2) (! (+ y$ (+ z$ x$)) :named @p_3)) :named @p_1))) (step t2 (cl (or @p_1 (! (not (! (<= @p_2 @p_3) :named @p_6)) :named @p_4) (! (not (! (<= @p_3 @p_2) :named @p_7)) :named @p_5))) :rule la_disequality) (step t3 (cl @p_1 @p_4 @p_5) :rule or :premises (t2)) (step t4 (cl @p_4 @p_5) :rule resolution :premises (t3 axiom0)) (step t5 (cl @p_6) :rule la_generic :args (1)) (step t6 (cl @p_5) :rule resolution :premises (t4 t5)) (step t7 (cl @p_7) :rule la_generic :args (1)) (step t8 (cl) :rule resolution :premises (t7 t6)) ffb6c06c5ee6006629798f020d7438a4445f818d 15 0 unsat (assume axiom0 (! (not (! (= (+ 3 1) 4) :named @p_1)) :named @p_3)) (step t2 (cl @p_1) :rule sum_simplify) (step t3 (cl (= @p_1 (! (= 4 4) :named @p_2))) :rule cong :premises (t2)) (step t4 (cl (= @p_2 true)) :rule eq_simplify) (step t5 (cl (= @p_1 true)) :rule trans :premises (t3 t4)) (step t6 (cl (= @p_3 (! (not true) :named @p_4))) :rule cong :premises (t5)) (step t7 (cl (= @p_4 false)) :rule not_simplify) (step t8 (cl (! (= @p_3 false) :named @p_5)) :rule trans :premises (t6 t7)) (step t9 (cl (! (not @p_5) :named @p_7) (! (not @p_3) :named @p_6) false) :rule equiv_pos2) (step t10 (cl (not @p_6) @p_1) :rule not_not) (step t11 (cl @p_7 @p_1 false) :rule th_resolution :premises (t10 t9)) (step t12 (cl false) :rule th_resolution :premises (axiom0 t8 t11)) (step t13 (cl (not false)) :rule false) (step t14 (cl) :rule resolution :premises (t12 t13)) dc0e836fbcd3ce41657a40fc20a29630682bbe52 18 0 unsat (assume axiom0 (! (not (! (< 5 (! (ite (! (<= 3 8) :named @p_1) 8 3) :named @p_2)) :named @p_4)) :named @p_6)) (step t2 (cl (= @p_1 true)) :rule comp_simplify) (step t3 (cl (= @p_2 (! (ite true 8 3) :named @p_3))) :rule cong :premises (t2)) (step t4 (cl (= 8 @p_3)) :rule ite_simplify) (step t5 (cl (= 8 @p_2)) :rule trans :premises (t3 t4)) (step t6 (cl (= @p_4 (! (< 5 8) :named @p_5))) :rule cong :premises (t5)) (step t7 (cl (= @p_5 true)) :rule comp_simplify) (step t8 (cl (= @p_4 true)) :rule trans :premises (t6 t7)) (step t9 (cl (= @p_6 (! (not true) :named @p_7))) :rule cong :premises (t8)) (step t10 (cl (= @p_7 false)) :rule not_simplify) (step t11 (cl (! (= @p_6 false) :named @p_8)) :rule trans :premises (t9 t10)) (step t12 (cl (! (not @p_8) :named @p_10) (! (not @p_6) :named @p_9) false) :rule equiv_pos2) (step t13 (cl (not @p_9) @p_4) :rule not_not) (step t14 (cl @p_10 @p_4 false) :rule th_resolution :premises (t13 t12)) (step t15 (cl false) :rule th_resolution :premises (axiom0 t11 t14)) (step t16 (cl (not false)) :rule false) (step t17 (cl) :rule resolution :premises (t15 t16)) 00e71b7773518acdf4ff42fd31ed0e8e2cc40c59 52 0 unsat (assume axiom0 (! (not (! (<= (! (ite (! (< (! (+ x$ y$) :named @p_1) 0.0) :named @p_6) (! (- @p_1) :named @p_7) @p_1) :named @p_3) (+ (! (ite (! (< x$ 0.0) :named @p_8) (! (- x$) :named @p_9) x$) :named @p_4) (! (ite (! (< y$ 0.0) :named @p_10) (! (- y$) :named @p_11) y$) :named @p_5))) :named @p_15)) :named @p_2)) (step t2 (cl (! (= @p_2 (! (and (! (not (! (<= @p_3 (+ @p_4 @p_5)) :named @p_28)) :named @p_17) (! (ite @p_6 (! (= @p_7 @p_3) :named @p_20) (! (= @p_1 @p_3) :named @p_19)) :named @p_18) (! (ite @p_8 (! (= @p_9 @p_4) :named @p_23) (! (= x$ @p_4) :named @p_22)) :named @p_21) (! (ite @p_10 (! (= @p_11 @p_5) :named @p_26) (! (= y$ @p_5) :named @p_25)) :named @p_24)) :named @p_13)) :named @p_12)) :rule ite_intro) (step t3 (cl (! (not @p_12) :named @p_16) (! (not @p_2) :named @p_14) @p_13) :rule equiv_pos2) (step t4 (cl (not @p_14) @p_15) :rule not_not) (step t5 (cl @p_16 @p_15 @p_13) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_13) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl @p_17) :rule and :premises (t6)) (step t8 (cl @p_18) :rule and :premises (t6)) (step t9 (cl @p_6 @p_19) :rule ite1 :premises (t8)) (step t10 (cl (! (not @p_6) :named @p_32) @p_20) :rule ite2 :premises (t8)) (step t11 (cl @p_21) :rule and :premises (t6)) (step t12 (cl @p_8 @p_22) :rule ite1 :premises (t11)) (step t13 (cl (! (not @p_8) :named @p_36) @p_23) :rule ite2 :premises (t11)) (step t14 (cl @p_24) :rule and :premises (t6)) (step t15 (cl @p_10 @p_25) :rule ite1 :premises (t14)) (step t16 (cl (! (not @p_10) :named @p_41) @p_26) :rule ite2 :premises (t14)) (step t17 (cl (! (<= @p_11 @p_5) :named @p_27) (! (<= @p_5 @p_11) :named @p_39)) :rule la_generic :args (1.0 1.0)) (step t18 (cl @p_27 @p_10 (! (not @p_25) :named @p_29)) :rule la_generic :args (1.0 2.0 (- 1))) (step t19 (cl (! (= @p_11 @p_11) :named @p_37)) :rule eq_reflexive) (step t20 (cl @p_28 @p_29 (! (not @p_22) :named @p_30) (! (not @p_19) :named @p_31)) :rule la_generic :args (1.0 (- 1) (- 1) 1)) (step t21 (cl @p_29 @p_30 @p_31) :rule resolution :premises (t20 t7)) (step t22 (cl @p_8 @p_32 @p_10) :rule la_generic :args (1.0 1.0 1.0)) (step t23 (cl @p_8 @p_10 @p_29) :rule resolution :premises (t22 t9 t21 t12)) (step t24 (cl @p_28 (! (not @p_27) :named @p_33) (! (not @p_23) :named @p_34) (! (not @p_20) :named @p_35)) :rule la_generic :args (1.0 1.0 (- 1) 1)) (step t25 (cl @p_33 @p_34 @p_35) :rule resolution :premises (t24 t7)) (step t26 (cl @p_36 @p_6 (not (! (<= y$ @p_11) :named @p_38))) :rule la_generic :args (2.0 2.0 1.0)) (step t27 (cl @p_29 (not @p_37) @p_38 (! (not @p_39) :named @p_40)) :rule eq_congruent_pred) (step t28 (cl @p_29 @p_38 @p_40) :rule th_resolution :premises (t27 t19)) (step t29 (cl @p_33 @p_10 @p_40) :rule resolution :premises (t26 t10 t25 t13 t23 t28 t15)) (step t30 (cl @p_41 @p_6 @p_36) :rule la_generic :args (1.0 1.0 1.0)) (step t31 (cl @p_41 @p_36 @p_33) :rule resolution :premises (t30 t10 t25 t13)) (step t32 (cl @p_28 @p_41 @p_33 @p_30 @p_31) :rule la_generic :args (1.0 2.0 1.0 (- 1) 1)) (step t33 (cl @p_41 @p_33 @p_30 @p_31) :rule resolution :premises (t32 t7)) (step t34 (cl @p_28 @p_33 @p_8 @p_30 @p_35) :rule la_generic :args (1.0 1.0 2.0 (- 1) 1)) (step t35 (cl @p_33 @p_8 @p_30 @p_35) :rule resolution :premises (t34 t7)) (step t36 (cl (! (not @p_26) :named @p_42) @p_42 @p_27 @p_40) :rule eq_congruent_pred) (step t37 (cl @p_42 @p_27 @p_40) :rule contraction :premises (t36)) (step t38 (cl @p_40 @p_42) :rule resolution :premises (t35 t10 t9 t33 t12 t31 t29 t37)) (step t39 (cl @p_42 @p_42 @p_33 @p_39) :rule eq_congruent_pred) (step t40 (cl @p_42 @p_33 @p_39) :rule contraction :premises (t39)) (step t41 (cl @p_42) :rule resolution :premises (t40 t17 t38)) (step t42 (cl @p_41) :rule resolution :premises (t16 t41)) (step t43 (cl @p_25) :rule resolution :premises (t15 t42)) (step t44 (cl @p_27) :rule resolution :premises (t18 t42 t43)) (step t45 (cl @p_8) :rule resolution :premises (t23 t42 t43)) (step t46 (cl @p_23) :rule resolution :premises (t13 t45)) (step t47 (cl @p_35) :rule resolution :premises (t25 t46 t44)) (step t48 (cl @p_32) :rule resolution :premises (t10 t47)) (step t49 (cl @p_19) :rule resolution :premises (t9 t48)) (step t50 (cl @p_28 @p_36 @p_31 @p_29 @p_34) :rule la_generic :args (1.0 2.0 1 (- 1) (- 1))) (step t51 (cl) :rule resolution :premises (t50 t7 t49 t45 t46 t43)) 402dc6e98a3b5a465e12e6bcdd5fca0ae68d74d8 19 0 unsat (assume axiom0 (not (= (p$ (! (< 2 3) :named @p_2)) (! (p$ true) :named @p_1)))) (step t2 (cl (! (not (! (= @p_1 (! (ite @p_2 @p_1 (! (p$ false) :named @p_4)) :named @p_3)) :named @p_6)) :named @p_8)) :rule bfun_elim :premises (axiom0)) (step t3 (cl (= @p_2 true)) :rule comp_simplify) (step t4 (cl (= @p_3 (! (ite true @p_1 @p_4) :named @p_5))) :rule cong :premises (t3)) (step t5 (cl (= @p_1 @p_5)) :rule ite_simplify) (step t6 (cl @p_6) :rule trans :premises (t4 t5)) (step t7 (cl (= @p_6 (! (= @p_1 @p_1) :named @p_7))) :rule cong :premises (t6)) (step t8 (cl (= @p_7 true)) :rule eq_simplify) (step t9 (cl (= @p_6 true)) :rule trans :premises (t7 t8)) (step t10 (cl (= @p_8 (! (not true) :named @p_9))) :rule cong :premises (t9)) (step t11 (cl (= @p_9 false)) :rule not_simplify) (step t12 (cl (! (= @p_8 false) :named @p_10)) :rule trans :premises (t10 t11)) (step t13 (cl (! (not @p_10) :named @p_12) (! (not @p_8) :named @p_11) false) :rule equiv_pos2) (step t14 (cl (not @p_11) @p_6) :rule not_not) (step t15 (cl @p_12 @p_6 false) :rule th_resolution :premises (t14 t13)) (step t16 (cl false) :rule th_resolution :premises (t2 t12 t15)) (step t17 (cl (not false)) :rule false) (step t18 (cl) :rule resolution :premises (t16 t17)) ed3db2a47b567944ed8f84b91963c7783c1f6052 14 0 unsat (assume axiom0 (! (not (! (or (! (<= 4 (! (+ x$ 3) :named @p_1)) :named @p_2) (! (< x$ 1) :named @p_6)) :named @p_4)) :named @p_7)) (step t2 (cl (= @p_1 (! (+ 3 x$) :named @p_3))) :rule sum_simplify) (step t3 (cl (= @p_2 (! (<= 4 @p_3) :named @p_5))) :rule cong :premises (t2)) (step t4 (cl (= @p_4 (! (or @p_5 @p_6) :named @p_8))) :rule cong :premises (t3)) (step t5 (cl (! (= @p_7 (! (not @p_8) :named @p_10)) :named @p_9)) :rule cong :premises (t4)) (step t6 (cl (! (not @p_9) :named @p_12) (! (not @p_7) :named @p_11) @p_10) :rule equiv_pos2) (step t7 (cl (not @p_11) @p_4) :rule not_not) (step t8 (cl @p_12 @p_4 @p_10) :rule th_resolution :premises (t7 t6)) (step t9 (cl @p_10) :rule th_resolution :premises (axiom0 t5 t8)) (step t10 (cl (not @p_5)) :rule not_or :premises (t9)) (step t11 (cl (not @p_6)) :rule not_or :premises (t9)) (step t12 (cl @p_6 @p_5) :rule la_generic :args (1 1)) (step t13 (cl) :rule resolution :premises (t12 t10 t11)) d2678f4fb6d818fd1c6e93a61560f2dfba8a9409 9 0 unsat (assume axiom1 (! (= y$ (! (+ x$ 4) :named @p_1)) :named @p_2)) (assume axiom2 (not (! (< 0 (- y$ x$)) :named @p_6))) (step t3 (cl (= @p_1 (! (+ 4 x$) :named @p_3))) :rule sum_simplify) (step t4 (cl (! (= @p_2 (! (= y$ @p_3) :named @p_5)) :named @p_4)) :rule cong :premises (t3)) (step t5 (cl (not @p_4) (not @p_2) @p_5) :rule equiv_pos2) (step t6 (cl @p_5) :rule th_resolution :premises (axiom1 t4 t5)) (step t7 (cl @p_6 (not @p_5)) :rule la_generic :args (1 1)) (step t8 (cl) :rule resolution :premises (t7 t6 axiom2)) 4e88c3cf4e2da31d65e91a3805e3cea4a5a2813d 20 0 unsat (assume axiom0 (! (not (! (not (! (= (! (+ 2 2) :named @p_3) 5) :named @p_2)) :named @p_5)) :named @p_1)) (step t2 (cl (! (not @p_1) :named @p_9) @p_2) :rule not_not) (step t3 (cl @p_2) :rule th_resolution :premises (t2 axiom0)) (step t4 (cl (= @p_3 4)) :rule sum_simplify) (step t5 (cl (= @p_2 (! (= 5 4) :named @p_4))) :rule cong :premises (t4)) (step t6 (cl (= @p_4 false)) :rule eq_simplify) (step t7 (cl (= @p_2 false)) :rule trans :premises (t5 t6)) (step t8 (cl (= @p_5 (! (not false) :named @p_6))) :rule cong :premises (t7)) (step t9 (cl (= @p_6 true)) :rule not_simplify) (step t10 (cl (= @p_5 true)) :rule trans :premises (t8 t9)) (step t11 (cl (= @p_1 (! (not true) :named @p_7))) :rule cong :premises (t10)) (step t12 (cl (= @p_7 false)) :rule not_simplify) (step t13 (cl (! (= @p_1 false) :named @p_8)) :rule trans :premises (t11 t12)) (step t14 (cl (! (not @p_8) :named @p_10) @p_9 false) :rule equiv_pos2) (step t15 (cl (not @p_9) @p_5) :rule not_not) (step t16 (cl @p_10 @p_5 false) :rule th_resolution :premises (t15 t14)) (step t17 (cl false) :rule th_resolution :premises (t3 t13 t16)) (step t18 (cl @p_6) :rule false) (step t19 (cl) :rule resolution :premises (t17 t18)) 3bbfe9f2086fe18bc1f9eeb965fef30f8b2c9daa 6 0 unsat (assume axiom0 (! (< (+ (* 3 x$) (* 7 a$)) 4) :named @p_3)) (assume axiom1 (! (< 3 (* 2 x$)) :named @p_1)) (assume axiom2 (not (! (< a$ 0) :named @p_2))) (step t4 (cl (not @p_1) @p_2 (not @p_3)) :rule la_generic :args ((div 3 2) 7 1)) (step t5 (cl) :rule resolution :premises (t4 axiom0 axiom1 axiom2)) dbde8d1b71dac8b258507f86acdbe195bf64b2b7 29 0 unsat (assume axiom0 (! (not (! (= (! (or (! (<= 0 (! (+ y$ (! (* (! (- 1) :named @p_14) x$) :named @p_15)) :named @p_16)) :named @p_3) (or (! (not (! (<= 0 x$) :named @p_1)) :named @p_4) @p_1)) :named @p_2) (! (not false) :named @p_7)) :named @p_5)) :named @p_8)) (step t2 (cl (= @p_2 (! (or @p_3 @p_4 @p_1) :named @p_6))) :rule ac_simp) (step t3 (cl (= @p_5 (! (= @p_6 @p_7) :named @p_9))) :rule cong :premises (t2)) (step t4 (cl (! (= @p_8 (! (not @p_9) :named @p_11)) :named @p_10)) :rule cong :premises (t3)) (step t5 (cl (! (not @p_10) :named @p_13) (! (not @p_8) :named @p_12) @p_11) :rule equiv_pos2) (step t6 (cl (not @p_12) @p_5) :rule not_not) (step t7 (cl @p_13 @p_5 @p_11) :rule th_resolution :premises (t6 t5)) (step t8 (cl @p_11) :rule th_resolution :premises (axiom0 t4 t7)) (step t9 (cl (= @p_14 (- 1))) :rule minus_simplify) (step t10 (cl (= @p_15 (! (* (- 1) x$) :named @p_17))) :rule cong :premises (t9)) (step t11 (cl (= @p_16 (! (+ y$ @p_17) :named @p_18))) :rule cong :premises (t10)) (step t12 (cl (= @p_3 (! (<= 0 @p_18) :named @p_19))) :rule cong :premises (t11)) (step t13 (cl (= @p_6 (! (or @p_19 @p_4 @p_1) :named @p_20))) :rule cong :premises (t12)) (step t14 (cl (= @p_20 true)) :rule or_simplify) (step t15 (cl (= @p_6 true)) :rule trans :premises (t13 t14)) (step t16 (cl (= @p_7 true)) :rule not_simplify) (step t17 (cl (= @p_9 (! (= true true) :named @p_21))) :rule cong :premises (t15 t16)) (step t18 (cl (= @p_21 true)) :rule equiv_simplify) (step t19 (cl (= @p_9 true)) :rule trans :premises (t17 t18)) (step t20 (cl (= @p_11 (! (not true) :named @p_22))) :rule cong :premises (t19)) (step t21 (cl (= @p_22 false)) :rule not_simplify) (step t22 (cl (! (= @p_11 false) :named @p_23)) :rule trans :premises (t20 t21)) (step t23 (cl (! (not @p_23) :named @p_25) (! (not @p_11) :named @p_24) false) :rule equiv_pos2) (step t24 (cl (not @p_24) @p_9) :rule not_not) (step t25 (cl @p_25 @p_9 false) :rule th_resolution :premises (t24 t23)) (step t26 (cl false) :rule th_resolution :premises (t8 t22 t25)) (step t27 (cl @p_7) :rule false) (step t28 (cl) :rule resolution :premises (t26 t27)) d6ba634ec42000181b77274dd94c91d844cf476e 21 0 unsat (assume axiom0 (! (not (! (or (! (< (! (+ x$ x$) :named @p_11) (! (+ (! (* 2.0 x$) :named @p_10) 1.0) :named @p_9)) :named @p_1) (or false @p_1)) :named @p_2)) :named @p_3)) (step t2 (cl (= @p_2 (! (or @p_1 false) :named @p_4))) :rule ac_simp) (step t3 (cl (! (= @p_3 (! (not @p_4) :named @p_6)) :named @p_5)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_5) :named @p_8) (! (not @p_3) :named @p_7) @p_6) :rule equiv_pos2) (step t5 (cl (not @p_7) @p_2) :rule not_not) (step t6 (cl @p_8 @p_2 @p_6) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_6) :rule th_resolution :premises (axiom0 t3 t6)) (step t8 (cl (= @p_9 (! (+ 1.0 @p_10) :named @p_12))) :rule sum_simplify) (step t9 (cl (= @p_1 (! (< @p_11 @p_12) :named @p_13))) :rule cong :premises (t8)) (step t10 (cl (= @p_4 (! (or @p_13 false) :named @p_14))) :rule cong :premises (t9)) (step t11 (cl (= @p_14 (! (or @p_13) :named @p_15))) :rule or_simplify) (step t12 (cl (= @p_15 @p_13)) :rule or_simplify) (step t13 (cl (= @p_4 @p_13)) :rule trans :premises (t10 t11 t12)) (step t14 (cl (! (= @p_6 (! (not @p_13) :named @p_17)) :named @p_16)) :rule cong :premises (t13)) (step t15 (cl (! (not @p_16) :named @p_19) (! (not @p_6) :named @p_18) @p_17) :rule equiv_pos2) (step t16 (cl (not @p_18) @p_4) :rule not_not) (step t17 (cl @p_19 @p_4 @p_17) :rule th_resolution :premises (t16 t15)) (step t18 (cl @p_17) :rule th_resolution :premises (t7 t14 t17)) (step t19 (cl @p_13) :rule la_tautology) (step t20 (cl) :rule resolution :premises (t19 t18)) 95657842b84a23b6c0020065f8c8597823332a77 62 0 unsat (assume axiom0 (! (not (! (or (! (and (! (< n$ m$) :named @p_1) (! (< m$ n$a) :named @p_3)) :named @p_11) (or (! (and @p_1 (! (= m$ n$a) :named @p_6)) :named @p_12) (or (! (and (! (< n$ n$a) :named @p_8) (! (< n$a m$) :named @p_2)) :named @p_13) (or (! (and (! (= n$ n$a) :named @p_5) @p_2) :named @p_14) (or (! (and (! (= n$ m$) :named @p_4) @p_3) :named @p_15) (or (! (and @p_2 (! (< m$ n$) :named @p_7)) :named @p_16) (or (! (and @p_2 @p_4) :named @p_17) (or (! (and (! (< n$a n$) :named @p_9) @p_1) :named @p_18) (or (! (and @p_5 @p_1) :named @p_19) (or (! (and @p_6 @p_7) :named @p_20) (or (! (and @p_7 @p_8) :named @p_21) (or (! (and @p_7 @p_5) :named @p_22) (or (! (and @p_3 @p_9) :named @p_23) (or (! (and @p_4 @p_8) :named @p_24) (or (! (and @p_6 @p_9) :named @p_25) (! (and @p_6 @p_4) :named @p_26)))))))))))))))) :named @p_10)) :named @p_27)) (step t2 (cl (= @p_10 (! (or @p_11 @p_12 @p_13 @p_14 @p_15 @p_16 @p_17 @p_18 @p_19 @p_20 @p_21 @p_22 @p_23 @p_24 @p_25 @p_26) :named @p_28))) :rule ac_simp) (step t3 (cl (! (= @p_27 (! (not @p_28) :named @p_30)) :named @p_29)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_29) :named @p_32) (! (not @p_27) :named @p_31) @p_30) :rule equiv_pos2) (step t5 (cl (not @p_31) @p_10) :rule not_not) (step t6 (cl @p_32 @p_10 @p_30) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_30) :rule th_resolution :premises (axiom0 t3 t6)) (step t8 (cl (not @p_11)) :rule not_or :premises (t7)) (step t9 (cl (! (not @p_1) :named @p_33) (! (not @p_3) :named @p_38)) :rule not_and :premises (t8)) (step t10 (cl (not @p_12)) :rule not_or :premises (t7)) (step t11 (cl @p_33 (! (not @p_6) :named @p_41)) :rule not_and :premises (t10)) (step t12 (cl (not @p_13)) :rule not_or :premises (t7)) (step t13 (cl (! (not @p_8) :named @p_36) (! (not @p_2) :named @p_34)) :rule not_and :premises (t12)) (step t14 (cl (not @p_16)) :rule not_or :premises (t7)) (step t15 (cl @p_34 (! (not @p_7) :named @p_35)) :rule not_and :premises (t14)) (step t16 (cl (not @p_17)) :rule not_or :premises (t7)) (step t17 (cl @p_34 (! (not @p_4) :named @p_40)) :rule not_and :premises (t16)) (step t18 (cl (not @p_18)) :rule not_or :premises (t7)) (step t19 (cl (! (not @p_9) :named @p_39) @p_33) :rule not_and :premises (t18)) (step t20 (cl (not @p_19)) :rule not_or :premises (t7)) (step t21 (cl (! (not @p_5) :named @p_37) @p_33) :rule not_and :premises (t20)) (step t22 (cl (not @p_21)) :rule not_or :premises (t7)) (step t23 (cl @p_35 @p_36) :rule not_and :premises (t22)) (step t24 (cl (not @p_22)) :rule not_or :premises (t7)) (step t25 (cl @p_35 @p_37) :rule not_and :premises (t24)) (step t26 (cl (not @p_23)) :rule not_or :premises (t7)) (step t27 (cl @p_38 @p_39) :rule not_and :premises (t26)) (step t28 (cl (not @p_24)) :rule not_or :premises (t7)) (step t29 (cl @p_40 @p_36) :rule not_and :premises (t28)) (step t30 (cl (not @p_25)) :rule not_or :premises (t7)) (step t31 (cl @p_41 @p_39) :rule not_and :premises (t30)) (step t32 (cl (not @p_26)) :rule not_or :premises (t7)) (step t33 (cl @p_41 @p_40) :rule not_and :premises (t32)) (step t34 (cl (or @p_4 (! (not (! (<= n$ m$) :named @p_50)) :named @p_42) (! (not (! (<= m$ n$) :named @p_46)) :named @p_43))) :rule la_disequality) (step t35 (cl @p_4 @p_42 @p_43) :rule or :premises (t34)) (step t36 (cl (or @p_5 (! (not (! (<= n$ n$a) :named @p_52)) :named @p_44) (! (not (! (<= n$a n$) :named @p_51)) :named @p_45))) :rule la_disequality) (step t37 (cl @p_5 @p_44 @p_45) :rule or :premises (t36)) (step t38 (cl @p_1 @p_46) :rule la_generic :args (1 1)) (step t39 (cl (or @p_6 (! (not (! (<= m$ n$a) :named @p_53)) :named @p_47) (! (not (! (<= n$a m$) :named @p_49)) :named @p_48))) :rule la_disequality) (step t40 (cl @p_6 @p_47 @p_48) :rule or :premises (t39)) (step t41 (cl @p_3 @p_49) :rule la_generic :args (1 1)) (step t42 (cl @p_7 @p_50) :rule la_generic :args (1 1)) (step t43 (cl @p_51 @p_8) :rule la_generic :args (1 1)) (step t44 (cl @p_52 @p_9) :rule la_generic :args (1 1)) (step t45 (cl @p_44 @p_35) :rule resolution :premises (t37 t43 t25 t23)) (step t46 (cl @p_40 @p_37 @p_6) :rule eq_transitive) (step t47 (cl @p_37 @p_42) :rule resolution :premises (t46 t33 t35 t38 t21)) (step t48 (cl @p_53 @p_2) :rule la_generic :args (1 1)) (step t49 (cl @p_44) :rule resolution :premises (t48 t40 t41 t11 t9 t38 t35 t29 t13 t43 t37 t47 t42 t45)) (step t50 (cl @p_9) :rule resolution :premises (t44 t49)) (step t51 (cl @p_33) :rule resolution :premises (t19 t50)) (step t52 (cl @p_38) :rule resolution :premises (t27 t50)) (step t53 (cl @p_41) :rule resolution :premises (t31 t50)) (step t54 (cl @p_46) :rule resolution :premises (t38 t51)) (step t55 (cl @p_49) :rule resolution :premises (t41 t52)) (step t56 (cl @p_47) :rule resolution :premises (t40 t53 t55)) (step t57 (cl @p_2) :rule resolution :premises (t48 t56)) (step t58 (cl @p_35) :rule resolution :premises (t15 t57)) (step t59 (cl @p_40) :rule resolution :premises (t17 t57)) (step t60 (cl @p_50) :rule resolution :premises (t42 t58)) (step t61 (cl) :rule resolution :premises (t35 t59 t54 t60)) ce9ae392bbb004278cf4005c9cdc4ec6dc2c6c3e 16 0 unsat (assume axiom0 (! (not (! (not (! (exists ((?v0 Real)) false) :named @p_2)) :named @p_3)) :named @p_1)) (step t2 (cl (! (not @p_1) :named @p_6) @p_2) :rule not_not) (step t3 (cl @p_2) :rule th_resolution :premises (t2 axiom0)) (step t4 (cl (= @p_2 false)) :rule qnt_rm_unused) (step t5 (cl (= @p_3 (! (not false) :named @p_4))) :rule cong :premises (t4)) (step t6 (cl (! (= @p_1 (! (not @p_4) :named @p_7)) :named @p_5)) :rule cong :premises (t5)) (step t7 (cl (! (not @p_5) :named @p_8) @p_6 @p_7) :rule equiv_pos2) (step t8 (cl (not @p_6) @p_3) :rule not_not) (step t9 (cl @p_8 @p_3 @p_7) :rule th_resolution :premises (t8 t7)) (step t10 (cl (not @p_7) false) :rule not_not) (step t11 (cl @p_8 @p_3 false) :rule th_resolution :premises (t10 t9)) (step t12 (cl @p_7) :rule th_resolution :premises (t3 t6 t11)) (step t13 (cl false) :rule th_resolution :premises (t10 t12)) (step t14 (cl @p_4) :rule false) (step t15 (cl) :rule resolution :premises (t13 t14)) be1ac1d3f1b32c69d4c86a7bb2fc544d58daa87d 16 0 unsat (assume axiom0 (! (not (! (not (! (exists ((?v0 Int)) false) :named @p_2)) :named @p_3)) :named @p_1)) (step t2 (cl (! (not @p_1) :named @p_6) @p_2) :rule not_not) (step t3 (cl @p_2) :rule th_resolution :premises (t2 axiom0)) (step t4 (cl (= @p_2 false)) :rule qnt_rm_unused) (step t5 (cl (= @p_3 (! (not false) :named @p_4))) :rule cong :premises (t4)) (step t6 (cl (! (= @p_1 (! (not @p_4) :named @p_7)) :named @p_5)) :rule cong :premises (t5)) (step t7 (cl (! (not @p_5) :named @p_8) @p_6 @p_7) :rule equiv_pos2) (step t8 (cl (not @p_6) @p_3) :rule not_not) (step t9 (cl @p_8 @p_3 @p_7) :rule th_resolution :premises (t8 t7)) (step t10 (cl (not @p_7) false) :rule not_not) (step t11 (cl @p_8 @p_3 false) :rule th_resolution :premises (t10 t9)) (step t12 (cl @p_7) :rule th_resolution :premises (t3 t6 t11)) (step t13 (cl false) :rule th_resolution :premises (t10 t12)) (step t14 (cl @p_4) :rule false) (step t15 (cl) :rule resolution :premises (t13 t14)) 7c2862b95580cb7e81a355724e88267dce92987c 44 0 unsat (assume axiom0 (! (not (! (forall ((?v0 Int) (?v1 Int)) (! (=> (! (and (! (= 0 ?v0) :named @p_2) (! (= 1 ?v1) :named @p_4)) :named @p_6) (! (not (! (= ?v0 ?v1) :named @p_8)) :named @p_10)) :named @p_12)) :named @p_1)) :named @p_14)) (anchor :step t2 :args ((:= (?v0 Int) 0) (:= (?v1 Int) 1))) (step t2.t1 (cl @p_2) :rule refl) (step t2.t2 (cl (= @p_2 (! (= 0 0) :named @p_3))) :rule cong :premises (t2.t1)) (step t2.t3 (cl @p_4) :rule refl) (step t2.t4 (cl (= @p_4 (! (= 1 1) :named @p_5))) :rule cong :premises (t2.t3)) (step t2.t5 (cl (= @p_6 (! (and @p_3 @p_5) :named @p_7))) :rule cong :premises (t2.t2 t2.t4)) (step t2.t6 (cl @p_2) :rule refl) (step t2.t7 (cl @p_4) :rule refl) (step t2.t8 (cl (= @p_8 (! (= 0 1) :named @p_9))) :rule cong :premises (t2.t6 t2.t7)) (step t2.t9 (cl (= @p_10 (! (not @p_9) :named @p_11))) :rule cong :premises (t2.t8)) (step t2.t10 (cl (= @p_12 (! (=> @p_7 @p_11) :named @p_13))) :rule cong :premises (t2.t5 t2.t9)) (step t2 (cl (= @p_1 @p_13)) :rule onepoint) (step t3 (cl (! (= @p_14 (! (not @p_13) :named @p_16)) :named @p_15)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_15) :named @p_18) (! (not @p_14) :named @p_17) @p_16) :rule equiv_pos2) (step t5 (cl (not @p_17) @p_1) :rule not_not) (step t6 (cl @p_18 @p_1 @p_16) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_16) :rule th_resolution :premises (axiom0 t3 t6)) (step t8 (cl (! (= @p_16 (! (and @p_7 (! (not @p_11) :named @p_23)) :named @p_20)) :named @p_19)) :rule bool_simplify) (step t9 (cl (! (not @p_19) :named @p_22) (! (not @p_16) :named @p_21) @p_20) :rule equiv_pos2) (step t10 (cl (not @p_21) @p_13) :rule not_not) (step t11 (cl @p_22 @p_13 @p_20) :rule th_resolution :premises (t10 t9)) (step t12 (cl @p_20) :rule th_resolution :premises (t7 t8 t11)) (step t13 (cl (! (= @p_20 (! (and @p_3 @p_5 @p_23) :named @p_25)) :named @p_24)) :rule ac_simp) (step t14 (cl (not @p_24) (not @p_20) @p_25) :rule equiv_pos2) (step t15 (cl @p_25) :rule th_resolution :premises (t12 t13 t14)) (step t16 (cl (= @p_3 true)) :rule eq_simplify) (step t17 (cl (= @p_5 true)) :rule eq_simplify) (step t18 (cl (= @p_9 false)) :rule eq_simplify) (step t19 (cl (= @p_11 (! (not false) :named @p_26))) :rule cong :premises (t18)) (step t20 (cl (= @p_26 true)) :rule not_simplify) (step t21 (cl (= @p_11 true)) :rule trans :premises (t19 t20)) (step t22 (cl (= @p_23 (! (not true) :named @p_27))) :rule cong :premises (t21)) (step t23 (cl (= @p_27 false)) :rule not_simplify) (step t24 (cl (= @p_23 false)) :rule trans :premises (t22 t23)) (step t25 (cl (= @p_25 (! (and true true false) :named @p_28))) :rule cong :premises (t16 t17 t24)) (step t26 (cl (= @p_28 (! (and false) :named @p_29))) :rule and_simplify) (step t27 (cl (= @p_29 false)) :rule and_simplify) (step t28 (cl (! (= @p_25 false) :named @p_30)) :rule trans :premises (t25 t26 t27)) (step t29 (cl (not @p_30) (not @p_25) false) :rule equiv_pos2) (step t30 (cl false) :rule th_resolution :premises (t15 t28 t29)) (step t31 (cl @p_26) :rule false) (step t32 (cl) :rule resolution :premises (t30 t31)) 632d72943748607b3a6a192dc119caef23e9e71d 74 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr2 Int)) (not (forall ((veriT_vr3 Int)) (! (=> (! (< veriT_vr2 veriT_vr3) :named @p_30) (! (< (! (+ 1 (! (* 2 veriT_vr2) :named @p_32)) :named @p_33) (! (* 2 veriT_vr3) :named @p_35)) :named @p_36)) :named @p_37)))) :named @p_43)) (define-fun veriT_sk1 () Int (! (choice ((veriT_vr3 Int)) (not (=> (< @p_43 veriT_vr3) (< (+ 1 (* 2 @p_43)) @p_35)))) :named @p_44)) (assume axiom0 (! (not (! (forall ((?v0 Int) (?v1 Int)) (! (=> (! (< ?v0 ?v1) :named @p_2) (! (< (! (+ (! (* 2 ?v0) :named @p_5) 1) :named @p_7) (! (* 2 ?v1) :named @p_10)) :named @p_12)) :named @p_14)) :named @p_1)) :named @p_16)) (anchor :step t2 :args ((:= (?v0 Int) veriT_vr0) (:= (?v1 Int) veriT_vr1))) (step t2.t1 (cl (! (= ?v0 veriT_vr0) :named @p_4)) :rule refl) (step t2.t2 (cl (! (= ?v1 veriT_vr1) :named @p_9)) :rule refl) (step t2.t3 (cl (= @p_2 (! (< veriT_vr0 veriT_vr1) :named @p_3))) :rule cong :premises (t2.t1 t2.t2)) (step t2.t4 (cl @p_4) :rule refl) (step t2.t5 (cl (= @p_5 (! (* 2 veriT_vr0) :named @p_6))) :rule cong :premises (t2.t4)) (step t2.t6 (cl (= @p_7 (! (+ @p_6 1) :named @p_8))) :rule cong :premises (t2.t5)) (step t2.t7 (cl @p_9) :rule refl) (step t2.t8 (cl (= @p_10 (! (* 2 veriT_vr1) :named @p_11))) :rule cong :premises (t2.t7)) (step t2.t9 (cl (= @p_12 (! (< @p_8 @p_11) :named @p_13))) :rule cong :premises (t2.t6 t2.t8)) (step t2.t10 (cl (= @p_14 (! (=> @p_3 @p_13) :named @p_15))) :rule cong :premises (t2.t3 t2.t9)) (step t2 (cl (= @p_1 (! (forall ((veriT_vr0 Int) (veriT_vr1 Int)) @p_15) :named @p_17))) :rule bind) (step t3 (cl (! (= @p_16 (! (not @p_17) :named @p_19)) :named @p_18)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_18) :named @p_21) (! (not @p_16) :named @p_20) @p_19) :rule equiv_pos2) (step t5 (cl (not @p_20) @p_1) :rule not_not) (step t6 (cl @p_21 @p_1 @p_19) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_19) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((veriT_vr0 Int) (veriT_vr1 Int))) (step t8.t1 (cl (= @p_8 (! (+ 1 @p_6) :named @p_22))) :rule sum_simplify) (step t8.t2 (cl (= @p_13 (! (< @p_22 @p_11) :named @p_23))) :rule cong :premises (t8.t1)) (step t8.t3 (cl (= @p_15 (! (=> @p_3 @p_23) :named @p_24))) :rule cong :premises (t8.t2)) (step t8 (cl (= @p_17 (! (forall ((veriT_vr0 Int) (veriT_vr1 Int)) @p_24) :named @p_25))) :rule bind) (step t9 (cl (! (= @p_19 (! (not @p_25) :named @p_27)) :named @p_26)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_26) :named @p_29) (! (not @p_19) :named @p_28) @p_27) :rule equiv_pos2) (step t11 (cl (not @p_28) @p_17) :rule not_not) (step t12 (cl @p_29 @p_17 @p_27) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_27) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((:= (veriT_vr0 Int) veriT_vr2) (:= (veriT_vr1 Int) veriT_vr3))) (step t14.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_31)) :rule refl) (step t14.t2 (cl (! (= veriT_vr1 veriT_vr3) :named @p_34)) :rule refl) (step t14.t3 (cl (= @p_3 @p_30)) :rule cong :premises (t14.t1 t14.t2)) (step t14.t4 (cl @p_31) :rule refl) (step t14.t5 (cl (= @p_6 @p_32)) :rule cong :premises (t14.t4)) (step t14.t6 (cl (= @p_22 @p_33)) :rule cong :premises (t14.t5)) (step t14.t7 (cl @p_34) :rule refl) (step t14.t8 (cl (= @p_11 @p_35)) :rule cong :premises (t14.t7)) (step t14.t9 (cl (= @p_23 @p_36)) :rule cong :premises (t14.t6 t14.t8)) (step t14.t10 (cl (= @p_24 @p_37)) :rule cong :premises (t14.t3 t14.t9)) (step t14 (cl (= @p_25 (! (forall ((veriT_vr2 Int) (veriT_vr3 Int)) @p_37) :named @p_38))) :rule bind) (step t15 (cl (! (= @p_27 (! (not @p_38) :named @p_40)) :named @p_39)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_39) :named @p_42) (! (not @p_27) :named @p_41) @p_40) :rule equiv_pos2) (step t17 (cl (not @p_41) @p_25) :rule not_not) (step t18 (cl @p_42 @p_25 @p_40) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_40) :rule th_resolution :premises (t13 t15 t18)) (anchor :step t20 :args ((:= (veriT_vr2 Int) veriT_sk0) (:= (veriT_vr3 Int) veriT_sk1))) (step t20.t1 (cl (! (= veriT_vr2 veriT_sk0) :named @p_46)) :rule refl) (step t20.t2 (cl (! (= veriT_vr3 veriT_sk1) :named @p_49)) :rule refl) (step t20.t3 (cl (= @p_30 (! (< veriT_sk0 veriT_sk1) :named @p_45))) :rule cong :premises (t20.t1 t20.t2)) (step t20.t4 (cl @p_46) :rule refl) (step t20.t5 (cl (= @p_32 (! (* 2 veriT_sk0) :named @p_47))) :rule cong :premises (t20.t4)) (step t20.t6 (cl (= @p_33 (! (+ 1 @p_47) :named @p_48))) :rule cong :premises (t20.t5)) (step t20.t7 (cl @p_49) :rule refl) (step t20.t8 (cl (= @p_35 (! (* 2 veriT_sk1) :named @p_50))) :rule cong :premises (t20.t7)) (step t20.t9 (cl (= @p_36 (! (< @p_48 @p_50) :named @p_51))) :rule cong :premises (t20.t6 t20.t8)) (step t20.t10 (cl (= @p_37 (! (=> @p_45 @p_51) :named @p_52))) :rule cong :premises (t20.t3 t20.t9)) (step t20 (cl (= @p_38 @p_52)) :rule sko_forall) (step t21 (cl (! (= @p_40 (! (not @p_52) :named @p_54)) :named @p_53)) :rule cong :premises (t20)) (step t22 (cl (! (not @p_53) :named @p_56) (! (not @p_40) :named @p_55) @p_54) :rule equiv_pos2) (step t23 (cl (not @p_55) @p_38) :rule not_not) (step t24 (cl @p_56 @p_38 @p_54) :rule th_resolution :premises (t23 t22)) (step t25 (cl @p_54) :rule th_resolution :premises (t19 t21 t24)) (step t26 (cl (! (= @p_54 (! (and @p_45 (! (not @p_51) :named @p_61)) :named @p_58)) :named @p_57)) :rule bool_simplify) (step t27 (cl (! (not @p_57) :named @p_60) (! (not @p_54) :named @p_59) @p_58) :rule equiv_pos2) (step t28 (cl (not @p_59) @p_52) :rule not_not) (step t29 (cl @p_60 @p_52 @p_58) :rule th_resolution :premises (t28 t27)) (step t30 (cl @p_58) :rule th_resolution :premises (t25 t26 t29)) (step t31 (cl @p_45) :rule and :premises (t30)) (step t32 (cl @p_61) :rule and :premises (t30)) (step t33 (cl @p_51 (not @p_45)) :rule la_generic :args ((div 1 2) 1)) (step t34 (cl) :rule resolution :premises (t33 t31 t32)) b3000c2a2d0d57028ec3d5228440616f9e6398d4 84 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr2 Int)) (not (forall ((veriT_vr3 Int)) (! (or (! (< 2 (! (+ veriT_vr2 veriT_vr3) :named @p_29)) :named @p_30) (! (= 2 @p_29) :named @p_34) (! (< @p_29 2) :named @p_35)) :named @p_36)))) :named @p_42)) (define-fun veriT_sk1 () Int (! (choice ((veriT_vr3 Int)) (not (or (< 2 (! (+ @p_42 veriT_vr3) :named @p_43)) (= 2 @p_43) (< @p_43 2)))) :named @p_45)) (assume axiom0 (! (not (! (forall ((?v0 Int) (?v1 Int)) (! (or (! (< 2 (! (+ ?v0 ?v1) :named @p_1)) :named @p_4) (! (or (! (= 2 @p_1) :named @p_9) (! (< @p_1 2) :named @p_11)) :named @p_13)) :named @p_15)) :named @p_2)) :named @p_17)) (anchor :step t2 :args ((:= (?v0 Int) veriT_vr0) (:= (?v1 Int) veriT_vr1))) (step t2.t1 (cl (! (= ?v0 veriT_vr0) :named @p_6)) :rule refl) (step t2.t2 (cl (! (= ?v1 veriT_vr1) :named @p_7)) :rule refl) (step t2.t3 (cl (! (= @p_1 (! (+ veriT_vr0 veriT_vr1) :named @p_3)) :named @p_8)) :rule cong :premises (t2.t1 t2.t2)) (step t2.t4 (cl (= @p_4 (! (< 2 @p_3) :named @p_5))) :rule cong :premises (t2.t3)) (step t2.t5 (cl @p_6) :rule refl) (step t2.t6 (cl @p_7) :rule refl) (step t2.t7 (cl @p_8) :rule cong :premises (t2.t5 t2.t6)) (step t2.t8 (cl (= @p_9 (! (= 2 @p_3) :named @p_10))) :rule cong :premises (t2.t7)) (step t2.t9 (cl @p_6) :rule refl) (step t2.t10 (cl @p_7) :rule refl) (step t2.t11 (cl @p_8) :rule cong :premises (t2.t9 t2.t10)) (step t2.t12 (cl (= @p_11 (! (< @p_3 2) :named @p_12))) :rule cong :premises (t2.t11)) (step t2.t13 (cl (= @p_13 (! (or @p_10 @p_12) :named @p_14))) :rule cong :premises (t2.t8 t2.t12)) (step t2.t14 (cl (= @p_15 (! (or @p_5 @p_14) :named @p_16))) :rule cong :premises (t2.t4 t2.t13)) (step t2 (cl (= @p_2 (! (forall ((veriT_vr0 Int) (veriT_vr1 Int)) @p_16) :named @p_18))) :rule bind) (step t3 (cl (! (= @p_17 (! (not @p_18) :named @p_20)) :named @p_19)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_19) :named @p_22) (! (not @p_17) :named @p_21) @p_20) :rule equiv_pos2) (step t5 (cl (not @p_21) @p_2) :rule not_not) (step t6 (cl @p_22 @p_2 @p_20) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_20) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((veriT_vr0 Int) (veriT_vr1 Int))) (step t8.t1 (cl (= @p_16 (! (or @p_5 @p_10 @p_12) :named @p_23))) :rule ac_simp) (step t8 (cl (= @p_18 (! (forall ((veriT_vr0 Int) (veriT_vr1 Int)) @p_23) :named @p_24))) :rule bind) (step t9 (cl (! (= @p_20 (! (not @p_24) :named @p_26)) :named @p_25)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_25) :named @p_28) (! (not @p_20) :named @p_27) @p_26) :rule equiv_pos2) (step t11 (cl (not @p_27) @p_18) :rule not_not) (step t12 (cl @p_28 @p_18 @p_26) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_26) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((:= (veriT_vr0 Int) veriT_vr2) (:= (veriT_vr1 Int) veriT_vr3))) (step t14.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_31)) :rule refl) (step t14.t2 (cl (! (= veriT_vr1 veriT_vr3) :named @p_32)) :rule refl) (step t14.t3 (cl (! (= @p_3 @p_29) :named @p_33)) :rule cong :premises (t14.t1 t14.t2)) (step t14.t4 (cl (= @p_5 @p_30)) :rule cong :premises (t14.t3)) (step t14.t5 (cl @p_31) :rule refl) (step t14.t6 (cl @p_32) :rule refl) (step t14.t7 (cl @p_33) :rule cong :premises (t14.t5 t14.t6)) (step t14.t8 (cl (= @p_10 @p_34)) :rule cong :premises (t14.t7)) (step t14.t9 (cl @p_31) :rule refl) (step t14.t10 (cl @p_32) :rule refl) (step t14.t11 (cl @p_33) :rule cong :premises (t14.t9 t14.t10)) (step t14.t12 (cl (= @p_12 @p_35)) :rule cong :premises (t14.t11)) (step t14.t13 (cl (= @p_23 @p_36)) :rule cong :premises (t14.t4 t14.t8 t14.t12)) (step t14 (cl (= @p_24 (! (forall ((veriT_vr2 Int) (veriT_vr3 Int)) @p_36) :named @p_37))) :rule bind) (step t15 (cl (! (= @p_26 (! (not @p_37) :named @p_39)) :named @p_38)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_38) :named @p_41) (! (not @p_26) :named @p_40) @p_39) :rule equiv_pos2) (step t17 (cl (not @p_40) @p_24) :rule not_not) (step t18 (cl @p_41 @p_24 @p_39) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_39) :rule th_resolution :premises (t13 t15 t18)) (anchor :step t20 :args ((:= (veriT_vr2 Int) veriT_sk0) (:= (veriT_vr3 Int) veriT_sk1))) (step t20.t1 (cl (! (= veriT_vr2 veriT_sk0) :named @p_47)) :rule refl) (step t20.t2 (cl (! (= veriT_vr3 veriT_sk1) :named @p_48)) :rule refl) (step t20.t3 (cl (! (= @p_29 (! (+ veriT_sk0 veriT_sk1) :named @p_44)) :named @p_49)) :rule cong :premises (t20.t1 t20.t2)) (step t20.t4 (cl (= @p_30 (! (< 2 @p_44) :named @p_46))) :rule cong :premises (t20.t3)) (step t20.t5 (cl @p_47) :rule refl) (step t20.t6 (cl @p_48) :rule refl) (step t20.t7 (cl @p_49) :rule cong :premises (t20.t5 t20.t6)) (step t20.t8 (cl (= @p_34 (! (= 2 @p_44) :named @p_50))) :rule cong :premises (t20.t7)) (step t20.t9 (cl @p_47) :rule refl) (step t20.t10 (cl @p_48) :rule refl) (step t20.t11 (cl @p_49) :rule cong :premises (t20.t9 t20.t10)) (step t20.t12 (cl (= @p_35 (! (< @p_44 2) :named @p_51))) :rule cong :premises (t20.t11)) (step t20.t13 (cl (= @p_36 (! (or @p_46 @p_50 @p_51) :named @p_52))) :rule cong :premises (t20.t4 t20.t8 t20.t12)) (step t20 (cl (= @p_37 @p_52)) :rule sko_forall) (step t21 (cl (! (= @p_39 (! (not @p_52) :named @p_54)) :named @p_53)) :rule cong :premises (t20)) (step t22 (cl (! (not @p_53) :named @p_56) (! (not @p_39) :named @p_55) @p_54) :rule equiv_pos2) (step t23 (cl (not @p_55) @p_37) :rule not_not) (step t24 (cl @p_56 @p_37 @p_54) :rule th_resolution :premises (t23 t22)) (step t25 (cl @p_54) :rule th_resolution :premises (t19 t21 t24)) (step t26 (cl (not @p_46)) :rule not_or :premises (t25)) (step t27 (cl (not @p_50)) :rule not_or :premises (t25)) (step t28 (cl (not @p_51)) :rule not_or :premises (t25)) (step t29 (cl (or @p_50 (! (not (! (<= 2 @p_44) :named @p_59)) :named @p_57) (! (not (! (<= @p_44 2) :named @p_60)) :named @p_58))) :rule la_disequality) (step t30 (cl @p_50 @p_57 @p_58) :rule or :premises (t29)) (step t31 (cl @p_57 @p_58) :rule resolution :premises (t30 t27)) (step t32 (cl @p_59 @p_51) :rule la_generic :args (1 1)) (step t33 (cl @p_59) :rule resolution :premises (t32 t28)) (step t34 (cl @p_58) :rule resolution :premises (t31 t33)) (step t35 (cl @p_60 @p_46) :rule la_generic :args (1 1)) (step t36 (cl) :rule resolution :premises (t35 t26 t34)) ea93df392983d8a481a89ee529b9a9f976280447 371 0 unsat (assume axiom0 (! (not (! (=> (! (and (! (= x3$ (- (! (ite (! (< x2$ 0) :named @p_30) (! (- x2$) :named @p_31) x2$) :named @p_21) x1$)) :named @p_9) (and (! (= x4$ (- (! (ite (! (< x3$ 0) :named @p_32) (! (- x3$) :named @p_33) x3$) :named @p_22) x2$)) :named @p_10) (and (! (= x5$ (- (! (ite (! (< x4$ 0) :named @p_34) (! (- x4$) :named @p_35) x4$) :named @p_23) x3$)) :named @p_11) (and (! (= x6$ (- (! (ite (! (< x5$ 0) :named @p_36) (! (- x5$) :named @p_37) x5$) :named @p_24) x4$)) :named @p_12) (and (! (= x7$ (- (! (ite (! (< x6$ 0) :named @p_38) (! (- x6$) :named @p_39) x6$) :named @p_25) x5$)) :named @p_13) (and (! (= x8$ (- (! (ite (! (< x7$ 0) :named @p_40) (! (- x7$) :named @p_41) x7$) :named @p_26) x6$)) :named @p_14) (and (! (= x9$ (- (! (ite (! (< x8$ 0) :named @p_42) (! (- x8$) :named @p_43) x8$) :named @p_27) x7$)) :named @p_15) (and (! (= x10$ (- (! (ite (! (< x9$ 0) :named @p_44) (! (- x9$) :named @p_45) x9$) :named @p_28) x8$)) :named @p_16) (! (= x11$ (- (! (ite (! (< x10$ 0) :named @p_46) (! (- x10$) :named @p_47) x10$) :named @p_29) x9$)) :named @p_17))))))))) :named @p_2) (! (and (! (= x1$ x10$) :named @p_70) (! (= x2$ x11$) :named @p_71)) :named @p_3)) :named @p_7)) :named @p_1)) (step t2 (cl (! (= @p_1 (! (and @p_2 (! (not @p_3) :named @p_18)) :named @p_5)) :named @p_4)) :rule bool_simplify) (step t3 (cl (! (not @p_4) :named @p_8) (! (not @p_1) :named @p_6) @p_5) :rule equiv_pos2) (step t4 (cl (not @p_6) @p_7) :rule not_not) (step t5 (cl @p_8 @p_7 @p_5) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_5) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl (! (= @p_5 (! (and @p_9 @p_10 @p_11 @p_12 @p_13 @p_14 @p_15 @p_16 @p_17 @p_18) :named @p_20)) :named @p_19)) :rule ac_simp) (step t8 (cl (not @p_19) (not @p_5) @p_20) :rule equiv_pos2) (step t9 (cl @p_20) :rule th_resolution :premises (t6 t7 t8)) (step t10 (cl (! (= @p_20 (! (and (and (! (= x3$ (- @p_21 x1$)) :named @p_50) (! (= x4$ (- @p_22 x2$)) :named @p_51) (! (= x5$ (- @p_23 x3$)) :named @p_52) (! (= x6$ (- @p_24 x4$)) :named @p_53) (! (= x7$ (- @p_25 x5$)) :named @p_54) (! (= x8$ (- @p_26 x6$)) :named @p_55) (! (= x9$ (- @p_27 x7$)) :named @p_56) (! (= x10$ (- @p_28 x8$)) :named @p_57) (! (= x11$ (! (- @p_29 x9$) :named @p_107)) :named @p_58) @p_18) (! (ite @p_30 (! (= @p_31 @p_21) :named @p_73) (! (= x2$ @p_21) :named @p_72)) :named @p_59) (! (ite @p_32 (! (= @p_33 @p_22) :named @p_75) (! (= x3$ @p_22) :named @p_74)) :named @p_60) (! (ite @p_34 (! (= @p_35 @p_23) :named @p_77) (! (= x4$ @p_23) :named @p_76)) :named @p_61) (! (ite @p_36 (! (= @p_37 @p_24) :named @p_79) (! (= x5$ @p_24) :named @p_78)) :named @p_62) (! (ite @p_38 (! (= @p_39 @p_25) :named @p_81) (! (= x6$ @p_25) :named @p_80)) :named @p_63) (! (ite @p_40 (! (= @p_41 @p_26) :named @p_83) (! (= x7$ @p_26) :named @p_82)) :named @p_64) (! (ite @p_42 (! (= @p_43 @p_27) :named @p_85) (! (= x8$ @p_27) :named @p_84)) :named @p_65) (! (ite @p_44 (! (= @p_45 @p_28) :named @p_87) (! (= x9$ @p_28) :named @p_86)) :named @p_66) (! (ite @p_46 (! (= @p_47 @p_29) :named @p_89) (! (= x10$ @p_29) :named @p_88)) :named @p_67)) :named @p_49)) :named @p_48)) :rule ite_intro) (step t11 (cl (not @p_48) (not @p_20) @p_49) :rule equiv_pos2) (step t12 (cl @p_49) :rule th_resolution :premises (t9 t10 t11)) (step t13 (cl (! (= @p_49 (! (and @p_50 @p_51 @p_52 @p_53 @p_54 @p_55 @p_56 @p_57 @p_58 @p_18 @p_59 @p_60 @p_61 @p_62 @p_63 @p_64 @p_65 @p_66 @p_67) :named @p_69)) :named @p_68)) :rule ac_simp) (step t14 (cl (not @p_68) (not @p_49) @p_69) :rule equiv_pos2) (step t15 (cl @p_69) :rule th_resolution :premises (t12 t13 t14)) (step t16 (cl @p_50) :rule and :premises (t15)) (step t17 (cl @p_51) :rule and :premises (t15)) (step t18 (cl @p_52) :rule and :premises (t15)) (step t19 (cl @p_53) :rule and :premises (t15)) (step t20 (cl @p_54) :rule and :premises (t15)) (step t21 (cl @p_55) :rule and :premises (t15)) (step t22 (cl @p_56) :rule and :premises (t15)) (step t23 (cl @p_57) :rule and :premises (t15)) (step t24 (cl @p_58) :rule and :premises (t15)) (step t25 (cl @p_18) :rule and :premises (t15)) (step t26 (cl (! (not @p_70) :named @p_165) (not @p_71)) :rule not_and :premises (t25)) (step t27 (cl @p_59) :rule and :premises (t15)) (step t28 (cl @p_30 @p_72) :rule ite1 :premises (t27)) (step t29 (cl (! (not @p_30) :named @p_140) @p_73) :rule ite2 :premises (t27)) (step t30 (cl @p_60) :rule and :premises (t15)) (step t31 (cl @p_32 @p_74) :rule ite1 :premises (t30)) (step t32 (cl (! (not @p_32) :named @p_137) @p_75) :rule ite2 :premises (t30)) (step t33 (cl @p_61) :rule and :premises (t15)) (step t34 (cl @p_34 @p_76) :rule ite1 :premises (t33)) (step t35 (cl (! (not @p_34) :named @p_144) @p_77) :rule ite2 :premises (t33)) (step t36 (cl @p_62) :rule and :premises (t15)) (step t37 (cl @p_36 @p_78) :rule ite1 :premises (t36)) (step t38 (cl (! (not @p_36) :named @p_125) @p_79) :rule ite2 :premises (t36)) (step t39 (cl @p_63) :rule and :premises (t15)) (step t40 (cl @p_38 @p_80) :rule ite1 :premises (t39)) (step t41 (cl (! (not @p_38) :named @p_128) @p_81) :rule ite2 :premises (t39)) (step t42 (cl @p_64) :rule and :premises (t15)) (step t43 (cl @p_40 @p_82) :rule ite1 :premises (t42)) (step t44 (cl (! (not @p_40) :named @p_136) @p_83) :rule ite2 :premises (t42)) (step t45 (cl @p_65) :rule and :premises (t15)) (step t46 (cl @p_42 @p_84) :rule ite1 :premises (t45)) (step t47 (cl (! (not @p_42) :named @p_101) @p_85) :rule ite2 :premises (t45)) (step t48 (cl @p_66) :rule and :premises (t15)) (step t49 (cl @p_44 @p_86) :rule ite1 :premises (t48)) (step t50 (cl (! (not @p_44) :named @p_169) @p_87) :rule ite2 :premises (t48)) (step t51 (cl @p_67) :rule and :premises (t15)) (step t52 (cl @p_46 @p_88) :rule ite1 :premises (t51)) (step t53 (cl (! (not @p_46) :named @p_112) @p_89) :rule ite2 :premises (t51)) (step t54 (cl (or @p_70 (! (not (! (<= x1$ x10$) :named @p_160)) :named @p_90) (! (not (! (<= x10$ x1$) :named @p_156)) :named @p_91))) :rule la_disequality) (step t55 (cl @p_70 @p_90 @p_91) :rule or :premises (t54)) (step t56 (cl (or @p_87 (! (not (! (<= @p_45 @p_28) :named @p_94)) :named @p_92) (! (not (! (<= @p_28 @p_45) :named @p_95)) :named @p_93))) :rule la_disequality) (step t57 (cl @p_87 @p_92 @p_93) :rule or :premises (t56)) (step t58 (cl @p_94 @p_95) :rule la_generic :args (1 1)) (step t59 (cl @p_94 @p_44 (! (not @p_86) :named @p_96)) :rule la_generic :args (1 2 (- 1))) (step t60 (cl (! (not @p_87) :named @p_97) (not (! (= x9$ @p_45) :named @p_98)) @p_93 (! (<= @p_45 x9$) :named @p_99)) :rule eq_congruent_pred) (step t61 (cl @p_96 @p_97 @p_98) :rule eq_transitive) (step t62 (cl @p_97 @p_93 @p_99 @p_96 @p_97) :rule th_resolution :premises (t60 t61)) (step t63 (cl @p_97 @p_93 @p_99 @p_96) :rule contraction :premises (t62)) (step t64 (cl (! (not (! (= @p_45 @p_45) :named @p_100)) :named @p_174) @p_96 @p_92 @p_99) :rule eq_congruent_pred) (step t65 (cl @p_100) :rule eq_reflexive) (step t66 (cl @p_96 @p_92 @p_99) :rule th_resolution :premises (t64 t65)) (step t67 (cl @p_99 @p_96) :rule resolution :premises (t63 t50 t59 t58 t66)) (step t68 (cl @p_101 (! (not (! (= x8$ @p_43) :named @p_102)) :named @p_167)) :rule la_generic :args (2 1)) (step t69 (cl (! (not @p_84) :named @p_103) (! (not @p_85) :named @p_104) @p_102) :rule eq_transitive) (step t70 (cl @p_101 @p_103 @p_104) :rule th_resolution :premises (t68 t69)) (step t71 (cl (or @p_102 (! (not (! (<= x8$ @p_43) :named @p_111)) :named @p_105) (! (not (! (<= @p_43 x8$) :named @p_147)) :named @p_106))) :rule la_disequality) (step t72 (cl @p_102 @p_105 @p_106) :rule or :premises (t71)) (step t73 (cl (or (! (= x2$ @p_107) :named @p_108) (! (not (! (<= x2$ @p_107) :named @p_158)) :named @p_109) (! (not (! (<= @p_107 x2$) :named @p_151)) :named @p_110))) :rule la_disequality) (step t74 (cl @p_108 @p_109 @p_110) :rule or :premises (t73)) (step t75 (cl @p_111 @p_101) :rule la_generic :args (1 2)) (step t76 (cl @p_112 (not (! (= x10$ @p_47) :named @p_113))) :rule la_generic :args (2 1)) (step t77 (cl (! (not @p_88) :named @p_114) (! (not @p_89) :named @p_115) @p_113) :rule eq_transitive) (step t78 (cl @p_112 @p_114 @p_115) :rule th_resolution :premises (t76 t77)) (step t79 (cl (or @p_77 (! (not (! (<= @p_35 @p_23) :named @p_122)) :named @p_116) (! (not (! (<= @p_23 @p_35) :named @p_146)) :named @p_117))) :rule la_disequality) (step t80 (cl @p_77 @p_116 @p_117) :rule or :premises (t79)) (step t81 (cl (! (<= @p_21 @p_31) :named @p_120) (! (<= @p_31 @p_21) :named @p_118)) :rule la_generic :args (1 1)) (step t82 (cl @p_30 @p_118 (! (not @p_72) :named @p_171)) :rule la_generic :args (2 1 (- 1))) (step t83 (cl @p_30 @p_118) :rule resolution :premises (t82 t28)) (step t84 (cl (! (not @p_73) :named @p_119) @p_119 (! (not @p_120) :named @p_121) @p_118) :rule eq_congruent_pred) (step t85 (cl @p_119 @p_121 @p_118) :rule contraction :premises (t84)) (step t86 (cl @p_118) :rule resolution :premises (t85 t29 t83 t81)) (step t87 (cl @p_122 @p_34 (! (not @p_76) :named @p_142)) :rule la_generic :args (1 2 (- 1))) (step t88 (cl (! (= x10$ x10$) :named @p_162)) :rule eq_reflexive) (step t89 (cl @p_114 @p_114 (! (not (! (<= x10$ @p_29) :named @p_163)) :named @p_123) (! (<= @p_29 x10$) :named @p_124)) :rule eq_congruent_pred) (step t90 (cl @p_114 @p_123 @p_124) :rule contraction :premises (t89)) (step t91 (cl (! (<= @p_39 @p_25) :named @p_129) (! (<= @p_25 @p_39) :named @p_183)) :rule la_generic :args (2 2)) (step t92 (cl @p_101 @p_125 (! (not @p_54) :named @p_130) (! (not @p_55) :named @p_131) (! (not @p_82) :named @p_126) (! (not @p_80) :named @p_127)) :rule la_generic :args (1 1 1 1 (- 1) (- 1))) (step t93 (cl @p_101 @p_125 @p_126 @p_127) :rule resolution :premises (t92 t20 t21)) (step t94 (cl @p_125 @p_128 (! (not @p_99) :named @p_132) (! (not @p_129) :named @p_133) @p_130 @p_131 (! (not @p_56) :named @p_134) @p_104 @p_126) :rule la_generic :args (4 6 1 4 4 2 (- 2) 2 (- 2))) (step t95 (cl @p_125 @p_128 @p_132 @p_133 @p_104 @p_126) :rule resolution :premises (t94 t20 t21 t22)) (step t96 (cl @p_101 (! (not (! (<= x9$ @p_45) :named @p_173)) :named @p_135) @p_128 @p_131 @p_134 @p_104 @p_126) :rule la_generic :args (4 1 2 2 2 (- 2) (- 2))) (step t97 (cl @p_101 @p_135 @p_128 @p_104 @p_126) :rule resolution :premises (t96 t21 t22)) (step t98 (cl @p_136 @p_42 @p_135 @p_134 @p_103) :rule la_generic :args (2 2 1 2 (- 2))) (step t99 (cl @p_136 @p_42 @p_135 @p_103) :rule resolution :premises (t98 t22)) (step t100 (cl @p_137 @p_38 @p_116 (! (not @p_52) :named @p_141) (! (not @p_53) :named @p_148) (! (not @p_79) :named @p_138)) :rule la_generic :args (1 1 1 1 (- 1) 1)) (step t101 (cl @p_137 @p_38 @p_116 @p_138) :rule resolution :premises (t100 t18 t19)) (step t102 (cl (! (not @p_77) :named @p_139) @p_139 @p_117 @p_122) :rule eq_congruent_pred) (step t103 (cl @p_139 @p_117 @p_122) :rule contraction :premises (t102)) (step t104 (cl @p_140 @p_125 (! (not @p_51) :named @p_145) @p_141 @p_142 (! (not @p_74) :named @p_143)) :rule la_generic :args (1 1 1 1 (- 1) (- 1))) (step t105 (cl @p_140 @p_125 @p_142 @p_143) :rule resolution :premises (t104 t17 t18)) (step t106 (cl @p_140 @p_144 @p_32 @p_145 @p_143) :rule la_generic :args (1 1 1 1 (- 1))) (step t107 (cl @p_140 @p_144 @p_32 @p_143) :rule resolution :premises (t106 t17)) (step t108 (cl @p_140 @p_32 @p_117 @p_145 @p_142 @p_143) :rule la_generic :args (2 2 1 2 (- 1) (- 2))) (step t109 (cl @p_140 @p_32 @p_117 @p_142 @p_143) :rule resolution :premises (t108 t17)) (step t110 (cl @p_139 @p_139 @p_116 @p_146) :rule eq_congruent_pred) (step t111 (cl @p_139 @p_116 @p_146) :rule contraction :premises (t110)) (step t112 (cl (! (not @p_78) :named @p_149) @p_138 (! (= x5$ @p_37) :named @p_170)) :rule eq_transitive) (step t113 (cl @p_140 @p_32 @p_147 @p_145 @p_141 @p_148 @p_130 @p_131 @p_149 @p_142 @p_127 @p_143 (! (not @p_83) :named @p_150)) :rule la_generic :args (2 4 1 2 (- 2) (- 4) (- 2) 2 4 2 2 (- 2) (- 2))) (step t114 (cl @p_140 @p_32 @p_147 @p_149 @p_142 @p_127 @p_143 @p_150) :rule resolution :premises (t113 t17 t18 t19 t20 t21)) (step t115 (cl @p_32 (! (not @p_124) :named @p_152) @p_151 @p_145 @p_141 @p_148 @p_130 @p_131 (! (not @p_57) :named @p_157) @p_96 @p_149 @p_142 @p_127 @p_143 @p_150) :rule la_generic :args (2 1 1 1 (- 1) (- 2) (- 1) 1 (- 1) 1 2 1 1 (- 1) (- 1))) (step t116 (cl @p_32 @p_152 @p_151 @p_96 @p_149 @p_142 @p_127 @p_143 @p_150) :rule resolution :premises (t115 t17 t18 t19 t20 t21 t23)) (step t117 (cl (! (<= @p_26 @p_41) :named @p_153) (! (<= @p_41 @p_26) :named @p_154)) :rule la_generic :args (1 1)) (step t118 (cl @p_150 @p_150 (! (not @p_153) :named @p_155) @p_154) :rule eq_congruent_pred) (step t119 (cl @p_150 @p_155 @p_154) :rule contraction :premises (t118)) (step t120 (cl @p_154 @p_150) :rule resolution :premises (t117 t119)) (step t121 (cl @p_156 (! (not @p_118) :named @p_168) @p_38 (! (not @p_50) :named @p_159) @p_145 @p_141 @p_148 @p_130 @p_134 @p_157 @p_96 @p_149 @p_142 @p_103 @p_127 @p_143) :rule la_generic :args (1 1 2 1 (- 1) (- 2) (- 1) 1 (- 1) (- 1) 1 1 2 1 (- 1) 1)) (step t122 (cl @p_156 @p_38 @p_96 @p_149 @p_142 @p_103 @p_127 @p_143) :rule resolution :premises (t121 t16 t17 t18 t19 t20 t22 t23 t86)) (step t123 (cl @p_91 @p_121 @p_123 @p_158 @p_159 @p_141 @p_148 @p_131 @p_134 @p_157 @p_96 @p_149 @p_142 @p_103 @p_150) :rule la_generic :args (1 1 1 1 (- 1) 1 1 (- 1) 1 2 (- 2) (- 1) (- 1) (- 1) 1)) (step t124 (cl @p_91 @p_121 @p_123 @p_158 @p_96 @p_149 @p_142 @p_103 @p_150) :rule resolution :premises (t123 t16 t18 t19 t21 t22 t23)) (step t125 (cl @p_160 @p_121 @p_32 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_96 @p_149 @p_103 @p_127 @p_143) :rule la_generic :args (1 1 2 (- 1) 1 (- 1) (- 1) 1 1 (- 1) 1 (- 1) 1 (- 1))) (step t126 (cl @p_160 @p_121 @p_32 @p_96 @p_149 @p_103 @p_127 @p_143) :rule resolution :premises (t125 t16 t17 t19 t20 t22 t23)) (step t127 (cl (! (not @p_108) :named @p_161) (not @p_58) @p_71) :rule eq_transitive) (step t128 (cl @p_161 @p_71) :rule resolution :premises (t127 t24)) (step t129 (cl (not @p_162) (! (not (! (= x1$ @p_29) :named @p_166)) :named @p_164) @p_91 @p_163) :rule eq_congruent_pred) (step t130 (cl @p_164 @p_91 @p_163) :rule th_resolution :premises (t129 t88)) (step t131 (cl @p_165 @p_114 @p_166) :rule eq_transitive) (step t132 (cl @p_91 @p_163 @p_165 @p_114) :rule th_resolution :premises (t130 t131)) (step t133 (cl @p_140 @p_32 @p_112 @p_145 @p_148 @p_130 @p_134 @p_157 @p_96 @p_149 @p_103 @p_127 @p_143) :rule la_generic :args (1 1 1 1 (- 1) (- 1) 1 1 (- 1) 1 (- 1) 1 (- 1))) (step t134 (cl @p_140 @p_32 @p_112 @p_96 @p_149 @p_103 @p_127 @p_143) :rule resolution :premises (t133 t17 t19 t20 t22 t23)) (step t135 (cl @p_167 @p_104 @p_84) :rule eq_transitive) (step t136 (cl @p_140 @p_32 @p_40 @p_145 @p_148 @p_130 @p_149 @p_127 @p_143) :rule la_generic :args (1 1 1 1 (- 1) (- 1) 1 1 (- 1))) (step t137 (cl @p_140 @p_32 @p_40 @p_149 @p_127 @p_143) :rule resolution :premises (t136 t17 t19 t20)) (step t138 (cl @p_119 @p_119 @p_120 @p_168) :rule eq_congruent_pred) (step t139 (cl @p_119 @p_120 @p_168) :rule contraction :premises (t138)) (step t140 (cl @p_119 @p_120) :rule resolution :premises (t139 t86)) (step t141 (cl @p_30 @p_136 @p_38 @p_145 @p_141 @p_130 @p_142 @p_127 @p_143) :rule la_generic :args (1 1 1 (- 1) (- 1) 1 1 (- 1) 1)) (step t142 (cl @p_30 @p_136 @p_38 @p_142 @p_127 @p_143) :rule resolution :premises (t141 t17 t18 t20)) (step t143 (cl @p_40 @p_112 @p_90 @p_168 @p_36 @p_159 @p_145 @p_141 @p_148 @p_130 @p_149 @p_142 @p_127 @p_143) :rule la_generic :args (1 1 1 1 2 1 (- 1) (- 2) (- 1) (- 1) 1 2 1 1)) (step t144 (cl @p_40 @p_112 @p_90 @p_36 @p_149 @p_142 @p_127 @p_143) :rule resolution :premises (t143 t16 t17 t18 t19 t20 t86)) (step t145 (cl @p_36 @p_42 @p_169 @p_32 @p_141 @p_148 @p_130 @p_134 @p_149 @p_142 @p_103 @p_127) :rule la_generic :args (1 1 1 1 (- 1) (- 1) (- 1) 1 1 1 (- 1) 1)) (step t146 (cl @p_36 @p_42 @p_169 @p_32 @p_149 @p_142 @p_103 @p_127) :rule resolution :premises (t145 t18 t19 t20 t22)) (step t147 (cl @p_125 (not @p_170)) :rule la_generic :args (2 1)) (step t148 (cl @p_125 @p_149 @p_138) :rule th_resolution :premises (t147 t112)) (step t149 (cl @p_125 @p_149) :rule resolution :premises (t148 t38)) (step t150 (cl @p_120 @p_36 @p_145 @p_141 @p_171 @p_142 @p_143) :rule la_generic :args (1 2 (- 2) (- 2) 1 2 2)) (step t151 (cl @p_120 @p_36 @p_171 @p_142 @p_143) :rule resolution :premises (t150 t17 t18)) (step t152 (cl (! (= @p_41 @p_41) :named @p_176)) :rule eq_reflexive) (step t153 (cl @p_144 @p_132 @p_36 @p_148 @p_131 @p_134 @p_149 @p_126 @p_103) :rule la_generic :args (2 1 2 2 (- 2) (- 2) (- 2) 2 2)) (step t154 (cl @p_144 @p_132 @p_36 @p_149 @p_126 @p_103) :rule resolution :premises (t153 t19 t21 t22)) (step t155 (cl @p_144 (! (<= @p_41 x7$) :named @p_172) @p_148 @p_130 @p_149 @p_127) :rule la_generic :args (2 1 2 2 (- 2) (- 2))) (step t156 (cl @p_144 @p_172 @p_149 @p_127) :rule resolution :premises (t155 t19 t20)) (step t157 (cl @p_173 @p_99) :rule la_generic :args ((- 2) 2)) (step t158 (cl @p_42 @p_151 @p_152 @p_91 @p_159 @p_141 @p_148 @p_130 @p_131 @p_134 @p_171 @p_149 @p_126 @p_103 @p_127 @p_139) :rule la_generic :args (2 1 1 1 (- 1) 1 (- 1) (- 2) (- 1) 1 1 1 1 (- 1) 2 (- 1))) (step t159 (cl @p_42 @p_151 @p_152 @p_91 @p_171 @p_149 @p_126 @p_103 @p_127 @p_139) :rule resolution :premises (t158 t16 t18 t19 t20 t21 t22)) (step t160 (cl @p_90 @p_158 @p_123 @p_36 @p_159 @p_141 @p_148 @p_131 @p_134 @p_171 @p_149 @p_126 @p_103 @p_139) :rule la_generic :args (1 1 1 2 1 (- 1) 1 (- 1) (- 1) (- 1) (- 1) 1 1 1)) (step t161 (cl @p_90 @p_158 @p_123 @p_36 @p_171 @p_149 @p_126 @p_103 @p_139) :rule resolution :premises (t160 t16 t18 t19 t21 t22)) (step t162 (cl @p_156 @p_42 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_171 @p_149 @p_97 @p_103 @p_127 @p_143) :rule la_generic :args (1 2 1 1 (- 1) (- 1) 1 (- 1) (- 1) 1 1 (- 1) 1 (- 1))) (step t163 (cl @p_156 @p_42 @p_171 @p_149 @p_97 @p_103 @p_127 @p_143) :rule resolution :premises (t162 t16 t17 t19 t20 t22 t23)) (step t164 (cl @p_144 @p_90 @p_112 @p_159 @p_145 @p_171 @p_143) :rule la_generic :args (1 1 1 1 1 (- 1) (- 1))) (step t165 (cl @p_144 @p_90 @p_112 @p_171 @p_143) :rule resolution :premises (t164 t16 t17)) (step t166 (cl @p_160 @p_36 @p_159 @p_145 @p_148 @p_130 @p_131 @p_134 @p_157 @p_171 @p_149 @p_97 @p_126 @p_103 @p_127 @p_143) :rule la_generic :args (1 2 (- 1) (- 1) 1 (- 1) (- 2) (- 1) 1 1 (- 1) (- 1) 2 1 1 1)) (step t167 (cl @p_160 @p_36 @p_171 @p_149 @p_97 @p_126 @p_103 @p_127 @p_143) :rule resolution :premises (t166 t16 t17 t19 t20 t21 t22 t23)) (step t168 (cl @p_160 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_171 @p_149 @p_104 @p_127 @p_143 @p_97) :rule la_generic :args (1 (- 1) (- 1) 1 1 (- 1) 1 1 (- 1) 1 (- 1) 1 (- 1))) (step t169 (cl @p_160 @p_171 @p_149 @p_104 @p_127 @p_143 @p_97) :rule resolution :premises (t168 t16 t17 t19 t20 t22 t23)) (step t170 (cl @p_160 @p_95 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_171 @p_149 @p_104 @p_127 @p_143) :rule la_generic :args (1 1 (- 1) (- 1) 1 1 (- 1) 1 1 (- 1) 1 (- 1) 1)) (step t171 (cl @p_160 @p_95 @p_171 @p_149 @p_104 @p_127 @p_143) :rule resolution :premises (t170 t16 t17 t19 t20 t22 t23)) (step t172 (cl @p_91 @p_151 @p_159 @p_141 @p_148 @p_130 @p_131 @p_134 @p_114 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139) :rule la_generic :args (1 1 (- 1) 1 (- 1) (- 2) (- 1) 1 1 1 1 (- 1) 1 2 (- 1))) (step t173 (cl @p_91 @p_151 @p_114 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139) :rule resolution :premises (t172 t16 t18 t19 t20 t21 t22)) (step t174 (cl @p_90 @p_158 @p_159 @p_141 @p_148 @p_130 @p_131 @p_134 @p_114 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139) :rule la_generic :args (1 1 1 (- 1) 1 2 1 (- 1) (- 1) (- 1) (- 1) 1 (- 1) (- 2) 1)) (step t175 (cl @p_90 @p_158 @p_114 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139) :rule resolution :premises (t174 t16 t18 t19 t20 t21 t22)) (step t176 (cl @p_156 @p_32 @p_159 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_157 @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_143) :rule la_generic :args (1 2 1 1 (- 2) 1 3 2 (- 1) (- 1) 1 (- 1) (- 1) 1 (- 2) (- 3) 2 (- 1))) (step t177 (cl @p_156 @p_32 @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_143) :rule resolution :premises (t176 t16 t17 t18 t19 t20 t21 t22 t23)) (step t178 (cl @p_156 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_171 @p_149 @p_104 @p_127 @p_143 @p_97) :rule la_generic :args (1 1 1 (- 1) (- 1) 1 (- 1) (- 1) 1 (- 1) 1 (- 1) 1)) (step t179 (cl @p_156 @p_171 @p_149 @p_104 @p_127 @p_143 @p_97) :rule resolution :premises (t178 t16 t17 t19 t20 t22 t23)) (step t180 (cl @p_144 @p_125 @p_136 @p_148 @p_130 @p_127 @p_138) :rule la_generic :args (1 2 1 1 1 (- 1) (- 1))) (step t181 (cl @p_144 @p_125 @p_136 @p_127 @p_138) :rule resolution :premises (t180 t19 t20)) (step t182 (cl @p_144 @p_125 @p_132 @p_148 @p_131 @p_134 @p_126 @p_103 @p_138) :rule la_generic :args (2 2 1 2 (- 2) (- 2) 2 2 (- 2))) (step t183 (cl @p_144 @p_125 @p_132 @p_126 @p_103 @p_138) :rule resolution :premises (t182 t19 t21 t22)) (step t184 (cl @p_91 @p_151 @p_159 @p_141 @p_148 @p_131 @p_134 @p_114 @p_171 @p_126 @p_103 @p_138 @p_139) :rule la_generic :args (1 1 (- 1) 1 (- 1) 1 1 1 1 (- 1) (- 1) 1 (- 1))) (step t185 (cl @p_91 @p_151 @p_114 @p_171 @p_126 @p_103 @p_138 @p_139) :rule resolution :premises (t184 t16 t18 t19 t21 t22)) (step t186 (cl @p_90 @p_158 @p_159 @p_141 @p_148 @p_131 @p_134 @p_114 @p_171 @p_126 @p_103 @p_138 @p_139) :rule la_generic :args (1 1 1 (- 1) 1 (- 1) (- 1) (- 1) (- 1) 1 1 (- 1) 1)) (step t187 (cl @p_90 @p_158 @p_114 @p_171 @p_126 @p_103 @p_138 @p_139) :rule resolution :premises (t186 t16 t18 t19 t21 t22)) (step t188 (cl @p_109 @p_91 @p_114 @p_171 @p_126 @p_103 @p_138 @p_139 @p_90) :rule resolution :premises (t128 t74 t26 t185 t55)) (step t189 (cl @p_156 @p_159 @p_145 @p_148 @p_130 @p_131 @p_134 @p_157 @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_97) :rule la_generic :args (1 1 1 (- 1) 1 2 1 (- 1) (- 1) (- 2) (- 1) (- 1) 1 (- 1) 1)) (step t190 (cl @p_156 @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_97) :rule resolution :premises (t189 t16 t17 t19 t20 t21 t22 t23)) (step t191 (cl @p_160 @p_159 @p_145 @p_148 @p_130 @p_131 @p_134 @p_157 @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_97) :rule la_generic :args (1 (- 1) (- 1) 1 (- 1) (- 2) (- 1) 1 1 2 1 1 (- 1) 1 (- 1))) (step t192 (cl @p_160 @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_97) :rule resolution :premises (t191 t16 t17 t19 t20 t21 t22 t23)) (step t193 (cl @p_160 @p_95 @p_159 @p_145 @p_148 @p_130 @p_131 @p_134 @p_157 @p_171 @p_126 @p_103 @p_127 @p_138 @p_143) :rule la_generic :args (1 1 (- 1) (- 1) 1 (- 1) (- 2) (- 1) 1 1 2 1 1 (- 1) 1)) (step t194 (cl @p_160 @p_95 @p_171 @p_126 @p_103 @p_127 @p_138 @p_143) :rule resolution :premises (t193 t16 t17 t19 t20 t21 t22 t23)) (step t195 (cl @p_125 @p_117 @p_132 @p_148 @p_131 @p_134 @p_142 @p_126 @p_103 @p_138) :rule la_generic :args (2 1 1 2 (- 2) (- 2) (- 1) 2 2 (- 2))) (step t196 (cl @p_125 @p_117 @p_132 @p_142 @p_126 @p_103 @p_138) :rule resolution :premises (t195 t19 t21 t22)) (step t197 (cl @p_90 @p_117 @p_112 @p_159 @p_145 @p_171 @p_142 @p_143) :rule la_generic :args (2 1 2 2 2 (- 2) (- 1) (- 2))) (step t198 (cl @p_90 @p_117 @p_112 @p_171 @p_142 @p_143) :rule resolution :premises (t197 t16 t17)) (step t199 (cl @p_90 @p_151 @p_159 @p_145 @p_141 @p_148 @p_131 @p_134 @p_171 @p_142 @p_126 @p_103 @p_115 @p_138 @p_143) :rule la_generic :args (1 1 1 2 1 (- 1) 1 1 (- 1) (- 1) (- 1) (- 1) 1 1 (- 2))) (step t200 (cl @p_90 @p_151 @p_171 @p_142 @p_126 @p_103 @p_115 @p_138 @p_143) :rule resolution :premises (t199 t16 t17 t18 t19 t21 t22)) (step t201 (cl @p_91 @p_158 @p_159 @p_145 @p_141 @p_148 @p_131 @p_134 @p_171 @p_142 @p_126 @p_103 @p_115 @p_138 @p_143) :rule la_generic :args (1 1 (- 1) (- 2) (- 1) 1 (- 1) (- 1) 1 1 1 1 (- 1) (- 1) 2)) (step t202 (cl @p_91 @p_158 @p_171 @p_142 @p_126 @p_103 @p_115 @p_138 @p_143) :rule resolution :premises (t201 t16 t17 t18 t19 t21 t22)) (step t203 (cl @p_46 @p_91 @p_146 @p_159 @p_145 @p_171 @p_142 @p_143) :rule la_generic :args (2 2 1 (- 2) (- 2) 2 1 2)) (step t204 (cl @p_46 @p_91 @p_146 @p_171 @p_142 @p_143) :rule resolution :premises (t203 t16 t17)) (step t205 (cl @p_173 @p_38 @p_131 @p_134 @p_126 @p_103) :rule la_generic :args (1 2 (- 2) (- 2) 2 2)) (step t206 (cl @p_173 @p_38 @p_126 @p_103) :rule resolution :premises (t205 t21 t22)) (step t207 (cl @p_96 @p_174 @p_135 @p_95) :rule eq_congruent_pred) (step t208 (cl @p_96 @p_135 @p_95) :rule th_resolution :premises (t207 t65)) (step t209 (cl @p_142 @p_77 @p_117) :rule resolution :premises (t87 t80 t35)) (step t210 (cl @p_137 @p_116 @p_42 @p_141 @p_130 @p_131 @p_142 @p_126 @p_127) :rule la_generic :args (2 1 2 2 (- 2) (- 2) (- 1) 2 2)) (step t211 (cl @p_137 @p_116 @p_42 @p_142 @p_126 @p_127) :rule resolution :premises (t210 t18 t20 t21)) (step t212 (cl @p_101 @p_103) :rule resolution :premises (t70 t47)) (step t213 (cl @p_146 @p_122) :rule la_generic :args (2 2)) (step t214 (cl @p_142 @p_137 @p_42 @p_126 @p_127) :rule resolution :premises (t213 t103 t35 t87 t211)) (step t215 (cl @p_137 @p_42 @p_144 @p_141 @p_130 @p_131 @p_126 @p_127 @p_139) :rule la_generic :args (1 1 1 1 (- 1) (- 1) 1 1 (- 1))) (step t216 (cl @p_137 @p_42 @p_144 @p_126 @p_127 @p_139) :rule resolution :premises (t215 t18 t20 t21)) (step t217 (cl @p_137 @p_144 @p_128 @p_141 @p_148 @p_149 @p_139) :rule la_generic :args (1 2 1 1 1 (- 1) (- 1))) (step t218 (cl @p_137 @p_144 @p_128 @p_149 @p_139) :rule resolution :premises (t217 t18 t19)) (step t219 (cl @p_137 @p_169 @p_141 @p_148 @p_130 @p_131 @p_134 @p_149 @p_104 @p_126 @p_127 @p_139) :rule la_generic :args (1 1 1 (- 1) (- 2) (- 1) 1 1 (- 1) 1 2 (- 1))) (step t220 (cl @p_137 @p_169 @p_149 @p_104 @p_126 @p_127 @p_139) :rule resolution :premises (t219 t18 t19 t20 t21 t22)) (step t221 (cl @p_156 @p_159 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_157 @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 (! (not @p_75) :named @p_175)) :rule la_generic :args (1 1 1 (- 2) 1 3 2 (- 1) (- 1) 1 (- 1) (- 1) 1 (- 2) (- 3) 2 (- 1))) (step t222 (cl @p_156 @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_175) :rule resolution :premises (t221 t16 t17 t18 t19 t20 t21 t22 t23)) (step t223 (cl @p_108 @p_91 @p_114 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_90) :rule resolution :premises (t74 t173 t175)) (step t224 (cl @p_91 @p_114 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_90) :rule resolution :premises (t128 t26 t223 t55)) (step t225 (cl @p_160 @p_159 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_157 @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_175) :rule la_generic :args (1 (- 1) (- 1) 2 (- 1) (- 3) (- 2) 1 1 (- 1) 1 1 (- 1) 2 3 (- 2) 1)) (step t226 (cl @p_160 @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_175) :rule resolution :premises (t225 t16 t17 t18 t19 t20 t21 t22 t23)) (step t227 (cl @p_96 @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_175 @p_114) :rule resolution :premises (t226 t224 t222)) (step t228 (cl @p_140 @p_137 @p_40 @p_145 @p_148 @p_130 @p_149 @p_127 @p_175) :rule la_generic :args (1 1 1 1 (- 1) (- 1) 1 1 (- 1))) (step t229 (cl @p_140 @p_137 @p_40 @p_149 @p_127 @p_175) :rule resolution :premises (t228 t17 t19 t20)) (step t230 (cl (not @p_176) @p_126 @p_172 (! (not @p_154) :named @p_177)) :rule eq_congruent_pred) (step t231 (cl @p_126 @p_172 @p_177) :rule th_resolution :premises (t230 t152)) (step t232 (cl (not @p_172) @p_136) :rule la_generic :args (1 2)) (step t233 (cl @p_136 @p_126) :rule resolution :premises (t232 t231 t120 t44)) (step t234 (cl (! (<= x4$ @p_35) :named @p_178) @p_40 @p_148 @p_130 @p_149 @p_127) :rule la_generic :args (1 2 (- 2) (- 2) 2 2)) (step t235 (cl @p_178 @p_40 @p_149 @p_127) :rule resolution :premises (t234 t19 t20)) (step t236 (cl @p_142 (not (! (= @p_35 @p_35) :named @p_179)) (! (not @p_178) :named @p_180) @p_146) :rule eq_congruent_pred) (step t237 (cl @p_179) :rule eq_reflexive) (step t238 (cl @p_142 @p_180 @p_146) :rule th_resolution :premises (t236 t237)) (step t239 (cl @p_180 @p_77) :rule resolution :premises (t238 t209 t34 t35)) (step t240 (cl @p_137 @p_112 @p_180 @p_90 @p_159 @p_145 @p_171 @p_175) :rule la_generic :args (4 2 1 2 2 2 (- 2) (- 2))) (step t241 (cl @p_137 @p_112 @p_180 @p_90 @p_171 @p_175) :rule resolution :premises (t240 t16 t17)) (step t242 (cl @p_36 @p_128 @p_30 @p_145 @p_141 @p_148 @p_149 @p_139 @p_143) :rule la_generic :args (3 2 1 (- 1) (- 1) 2 (- 2) 1 1)) (step t243 (cl @p_36 @p_128 @p_30 @p_149 @p_139 @p_143) :rule resolution :premises (t242 t17 t18 t19)) (step t244 (cl @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_139 @p_90 @p_117 @p_142 @p_125) :rule resolution :premises (t190 t188 t50 t187 t49 t52 t67 t198 t196)) (step t245 (cl @p_146 @p_139) :rule resolution :premises (t213 t111)) (step t246 (cl @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_139 @p_90 @p_125 @p_117) :rule resolution :premises (t190 t188 t50 t187 t49 t52 t67 t165 t183 t34 t244)) (step t247 (cl @p_171 @p_126 @p_103 @p_127 @p_138 @p_143 @p_139 @p_125 @p_117) :rule resolution :premises (t59 t49 t57 t50 t194 t192 t246)) (step t248 (cl @p_128 (not (! (= x6$ @p_39) :named @p_181))) :rule la_generic :args (2 1)) (step t249 (cl @p_127 (! (not @p_81) :named @p_182) @p_181) :rule eq_transitive) (step t250 (cl @p_128 @p_127 @p_182) :rule th_resolution :premises (t248 t249)) (step t251 (cl @p_128 @p_127) :rule resolution :premises (t250 t41)) (step t252 (cl @p_87 @p_135) :rule resolution :premises (t57 t208 t59 t49 t50)) (step t253 (cl @p_171 @p_149 @p_104 @p_126 @p_127 @p_139 @p_143 @p_114 @p_90) :rule resolution :premises (t32 t177 t227 t49 t50 t179 t224)) (step t254 (cl @p_90 @p_171 @p_143 @p_40 @p_36 @p_149 @p_127 @p_104 @p_126 @p_139) :rule resolution :premises (t34 t165 t144 t52 t253)) (step t255 (cl @p_171 @p_127 @p_143 @p_40 @p_126 @p_101) :rule resolution :premises (t59 t49 t50 t57 t169 t171 t254 t239 t235 t37 t93 t47)) (step t256 (cl @p_160 @p_171 @p_97 @p_126 @p_103 @p_127 @p_143) :rule resolution :premises (t167 t37 t38 t192)) (step t257 (cl @p_42 @p_91 @p_171 @p_149 @p_126 @p_103 @p_127 @p_139 @p_108 @p_114 @p_90 @p_36 @p_165) :rule resolution :premises (t159 t74 t90 t161 t132)) (step t258 (cl @p_90 @p_171 @p_143 @p_40 @p_149 @p_127 @p_42 @p_126 @p_103 @p_139 @p_97) :rule resolution :premises (t34 t165 t144 t52 t257 t128 t26 t55 t163 t149)) (step t259 (cl @p_171 @p_126 @p_103 @p_127 @p_143 @p_139 @p_90 @p_40 @p_42 @p_97) :rule resolution :premises (t247 t38 t37 t258 t245)) (step t260 (cl @p_171 @p_126 @p_143 @p_127 @p_40 @p_38) :rule resolution :premises (t74 t128 t202 t200 t26 t53 t55 t204 t190 t38 t37 t209 t235 t34 t239 t35 t259 t256 t252 t206 t46 t255)) (step t261 (cl @p_126 @p_127 @p_137 @p_40 @p_77) :rule resolution :premises (t93 t37 t214 t235 t34 t239 t35)) (step t262 (cl @p_137 @p_42 @p_126 @p_127 @p_139) :rule resolution :premises (t34 t216 t214)) (step t263 (cl @p_171 @p_126 @p_127 @p_40 @p_38) :rule resolution :premises (t52 t241 t227 t226 t49 t220 t47 t262 t235 t37 t38 t101 t103 t245 t261 t32 t31 t260)) (step t264 (cl @p_140 @p_137 @p_40 @p_149 @p_127) :rule resolution :premises (t229 t32)) (step t265 (cl @p_140 @p_40 @p_149 @p_127) :rule resolution :premises (t137 t31 t264)) (step t266 (cl @p_140 @p_42 @p_126 @p_127 @p_139) :rule resolution :premises (t34 t109 t107 t31 t262 t245)) (step t267 (cl @p_126 @p_127 @p_38) :rule resolution :premises (t31 t105 t214 t34 t35 t266 t93 t37 t265 t28 t263 t233)) (step t268 (cl @p_140 @p_169 @p_177 @p_137 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_149 @p_142 @p_103 @p_127 @p_175) :rule la_generic :args (2 1 1 1 2 (- 1) (- 3) (- 2) 1 1 3 1 (- 1) 2 (- 2))) (step t269 (cl @p_140 @p_169 @p_177 @p_137 @p_149 @p_142 @p_103 @p_127 @p_175) :rule resolution :premises (t268 t17 t18 t19 t20 t21 t22)) (step t270 (cl @p_140 @p_137 @p_90 @p_168 @p_112 @p_159) :rule la_generic :args (1 1 1 1 1 1)) (step t271 (cl @p_140 @p_137 @p_90 @p_112) :rule resolution :premises (t270 t16 t86)) (step t272 (cl @p_91 @p_121 @p_177 @p_151 @p_159 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_114 @p_149 @p_142 @p_103 @p_127 @p_175) :rule la_generic :args (1 1 1 1 (- 1) 2 (- 1) (- 3) (- 2) 1 1 1 3 1 (- 1) 2 (- 2))) (step t273 (cl @p_91 @p_121 @p_177 @p_151 @p_114 @p_149 @p_142 @p_103 @p_127 @p_175) :rule resolution :premises (t272 t16 t17 t18 t19 t20 t21 t22)) (step t274 (cl @p_91 @p_121 @p_177 @p_114 @p_149 @p_142 @p_103 @p_127 @p_175 @p_96 @p_150 @p_90) :rule resolution :premises (t273 t74 t128 t124 t26 t132 t55)) (step t275 (cl @p_156 @p_168 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_149 @p_103 @p_127 @p_175 @p_96) :rule la_generic :args (1 1 1 (- 1) 1 1 (- 1) (- 1) (- 1) 1 (- 1) 1 1)) (step t276 (cl @p_156 @p_149 @p_103 @p_127 @p_175 @p_96) :rule resolution :premises (t275 t16 t17 t19 t20 t22 t23 t86)) (step t277 (cl @p_91 @p_121 @p_177 @p_149 @p_142 @p_103 @p_127 @p_175 @p_96 @p_150 @p_90 @p_140 @p_137) :rule resolution :premises (t52 t274 t271)) (step t278 (cl @p_160 @p_121 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_149 @p_103 @p_127 @p_175 @p_96) :rule la_generic :args (1 1 (- 1) 1 (- 1) (- 1) 1 1 1 (- 1) 1 (- 1) (- 1))) (step t279 (cl @p_160 @p_121 @p_149 @p_103 @p_127 @p_175 @p_96) :rule resolution :premises (t278 t16 t17 t19 t20 t22 t23)) (step t280 (cl @p_121 @p_149 @p_103 @p_127 @p_175 @p_177 @p_142 @p_150 @p_140 @p_137) :rule resolution :premises (t279 t277 t276 t49 t269)) (step t281 (cl @p_140 @p_177 @p_101 @p_145 @p_141 @p_148 @p_130 @p_131 @p_149 @p_142 @p_127 @p_175) :rule la_generic :args (1 1 1 1 (- 1) (- 2) (- 1) 1 2 1 1 (- 1))) (step t282 (cl @p_140 @p_177 @p_101 @p_149 @p_142 @p_127 @p_175) :rule resolution :premises (t281 t17 t18 t19 t20 t21)) (step t283 (cl @p_140 @p_177 @p_149 @p_142 @p_127 @p_175 @p_150 @p_137) :rule resolution :premises (t282 t46 t280 t140 t29)) (step t284 (cl @p_160 @p_155 @p_159 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_157 @p_171 @p_149 @p_104 @p_142 @p_127 @p_175 @p_96) :rule la_generic :args (1 2 (- 1) (- 1) 2 3 1 (- 2) 1 1 1 (- 3) (- 1) (- 2) (- 1) 1 (- 1))) (step t285 (cl @p_160 @p_155 @p_171 @p_149 @p_104 @p_142 @p_127 @p_175 @p_96) :rule resolution :premises (t284 t16 t17 t18 t19 t20 t21 t22 t23)) (step t286 (cl @p_136 @p_101 @p_169 @p_134 @p_104) :rule la_generic :args (1 1 1 1 (- 1))) (step t287 (cl @p_136 @p_101 @p_169 @p_104) :rule resolution :premises (t286 t22)) (step t288 (cl @p_136 @p_101 @p_160 @p_155 @p_171 @p_149 @p_142 @p_127 @p_175) :rule resolution :premises (t287 t49 t285 t47)) (step t289 (cl @p_169 @p_173) :rule la_generic :args (2 1)) (step t290 (cl @p_42 @p_155 @p_120 @p_145 @p_141 @p_148 @p_130 @p_131 @p_171 @p_149 @p_142 @p_127 @p_175) :rule la_generic :args (2 2 1 (- 2) 2 4 2 (- 2) 1 (- 4) (- 2) (- 2) 2)) (step t291 (cl @p_42 @p_155 @p_120 @p_171 @p_149 @p_142 @p_127 @p_175) :rule resolution :premises (t290 t17 t18 t19 t20 t21)) (step t292 (cl @p_155 @p_171 @p_149 @p_142 @p_127 @p_175 @p_160 @p_136) :rule resolution :premises (t291 t279 t49 t289 t99 t46 t288)) (step t293 (cl @p_136 @p_112 @p_90 @p_30 @p_159 @p_145 @p_148 @p_130 @p_171 @p_149 @p_127 @p_175) :rule la_generic :args (1 1 1 2 1 (- 1) 1 1 (- 1) (- 1) (- 1) 1)) (step t294 (cl @p_136 @p_112 @p_90 @p_30 @p_171 @p_149 @p_127 @p_175) :rule resolution :premises (t293 t16 t17 t19 t20)) (step t295 (cl @p_177 @p_114 @p_149 @p_142 @p_103 @p_127 @p_175 @p_150 @p_90 @p_155 @p_171 @p_136) :rule resolution :premises (t274 t276 t49 t289 t291 t99 t212)) (step t296 (cl @p_177 @p_151 @p_145 @p_141 @p_148 @p_130 @p_131 @p_157 @p_114 @p_96 @p_149 @p_142 @p_127 @p_175) :rule la_generic :args (1 1 1 (- 1) (- 2) (- 1) 1 (- 1) 1 1 2 1 1 (- 1))) (step t297 (cl @p_177 @p_151 @p_114 @p_96 @p_149 @p_142 @p_127 @p_175) :rule resolution :premises (t296 t17 t18 t19 t20 t21 t23)) (step t298 (cl @p_155 @p_158 @p_145 @p_141 @p_148 @p_130 @p_131 @p_157 @p_114 @p_96 @p_149 @p_142 @p_127 @p_175) :rule la_generic :args (1 1 (- 1) 1 2 1 (- 1) 1 (- 1) (- 1) (- 2) (- 1) (- 1) 1)) (step t299 (cl @p_155 @p_158 @p_114 @p_96 @p_149 @p_142 @p_127 @p_175) :rule resolution :premises (t298 t17 t18 t19 t20 t21 t23)) (step t300 (cl @p_156 @p_177 @p_159 @p_145 @p_141 @p_148 @p_130 @p_131 @p_134 @p_157 @p_96 @p_171 @p_149 @p_142 @p_127 @p_175 @p_104) :rule la_generic :args (1 2 1 1 (- 2) (- 3) (- 1) 2 (- 1) (- 1) 1 (- 1) 3 2 1 (- 1) 1)) (step t301 (cl @p_156 @p_177 @p_96 @p_171 @p_149 @p_142 @p_127 @p_175 @p_104) :rule resolution :premises (t300 t16 t17 t18 t19 t20 t21 t22 t23)) (step t302 (cl @p_177 @p_142 @p_127 @p_155 @p_136 @p_150 @p_137 @p_38 @p_116) :rule resolution :premises (t301 t55 t26 t128 t74 t299 t297 t49 t287 t47 t46 t295 t52 t294 t292 t28 t283 t37 t38 t101 t32)) (step t303 (cl @p_106 @p_84) :rule resolution :premises (t72 t135 t75 t47 t46)) (step t304 (cl @p_127 @p_150 @p_38 @p_136 @p_177 @p_155 @p_34) :rule resolution :premises (t74 t116 t124 t90 t128 t132 t26 t52 t55 t134 t122 t126 t49 t146 t212 t303 t114 t37 t140 t105 t29 t142 t31 t302 t87 t34)) (step t305 (cl @p_150 @p_150 @p_153 @p_177) :rule eq_congruent_pred) (step t306 (cl @p_150 @p_153 @p_177) :rule contraction :premises (t305)) (step t307 (cl @p_127) :rule resolution :premises (t181 t38 t37 t156 t304 t306 t120 t232 t44 t43 t267 t251)) (step t308 (cl @p_38) :rule resolution :premises (t40 t307)) (step t309 (cl @p_81) :rule resolution :premises (t41 t308)) (step t310 (cl @p_125 @p_128 @p_136 @p_130 @p_182) :rule la_generic :args (1 1 1 1 (- 1))) (step t311 (cl @p_125 @p_136) :rule resolution :premises (t310 t20 t308 t309)) (step t312 (cl @p_125 @p_128 @p_137 @p_141 @p_148 @p_142 @p_138) :rule la_generic :args (2 1 1 1 1 (- 1) (- 1))) (step t313 (cl @p_125 @p_137 @p_142 @p_138) :rule resolution :premises (t312 t18 t19 t308)) (step t314 (cl @p_156 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_171 @p_103 @p_182 @p_138 @p_143 @p_96) :rule la_generic :args (1 1 1 (- 1) 1 (- 1) (- 1) (- 1) 1 (- 1) 1 (- 1) 1)) (step t315 (cl @p_156 @p_171 @p_103 @p_138 @p_143 @p_96) :rule resolution :premises (t314 t16 t17 t19 t20 t22 t23 t309)) (step t316 (cl @p_128 @p_169 @p_131 @p_134 @p_126 @p_103) :rule la_generic :args (1 1 1 1 (- 1) (- 1))) (step t317 (cl @p_169 @p_126 @p_103) :rule resolution :premises (t316 t21 t22 t308)) (step t318 (cl @p_182 @p_182 (! (not @p_183) :named @p_184) @p_129) :rule eq_congruent_pred) (step t319 (cl @p_182 @p_184 @p_129) :rule contraction :premises (t318)) (step t320 (cl @p_184 @p_129) :rule resolution :premises (t319 t309)) (step t321 (cl @p_129) :rule resolution :premises (t320 t91)) (step t322 (cl @p_125 @p_126 @p_101) :rule resolution :premises (t157 t95 t97 t47 t308 t321)) (step t323 (cl @p_109 @p_90 @p_171 @p_142 @p_126 @p_103 @p_115 @p_138 @p_143 @p_91) :rule resolution :premises (t128 t74 t26 t200 t55)) (step t324 (cl @p_160 @p_159 @p_145 @p_148 @p_130 @p_134 @p_157 @p_171 @p_103 @p_182 @p_138 @p_143 @p_96) :rule la_generic :args (1 (- 1) (- 1) 1 (- 1) 1 1 1 (- 1) 1 (- 1) 1 (- 1))) (step t325 (cl @p_160 @p_171 @p_103 @p_138 @p_143 @p_96) :rule resolution :premises (t324 t16 t17 t19 t20 t22 t23 t309)) (step t326 (cl @p_125) :rule resolution :premises (t325 t323 t202 t53 t204 t315 t28 t105 t31 t196 t313 t34 t183 t67 t49 t317 t46 t322 t43 t311 t38)) (step t327 (cl @p_78) :rule resolution :premises (t37 t326)) (step t328 (cl @p_36 @p_160 @p_121 @p_159 @p_141 @p_148 @p_130 @p_134 @p_157 @p_96 @p_171 @p_142 @p_103 @p_182 @p_149) :rule la_generic :args (2 2 1 (- 2) 2 2 (- 2) 2 2 (- 2) 1 (- 2) (- 2) 2 (- 2))) (step t329 (cl @p_160 @p_121 @p_96 @p_171 @p_142 @p_103) :rule resolution :premises (t328 t16 t18 t19 t20 t22 t23 t326 t327 t309)) (step t330 (cl @p_137 @p_128 @p_141 @p_148 @p_142 @p_149) :rule la_generic :args (1 1 1 1 (- 1) (- 1))) (step t331 (cl @p_137 @p_142) :rule resolution :premises (t330 t18 t19 t327 t308)) (step t332 (cl @p_171 @p_142 @p_143 @p_160 @p_96 @p_103) :rule resolution :premises (t151 t329 t326)) (step t333 (cl @p_160 @p_121 @p_159 @p_145 @p_141 @p_148 @p_130 @p_134 @p_157 @p_96 @p_149 @p_142 @p_103 @p_182 @p_143) :rule la_generic :args (1 1 (- 1) 1 2 1 (- 1) 1 1 (- 1) (- 1) (- 2) (- 1) 1 (- 1))) (step t334 (cl @p_160 @p_121 @p_96 @p_142 @p_103 @p_143) :rule resolution :premises (t333 t16 t17 t18 t19 t20 t22 t23 t327 t309)) (step t335 (cl @p_160 @p_142 @p_103 @p_143 @p_126) :rule resolution :premises (t334 t140 t29 t28 t332 t49 t317)) (step t336 (cl @p_128 @p_101 @p_40 @p_131 @p_126) :rule la_generic :args (1 1 1 1 (- 1))) (step t337 (cl @p_101 @p_40 @p_126) :rule resolution :premises (t336 t21 t308)) (step t338 (cl @p_151 @p_145 @p_141 @p_130 @p_131 @p_134 @p_157 @p_96 @p_142 @p_126 @p_103 @p_182 @p_143 @p_115) :rule la_generic :args (1 1 1 (- 1) 1 2 1 (- 1) (- 1) (- 1) (- 2) 1 (- 1) 1)) (step t339 (cl @p_151 @p_96 @p_142 @p_126 @p_103 @p_143 @p_115) :rule resolution :premises (t338 t17 t18 t20 t21 t22 t23 t309)) (step t340 (cl @p_151 @p_152 @p_40 @p_145 @p_141 @p_130 @p_131 @p_157 @p_96 @p_142 @p_126 @p_182 @p_143) :rule la_generic :args (1 1 2 1 1 (- 1) 1 (- 1) 1 (- 1) (- 1) 1 (- 1))) (step t341 (cl @p_151 @p_152 @p_40 @p_96 @p_142 @p_126 @p_143) :rule resolution :premises (t340 t17 t18 t20 t21 t23 t309)) (step t342 (cl @p_151 @p_40 @p_96 @p_142 @p_126 @p_143 @p_91 @p_165 @p_103) :rule resolution :premises (t341 t90 t132 t52 t53 t339)) (step t343 (cl @p_46 @p_123 @p_158 @p_145 @p_141 @p_130 @p_131 @p_134 @p_157 @p_96 @p_142 @p_126 @p_103 @p_182 @p_143) :rule la_generic :args (2 1 1 (- 1) (- 1) 1 (- 1) (- 2) (- 1) 1 1 1 2 (- 1) 1)) (step t344 (cl @p_46 @p_123 @p_158 @p_96 @p_142 @p_126 @p_103 @p_143) :rule resolution :premises (t343 t17 t18 t20 t21 t22 t23 t309)) (step t345 (cl @p_114 @p_158 @p_96 @p_142 @p_126 @p_103 @p_143 @p_91 @p_165) :rule resolution :premises (t78 t53 t344 t132)) (step t346 (cl @p_158 @p_145 @p_141 @p_130 @p_131 @p_134 @p_157 @p_96 @p_142 @p_126 @p_103 @p_115 @p_182 @p_143) :rule la_generic :args (1 (- 1) (- 1) 1 (- 1) (- 2) (- 1) 1 1 1 2 (- 1) (- 1) 1)) (step t347 (cl @p_158 @p_96 @p_142 @p_126 @p_103 @p_115 @p_143) :rule resolution :premises (t346 t17 t18 t20 t21 t22 t23 t309)) (step t348 (cl @p_96 @p_142 @p_126 @p_103 @p_143 @p_91 @p_40 @p_90) :rule resolution :premises (t347 t53 t52 t345 t74 t342 t128 t26 t55)) (step t349 (cl @p_156 @p_168 @p_159 @p_145 @p_141 @p_148 @p_130 @p_134 @p_157 @p_96 @p_149 @p_142 @p_103 @p_182 @p_143) :rule la_generic :args (1 1 1 (- 1) (- 2) (- 1) 1 (- 1) (- 1) 1 1 2 1 (- 1) 1)) (step t350 (cl @p_156 @p_96 @p_142 @p_103 @p_143) :rule resolution :premises (t349 t16 t17 t18 t19 t20 t22 t23 t327 t309 t86)) (step t351 (cl @p_126) :rule resolution :premises (t350 t348 t335 t31 t331 t34 t154 t67 t49 t317 t46 t337 t233 t327 t326)) (step t352 (cl @p_40) :rule resolution :premises (t43 t351)) (step t353 (cl @p_83) :rule resolution :premises (t44 t352)) (step t354 (cl @p_154) :rule resolution :premises (t120 t353)) (step t355 (cl @p_136 @p_112 @p_134 @p_157 @p_96 @p_103) :rule la_generic :args (1 1 1 1 (- 1) (- 1))) (step t356 (cl @p_112 @p_96 @p_103) :rule resolution :premises (t355 t22 t23 t352)) (step t357 (cl @p_177 @p_151 @p_145 @p_141 @p_130 @p_131 @p_157 @p_114 @p_96 @p_142 @p_182 @p_143) :rule la_generic :args (1 1 1 1 (- 1) 1 (- 1) 1 1 (- 1) 1 (- 1))) (step t358 (cl @p_151 @p_114 @p_96 @p_142 @p_143) :rule resolution :premises (t357 t17 t18 t20 t21 t23 t309 t354)) (step t359 (cl @p_120 @p_142 @p_143) :rule resolution :premises (t28 t29 t151 t140 t326)) (step t360 (cl @p_121 @p_142 @p_103 @p_143) :rule resolution :premises (t74 t124 t128 t132 t26 t358 t55 t52 t334 t356 t350 t49 t289 t99 t212 t327 t353 t352)) (step t361 (cl @p_128 @p_136 @p_177 @p_101 @p_131) :rule la_generic :args (1 1 1 1 1)) (step t362 (cl @p_101) :rule resolution :premises (t361 t21 t308 t352 t354)) (step t363 (cl @p_84) :rule resolution :premises (t46 t362)) (step t364 (cl @p_142) :rule resolution :premises (t360 t359 t31 t331 t363)) (step t365 (cl @p_34) :rule resolution :premises (t34 t364)) (step t366 (cl @p_77) :rule resolution :premises (t35 t365)) (step t367 (cl @p_137) :rule resolution :premises (t218 t366 t308 t327 t365)) (step t368 (cl @p_74) :rule resolution :premises (t31 t367)) (step t369 (cl @p_140) :rule resolution :premises (t107 t367 t368 t365)) (step t370 (cl) :rule resolution :premises (t243 t368 t327 t326 t366 t308 t369)) eddc1b21e40e4136745e32bdb4d4d38dda531d49 67 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr1 Int)) (not (! (ite (! (< 0 veriT_vr1) :named @p_27) (! (< 0 (! (+ 1 veriT_vr1) :named @p_29)) :named @p_30) (! (< veriT_vr1 1) :named @p_31)) :named @p_32))) :named @p_38)) (assume axiom0 (! (not (! (forall ((?v0 Int)) (! (ite (! (< 0 ?v0) :named @p_2) (! (< 0 (! (+ ?v0 1) :named @p_5)) :named @p_7) (! (< ?v0 1) :named @p_9)) :named @p_11)) :named @p_1)) :named @p_13)) (anchor :step t2 :args ((:= (?v0 Int) veriT_vr0))) (step t2.t1 (cl (! (= ?v0 veriT_vr0) :named @p_4)) :rule refl) (step t2.t2 (cl (= @p_2 (! (< 0 veriT_vr0) :named @p_3))) :rule cong :premises (t2.t1)) (step t2.t3 (cl @p_4) :rule refl) (step t2.t4 (cl (= @p_5 (! (+ veriT_vr0 1) :named @p_6))) :rule cong :premises (t2.t3)) (step t2.t5 (cl (= @p_7 (! (< 0 @p_6) :named @p_8))) :rule cong :premises (t2.t4)) (step t2.t6 (cl @p_4) :rule refl) (step t2.t7 (cl (= @p_9 (! (< veriT_vr0 1) :named @p_10))) :rule cong :premises (t2.t6)) (step t2.t8 (cl (= @p_11 (! (ite @p_3 @p_8 @p_10) :named @p_12))) :rule cong :premises (t2.t2 t2.t5 t2.t7)) (step t2 (cl (= @p_1 (! (forall ((veriT_vr0 Int)) @p_12) :named @p_14))) :rule bind) (step t3 (cl (! (= @p_13 (! (not @p_14) :named @p_16)) :named @p_15)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_15) :named @p_18) (! (not @p_13) :named @p_17) @p_16) :rule equiv_pos2) (step t5 (cl (not @p_17) @p_1) :rule not_not) (step t6 (cl @p_18 @p_1 @p_16) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_16) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((veriT_vr0 Int))) (step t8.t1 (cl (= @p_6 (! (+ 1 veriT_vr0) :named @p_19))) :rule sum_simplify) (step t8.t2 (cl (= @p_8 (! (< 0 @p_19) :named @p_20))) :rule cong :premises (t8.t1)) (step t8.t3 (cl (= @p_12 (! (ite @p_3 @p_20 @p_10) :named @p_21))) :rule cong :premises (t8.t2)) (step t8 (cl (= @p_14 (! (forall ((veriT_vr0 Int)) @p_21) :named @p_22))) :rule bind) (step t9 (cl (! (= @p_16 (! (not @p_22) :named @p_24)) :named @p_23)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_23) :named @p_26) (! (not @p_16) :named @p_25) @p_24) :rule equiv_pos2) (step t11 (cl (not @p_25) @p_14) :rule not_not) (step t12 (cl @p_26 @p_14 @p_24) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_24) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t14.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_28)) :rule refl) (step t14.t2 (cl (= @p_3 @p_27)) :rule cong :premises (t14.t1)) (step t14.t3 (cl @p_28) :rule refl) (step t14.t4 (cl (= @p_19 @p_29)) :rule cong :premises (t14.t3)) (step t14.t5 (cl (= @p_20 @p_30)) :rule cong :premises (t14.t4)) (step t14.t6 (cl @p_28) :rule refl) (step t14.t7 (cl (= @p_10 @p_31)) :rule cong :premises (t14.t6)) (step t14.t8 (cl (= @p_21 @p_32)) :rule cong :premises (t14.t2 t14.t5 t14.t7)) (step t14 (cl (= @p_22 (! (forall ((veriT_vr1 Int)) @p_32) :named @p_33))) :rule bind) (step t15 (cl (! (= @p_24 (! (not @p_33) :named @p_35)) :named @p_34)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_34) :named @p_37) (! (not @p_24) :named @p_36) @p_35) :rule equiv_pos2) (step t17 (cl (not @p_36) @p_22) :rule not_not) (step t18 (cl @p_37 @p_22 @p_35) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_35) :rule th_resolution :premises (t13 t15 t18)) (anchor :step t20 :args ((:= (veriT_vr1 Int) veriT_sk0))) (step t20.t1 (cl (! (= veriT_vr1 veriT_sk0) :named @p_40)) :rule refl) (step t20.t2 (cl (= @p_27 (! (< 0 veriT_sk0) :named @p_39))) :rule cong :premises (t20.t1)) (step t20.t3 (cl @p_40) :rule refl) (step t20.t4 (cl (= @p_29 (! (+ 1 veriT_sk0) :named @p_41))) :rule cong :premises (t20.t3)) (step t20.t5 (cl (= @p_30 (! (< 0 @p_41) :named @p_42))) :rule cong :premises (t20.t4)) (step t20.t6 (cl @p_40) :rule refl) (step t20.t7 (cl (= @p_31 (! (< veriT_sk0 1) :named @p_43))) :rule cong :premises (t20.t6)) (step t20.t8 (cl (= @p_32 (! (ite @p_39 @p_42 @p_43) :named @p_44))) :rule cong :premises (t20.t2 t20.t5 t20.t7)) (step t20 (cl (= @p_33 @p_44)) :rule sko_forall) (step t21 (cl (! (= @p_35 (! (not @p_44) :named @p_46)) :named @p_45)) :rule cong :premises (t20)) (step t22 (cl (! (not @p_45) :named @p_48) (! (not @p_35) :named @p_47) @p_46) :rule equiv_pos2) (step t23 (cl (not @p_47) @p_33) :rule not_not) (step t24 (cl @p_48 @p_33 @p_46) :rule th_resolution :premises (t23 t22)) (step t25 (cl @p_46) :rule th_resolution :premises (t19 t21 t24)) (step t26 (cl @p_39 (! (not @p_43) :named @p_50)) :rule not_ite1 :premises (t25)) (step t27 (cl (! (not @p_39) :named @p_49) (not @p_42)) :rule not_ite2 :premises (t25)) (step t28 (cl @p_43 @p_42) :rule la_generic :args (1 1)) (step t29 (cl @p_49 @p_42) :rule la_generic :args (1 1)) (step t30 (cl @p_42) :rule resolution :premises (t29 t26 t28)) (step t31 (cl @p_49) :rule resolution :premises (t27 t30)) (step t32 (cl @p_50) :rule resolution :premises (t26 t31)) (step t33 (cl @p_43 @p_39) :rule la_generic :args (1 1)) (step t34 (cl) :rule resolution :premises (t33 t31 t32)) 3b346287f457d2c0655eb09f3d5280a06ca3434a 107 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr1 Int)) (not (! (or (! (< veriT_vr1 0) :named @p_37) (! (< 0 veriT_vr1) :named @p_39)) :named @p_40))) :named @p_51)) (assume axiom0 (! (not (! (< 0 (! (ite (! (forall ((?v0 Int)) (! (or (! (< ?v0 0) :named @p_2) (! (< 0 ?v0) :named @p_5)) :named @p_7)) :named @p_1) (! (- 1) :named @p_11) 3) :named @p_9)) :named @p_12)) :named @p_14)) (anchor :step t2 :args ((:= (?v0 Int) veriT_vr0))) (step t2.t1 (cl (! (= ?v0 veriT_vr0) :named @p_4)) :rule refl) (step t2.t2 (cl (= @p_2 (! (< veriT_vr0 0) :named @p_3))) :rule cong :premises (t2.t1)) (step t2.t3 (cl @p_4) :rule refl) (step t2.t4 (cl (= @p_5 (! (< 0 veriT_vr0) :named @p_6))) :rule cong :premises (t2.t3)) (step t2.t5 (cl (= @p_7 (! (or @p_3 @p_6) :named @p_8))) :rule cong :premises (t2.t2 t2.t4)) (step t2 (cl (= @p_1 (! (forall ((veriT_vr0 Int)) @p_8) :named @p_10))) :rule bind) (step t3 (cl (= @p_9 (! (ite @p_10 @p_11 3) :named @p_13))) :rule cong :premises (t2)) (step t4 (cl (= @p_12 (! (< 0 @p_13) :named @p_15))) :rule cong :premises (t3)) (step t5 (cl (! (= @p_14 (! (not @p_15) :named @p_17)) :named @p_16)) :rule cong :premises (t4)) (step t6 (cl (! (not @p_16) :named @p_19) (! (not @p_14) :named @p_18) @p_17) :rule equiv_pos2) (step t7 (cl (not @p_18) @p_12) :rule not_not) (step t8 (cl @p_19 @p_12 @p_17) :rule th_resolution :premises (t7 t6)) (step t9 (cl @p_17) :rule th_resolution :premises (axiom0 t5 t8)) (step t10 (cl (= @p_11 (- 1))) :rule minus_simplify) (step t11 (cl (= @p_13 (! (ite @p_10 (- 1) 3) :named @p_20))) :rule cong :premises (t10)) (step t12 (cl (= @p_15 (! (< 0 @p_20) :named @p_21))) :rule cong :premises (t11)) (step t13 (cl (! (= @p_17 (! (not @p_21) :named @p_23)) :named @p_22)) :rule cong :premises (t12)) (step t14 (cl (! (not @p_22) :named @p_25) (! (not @p_17) :named @p_24) @p_23) :rule equiv_pos2) (step t15 (cl (not @p_24) @p_15) :rule not_not) (step t16 (cl @p_25 @p_15 @p_23) :rule th_resolution :premises (t15 t14)) (step t17 (cl @p_23) :rule th_resolution :premises (t9 t13 t16)) (step t18 (cl (! (= @p_23 (! (and (! (not (! (< 0 @p_20) :named @p_73)) :named @p_33) (! (ite @p_10 (! (= (- 1) @p_20) :named @p_31) (! (= 3 @p_20) :named @p_32)) :named @p_30)) :named @p_27)) :named @p_26)) :rule ite_intro) (step t19 (cl (! (not @p_26) :named @p_29) (! (not @p_23) :named @p_28) @p_27) :rule equiv_pos2) (step t20 (cl (not @p_28) @p_21) :rule not_not) (step t21 (cl @p_29 @p_21 @p_27) :rule th_resolution :premises (t20 t19)) (step t22 (cl @p_27) :rule th_resolution :premises (t17 t18 t21)) (step t23 (cl (= @p_30 (! (and (! (=> @p_10 @p_31) :named @p_41) (! (=> (! (not @p_10) :named @p_43) @p_32) :named @p_44)) :named @p_34))) :rule connective_def) (step t24 (cl (! (= @p_27 (! (and @p_33 @p_34) :named @p_36)) :named @p_35)) :rule cong :premises (t23)) (step t25 (cl (not @p_35) (not @p_27) @p_36) :rule equiv_pos2) (step t26 (cl @p_36) :rule th_resolution :premises (t22 t24 t25)) (anchor :step t27 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t27.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_38)) :rule refl) (step t27.t2 (cl (= @p_3 @p_37)) :rule cong :premises (t27.t1)) (step t27.t3 (cl @p_38) :rule refl) (step t27.t4 (cl (= @p_6 @p_39)) :rule cong :premises (t27.t3)) (step t27.t5 (cl (= @p_8 @p_40)) :rule cong :premises (t27.t2 t27.t4)) (step t27 (cl (= @p_10 (! (forall ((veriT_vr1 Int)) @p_40) :named @p_42))) :rule bind) (step t28 (cl (= @p_41 (! (=> @p_42 @p_31) :named @p_46))) :rule cong :premises (t27)) (step t29 (cl (= @p_43 (! (not @p_42) :named @p_45))) :rule cong :premises (t27)) (step t30 (cl (= @p_44 (! (=> @p_45 @p_32) :named @p_47))) :rule cong :premises (t29)) (step t31 (cl (= @p_34 (! (and @p_46 @p_47) :named @p_48))) :rule cong :premises (t28 t30)) (step t32 (cl (! (= @p_36 (! (and @p_33 @p_48) :named @p_50)) :named @p_49)) :rule cong :premises (t31)) (step t33 (cl (not @p_49) (not @p_36) @p_50) :rule equiv_pos2) (step t34 (cl @p_50) :rule th_resolution :premises (t26 t32 t33)) (anchor :step t35 :args ((:= (veriT_vr1 Int) veriT_sk0))) (step t35.t1 (cl (! (= veriT_vr1 veriT_sk0) :named @p_53)) :rule refl) (step t35.t2 (cl (= @p_37 (! (< veriT_sk0 0) :named @p_52))) :rule cong :premises (t35.t1)) (step t35.t3 (cl @p_53) :rule refl) (step t35.t4 (cl (= @p_39 (! (< 0 veriT_sk0) :named @p_54))) :rule cong :premises (t35.t3)) (step t35.t5 (cl (= @p_40 (! (or @p_52 @p_54) :named @p_55))) :rule cong :premises (t35.t2 t35.t4)) (step t35 (cl (= @p_42 @p_55)) :rule sko_forall) (step t36 (cl (= @p_46 (! (=> @p_55 @p_31) :named @p_56))) :rule cong :premises (t35)) (step t37 (cl (= @p_48 (! (and @p_56 @p_47) :named @p_57))) :rule cong :premises (t36)) (step t38 (cl (! (= @p_50 (! (and @p_33 @p_57) :named @p_59)) :named @p_58)) :rule cong :premises (t37)) (step t39 (cl (not @p_58) (not @p_50) @p_59) :rule equiv_pos2) (step t40 (cl @p_59) :rule th_resolution :premises (t34 t38 t39)) (anchor :step t41 :args ((:= (veriT_vr1 Int) veriT_vr2))) (step t41.t1 (cl (! (= veriT_vr1 veriT_vr2) :named @p_61)) :rule refl) (step t41.t2 (cl (= @p_37 (! (< veriT_vr2 0) :named @p_60))) :rule cong :premises (t41.t1)) (step t41.t3 (cl @p_61) :rule refl) (step t41.t4 (cl (= @p_39 (! (< 0 veriT_vr2) :named @p_62))) :rule cong :premises (t41.t3)) (step t41.t5 (cl (= @p_40 (! (or @p_60 @p_62) :named @p_63))) :rule cong :premises (t41.t2 t41.t4)) (step t41 (cl (= @p_42 (! (forall ((veriT_vr2 Int)) @p_63) :named @p_64))) :rule bind) (step t42 (cl (= @p_45 (! (not @p_64) :named @p_65))) :rule cong :premises (t41)) (step t43 (cl (= @p_47 (! (=> @p_65 @p_32) :named @p_66))) :rule cong :premises (t42)) (step t44 (cl (= @p_57 (! (and @p_56 @p_66) :named @p_67))) :rule cong :premises (t43)) (step t45 (cl (! (= @p_59 (! (and @p_33 @p_67) :named @p_69)) :named @p_68)) :rule cong :premises (t44)) (step t46 (cl (not @p_68) (not @p_59) @p_69) :rule equiv_pos2) (step t47 (cl @p_69) :rule th_resolution :premises (t40 t45 t46)) (step t48 (cl (! (= @p_69 (! (and @p_33 @p_56 @p_66) :named @p_71)) :named @p_70)) :rule ac_simp) (step t49 (cl (not @p_70) (not @p_69) @p_71) :rule equiv_pos2) (step t50 (cl @p_71) :rule th_resolution :premises (t47 t48 t49)) (step t51 (cl @p_33) :rule and :premises (t50)) (step t52 (cl @p_66) :rule and :premises (t50)) (step t53 (cl (! (not @p_65) :named @p_72) @p_32) :rule implies :premises (t52)) (step t54 (cl (not @p_72) @p_64) :rule not_not) (step t55 (cl @p_64 @p_32) :rule th_resolution :premises (t54 t53)) (step t56 (cl @p_73 (! (<= @p_20 3) :named @p_74)) :rule la_generic :args (1 1)) (step t57 (cl @p_74) :rule resolution :premises (t56 t51)) (step t58 (cl @p_73 (! (not (! (<= 3 @p_20) :named @p_77)) :named @p_75)) :rule la_generic :args (1 1)) (step t59 (cl @p_75) :rule resolution :premises (t58 t51)) (step t60 (cl (! (not @p_32) :named @p_76) @p_76 (! (not @p_74) :named @p_78) @p_77) :rule eq_congruent_pred) (step t61 (cl @p_76 @p_78 @p_77) :rule contraction :premises (t60)) (step t62 (cl @p_76) :rule resolution :premises (t61 t59 t57)) (step t63 (cl @p_64) :rule resolution :premises (t55 t62)) (step t64 (cl (or @p_65 (! (or (! (< 0 0) :named @p_79) @p_79) :named @p_80))) :rule forall_inst :args ((:= veriT_vr2 0))) (anchor :step t65) (assume t65.h1 @p_80) (step t65.t2 (cl (! (= @p_80 @p_79) :named @p_81)) :rule ac_simp) (step t65.t3 (cl (not @p_81) (! (not @p_80) :named @p_82) @p_79) :rule equiv_pos2) (step t65.t4 (cl @p_79) :rule th_resolution :premises (t65.h1 t65.t2 t65.t3)) (step t65.t5 (cl (! (= @p_79 false) :named @p_83)) :rule comp_simplify) (step t65.t6 (cl (not @p_83) (not @p_79) false) :rule equiv_pos2) (step t65.t7 (cl false) :rule th_resolution :premises (t65.t4 t65.t5 t65.t6)) (step t65 (cl @p_82 false) :rule subproof :discharge (h1)) (step t66 (cl @p_65 @p_80) :rule or :premises (t64)) (step t67 (cl (! (or @p_65 false) :named @p_84) @p_72) :rule or_neg) (step t68 (cl @p_84 @p_64) :rule th_resolution :premises (t54 t67)) (step t69 (cl @p_84 (! (not false) :named @p_85)) :rule or_neg) (step t70 (cl @p_84) :rule th_resolution :premises (t66 t65 t68 t69)) (step t71 (cl @p_85) :rule false) (step t72 (cl @p_65 false) :rule or :premises (t70)) (step t73 (cl) :rule resolution :premises (t72 t63 t71)) f2b4e98e9e25f5e157253b295acc06d5501334ba 74 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr2 Int)) (not (forall ((veriT_vr3 Int)) (! (=> (! (and (! (< 0 veriT_vr2) :named @p_27) (! (< 0 veriT_vr3) :named @p_28)) :named @p_29) (! (< 0 (! (+ veriT_vr2 veriT_vr3) :named @p_32)) :named @p_33)) :named @p_34)))) :named @p_40)) (define-fun veriT_sk1 () Int (! (choice ((veriT_vr3 Int)) (not (=> (and (< 0 @p_40) @p_28) (< 0 (+ @p_40 veriT_vr3))))) :named @p_41)) (assume axiom0 (! (not (! (exists ((?v0 Int)) (! (forall ((?v1 Int) (?v2 Int)) (! (=> (! (and (! (< 0 ?v1) :named @p_8) (! (< 0 ?v2) :named @p_10)) :named @p_12) (! (< 0 (! (+ ?v1 ?v2) :named @p_16)) :named @p_18)) :named @p_20)) :named @p_2)) :named @p_1)) :named @p_3)) (step t2 (cl (= @p_1 @p_2)) :rule qnt_rm_unused) (step t3 (cl (! (= @p_3 (! (not @p_2) :named @p_5)) :named @p_4)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_4) :named @p_7) (! (not @p_3) :named @p_6) @p_5) :rule equiv_pos2) (step t5 (cl (not @p_6) @p_1) :rule not_not) (step t6 (cl @p_7 @p_1 @p_5) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_5) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((:= (?v1 Int) veriT_vr0) (:= (?v2 Int) veriT_vr1))) (step t8.t1 (cl (! (= ?v1 veriT_vr0) :named @p_14)) :rule refl) (step t8.t2 (cl (= @p_8 (! (< 0 veriT_vr0) :named @p_9))) :rule cong :premises (t8.t1)) (step t8.t3 (cl (! (= ?v2 veriT_vr1) :named @p_15)) :rule refl) (step t8.t4 (cl (= @p_10 (! (< 0 veriT_vr1) :named @p_11))) :rule cong :premises (t8.t3)) (step t8.t5 (cl (= @p_12 (! (and @p_9 @p_11) :named @p_13))) :rule cong :premises (t8.t2 t8.t4)) (step t8.t6 (cl @p_14) :rule refl) (step t8.t7 (cl @p_15) :rule refl) (step t8.t8 (cl (= @p_16 (! (+ veriT_vr0 veriT_vr1) :named @p_17))) :rule cong :premises (t8.t6 t8.t7)) (step t8.t9 (cl (= @p_18 (! (< 0 @p_17) :named @p_19))) :rule cong :premises (t8.t8)) (step t8.t10 (cl (= @p_20 (! (=> @p_13 @p_19) :named @p_21))) :rule cong :premises (t8.t5 t8.t9)) (step t8 (cl (= @p_2 (! (forall ((veriT_vr0 Int) (veriT_vr1 Int)) @p_21) :named @p_22))) :rule bind) (step t9 (cl (! (= @p_5 (! (not @p_22) :named @p_24)) :named @p_23)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_23) :named @p_26) (! (not @p_5) :named @p_25) @p_24) :rule equiv_pos2) (step t11 (cl (not @p_25) @p_2) :rule not_not) (step t12 (cl @p_26 @p_2 @p_24) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_24) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((:= (veriT_vr0 Int) veriT_vr2) (:= (veriT_vr1 Int) veriT_vr3))) (step t14.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_30)) :rule refl) (step t14.t2 (cl (= @p_9 @p_27)) :rule cong :premises (t14.t1)) (step t14.t3 (cl (! (= veriT_vr1 veriT_vr3) :named @p_31)) :rule refl) (step t14.t4 (cl (= @p_11 @p_28)) :rule cong :premises (t14.t3)) (step t14.t5 (cl (= @p_13 @p_29)) :rule cong :premises (t14.t2 t14.t4)) (step t14.t6 (cl @p_30) :rule refl) (step t14.t7 (cl @p_31) :rule refl) (step t14.t8 (cl (= @p_17 @p_32)) :rule cong :premises (t14.t6 t14.t7)) (step t14.t9 (cl (= @p_19 @p_33)) :rule cong :premises (t14.t8)) (step t14.t10 (cl (= @p_21 @p_34)) :rule cong :premises (t14.t5 t14.t9)) (step t14 (cl (= @p_22 (! (forall ((veriT_vr2 Int) (veriT_vr3 Int)) @p_34) :named @p_35))) :rule bind) (step t15 (cl (! (= @p_24 (! (not @p_35) :named @p_37)) :named @p_36)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_36) :named @p_39) (! (not @p_24) :named @p_38) @p_37) :rule equiv_pos2) (step t17 (cl (not @p_38) @p_22) :rule not_not) (step t18 (cl @p_39 @p_22 @p_37) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_37) :rule th_resolution :premises (t13 t15 t18)) (anchor :step t20 :args ((:= (veriT_vr2 Int) veriT_sk0) (:= (veriT_vr3 Int) veriT_sk1))) (step t20.t1 (cl (! (= veriT_vr2 veriT_sk0) :named @p_45)) :rule refl) (step t20.t2 (cl (= @p_27 (! (< 0 veriT_sk0) :named @p_42))) :rule cong :premises (t20.t1)) (step t20.t3 (cl (! (= veriT_vr3 veriT_sk1) :named @p_46)) :rule refl) (step t20.t4 (cl (= @p_28 (! (< 0 veriT_sk1) :named @p_43))) :rule cong :premises (t20.t3)) (step t20.t5 (cl (= @p_29 (! (and @p_42 @p_43) :named @p_44))) :rule cong :premises (t20.t2 t20.t4)) (step t20.t6 (cl @p_45) :rule refl) (step t20.t7 (cl @p_46) :rule refl) (step t20.t8 (cl (= @p_32 (! (+ veriT_sk0 veriT_sk1) :named @p_47))) :rule cong :premises (t20.t6 t20.t7)) (step t20.t9 (cl (= @p_33 (! (< 0 @p_47) :named @p_48))) :rule cong :premises (t20.t8)) (step t20.t10 (cl (= @p_34 (! (=> @p_44 @p_48) :named @p_49))) :rule cong :premises (t20.t5 t20.t9)) (step t20 (cl (= @p_35 @p_49)) :rule sko_forall) (step t21 (cl (! (= @p_37 (! (not @p_49) :named @p_51)) :named @p_50)) :rule cong :premises (t20)) (step t22 (cl (! (not @p_50) :named @p_53) (! (not @p_37) :named @p_52) @p_51) :rule equiv_pos2) (step t23 (cl (not @p_52) @p_35) :rule not_not) (step t24 (cl @p_53 @p_35 @p_51) :rule th_resolution :premises (t23 t22)) (step t25 (cl @p_51) :rule th_resolution :premises (t19 t21 t24)) (step t26 (cl (! (= @p_51 (! (and @p_44 (! (not @p_48) :named @p_58)) :named @p_55)) :named @p_54)) :rule bool_simplify) (step t27 (cl (! (not @p_54) :named @p_57) (! (not @p_51) :named @p_56) @p_55) :rule equiv_pos2) (step t28 (cl (not @p_56) @p_49) :rule not_not) (step t29 (cl @p_57 @p_49 @p_55) :rule th_resolution :premises (t28 t27)) (step t30 (cl @p_55) :rule th_resolution :premises (t25 t26 t29)) (step t31 (cl (! (= @p_55 (! (and @p_42 @p_43 @p_58) :named @p_60)) :named @p_59)) :rule ac_simp) (step t32 (cl (not @p_59) (not @p_55) @p_60) :rule equiv_pos2) (step t33 (cl @p_60) :rule th_resolution :premises (t30 t31 t32)) (step t34 (cl @p_42) :rule and :premises (t33)) (step t35 (cl @p_43) :rule and :premises (t33)) (step t36 (cl @p_58) :rule and :premises (t33)) (step t37 (cl (not @p_43) @p_48 (not @p_42)) :rule la_generic :args (1 1 1)) (step t38 (cl) :rule resolution :premises (t37 t34 t35 t36)) 373457faab0c37c0d0d8b85464bbfee7e80cb98a 77 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr2 Int)) (not (forall ((veriT_vr3 Real)) (! (=> (! (and (! (< 0 veriT_vr2) :named @p_32) (! (< 0.0 veriT_vr3) :named @p_33)) :named @p_34) (! (< (- 1) veriT_vr2) :named @p_36)) :named @p_37)))) :named @p_43)) (define-fun veriT_sk1 () Real (! (choice ((veriT_vr3 Real)) (not (=> (and (< 0 @p_43) @p_33) (< (- 1) @p_43)))) :named @p_45)) (assume axiom0 (! (not (! (exists ((?v0 Int)) (! (forall ((?v1 Int) (?v2 Real)) (! (=> (! (and (! (< 0 ?v1) :named @p_9) (! (< 0.0 ?v2) :named @p_11)) :named @p_13) (! (< (! (- 1) :named @p_8) ?v1) :named @p_16)) :named @p_18)) :named @p_2)) :named @p_1)) :named @p_3)) (step t2 (cl (= @p_1 @p_2)) :rule qnt_rm_unused) (step t3 (cl (! (= @p_3 (! (not @p_2) :named @p_5)) :named @p_4)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_4) :named @p_7) (! (not @p_3) :named @p_6) @p_5) :rule equiv_pos2) (step t5 (cl (not @p_6) @p_1) :rule not_not) (step t6 (cl @p_7 @p_1 @p_5) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_5) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((:= (?v1 Int) veriT_vr0) (:= (?v2 Real) veriT_vr1))) (step t8.t1 (cl (! (= ?v1 veriT_vr0) :named @p_15)) :rule refl) (step t8.t2 (cl (= @p_9 (! (< 0 veriT_vr0) :named @p_10))) :rule cong :premises (t8.t1)) (step t8.t3 (cl (= ?v2 veriT_vr1)) :rule refl) (step t8.t4 (cl (= @p_11 (! (< 0.0 veriT_vr1) :named @p_12))) :rule cong :premises (t8.t3)) (step t8.t5 (cl (= @p_13 (! (and @p_10 @p_12) :named @p_14))) :rule cong :premises (t8.t2 t8.t4)) (step t8.t6 (cl @p_15) :rule refl) (step t8.t7 (cl (= @p_16 (! (< @p_8 veriT_vr0) :named @p_17))) :rule cong :premises (t8.t6)) (step t8.t8 (cl (= @p_18 (! (=> @p_14 @p_17) :named @p_19))) :rule cong :premises (t8.t5 t8.t7)) (step t8 (cl (= @p_2 (! (forall ((veriT_vr0 Int) (veriT_vr1 Real)) @p_19) :named @p_20))) :rule bind) (step t9 (cl (! (= @p_5 (! (not @p_20) :named @p_22)) :named @p_21)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_21) :named @p_24) (! (not @p_5) :named @p_23) @p_22) :rule equiv_pos2) (step t11 (cl (not @p_23) @p_2) :rule not_not) (step t12 (cl @p_24 @p_2 @p_22) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_22) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((veriT_vr0 Int) (veriT_vr1 Real))) (step t14.t1 (cl (= @p_8 (- 1))) :rule minus_simplify) (step t14.t2 (cl (= @p_17 (! (< (- 1) veriT_vr0) :named @p_25))) :rule cong :premises (t14.t1)) (step t14.t3 (cl (= @p_19 (! (=> @p_14 @p_25) :named @p_26))) :rule cong :premises (t14.t2)) (step t14 (cl (= @p_20 (! (forall ((veriT_vr0 Int) (veriT_vr1 Real)) @p_26) :named @p_27))) :rule bind) (step t15 (cl (! (= @p_22 (! (not @p_27) :named @p_29)) :named @p_28)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_28) :named @p_31) (! (not @p_22) :named @p_30) @p_29) :rule equiv_pos2) (step t17 (cl (not @p_30) @p_20) :rule not_not) (step t18 (cl @p_31 @p_20 @p_29) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_29) :rule th_resolution :premises (t13 t15 t18)) (anchor :step t20 :args ((:= (veriT_vr0 Int) veriT_vr2) (:= (veriT_vr1 Real) veriT_vr3))) (step t20.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_35)) :rule refl) (step t20.t2 (cl (= @p_10 @p_32)) :rule cong :premises (t20.t1)) (step t20.t3 (cl (= veriT_vr1 veriT_vr3)) :rule refl) (step t20.t4 (cl (= @p_12 @p_33)) :rule cong :premises (t20.t3)) (step t20.t5 (cl (= @p_14 @p_34)) :rule cong :premises (t20.t2 t20.t4)) (step t20.t6 (cl @p_35) :rule refl) (step t20.t7 (cl (= @p_25 @p_36)) :rule cong :premises (t20.t6)) (step t20.t8 (cl (= @p_26 @p_37)) :rule cong :premises (t20.t5 t20.t7)) (step t20 (cl (= @p_27 (! (forall ((veriT_vr2 Int) (veriT_vr3 Real)) @p_37) :named @p_38))) :rule bind) (step t21 (cl (! (= @p_29 (! (not @p_38) :named @p_40)) :named @p_39)) :rule cong :premises (t20)) (step t22 (cl (! (not @p_39) :named @p_42) (! (not @p_29) :named @p_41) @p_40) :rule equiv_pos2) (step t23 (cl (not @p_41) @p_27) :rule not_not) (step t24 (cl @p_42 @p_27 @p_40) :rule th_resolution :premises (t23 t22)) (step t25 (cl @p_40) :rule th_resolution :premises (t19 t21 t24)) (anchor :step t26 :args ((:= (veriT_vr2 Int) veriT_sk0) (:= (veriT_vr3 Real) veriT_sk1))) (step t26.t1 (cl (! (= veriT_vr2 veriT_sk0) :named @p_48)) :rule refl) (step t26.t2 (cl (= @p_32 (! (< 0 veriT_sk0) :named @p_44))) :rule cong :premises (t26.t1)) (step t26.t3 (cl (= veriT_vr3 veriT_sk1)) :rule refl) (step t26.t4 (cl (= @p_33 (! (< 0.0 veriT_sk1) :named @p_46))) :rule cong :premises (t26.t3)) (step t26.t5 (cl (= @p_34 (! (and @p_44 @p_46) :named @p_47))) :rule cong :premises (t26.t2 t26.t4)) (step t26.t6 (cl @p_48) :rule refl) (step t26.t7 (cl (= @p_36 (! (< (- 1) veriT_sk0) :named @p_49))) :rule cong :premises (t26.t6)) (step t26.t8 (cl (= @p_37 (! (=> @p_47 @p_49) :named @p_50))) :rule cong :premises (t26.t5 t26.t7)) (step t26 (cl (= @p_38 @p_50)) :rule sko_forall) (step t27 (cl (! (= @p_40 (! (not @p_50) :named @p_52)) :named @p_51)) :rule cong :premises (t26)) (step t28 (cl (! (not @p_51) :named @p_54) (! (not @p_40) :named @p_53) @p_52) :rule equiv_pos2) (step t29 (cl (not @p_53) @p_38) :rule not_not) (step t30 (cl @p_54 @p_38 @p_52) :rule th_resolution :premises (t29 t28)) (step t31 (cl @p_52) :rule th_resolution :premises (t25 t27 t30)) (step t32 (cl (! (= @p_52 (! (and @p_47 (! (not @p_49) :named @p_59)) :named @p_56)) :named @p_55)) :rule bool_simplify) (step t33 (cl (! (not @p_55) :named @p_58) (! (not @p_52) :named @p_57) @p_56) :rule equiv_pos2) (step t34 (cl (not @p_57) @p_50) :rule not_not) (step t35 (cl @p_58 @p_50 @p_56) :rule th_resolution :premises (t34 t33)) (step t36 (cl @p_56) :rule th_resolution :premises (t31 t32 t35)) (step t37 (cl (! (= @p_56 (! (and @p_44 @p_46 @p_59) :named @p_61)) :named @p_60)) :rule ac_simp) (step t38 (cl (not @p_60) (not @p_56) @p_61) :rule equiv_pos2) (step t39 (cl @p_61) :rule th_resolution :premises (t36 t37 t38)) (step t40 (cl @p_44) :rule and :premises (t39)) (step t41 (cl @p_59) :rule and :premises (t39)) (step t42 (cl @p_49 (not @p_44)) :rule la_generic :args (1.0 1.0)) (step t43 (cl) :rule resolution :premises (t42 t40 t41)) 3b8147fceb728295aa24bd5bef9bf1721184b75c 49 0 unsat (define-fun veriT_sk0 () Int (! (choice ((veriT_vr1 Int)) (not (! (or (! (< 0 veriT_vr1) :named @p_20) (! (< veriT_vr1 1) :named @p_22)) :named @p_23))) :named @p_29)) (assume axiom0 (! (not (! (forall ((?v0 Int) (?v1 Int)) (! (or (! (< 0 ?v1) :named @p_9) (! (< ?v1 1) :named @p_12)) :named @p_2)) :named @p_1)) :named @p_3)) (step t2 (cl (= @p_1 (! (forall ((?v1 Int)) @p_2) :named @p_4))) :rule qnt_rm_unused) (step t3 (cl (! (= @p_3 (! (not @p_4) :named @p_6)) :named @p_5)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_5) :named @p_8) (! (not @p_3) :named @p_7) @p_6) :rule equiv_pos2) (step t5 (cl (not @p_7) @p_1) :rule not_not) (step t6 (cl @p_8 @p_1 @p_6) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_6) :rule th_resolution :premises (axiom0 t3 t6)) (anchor :step t8 :args ((:= (?v1 Int) veriT_vr0))) (step t8.t1 (cl (! (= ?v1 veriT_vr0) :named @p_11)) :rule refl) (step t8.t2 (cl (= @p_9 (! (< 0 veriT_vr0) :named @p_10))) :rule cong :premises (t8.t1)) (step t8.t3 (cl @p_11) :rule refl) (step t8.t4 (cl (= @p_12 (! (< veriT_vr0 1) :named @p_13))) :rule cong :premises (t8.t3)) (step t8.t5 (cl (= @p_2 (! (or @p_10 @p_13) :named @p_14))) :rule cong :premises (t8.t2 t8.t4)) (step t8 (cl (= @p_4 (! (forall ((veriT_vr0 Int)) @p_14) :named @p_15))) :rule bind) (step t9 (cl (! (= @p_6 (! (not @p_15) :named @p_17)) :named @p_16)) :rule cong :premises (t8)) (step t10 (cl (! (not @p_16) :named @p_19) (! (not @p_6) :named @p_18) @p_17) :rule equiv_pos2) (step t11 (cl (not @p_18) @p_4) :rule not_not) (step t12 (cl @p_19 @p_4 @p_17) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_17) :rule th_resolution :premises (t7 t9 t12)) (anchor :step t14 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t14.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_21)) :rule refl) (step t14.t2 (cl (= @p_10 @p_20)) :rule cong :premises (t14.t1)) (step t14.t3 (cl @p_21) :rule refl) (step t14.t4 (cl (= @p_13 @p_22)) :rule cong :premises (t14.t3)) (step t14.t5 (cl (= @p_14 @p_23)) :rule cong :premises (t14.t2 t14.t4)) (step t14 (cl (= @p_15 (! (forall ((veriT_vr1 Int)) @p_23) :named @p_24))) :rule bind) (step t15 (cl (! (= @p_17 (! (not @p_24) :named @p_26)) :named @p_25)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_25) :named @p_28) (! (not @p_17) :named @p_27) @p_26) :rule equiv_pos2) (step t17 (cl (not @p_27) @p_15) :rule not_not) (step t18 (cl @p_28 @p_15 @p_26) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_26) :rule th_resolution :premises (t13 t15 t18)) (anchor :step t20 :args ((:= (veriT_vr1 Int) veriT_sk0))) (step t20.t1 (cl (! (= veriT_vr1 veriT_sk0) :named @p_31)) :rule refl) (step t20.t2 (cl (= @p_20 (! (< 0 veriT_sk0) :named @p_30))) :rule cong :premises (t20.t1)) (step t20.t3 (cl @p_31) :rule refl) (step t20.t4 (cl (= @p_22 (! (< veriT_sk0 1) :named @p_32))) :rule cong :premises (t20.t3)) (step t20.t5 (cl (= @p_23 (! (or @p_30 @p_32) :named @p_33))) :rule cong :premises (t20.t2 t20.t4)) (step t20 (cl (= @p_24 @p_33)) :rule sko_forall) (step t21 (cl (! (= @p_26 (! (not @p_33) :named @p_35)) :named @p_34)) :rule cong :premises (t20)) (step t22 (cl (! (not @p_34) :named @p_37) (! (not @p_26) :named @p_36) @p_35) :rule equiv_pos2) (step t23 (cl (not @p_36) @p_24) :rule not_not) (step t24 (cl @p_37 @p_24 @p_35) :rule th_resolution :premises (t23 t22)) (step t25 (cl @p_35) :rule th_resolution :premises (t19 t21 t24)) (step t26 (cl (not @p_30)) :rule not_or :premises (t25)) (step t27 (cl (not @p_32)) :rule not_or :premises (t25)) (step t28 (cl @p_32 @p_30) :rule la_generic :args (1 1)) (step t29 (cl) :rule resolution :premises (t28 t26 t27)) 467ecef4941df0c9404f2c8e1d81f077fb40e73a 7 0 unsat (assume axiom0 (! (not (! (not (! (= 1 (* 2 (of_nat$ x$))) :named @p_2)) :named @p_3)) :named @p_1)) (step t2 (cl (not @p_1) @p_2) :rule not_not) (step t3 (cl @p_2) :rule th_resolution :premises (t2 axiom0)) (step t4 (cl @p_3 @p_3) :rule lia_generic) (step t5 (cl @p_3) :rule contraction :premises (t4)) (step t6 (cl) :rule resolution :premises (t5 t3)) 3549c5e5446637c428a1dc6a809dddaffe6daeca 11 0 unsat (assume axiom0 (! (not (! (=> (! (< (! (of_nat$ a$) :named @p_1) 3) :named @p_3) (! (< (* 2 @p_1) 7) :named @p_4)) :named @p_8)) :named @p_2)) (step t2 (cl (! (= @p_2 (! (and @p_3 (! (not @p_4) :named @p_10)) :named @p_6)) :named @p_5)) :rule bool_simplify) (step t3 (cl (! (not @p_5) :named @p_9) (! (not @p_2) :named @p_7) @p_6) :rule equiv_pos2) (step t4 (cl (not @p_7) @p_8) :rule not_not) (step t5 (cl @p_9 @p_8 @p_6) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_6) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl @p_3) :rule and :premises (t6)) (step t8 (cl @p_10) :rule and :premises (t6)) (step t9 (cl @p_4 (not @p_3)) :rule la_generic :args ((div 1 2) 1)) (step t10 (cl) :rule resolution :premises (t9 t7 t8)) af4e96cd41efee9e27fd5c2ad7650835fd28bdc9 21 0 unsat (assume axiom0 (! (not (! (< (! (* 0 (! (+ 1 (! (of_nat$ y$) :named @p_2)) :named @p_1)) :named @p_3) (! (ite (! (< @p_1 @p_2) :named @p_12) 0 (! (- @p_1 @p_2) :named @p_13)) :named @p_5)) :named @p_4)) :named @p_6)) (step t2 (cl (= 0 @p_3)) :rule prod_simplify) (step t3 (cl (= @p_4 (! (< 0 @p_5) :named @p_7))) :rule cong :premises (t2)) (step t4 (cl (! (= @p_6 (! (not @p_7) :named @p_9)) :named @p_8)) :rule cong :premises (t3)) (step t5 (cl (! (not @p_8) :named @p_11) (! (not @p_6) :named @p_10) @p_9) :rule equiv_pos2) (step t6 (cl (not @p_10) @p_4) :rule not_not) (step t7 (cl @p_11 @p_4 @p_9) :rule th_resolution :premises (t6 t5)) (step t8 (cl @p_9) :rule th_resolution :premises (axiom0 t4 t7)) (step t9 (cl (! (= @p_9 (! (and (! (not (! (< 0 @p_5) :named @p_21)) :named @p_18) (! (ite @p_12 (= 0 @p_5) (! (= @p_13 @p_5) :named @p_20)) :named @p_19)) :named @p_15)) :named @p_14)) :rule ite_intro) (step t10 (cl (! (not @p_14) :named @p_17) (! (not @p_9) :named @p_16) @p_15) :rule equiv_pos2) (step t11 (cl (not @p_16) @p_7) :rule not_not) (step t12 (cl @p_17 @p_7 @p_15) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_15) :rule th_resolution :premises (t8 t9 t12)) (step t14 (cl @p_18) :rule and :premises (t13)) (step t15 (cl @p_19) :rule and :premises (t13)) (step t16 (cl @p_12 @p_20) :rule ite1 :premises (t15)) (step t17 (cl (not @p_12)) :rule la_tautology) (step t18 (cl @p_20) :rule resolution :premises (t16 t17)) (step t19 (cl @p_21 (not @p_20)) :rule la_generic :args (1 (- 1))) (step t20 (cl) :rule resolution :premises (t19 t14 t18)) efbe007514b68c6b40ac16694834314f7692069b 33 0 unsat (assume axiom0 (! (not (! (or false (or (! (= (! (ite (! (< 0 (! (+ 1 (! (of_nat$ y$) :named @p_1)) :named @p_2)) :named @p_13) true false) :named @p_3) (! (= @p_1 (! (ite (! (< @p_2 1) :named @p_27) 0 (! (- @p_2 1) :named @p_28)) :named @p_26)) :named @p_14)) :named @p_5) (! (=> (! (not @p_3) :named @p_15) false) :named @p_6))) :named @p_4)) :named @p_7)) (assume axiom1 (! (<= 0 @p_1) :named @p_34)) (step t3 (cl (= @p_4 (! (or false @p_5 @p_6) :named @p_8))) :rule ac_simp) (step t4 (cl (! (= @p_7 (! (not @p_8) :named @p_10)) :named @p_9)) :rule cong :premises (t3)) (step t5 (cl (! (not @p_9) :named @p_12) (! (not @p_7) :named @p_11) @p_10) :rule equiv_pos2) (step t6 (cl (not @p_11) @p_4) :rule not_not) (step t7 (cl @p_12 @p_4 @p_10) :rule th_resolution :premises (t6 t5)) (step t8 (cl @p_10) :rule th_resolution :premises (axiom0 t4 t7)) (step t9 (cl (= @p_3 @p_13)) :rule ite_simplify) (step t10 (cl (= @p_5 (! (= @p_13 @p_14) :named @p_19))) :rule cong :premises (t9)) (step t11 (cl (= @p_15 (! (not @p_13) :named @p_16))) :rule cong :premises (t9)) (step t12 (cl (= @p_6 (! (=> @p_16 false) :named @p_17))) :rule cong :premises (t11)) (step t13 (cl (= @p_17 (! (not @p_16) :named @p_18))) :rule implies_simplify) (step t14 (cl (= @p_18 @p_13)) :rule not_simplify) (step t15 (cl (= @p_6 @p_13)) :rule trans :premises (t12 t13 t14)) (step t16 (cl (= @p_8 (! (or false @p_19 @p_13) :named @p_20))) :rule cong :premises (t10 t15)) (step t17 (cl (= @p_20 (! (or @p_19 @p_13) :named @p_21))) :rule or_simplify) (step t18 (cl (= @p_8 @p_21)) :rule trans :premises (t16 t17)) (step t19 (cl (! (= @p_10 (! (not @p_21) :named @p_23)) :named @p_22)) :rule cong :premises (t18)) (step t20 (cl (! (not @p_22) :named @p_25) (! (not @p_10) :named @p_24) @p_23) :rule equiv_pos2) (step t21 (cl (not @p_24) @p_8) :rule not_not) (step t22 (cl @p_25 @p_8 @p_23) :rule th_resolution :premises (t21 t20)) (step t23 (cl @p_23) :rule th_resolution :premises (t8 t19 t22)) (step t24 (cl (! (= @p_23 (! (and (! (not (or (= @p_13 (= @p_1 @p_26)) @p_13)) :named @p_33) (ite @p_27 (= 0 @p_26) (= @p_28 @p_26))) :named @p_30)) :named @p_29)) :rule ite_intro) (step t25 (cl (! (not @p_29) :named @p_32) (! (not @p_23) :named @p_31) @p_30) :rule equiv_pos2) (step t26 (cl (not @p_31) @p_21) :rule not_not) (step t27 (cl @p_32 @p_21 @p_30) :rule th_resolution :premises (t26 t25)) (step t28 (cl @p_30) :rule th_resolution :premises (t23 t24 t27)) (step t29 (cl @p_33) :rule and :premises (t28)) (step t30 (cl @p_16) :rule not_or :premises (t29)) (step t31 (cl @p_13 (not @p_34)) :rule la_generic :args (1 1)) (step t32 (cl) :rule resolution :premises (t31 t30 axiom1)) d3c98c27318a98e589a892ffe99adffc556ea833 76 0 unsat (assume axiom4 (! (forall ((?v0 Int)) (! (= (! (of_nat$ (! (nat$ ?v0) :named @p_3)) :named @p_5) (! (ite (! (<= 0 ?v0) :named @p_8) ?v0 0) :named @p_10)) :named @p_12)) :named @p_2)) (assume axiom1 (! (not (! (= (! (ite (! (< x$ 0) :named @p_25) (! (- x$) :named @p_26) x$) :named @p_1) (of_nat$ (nat$ @p_1))) :named @p_30)) :named @p_24)) (anchor :step t3 :args ((:= (?v0 Int) veriT_vr0))) (step t3.t1 (cl (! (= ?v0 veriT_vr0) :named @p_7)) :rule refl) (step t3.t2 (cl (= @p_3 (! (nat$ veriT_vr0) :named @p_4))) :rule cong :premises (t3.t1)) (step t3.t3 (cl (= @p_5 (! (of_nat$ @p_4) :named @p_6))) :rule cong :premises (t3.t2)) (step t3.t4 (cl @p_7) :rule refl) (step t3.t5 (cl (= @p_8 (! (<= 0 veriT_vr0) :named @p_9))) :rule cong :premises (t3.t4)) (step t3.t6 (cl @p_7) :rule refl) (step t3.t7 (cl (= @p_10 (! (ite @p_9 veriT_vr0 0) :named @p_11))) :rule cong :premises (t3.t5 t3.t6)) (step t3.t8 (cl (= @p_12 (! (= @p_6 @p_11) :named @p_13))) :rule cong :premises (t3.t3 t3.t7)) (step t3 (cl (! (= @p_2 (! (forall ((veriT_vr0 Int)) @p_13) :named @p_15)) :named @p_14)) :rule bind) (step t4 (cl (not @p_14) (not @p_2) @p_15) :rule equiv_pos2) (step t5 (cl @p_15) :rule th_resolution :premises (axiom4 t3 t4)) (anchor :step t6 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t6.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_18)) :rule refl) (step t6.t2 (cl (= @p_4 (! (nat$ veriT_vr1) :named @p_16))) :rule cong :premises (t6.t1)) (step t6.t3 (cl (= @p_6 (! (of_nat$ @p_16) :named @p_17))) :rule cong :premises (t6.t2)) (step t6.t4 (cl @p_18) :rule refl) (step t6.t5 (cl (= @p_9 (! (<= 0 veriT_vr1) :named @p_19))) :rule cong :premises (t6.t4)) (step t6.t6 (cl @p_18) :rule refl) (step t6.t7 (cl (= @p_11 (! (ite @p_19 veriT_vr1 0) :named @p_20))) :rule cong :premises (t6.t5 t6.t6)) (step t6.t8 (cl (= @p_13 (! (= @p_17 @p_20) :named @p_21))) :rule cong :premises (t6.t3 t6.t7)) (step t6 (cl (! (= @p_15 (! (forall ((veriT_vr1 Int)) @p_21) :named @p_23)) :named @p_22)) :rule bind) (step t7 (cl (not @p_22) (not @p_15) @p_23) :rule equiv_pos2) (step t8 (cl @p_23) :rule th_resolution :premises (t5 t6 t7)) (step t9 (cl (! (= @p_24 (! (and (! (not (! (= @p_1 (! (of_nat$ (nat$ @p_1)) :named @p_36)) :named @p_53)) :named @p_32) (! (ite @p_25 (! (= @p_26 @p_1) :named @p_35) (! (= x$ @p_1) :named @p_34)) :named @p_33)) :named @p_28)) :named @p_27)) :rule ite_intro) (step t10 (cl (! (not @p_27) :named @p_31) (! (not @p_24) :named @p_29) @p_28) :rule equiv_pos2) (step t11 (cl (not @p_29) @p_30) :rule not_not) (step t12 (cl @p_31 @p_30 @p_28) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_28) :rule th_resolution :premises (axiom1 t9 t12)) (step t14 (cl @p_32) :rule and :premises (t13)) (step t15 (cl @p_33) :rule and :premises (t13)) (step t16 (cl @p_25 @p_34) :rule ite1 :premises (t15)) (step t17 (cl (! (not @p_25) :named @p_64) @p_35) :rule ite2 :premises (t15)) (step t18 (cl (or (! (not @p_23) :named @p_43) (! (= @p_36 (! (ite (! (<= 0 @p_1) :named @p_39) @p_1 0) :named @p_38)) :named @p_37))) :rule forall_inst :args ((:= veriT_vr1 @p_1))) (anchor :step t19) (assume t19.h1 @p_37) (step t19.t2 (cl (! (= @p_37 (! (and (! (= @p_36 @p_38) :named @p_47) (! (ite @p_39 (! (= @p_1 @p_38) :named @p_51) (! (= 0 @p_38) :named @p_49)) :named @p_48)) :named @p_40)) :named @p_41)) :rule ite_intro) (step t19.t3 (cl (not @p_41) (! (not @p_37) :named @p_42) @p_40) :rule equiv_pos2) (step t19.t4 (cl @p_40) :rule th_resolution :premises (t19.h1 t19.t2 t19.t3)) (step t19 (cl @p_42 @p_40) :rule subproof :discharge (h1)) (step t20 (cl @p_43 @p_37) :rule or :premises (t18)) (step t21 (cl (! (or @p_43 @p_40) :named @p_45) (! (not @p_43) :named @p_44)) :rule or_neg) (step t22 (cl (not @p_44) @p_23) :rule not_not) (step t23 (cl @p_45 @p_23) :rule th_resolution :premises (t22 t21)) (step t24 (cl @p_45 (! (not @p_40) :named @p_46)) :rule or_neg) (step t25 (cl @p_45) :rule th_resolution :premises (t20 t19 t23 t24)) (step t26 (cl @p_46 @p_47) :rule and_pos) (step t27 (cl (! (not @p_48) :named @p_50) @p_39 @p_49) :rule ite_pos1) (step t28 (cl @p_50 (! (not @p_39) :named @p_56) @p_51) :rule ite_pos2) (step t29 (cl @p_46 @p_48) :rule and_pos) (step t30 (cl @p_43 @p_40) :rule or :premises (t25)) (step t31 (cl @p_40) :rule resolution :premises (t30 t8)) (step t32 (cl @p_47) :rule resolution :premises (t26 t31)) (step t33 (cl @p_48) :rule resolution :premises (t29 t31)) (step t34 (cl (! (= @p_1 @p_1) :named @p_52)) :rule eq_reflexive) (step t35 (cl (! (not @p_52) :named @p_57) (! (not @p_51) :named @p_54) (! (not @p_47) :named @p_55) @p_53) :rule eq_transitive) (step t36 (cl @p_54 @p_55 @p_53) :rule th_resolution :premises (t35 t34)) (step t37 (cl @p_54) :rule resolution :premises (t36 t14 t32)) (step t38 (cl @p_56) :rule resolution :premises (t28 t37 t33)) (step t39 (cl @p_49) :rule resolution :premises (t27 t38 t33)) (step t40 (cl (! (not (! (= 0 @p_36) :named @p_60)) :named @p_58) @p_57 (! (not (! (<= @p_36 @p_1) :named @p_62)) :named @p_59) @p_39) :rule eq_congruent_pred) (step t41 (cl @p_58 @p_59 @p_39) :rule th_resolution :premises (t40 t34)) (step t42 (cl (! (not @p_49) :named @p_61) @p_55 @p_60) :rule eq_transitive) (step t43 (cl @p_59 @p_39 @p_61 @p_55) :rule th_resolution :premises (t41 t42)) (step t44 (cl @p_59) :rule resolution :premises (t43 t32 t38 t39)) (step t45 (cl @p_62 @p_25 @p_58 (! (not @p_34) :named @p_63)) :rule la_generic :args (1 1 1 (- 1))) (step t46 (cl @p_62 @p_25 @p_63 @p_61 @p_55) :rule th_resolution :premises (t45 t42)) (step t47 (cl @p_25 @p_63) :rule resolution :premises (t46 t32 t39 t44)) (step t48 (cl @p_25) :rule resolution :premises (t47 t16)) (step t49 (cl @p_35) :rule resolution :premises (t17 t48)) (step t50 (cl @p_62 @p_64 @p_58 (! (not @p_35) :named @p_65)) :rule la_generic :args (1 1 1 (- 1))) (step t51 (cl @p_62 @p_64 @p_65 @p_61 @p_55) :rule th_resolution :premises (t50 t42)) (step t52 (cl) :rule resolution :premises (t51 t48 t49 t32 t39 t44)) d4fe162ae425370cea445760001be7c859bc0288 337 0 unsat (define-fun veriT_sk1 () Nat$ (! (choice ((veriT_vr18 Nat$)) (not (! (=> (! (dvd$ veriT_vr18 (! (nat$ (! (+ 1 (! (* 4 (! (of_nat$ m$) :named @p_3)) :named @p_117)) :named @p_119)) :named @p_120)) :named @p_206) (! (or (! (= 1 (! (of_nat$ veriT_vr18) :named @p_171)) :named @p_208) (! (= (! (of_nat$ @p_120) :named @p_158) @p_171) :named @p_210)) :named @p_211)) :named @p_205))) :named @p_172)) (assume axiom0 (! (forall ((?v0 Nat$)) (! (= (! (prime_nat$ ?v0) :named @p_7) (! (and (! (< 1 (! (of_nat$ ?v0) :named @p_1)) :named @p_10) (! (forall ((?v1 Nat$)) (! (=> (! (dvd$ ?v1 ?v0) :named @p_14) (! (or (! (= 1 (! (of_nat$ ?v1) :named @p_2)) :named @p_17) (! (= @p_1 @p_2) :named @p_21)) :named @p_23)) :named @p_25)) :named @p_12)) :named @p_27)) :named @p_29)) :named @p_4)) (assume axiom1 (! (not (! (=> (! (prime_nat$ (! (nat$ (! (+ @p_117 1) :named @p_116)) :named @p_118)) :named @p_109) (! (<= 1 @p_3) :named @p_110)) :named @p_114)) :named @p_108)) (assume axiom2 (! (forall ((?v0 Nat$)) (! (<= 0 @p_1) :named @p_127)) :named @p_125)) (assume axiom4 (! (forall ((?v0 Int)) (! (= (! (of_nat$ (! (nat$ ?v0) :named @p_136)) :named @p_138) (! (ite (! (<= 0 ?v0) :named @p_141) ?v0 0) :named @p_143)) :named @p_145)) :named @p_135)) (anchor :step t5 :args ((:= (?v0 Nat$) veriT_vr0))) (step t5.t1 (cl (! (= ?v0 veriT_vr0) :named @p_9)) :rule refl) (step t5.t2 (cl (= @p_7 (! (prime_nat$ veriT_vr0) :named @p_8))) :rule cong :premises (t5.t1)) (step t5.t3 (cl @p_9) :rule refl) (step t5.t4 (cl (! (= @p_1 (! (of_nat$ veriT_vr0) :named @p_5)) :named @p_19)) :rule cong :premises (t5.t3)) (step t5.t5 (cl (= @p_10 (! (< 1 @p_5) :named @p_11))) :rule cong :premises (t5.t4)) (anchor :step t5.t6 :args ((:= (?v1 Nat$) veriT_vr1))) (step t5.t6.t1 (cl (! (= ?v1 veriT_vr1) :named @p_16)) :rule refl) (step t5.t6.t2 (cl @p_9) :rule refl) (step t5.t6.t3 (cl (= @p_14 (! (dvd$ veriT_vr1 veriT_vr0) :named @p_15))) :rule cong :premises (t5.t6.t1 t5.t6.t2)) (step t5.t6.t4 (cl @p_16) :rule refl) (step t5.t6.t5 (cl (! (= @p_2 (! (of_nat$ veriT_vr1) :named @p_6)) :named @p_20)) :rule cong :premises (t5.t6.t4)) (step t5.t6.t6 (cl (= @p_17 (! (= 1 @p_6) :named @p_18))) :rule cong :premises (t5.t6.t5)) (step t5.t6.t7 (cl @p_9) :rule refl) (step t5.t6.t8 (cl @p_19) :rule cong :premises (t5.t6.t7)) (step t5.t6.t9 (cl @p_16) :rule refl) (step t5.t6.t10 (cl @p_20) :rule cong :premises (t5.t6.t9)) (step t5.t6.t11 (cl (= @p_21 (! (= @p_5 @p_6) :named @p_22))) :rule cong :premises (t5.t6.t8 t5.t6.t10)) (step t5.t6.t12 (cl (= @p_23 (! (or @p_18 @p_22) :named @p_24))) :rule cong :premises (t5.t6.t6 t5.t6.t11)) (step t5.t6.t13 (cl (= @p_25 (! (=> @p_15 @p_24) :named @p_26))) :rule cong :premises (t5.t6.t3 t5.t6.t12)) (step t5.t6 (cl (= @p_12 (! (forall ((veriT_vr1 Nat$)) @p_26) :named @p_13))) :rule bind) (step t5.t7 (cl (= @p_27 (! (and @p_11 @p_13) :named @p_28))) :rule cong :premises (t5.t5 t5.t6)) (step t5.t8 (cl (= @p_29 (! (= @p_8 @p_28) :named @p_30))) :rule cong :premises (t5.t2 t5.t7)) (step t5 (cl (! (= @p_4 (! (forall ((veriT_vr0 Nat$)) @p_30) :named @p_32)) :named @p_31)) :rule bind) (step t6 (cl (not @p_31) (not @p_4) @p_32) :rule equiv_pos2) (step t7 (cl @p_32) :rule th_resolution :premises (axiom0 t5 t6)) (anchor :step t8 :args ((veriT_vr0 Nat$))) (step t8.t1 (cl (= @p_30 (! (and (! (=> @p_8 @p_28) :named @p_52) (! (=> @p_28 @p_8) :named @p_65)) :named @p_33))) :rule connective_def) (step t8 (cl (! (= @p_32 (! (forall ((veriT_vr0 Nat$)) @p_33) :named @p_35)) :named @p_34)) :rule bind) (step t9 (cl (not @p_34) (not @p_32) @p_35) :rule equiv_pos2) (step t10 (cl @p_35) :rule th_resolution :premises (t7 t8 t9)) (anchor :step t11 :args ((:= (veriT_vr0 Nat$) veriT_vr2))) (step t11.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_41)) :rule refl) (step t11.t2 (cl (! (= @p_8 (! (prime_nat$ veriT_vr2) :named @p_40)) :named @p_64)) :rule cong :premises (t11.t1)) (step t11.t3 (cl @p_41) :rule refl) (step t11.t4 (cl (! (= @p_5 (! (of_nat$ veriT_vr2) :named @p_36)) :named @p_46)) :rule cong :premises (t11.t3)) (step t11.t5 (cl (! (= @p_11 (! (< 1 @p_36) :named @p_38)) :named @p_54)) :rule cong :premises (t11.t4)) (anchor :step t11.t6 :args ((:= (veriT_vr1 Nat$) veriT_vr3))) (step t11.t6.t1 (cl (! (= veriT_vr1 veriT_vr3) :named @p_44)) :rule refl) (step t11.t6.t2 (cl @p_41) :rule refl) (step t11.t6.t3 (cl (= @p_15 (! (dvd$ veriT_vr3 veriT_vr2) :named @p_43))) :rule cong :premises (t11.t6.t1 t11.t6.t2)) (step t11.t6.t4 (cl @p_44) :rule refl) (step t11.t6.t5 (cl (! (= @p_6 (! (of_nat$ veriT_vr3) :named @p_37)) :named @p_47)) :rule cong :premises (t11.t6.t4)) (step t11.t6.t6 (cl (= @p_18 (! (= 1 @p_37) :named @p_45))) :rule cong :premises (t11.t6.t5)) (step t11.t6.t7 (cl @p_41) :rule refl) (step t11.t6.t8 (cl @p_46) :rule cong :premises (t11.t6.t7)) (step t11.t6.t9 (cl @p_44) :rule refl) (step t11.t6.t10 (cl @p_47) :rule cong :premises (t11.t6.t9)) (step t11.t6.t11 (cl (= @p_22 (! (= @p_36 @p_37) :named @p_48))) :rule cong :premises (t11.t6.t8 t11.t6.t10)) (step t11.t6.t12 (cl (= @p_24 (! (or @p_45 @p_48) :named @p_49))) :rule cong :premises (t11.t6.t6 t11.t6.t11)) (step t11.t6.t13 (cl (= @p_26 (! (=> @p_43 @p_49) :named @p_50))) :rule cong :premises (t11.t6.t3 t11.t6.t12)) (step t11.t6 (cl (= @p_13 (! (forall ((veriT_vr3 Nat$)) @p_50) :named @p_42))) :rule bind) (step t11.t7 (cl (= @p_28 (! (and @p_38 @p_42) :named @p_51))) :rule cong :premises (t11.t5 t11.t6)) (step t11.t8 (cl (= @p_52 (! (=> @p_40 @p_51) :named @p_53))) :rule cong :premises (t11.t2 t11.t7)) (step t11.t9 (cl @p_41) :rule refl) (step t11.t10 (cl @p_46) :rule cong :premises (t11.t9)) (step t11.t11 (cl @p_54) :rule cong :premises (t11.t10)) (anchor :step t11.t12 :args ((:= (veriT_vr1 Nat$) veriT_vr4))) (step t11.t12.t1 (cl (! (= veriT_vr1 veriT_vr4) :named @p_57)) :rule refl) (step t11.t12.t2 (cl @p_41) :rule refl) (step t11.t12.t3 (cl (= @p_15 (! (dvd$ veriT_vr4 veriT_vr2) :named @p_56))) :rule cong :premises (t11.t12.t1 t11.t12.t2)) (step t11.t12.t4 (cl @p_57) :rule refl) (step t11.t12.t5 (cl (! (= @p_6 (! (of_nat$ veriT_vr4) :named @p_39)) :named @p_59)) :rule cong :premises (t11.t12.t4)) (step t11.t12.t6 (cl (= @p_18 (! (= 1 @p_39) :named @p_58))) :rule cong :premises (t11.t12.t5)) (step t11.t12.t7 (cl @p_41) :rule refl) (step t11.t12.t8 (cl @p_46) :rule cong :premises (t11.t12.t7)) (step t11.t12.t9 (cl @p_57) :rule refl) (step t11.t12.t10 (cl @p_59) :rule cong :premises (t11.t12.t9)) (step t11.t12.t11 (cl (= @p_22 (! (= @p_36 @p_39) :named @p_60))) :rule cong :premises (t11.t12.t8 t11.t12.t10)) (step t11.t12.t12 (cl (= @p_24 (! (or @p_58 @p_60) :named @p_61))) :rule cong :premises (t11.t12.t6 t11.t12.t11)) (step t11.t12.t13 (cl (= @p_26 (! (=> @p_56 @p_61) :named @p_62))) :rule cong :premises (t11.t12.t3 t11.t12.t12)) (step t11.t12 (cl (= @p_13 (! (forall ((veriT_vr4 Nat$)) @p_62) :named @p_55))) :rule bind) (step t11.t13 (cl (= @p_28 (! (and @p_38 @p_55) :named @p_63))) :rule cong :premises (t11.t11 t11.t12)) (step t11.t14 (cl @p_41) :rule refl) (step t11.t15 (cl @p_64) :rule cong :premises (t11.t14)) (step t11.t16 (cl (= @p_65 (! (=> @p_63 @p_40) :named @p_66))) :rule cong :premises (t11.t13 t11.t15)) (step t11.t17 (cl (= @p_33 (! (and @p_53 @p_66) :named @p_67))) :rule cong :premises (t11.t8 t11.t16)) (step t11 (cl (! (= @p_35 (! (forall ((veriT_vr2 Nat$)) @p_67) :named @p_69)) :named @p_68)) :rule bind) (step t12 (cl (not @p_68) (not @p_35) @p_69) :rule equiv_pos2) (step t13 (cl @p_69) :rule th_resolution :premises (t10 t11 t12)) (anchor :step t14 :args ((:= (veriT_vr2 Nat$) veriT_vr5))) (step t14.t1 (cl (! (= veriT_vr2 veriT_vr5) :named @p_74)) :rule refl) (step t14.t2 (cl (! (= @p_40 (! (prime_nat$ veriT_vr5) :named @p_73)) :named @p_89)) :rule cong :premises (t14.t1)) (step t14.t3 (cl @p_74) :rule refl) (step t14.t4 (cl (! (= @p_36 (! (of_nat$ veriT_vr5) :named @p_70)) :named @p_80)) :rule cong :premises (t14.t3)) (step t14.t5 (cl (! (= @p_38 (! (< 1 @p_70) :named @p_75)) :named @p_86)) :rule cong :premises (t14.t4)) (anchor :step t14.t6 :args ((:= (veriT_vr3 Nat$) veriT_vr6))) (step t14.t6.t1 (cl (! (= veriT_vr3 veriT_vr6) :named @p_78)) :rule refl) (step t14.t6.t2 (cl @p_74) :rule refl) (step t14.t6.t3 (cl (= @p_43 (! (dvd$ veriT_vr6 veriT_vr5) :named @p_77))) :rule cong :premises (t14.t6.t1 t14.t6.t2)) (step t14.t6.t4 (cl @p_78) :rule refl) (step t14.t6.t5 (cl (! (= @p_37 (! (of_nat$ veriT_vr6) :named @p_71)) :named @p_81)) :rule cong :premises (t14.t6.t4)) (step t14.t6.t6 (cl (= @p_45 (! (= 1 @p_71) :named @p_79))) :rule cong :premises (t14.t6.t5)) (step t14.t6.t7 (cl @p_74) :rule refl) (step t14.t6.t8 (cl @p_80) :rule cong :premises (t14.t6.t7)) (step t14.t6.t9 (cl @p_78) :rule refl) (step t14.t6.t10 (cl @p_81) :rule cong :premises (t14.t6.t9)) (step t14.t6.t11 (cl (= @p_48 (! (= @p_70 @p_71) :named @p_82))) :rule cong :premises (t14.t6.t8 t14.t6.t10)) (step t14.t6.t12 (cl (= @p_49 (! (or @p_79 @p_82) :named @p_83))) :rule cong :premises (t14.t6.t6 t14.t6.t11)) (step t14.t6.t13 (cl (= @p_50 (! (=> @p_77 @p_83) :named @p_84))) :rule cong :premises (t14.t6.t3 t14.t6.t12)) (step t14.t6 (cl (= @p_42 (! (forall ((veriT_vr6 Nat$)) @p_84) :named @p_76))) :rule bind) (step t14.t7 (cl (= @p_51 (! (and @p_75 @p_76) :named @p_72))) :rule cong :premises (t14.t5 t14.t6)) (step t14.t8 (cl (= @p_53 (! (=> @p_73 @p_72) :named @p_85))) :rule cong :premises (t14.t2 t14.t7)) (step t14.t9 (cl @p_74) :rule refl) (step t14.t10 (cl @p_80) :rule cong :premises (t14.t9)) (step t14.t11 (cl @p_86) :rule cong :premises (t14.t10)) (anchor :step t14.t12 :args ((:= (veriT_vr4 Nat$) veriT_vr6))) (step t14.t12.t1 (cl (! (= veriT_vr4 veriT_vr6) :named @p_87)) :rule refl) (step t14.t12.t2 (cl @p_74) :rule refl) (step t14.t12.t3 (cl (= @p_56 @p_77)) :rule cong :premises (t14.t12.t1 t14.t12.t2)) (step t14.t12.t4 (cl @p_87) :rule refl) (step t14.t12.t5 (cl (! (= @p_39 @p_71) :named @p_88)) :rule cong :premises (t14.t12.t4)) (step t14.t12.t6 (cl (= @p_58 @p_79)) :rule cong :premises (t14.t12.t5)) (step t14.t12.t7 (cl @p_74) :rule refl) (step t14.t12.t8 (cl @p_80) :rule cong :premises (t14.t12.t7)) (step t14.t12.t9 (cl @p_87) :rule refl) (step t14.t12.t10 (cl @p_88) :rule cong :premises (t14.t12.t9)) (step t14.t12.t11 (cl (= @p_60 @p_82)) :rule cong :premises (t14.t12.t8 t14.t12.t10)) (step t14.t12.t12 (cl (= @p_61 @p_83)) :rule cong :premises (t14.t12.t6 t14.t12.t11)) (step t14.t12.t13 (cl (= @p_62 @p_84)) :rule cong :premises (t14.t12.t3 t14.t12.t12)) (step t14.t12 (cl (= @p_55 @p_76)) :rule bind) (step t14.t13 (cl (= @p_63 @p_72)) :rule cong :premises (t14.t11 t14.t12)) (step t14.t14 (cl @p_74) :rule refl) (step t14.t15 (cl @p_89) :rule cong :premises (t14.t14)) (step t14.t16 (cl (= @p_66 (! (=> @p_72 @p_73) :named @p_90))) :rule cong :premises (t14.t13 t14.t15)) (step t14.t17 (cl (= @p_67 (! (and @p_85 @p_90) :named @p_91))) :rule cong :premises (t14.t8 t14.t16)) (step t14 (cl (! (= @p_69 (! (forall ((veriT_vr5 Nat$)) @p_91) :named @p_93)) :named @p_92)) :rule bind) (step t15 (cl (not @p_92) (not @p_69) @p_93) :rule equiv_pos2) (step t16 (cl @p_93) :rule th_resolution :premises (t13 t14 t15)) (anchor :step t17 :args ((:= (veriT_vr5 Nat$) veriT_vr5))) (anchor :step t17.t1 :args ((:= (veriT_vr6 Nat$) veriT_vr7))) (step t17.t1.t1 (cl (! (= veriT_vr6 veriT_vr7) :named @p_97)) :rule refl) (step t17.t1.t2 (cl (= @p_77 (! (dvd$ veriT_vr7 veriT_vr5) :named @p_96))) :rule cong :premises (t17.t1.t1)) (step t17.t1.t3 (cl @p_97) :rule refl) (step t17.t1.t4 (cl (! (= @p_71 (! (of_nat$ veriT_vr7) :named @p_94)) :named @p_99)) :rule cong :premises (t17.t1.t3)) (step t17.t1.t5 (cl (= @p_79 (! (= 1 @p_94) :named @p_98))) :rule cong :premises (t17.t1.t4)) (step t17.t1.t6 (cl @p_97) :rule refl) (step t17.t1.t7 (cl @p_99) :rule cong :premises (t17.t1.t6)) (step t17.t1.t8 (cl (= @p_82 (! (= @p_70 @p_94) :named @p_100))) :rule cong :premises (t17.t1.t7)) (step t17.t1.t9 (cl (= @p_83 (! (or @p_98 @p_100) :named @p_101))) :rule cong :premises (t17.t1.t5 t17.t1.t8)) (step t17.t1.t10 (cl (= @p_84 (! (=> @p_96 @p_101) :named @p_102))) :rule cong :premises (t17.t1.t2 t17.t1.t9)) (step t17.t1 (cl (= @p_76 (! (forall ((veriT_vr7 Nat$)) @p_102) :named @p_95))) :rule bind) (step t17.t2 (cl (= @p_72 (! (and @p_75 @p_95) :named @p_103))) :rule cong :premises (t17.t1)) (step t17.t3 (cl (= @p_90 (! (=> @p_103 @p_73) :named @p_104))) :rule cong :premises (t17.t2)) (step t17.t4 (cl (= @p_91 (! (and @p_85 @p_104) :named @p_105))) :rule cong :premises (t17.t3)) (step t17 (cl (! (= @p_93 (! (forall ((veriT_vr5 Nat$)) @p_105) :named @p_107)) :named @p_106)) :rule bind) (step t18 (cl (not @p_106) (not @p_93) @p_107) :rule equiv_pos2) (step t19 (cl @p_107) :rule th_resolution :premises (t16 t17 t18)) (step t20 (cl (! (= @p_108 (! (and @p_109 (! (not @p_110) :named @p_122)) :named @p_112)) :named @p_111)) :rule bool_simplify) (step t21 (cl (! (not @p_111) :named @p_115) (! (not @p_108) :named @p_113) @p_112) :rule equiv_pos2) (step t22 (cl (not @p_113) @p_114) :rule not_not) (step t23 (cl @p_115 @p_114 @p_112) :rule th_resolution :premises (t22 t21)) (step t24 (cl @p_112) :rule th_resolution :premises (axiom1 t20 t23)) (step t25 (cl (= @p_116 @p_119)) :rule sum_simplify) (step t26 (cl (= @p_118 @p_120)) :rule cong :premises (t25)) (step t27 (cl (= @p_109 (! (prime_nat$ @p_120) :named @p_121))) :rule cong :premises (t26)) (step t28 (cl (! (= @p_112 (! (and @p_121 @p_122) :named @p_124)) :named @p_123)) :rule cong :premises (t27)) (step t29 (cl (not @p_123) (not @p_112) @p_124) :rule equiv_pos2) (step t30 (cl @p_124) :rule th_resolution :premises (t24 t28 t29)) (anchor :step t31 :args ((:= (?v0 Nat$) veriT_vr8))) (step t31.t1 (cl (= ?v0 veriT_vr8)) :rule refl) (step t31.t2 (cl (= @p_1 (! (of_nat$ veriT_vr8) :named @p_126))) :rule cong :premises (t31.t1)) (step t31.t3 (cl (= @p_127 (! (<= 0 @p_126) :named @p_128))) :rule cong :premises (t31.t2)) (step t31 (cl (! (= @p_125 (! (forall ((veriT_vr8 Nat$)) @p_128) :named @p_130)) :named @p_129)) :rule bind) (step t32 (cl (not @p_129) (not @p_125) @p_130) :rule equiv_pos2) (step t33 (cl @p_130) :rule th_resolution :premises (axiom2 t31 t32)) (anchor :step t34 :args ((:= (veriT_vr8 Nat$) veriT_vr9))) (step t34.t1 (cl (= veriT_vr8 veriT_vr9)) :rule refl) (step t34.t2 (cl (= @p_126 (! (of_nat$ veriT_vr9) :named @p_131))) :rule cong :premises (t34.t1)) (step t34.t3 (cl (= @p_128 (! (<= 0 @p_131) :named @p_132))) :rule cong :premises (t34.t2)) (step t34 (cl (! (= @p_130 (! (forall ((veriT_vr9 Nat$)) @p_132) :named @p_134)) :named @p_133)) :rule bind) (step t35 (cl (not @p_133) (not @p_130) @p_134) :rule equiv_pos2) (step t36 (cl @p_134) :rule th_resolution :premises (t33 t34 t35)) (anchor :step t37 :args ((:= (?v0 Int) veriT_vr12))) (step t37.t1 (cl (! (= ?v0 veriT_vr12) :named @p_140)) :rule refl) (step t37.t2 (cl (= @p_136 (! (nat$ veriT_vr12) :named @p_137))) :rule cong :premises (t37.t1)) (step t37.t3 (cl (= @p_138 (! (of_nat$ @p_137) :named @p_139))) :rule cong :premises (t37.t2)) (step t37.t4 (cl @p_140) :rule refl) (step t37.t5 (cl (= @p_141 (! (<= 0 veriT_vr12) :named @p_142))) :rule cong :premises (t37.t4)) (step t37.t6 (cl @p_140) :rule refl) (step t37.t7 (cl (= @p_143 (! (ite @p_142 veriT_vr12 0) :named @p_144))) :rule cong :premises (t37.t5 t37.t6)) (step t37.t8 (cl (= @p_145 (! (= @p_139 @p_144) :named @p_146))) :rule cong :premises (t37.t3 t37.t7)) (step t37 (cl (! (= @p_135 (! (forall ((veriT_vr12 Int)) @p_146) :named @p_148)) :named @p_147)) :rule bind) (step t38 (cl (not @p_147) (not @p_135) @p_148) :rule equiv_pos2) (step t39 (cl @p_148) :rule th_resolution :premises (axiom4 t37 t38)) (anchor :step t40 :args ((:= (veriT_vr12 Int) veriT_vr13))) (step t40.t1 (cl (! (= veriT_vr12 veriT_vr13) :named @p_151)) :rule refl) (step t40.t2 (cl (= @p_137 (! (nat$ veriT_vr13) :named @p_149))) :rule cong :premises (t40.t1)) (step t40.t3 (cl (= @p_139 (! (of_nat$ @p_149) :named @p_150))) :rule cong :premises (t40.t2)) (step t40.t4 (cl @p_151) :rule refl) (step t40.t5 (cl (= @p_142 (! (<= 0 veriT_vr13) :named @p_152))) :rule cong :premises (t40.t4)) (step t40.t6 (cl @p_151) :rule refl) (step t40.t7 (cl (= @p_144 (! (ite @p_152 veriT_vr13 0) :named @p_153))) :rule cong :premises (t40.t5 t40.t6)) (step t40.t8 (cl (= @p_146 (! (= @p_150 @p_153) :named @p_154))) :rule cong :premises (t40.t3 t40.t7)) (step t40 (cl (! (= @p_148 (! (forall ((veriT_vr13 Int)) @p_154) :named @p_156)) :named @p_155)) :rule bind) (step t41 (cl (not @p_155) (not @p_148) @p_156) :rule equiv_pos2) (step t42 (cl @p_156) :rule th_resolution :premises (t39 t40 t41)) (step t43 (cl @p_121) :rule and :premises (t30)) (step t44 (cl @p_122) :rule and :premises (t30)) (step t45 (cl (or (! (not @p_156) :named @p_164) (! (= @p_158 (! (ite (! (<= 0 @p_119) :named @p_160) @p_119 0) :named @p_159)) :named @p_157))) :rule forall_inst :args ((:= veriT_vr13 @p_119))) (anchor :step t46) (assume t46.h1 @p_157) (step t46.t2 (cl (! (= @p_157 (! (and (! (= @p_158 @p_159) :named @p_244) (! (ite @p_160 (! (= @p_119 @p_159) :named @p_246) (= 0 @p_159)) :named @p_245)) :named @p_161)) :named @p_162)) :rule ite_intro) (step t46.t3 (cl (not @p_162) (! (not @p_157) :named @p_163) @p_161) :rule equiv_pos2) (step t46.t4 (cl @p_161) :rule th_resolution :premises (t46.h1 t46.t2 t46.t3)) (step t46 (cl @p_163 @p_161) :rule subproof :discharge (h1)) (step t47 (cl @p_164 @p_157) :rule or :premises (t45)) (step t48 (cl (! (or @p_164 @p_161) :named @p_166) (! (not @p_164) :named @p_165)) :rule or_neg) (step t49 (cl (not @p_165) @p_156) :rule not_not) (step t50 (cl @p_166 @p_156) :rule th_resolution :premises (t49 t48)) (step t51 (cl @p_166 (! (not @p_161) :named @p_243)) :rule or_neg) (step t52 (cl @p_166) :rule th_resolution :premises (t47 t46 t50 t51)) (step t53 (cl (or (! (not @p_134) :named @p_247) (! (<= 0 @p_3) :named @p_248))) :rule forall_inst :args ((:= veriT_vr9 m$))) (step t54 (cl (not (! (not (! (not @p_107) :named @p_167)) :named @p_241)) @p_107) :rule not_not) (step t55 (cl (or @p_167 (! (and (! (=> @p_121 (! (and (! (< 1 @p_158) :named @p_168) (! (forall ((veriT_vr6 Nat$)) (! (=> (! (dvd$ veriT_vr6 @p_120) :named @p_176) (! (or @p_79 (! (= @p_71 @p_158) :named @p_181)) :named @p_183)) :named @p_185)) :named @p_174)) :named @p_187)) :named @p_189) (! (=> (! (and @p_168 (! (forall ((veriT_vr7 Nat$)) (! (=> (! (dvd$ veriT_vr7 @p_120) :named @p_192) (! (or @p_98 (! (= @p_94 @p_158) :named @p_195)) :named @p_196)) :named @p_197)) :named @p_191)) :named @p_198) @p_121) :named @p_199)) :named @p_169))) :rule forall_inst :args ((:= veriT_vr5 @p_120))) (anchor :step t56) (assume t56.h1 @p_169) (anchor :step t56.t2 :args ((:= (veriT_vr6 Nat$) veriT_vr17))) (step t56.t2.t1 (cl (! (= veriT_vr6 veriT_vr17) :named @p_178)) :rule refl) (step t56.t2.t2 (cl (= @p_176 (! (dvd$ veriT_vr17 @p_120) :named @p_177))) :rule cong :premises (t56.t2.t1)) (step t56.t2.t3 (cl @p_178) :rule refl) (step t56.t2.t4 (cl (! (= @p_71 (! (of_nat$ veriT_vr17) :named @p_175)) :named @p_180)) :rule cong :premises (t56.t2.t3)) (step t56.t2.t5 (cl (= @p_79 (! (= 1 @p_175) :named @p_179))) :rule cong :premises (t56.t2.t4)) (step t56.t2.t6 (cl @p_178) :rule refl) (step t56.t2.t7 (cl @p_180) :rule cong :premises (t56.t2.t6)) (step t56.t2.t8 (cl (= @p_181 (! (= @p_158 @p_175) :named @p_182))) :rule cong :premises (t56.t2.t7)) (step t56.t2.t9 (cl (= @p_183 (! (or @p_179 @p_182) :named @p_184))) :rule cong :premises (t56.t2.t5 t56.t2.t8)) (step t56.t2.t10 (cl (= @p_185 (! (=> @p_177 @p_184) :named @p_186))) :rule cong :premises (t56.t2.t2 t56.t2.t9)) (step t56.t2 (cl (= @p_174 (! (forall ((veriT_vr17 Nat$)) @p_186) :named @p_188))) :rule bind) (step t56.t3 (cl (= @p_187 (! (and @p_168 @p_188) :named @p_190))) :rule cong :premises (t56.t2)) (step t56.t4 (cl (= @p_189 (! (=> @p_121 @p_190) :named @p_200))) :rule cong :premises (t56.t3)) (anchor :step t56.t5 :args ((:= (veriT_vr7 Nat$) veriT_vr17))) (step t56.t5.t1 (cl (! (= veriT_vr7 veriT_vr17) :named @p_193)) :rule refl) (step t56.t5.t2 (cl (= @p_192 @p_177)) :rule cong :premises (t56.t5.t1)) (step t56.t5.t3 (cl @p_193) :rule refl) (step t56.t5.t4 (cl (! (= @p_94 @p_175) :named @p_194)) :rule cong :premises (t56.t5.t3)) (step t56.t5.t5 (cl (= @p_98 @p_179)) :rule cong :premises (t56.t5.t4)) (step t56.t5.t6 (cl @p_193) :rule refl) (step t56.t5.t7 (cl @p_194) :rule cong :premises (t56.t5.t6)) (step t56.t5.t8 (cl (= @p_195 @p_182)) :rule cong :premises (t56.t5.t7)) (step t56.t5.t9 (cl (= @p_196 @p_184)) :rule cong :premises (t56.t5.t5 t56.t5.t8)) (step t56.t5.t10 (cl (= @p_197 @p_186)) :rule cong :premises (t56.t5.t2 t56.t5.t9)) (step t56.t5 (cl (= @p_191 @p_188)) :rule bind) (step t56.t6 (cl (= @p_198 @p_190)) :rule cong :premises (t56.t5)) (step t56.t7 (cl (= @p_199 (! (=> @p_190 @p_121) :named @p_201))) :rule cong :premises (t56.t6)) (step t56.t8 (cl (! (= @p_169 (! (and @p_200 @p_201) :named @p_204)) :named @p_202)) :rule cong :premises (t56.t4 t56.t7)) (step t56.t9 (cl (not @p_202) (! (not @p_169) :named @p_203) @p_204) :rule equiv_pos2) (step t56.t10 (cl @p_204) :rule th_resolution :premises (t56.h1 t56.t8 t56.t9)) (anchor :step t56.t11 :args ((:= (veriT_vr17 Nat$) veriT_vr18))) (step t56.t11.t1 (cl (! (= veriT_vr17 veriT_vr18) :named @p_207)) :rule refl) (step t56.t11.t2 (cl (= @p_177 @p_206)) :rule cong :premises (t56.t11.t1)) (step t56.t11.t3 (cl @p_207) :rule refl) (step t56.t11.t4 (cl (! (= @p_175 @p_171) :named @p_209)) :rule cong :premises (t56.t11.t3)) (step t56.t11.t5 (cl (= @p_179 @p_208)) :rule cong :premises (t56.t11.t4)) (step t56.t11.t6 (cl @p_207) :rule refl) (step t56.t11.t7 (cl @p_209) :rule cong :premises (t56.t11.t6)) (step t56.t11.t8 (cl (= @p_182 @p_210)) :rule cong :premises (t56.t11.t7)) (step t56.t11.t9 (cl (= @p_184 @p_211)) :rule cong :premises (t56.t11.t5 t56.t11.t8)) (step t56.t11.t10 (cl (= @p_186 @p_205)) :rule cong :premises (t56.t11.t2 t56.t11.t9)) (step t56.t11 (cl (= @p_188 (! (forall ((veriT_vr18 Nat$)) @p_205) :named @p_212))) :rule bind) (step t56.t12 (cl (= @p_190 (! (and @p_168 @p_212) :named @p_213))) :rule cong :premises (t56.t11)) (step t56.t13 (cl (= @p_200 (! (=> @p_121 @p_213) :named @p_214))) :rule cong :premises (t56.t12)) (step t56.t14 (cl (= @p_201 (! (=> @p_213 @p_121) :named @p_215))) :rule cong :premises (t56.t12)) (step t56.t15 (cl (! (= @p_204 (! (and @p_214 @p_215) :named @p_217)) :named @p_216)) :rule cong :premises (t56.t13 t56.t14)) (step t56.t16 (cl (not @p_216) (not @p_204) @p_217) :rule equiv_pos2) (step t56.t17 (cl @p_217) :rule th_resolution :premises (t56.t10 t56.t15 t56.t16)) (anchor :step t56.t18 :args ((:= (veriT_vr18 Nat$) veriT_sk1))) (step t56.t18.t1 (cl (! (= veriT_vr18 veriT_sk1) :named @p_220)) :rule refl) (step t56.t18.t2 (cl (= @p_206 (! (dvd$ veriT_sk1 @p_120) :named @p_219))) :rule cong :premises (t56.t18.t1)) (step t56.t18.t3 (cl @p_220) :rule refl) (step t56.t18.t4 (cl (! (= @p_171 (! (of_nat$ veriT_sk1) :named @p_173)) :named @p_222)) :rule cong :premises (t56.t18.t3)) (step t56.t18.t5 (cl (= @p_208 (! (= 1 @p_173) :named @p_221))) :rule cong :premises (t56.t18.t4)) (step t56.t18.t6 (cl @p_220) :rule refl) (step t56.t18.t7 (cl @p_222) :rule cong :premises (t56.t18.t6)) (step t56.t18.t8 (cl (= @p_210 (! (= @p_158 @p_173) :named @p_223))) :rule cong :premises (t56.t18.t7)) (step t56.t18.t9 (cl (= @p_211 (! (or @p_221 @p_223) :named @p_224))) :rule cong :premises (t56.t18.t5 t56.t18.t8)) (step t56.t18.t10 (cl (= @p_205 (! (=> @p_219 @p_224) :named @p_218))) :rule cong :premises (t56.t18.t2 t56.t18.t9)) (step t56.t18 (cl (= @p_212 @p_218)) :rule sko_forall) (step t56.t19 (cl (= @p_213 (! (and @p_168 @p_218) :named @p_225))) :rule cong :premises (t56.t18)) (step t56.t20 (cl (= @p_215 (! (=> @p_225 @p_121) :named @p_226))) :rule cong :premises (t56.t19)) (step t56.t21 (cl (! (= @p_217 (! (and @p_214 @p_226) :named @p_228)) :named @p_227)) :rule cong :premises (t56.t20)) (step t56.t22 (cl (not @p_227) (not @p_217) @p_228) :rule equiv_pos2) (step t56.t23 (cl @p_228) :rule th_resolution :premises (t56.t17 t56.t21 t56.t22)) (anchor :step t56.t24 :args ((:= (veriT_vr18 Nat$) veriT_vr19))) (step t56.t24.t1 (cl (! (= veriT_vr18 veriT_vr19) :named @p_231)) :rule refl) (step t56.t24.t2 (cl (= @p_206 (! (dvd$ veriT_vr19 @p_120) :named @p_230))) :rule cong :premises (t56.t24.t1)) (step t56.t24.t3 (cl @p_231) :rule refl) (step t56.t24.t4 (cl (! (= @p_171 (! (of_nat$ veriT_vr19) :named @p_170)) :named @p_233)) :rule cong :premises (t56.t24.t3)) (step t56.t24.t5 (cl (= @p_208 (! (= 1 @p_170) :named @p_232))) :rule cong :premises (t56.t24.t4)) (step t56.t24.t6 (cl @p_231) :rule refl) (step t56.t24.t7 (cl @p_233) :rule cong :premises (t56.t24.t6)) (step t56.t24.t8 (cl (= @p_210 (! (= @p_158 @p_170) :named @p_234))) :rule cong :premises (t56.t24.t7)) (step t56.t24.t9 (cl (= @p_211 (! (or @p_232 @p_234) :named @p_235))) :rule cong :premises (t56.t24.t5 t56.t24.t8)) (step t56.t24.t10 (cl (= @p_205 (! (=> @p_230 @p_235) :named @p_236))) :rule cong :premises (t56.t24.t2 t56.t24.t9)) (step t56.t24 (cl (= @p_212 (! (forall ((veriT_vr19 Nat$)) @p_236) :named @p_229))) :rule bind) (step t56.t25 (cl (= @p_213 (! (and @p_168 @p_229) :named @p_237))) :rule cong :premises (t56.t24)) (step t56.t26 (cl (= @p_214 (! (=> @p_121 @p_237) :named @p_238))) :rule cong :premises (t56.t25)) (step t56.t27 (cl (! (= @p_228 (! (and @p_238 @p_226) :named @p_239)) :named @p_240)) :rule cong :premises (t56.t26)) (step t56.t28 (cl (not @p_240) (not @p_228) @p_239) :rule equiv_pos2) (step t56.t29 (cl @p_239) :rule th_resolution :premises (t56.t23 t56.t27 t56.t28)) (step t56 (cl @p_203 @p_239) :rule subproof :discharge (h1)) (step t57 (cl @p_167 @p_169) :rule or :premises (t55)) (step t58 (cl (! (or @p_167 @p_239) :named @p_242) @p_241) :rule or_neg) (step t59 (cl @p_242 @p_107) :rule th_resolution :premises (t54 t58)) (step t60 (cl @p_242 (! (not @p_239) :named @p_249)) :rule or_neg) (step t61 (cl @p_242) :rule th_resolution :premises (t57 t56 t59 t60)) (step t62 (cl @p_243 @p_244) :rule and_pos) (step t63 (cl (not @p_245) (not @p_160) @p_246) :rule ite_pos2) (step t64 (cl @p_243 @p_245) :rule and_pos) (step t65 (cl @p_164 @p_161) :rule or :premises (t52)) (step t66 (cl @p_161) :rule resolution :premises (t65 t42)) (step t67 (cl @p_244) :rule resolution :premises (t62 t66)) (step t68 (cl @p_245) :rule resolution :premises (t64 t66)) (step t69 (cl @p_247 @p_248) :rule or :premises (t53)) (step t70 (cl @p_248) :rule resolution :premises (t69 t36)) (step t71 (cl (not @p_237) @p_168) :rule and_pos) (step t72 (cl (! (not @p_238) :named @p_250) (not @p_121) @p_237) :rule implies_pos) (step t73 (cl @p_249 @p_238) :rule and_pos) (step t74 (cl @p_167 @p_239) :rule or :premises (t61)) (step t75 (cl @p_250 @p_237) :rule resolution :premises (t72 t43)) (step t76 (cl @p_239) :rule resolution :premises (t74 t19)) (step t77 (cl @p_238) :rule resolution :premises (t73 t76)) (step t78 (cl @p_237) :rule resolution :premises (t75 t77)) (step t79 (cl @p_168) :rule resolution :premises (t71 t78)) (step t80 (cl (not @p_248) @p_160) :rule la_generic :args (4 1)) (step t81 (cl @p_160) :rule resolution :premises (t80 t70)) (step t82 (cl @p_246) :rule resolution :premises (t63 t81 t68)) (step t83 (cl (! (not @p_168) :named @p_252) @p_110 (not (! (= @p_119 @p_158) :named @p_251))) :rule la_generic :args (1 4 1)) (step t84 (cl (! (not @p_246) :named @p_253) (! (not @p_244) :named @p_254) @p_251) :rule eq_transitive) (step t85 (cl @p_252 @p_110 @p_253 @p_254) :rule th_resolution :premises (t83 t84)) (step t86 (cl) :rule resolution :premises (t85 t44 t67 t82 t79)) 8d6d329bb2354ffa218e4b6256b283b3915c3682 7 0 unsat (assume axiom0 (! (not (! (not (! (= 1 (* 2 x$)) :named @p_2)) :named @p_3)) :named @p_1)) (step t2 (cl (not @p_1) @p_2) :rule not_not) (step t3 (cl @p_2) :rule th_resolution :premises (t2 axiom0)) (step t4 (cl @p_3 @p_3) :rule lia_generic) (step t5 (cl @p_3) :rule contraction :premises (t4)) (step t6 (cl) :rule resolution :premises (t5 t3)) 1cf2d299f7ccd81bab0d38f30ec8595b5f133da6 35 0 unsat (assume axiom0 (! (forall ((?v0 A$) (?v1 B$)) (! (= ?v0 (! (fst$ (! (pair$ ?v0 ?v1) :named @p_3)) :named @p_5)) :named @p_7)) :named @p_1)) (assume axiom1 (! (not (! (=> (! (= (! (fst$ (pair$ x$ y$)) :named @p_26) a$) :named @p_18) (! (= x$ a$) :named @p_19)) :named @p_23)) :named @p_17)) (anchor :step t3 :args ((:= (?v0 A$) veriT_vr0) (:= (?v1 B$) veriT_vr1))) (step t3.t1 (cl (! (= ?v0 veriT_vr0) :named @p_2)) :rule refl) (step t3.t2 (cl @p_2) :rule refl) (step t3.t3 (cl (= ?v1 veriT_vr1)) :rule refl) (step t3.t4 (cl (= @p_3 (! (pair$ veriT_vr0 veriT_vr1) :named @p_4))) :rule cong :premises (t3.t2 t3.t3)) (step t3.t5 (cl (= @p_5 (! (fst$ @p_4) :named @p_6))) :rule cong :premises (t3.t4)) (step t3.t6 (cl (= @p_7 (! (= veriT_vr0 @p_6) :named @p_8))) :rule cong :premises (t3.t1 t3.t5)) (step t3 (cl (! (= @p_1 (! (forall ((veriT_vr0 A$) (veriT_vr1 B$)) @p_8) :named @p_10)) :named @p_9)) :rule bind) (step t4 (cl (not @p_9) (not @p_1) @p_10) :rule equiv_pos2) (step t5 (cl @p_10) :rule th_resolution :premises (axiom0 t3 t4)) (anchor :step t6 :args ((:= (veriT_vr0 A$) veriT_vr2) (:= (veriT_vr1 B$) veriT_vr3))) (step t6.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_11)) :rule refl) (step t6.t2 (cl @p_11) :rule refl) (step t6.t3 (cl (= veriT_vr1 veriT_vr3)) :rule refl) (step t6.t4 (cl (= @p_4 (! (pair$ veriT_vr2 veriT_vr3) :named @p_12))) :rule cong :premises (t6.t2 t6.t3)) (step t6.t5 (cl (= @p_6 (! (fst$ @p_12) :named @p_13))) :rule cong :premises (t6.t4)) (step t6.t6 (cl (= @p_8 (! (= veriT_vr2 @p_13) :named @p_14))) :rule cong :premises (t6.t1 t6.t5)) (step t6 (cl (! (= @p_10 (! (forall ((veriT_vr2 A$) (veriT_vr3 B$)) @p_14) :named @p_16)) :named @p_15)) :rule bind) (step t7 (cl (not @p_15) (not @p_10) @p_16) :rule equiv_pos2) (step t8 (cl @p_16) :rule th_resolution :premises (t5 t6 t7)) (step t9 (cl (! (= @p_17 (! (and @p_18 (! (not @p_19) :named @p_25)) :named @p_21)) :named @p_20)) :rule bool_simplify) (step t10 (cl (! (not @p_20) :named @p_24) (! (not @p_17) :named @p_22) @p_21) :rule equiv_pos2) (step t11 (cl (not @p_22) @p_23) :rule not_not) (step t12 (cl @p_24 @p_23 @p_21) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_21) :rule th_resolution :premises (axiom1 t9 t12)) (step t14 (cl @p_18) :rule and :premises (t13)) (step t15 (cl @p_25) :rule and :premises (t13)) (step t16 (cl (or (! (not @p_16) :named @p_27) (! (= x$ @p_26) :named @p_28))) :rule forall_inst :args ((:= veriT_vr2 x$) (:= veriT_vr3 y$))) (step t17 (cl @p_27 @p_28) :rule or :premises (t16)) (step t18 (cl @p_28) :rule resolution :premises (t17 t8)) (step t19 (cl (not @p_28) (not @p_18) @p_19) :rule eq_transitive) (step t20 (cl) :rule resolution :premises (t19 t14 t15 t18)) 48c92486160c88d9d50f41cf6a3fd33cb769b2b7 67 0 unsat (assume axiom1 (! (forall ((?v0 A$) (?v1 B$)) (! (= ?v0 (! (fst$a (! (pair$a ?v0 ?v1) :named @p_3)) :named @p_5)) :named @p_7)) :named @p_1)) (assume axiom3 (! (forall ((?v0 B$) (?v1 A$)) (! (= ?v1 (! (snd$a (! (pair$ ?v0 ?v1) :named @p_19)) :named @p_21)) :named @p_23)) :named @p_17)) (assume axiom4 (! (not (! (=> (! (and (! (= p1$ (! (pair$a x$ y$) :named @p_47)) :named @p_41) (! (= p2$ (! (pair$ y$ x$) :named @p_46)) :named @p_42)) :named @p_34) (! (= (! (fst$a p1$) :named @p_52) (! (snd$a p2$) :named @p_54)) :named @p_35)) :named @p_39)) :named @p_33)) (anchor :step t4 :args ((:= (?v0 A$) veriT_vr4) (:= (?v1 B$) veriT_vr5))) (step t4.t1 (cl (! (= ?v0 veriT_vr4) :named @p_2)) :rule refl) (step t4.t2 (cl @p_2) :rule refl) (step t4.t3 (cl (= ?v1 veriT_vr5)) :rule refl) (step t4.t4 (cl (= @p_3 (! (pair$a veriT_vr4 veriT_vr5) :named @p_4))) :rule cong :premises (t4.t2 t4.t3)) (step t4.t5 (cl (= @p_5 (! (fst$a @p_4) :named @p_6))) :rule cong :premises (t4.t4)) (step t4.t6 (cl (= @p_7 (! (= veriT_vr4 @p_6) :named @p_8))) :rule cong :premises (t4.t1 t4.t5)) (step t4 (cl (! (= @p_1 (! (forall ((veriT_vr4 A$) (veriT_vr5 B$)) @p_8) :named @p_10)) :named @p_9)) :rule bind) (step t5 (cl (not @p_9) (not @p_1) @p_10) :rule equiv_pos2) (step t6 (cl @p_10) :rule th_resolution :premises (axiom1 t4 t5)) (anchor :step t7 :args ((:= (veriT_vr4 A$) veriT_vr6) (:= (veriT_vr5 B$) veriT_vr7))) (step t7.t1 (cl (! (= veriT_vr4 veriT_vr6) :named @p_11)) :rule refl) (step t7.t2 (cl @p_11) :rule refl) (step t7.t3 (cl (= veriT_vr5 veriT_vr7)) :rule refl) (step t7.t4 (cl (= @p_4 (! (pair$a veriT_vr6 veriT_vr7) :named @p_12))) :rule cong :premises (t7.t2 t7.t3)) (step t7.t5 (cl (= @p_6 (! (fst$a @p_12) :named @p_13))) :rule cong :premises (t7.t4)) (step t7.t6 (cl (= @p_8 (! (= veriT_vr6 @p_13) :named @p_14))) :rule cong :premises (t7.t1 t7.t5)) (step t7 (cl (! (= @p_10 (! (forall ((veriT_vr6 A$) (veriT_vr7 B$)) @p_14) :named @p_16)) :named @p_15)) :rule bind) (step t8 (cl (not @p_15) (not @p_10) @p_16) :rule equiv_pos2) (step t9 (cl @p_16) :rule th_resolution :premises (t6 t7 t8)) (anchor :step t10 :args ((:= (?v0 B$) veriT_vr12) (:= (?v1 A$) veriT_vr13))) (step t10.t1 (cl (! (= ?v1 veriT_vr13) :named @p_18)) :rule refl) (step t10.t2 (cl (= ?v0 veriT_vr12)) :rule refl) (step t10.t3 (cl @p_18) :rule refl) (step t10.t4 (cl (= @p_19 (! (pair$ veriT_vr12 veriT_vr13) :named @p_20))) :rule cong :premises (t10.t2 t10.t3)) (step t10.t5 (cl (= @p_21 (! (snd$a @p_20) :named @p_22))) :rule cong :premises (t10.t4)) (step t10.t6 (cl (= @p_23 (! (= veriT_vr13 @p_22) :named @p_24))) :rule cong :premises (t10.t1 t10.t5)) (step t10 (cl (! (= @p_17 (! (forall ((veriT_vr12 B$) (veriT_vr13 A$)) @p_24) :named @p_26)) :named @p_25)) :rule bind) (step t11 (cl (not @p_25) (not @p_17) @p_26) :rule equiv_pos2) (step t12 (cl @p_26) :rule th_resolution :premises (axiom3 t10 t11)) (anchor :step t13 :args ((:= (veriT_vr12 B$) veriT_vr14) (:= (veriT_vr13 A$) veriT_vr15))) (step t13.t1 (cl (! (= veriT_vr13 veriT_vr15) :named @p_27)) :rule refl) (step t13.t2 (cl (= veriT_vr12 veriT_vr14)) :rule refl) (step t13.t3 (cl @p_27) :rule refl) (step t13.t4 (cl (= @p_20 (! (pair$ veriT_vr14 veriT_vr15) :named @p_28))) :rule cong :premises (t13.t2 t13.t3)) (step t13.t5 (cl (= @p_22 (! (snd$a @p_28) :named @p_29))) :rule cong :premises (t13.t4)) (step t13.t6 (cl (= @p_24 (! (= veriT_vr15 @p_29) :named @p_30))) :rule cong :premises (t13.t1 t13.t5)) (step t13 (cl (! (= @p_26 (! (forall ((veriT_vr14 B$) (veriT_vr15 A$)) @p_30) :named @p_32)) :named @p_31)) :rule bind) (step t14 (cl (not @p_31) (not @p_26) @p_32) :rule equiv_pos2) (step t15 (cl @p_32) :rule th_resolution :premises (t12 t13 t14)) (step t16 (cl (! (= @p_33 (! (and @p_34 (! (not @p_35) :named @p_43)) :named @p_37)) :named @p_36)) :rule bool_simplify) (step t17 (cl (! (not @p_36) :named @p_40) (! (not @p_33) :named @p_38) @p_37) :rule equiv_pos2) (step t18 (cl (not @p_38) @p_39) :rule not_not) (step t19 (cl @p_40 @p_39 @p_37) :rule th_resolution :premises (t18 t17)) (step t20 (cl @p_37) :rule th_resolution :premises (axiom4 t16 t19)) (step t21 (cl (! (= @p_37 (! (and @p_41 @p_42 @p_43) :named @p_45)) :named @p_44)) :rule ac_simp) (step t22 (cl (not @p_44) (not @p_37) @p_45) :rule equiv_pos2) (step t23 (cl @p_45) :rule th_resolution :premises (t20 t21 t22)) (step t24 (cl @p_41) :rule and :premises (t23)) (step t25 (cl @p_42) :rule and :premises (t23)) (step t26 (cl @p_43) :rule and :premises (t23)) (step t27 (cl (or (! (not @p_32) :named @p_48) (! (= x$ (! (snd$a @p_46) :named @p_55)) :named @p_49))) :rule forall_inst :args ((:= veriT_vr14 y$) (:= veriT_vr15 x$))) (step t28 (cl (or (! (not @p_16) :named @p_50) (! (= x$ (! (fst$a @p_47) :named @p_53)) :named @p_51))) :rule forall_inst :args ((:= veriT_vr6 x$) (:= veriT_vr7 y$))) (step t29 (cl @p_48 @p_49) :rule or :premises (t27)) (step t30 (cl @p_49) :rule resolution :premises (t29 t15)) (step t31 (cl @p_50 @p_51) :rule or :premises (t28)) (step t32 (cl @p_51) :rule resolution :premises (t31 t9)) (step t33 (cl (! (not (! (= @p_52 @p_53) :named @p_61)) :named @p_57) (! (not @p_51) :named @p_58) (! (not @p_49) :named @p_59) (not (! (= @p_54 @p_55) :named @p_56)) @p_35) :rule eq_transitive) (step t34 (cl (! (not @p_42) :named @p_60) @p_56) :rule eq_congruent) (step t35 (cl @p_57 @p_58 @p_59 @p_35 @p_60) :rule th_resolution :premises (t33 t34)) (step t36 (cl (! (not @p_41) :named @p_62) @p_61) :rule eq_congruent) (step t37 (cl @p_58 @p_59 @p_35 @p_60 @p_62) :rule th_resolution :premises (t35 t36)) (step t38 (cl) :rule resolution :premises (t37 t24 t25 t26 t30 t32)) d98294078e51b7d929ca2e2c002b0b435565488d 25 0 unsat (assume axiom0 (! (not (! (or (! (= (! (f$ g$ x$) :named @p_1) (! (and (! (fun_app$ g$ x$) :named @p_2) true) :named @p_13)) :named @p_4) (or (! (= @p_1 true) :named @p_5) (! (= @p_2 true) :named @p_6))) :named @p_3)) :named @p_7)) (step t2 (cl (= @p_3 (! (or @p_4 @p_5 @p_6) :named @p_8))) :rule ac_simp) (step t3 (cl (! (= @p_7 (! (not @p_8) :named @p_10)) :named @p_9)) :rule cong :premises (t2)) (step t4 (cl (! (not @p_9) :named @p_12) (! (not @p_7) :named @p_11) @p_10) :rule equiv_pos2) (step t5 (cl (not @p_11) @p_3) :rule not_not) (step t6 (cl @p_12 @p_3 @p_10) :rule th_resolution :premises (t5 t4)) (step t7 (cl @p_10) :rule th_resolution :premises (axiom0 t3 t6)) (step t8 (cl (= @p_13 (! (and @p_2) :named @p_14))) :rule and_simplify) (step t9 (cl (= @p_14 @p_2)) :rule and_simplify) (step t10 (cl (= @p_13 @p_2)) :rule trans :premises (t8 t9)) (step t11 (cl (= @p_4 (! (= @p_1 @p_2) :named @p_15))) :rule cong :premises (t10)) (step t12 (cl (= @p_5 @p_1)) :rule equiv_simplify) (step t13 (cl (= @p_6 @p_2)) :rule equiv_simplify) (step t14 (cl (= @p_8 (! (or @p_15 @p_1 @p_2) :named @p_16))) :rule cong :premises (t11 t12 t13)) (step t15 (cl (! (= @p_10 (! (not @p_16) :named @p_18)) :named @p_17)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_17) :named @p_20) (! (not @p_10) :named @p_19) @p_18) :rule equiv_pos2) (step t17 (cl (not @p_19) @p_8) :rule not_not) (step t18 (cl @p_20 @p_8 @p_18) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_18) :rule th_resolution :premises (t7 t15 t18)) (step t20 (cl (not @p_15)) :rule not_or :premises (t19)) (step t21 (cl @p_1 @p_2) :rule not_equiv1 :premises (t20)) (step t22 (cl (not @p_1)) :rule not_or :premises (t19)) (step t23 (cl (not @p_2)) :rule not_or :premises (t19)) (step t24 (cl) :rule resolution :premises (t21 t22 t23)) a6d9c8f29a4df3b9eaa40764e38a8a41a8961f01 97 0 unsat (assume axiom1 (! (forall ((?v0 A_b_fun$) (?v1 A$) (?v2 B$) (?v3 A$)) (! (= (! (fun_app$ (! (fun_upd$ ?v0 ?v1 ?v2) :named @p_2) ?v3) :named @p_4) (! (ite (! (= ?v1 ?v3) :named @p_8) ?v2 (! (fun_app$ ?v0 ?v3) :named @p_12)) :named @p_14)) :named @p_16)) :named @p_1)) (assume axiom2 (! (not (! (=> (! (and (! (not (! (= i$ i1$) :named @p_62)) :named @p_40) (! (not (! (= i$ i2$) :named @p_46)) :named @p_41)) :named @p_33) (! (= (! (fun_app$ (fun_upd$ (! (fun_upd$ f$ i1$ v1$) :named @p_47) i2$ v2$) i$) :named @p_45) (! (fun_app$ f$ i$) :named @p_63)) :named @p_34)) :named @p_38)) :named @p_32)) (anchor :step t3 :args ((:= (?v0 A_b_fun$) veriT_vr6) (:= (?v1 A$) veriT_vr7) (:= (?v2 B$) veriT_vr8) (:= (?v3 A$) veriT_vr9))) (step t3.t1 (cl (! (= ?v0 veriT_vr6) :named @p_11)) :rule refl) (step t3.t2 (cl (! (= ?v1 veriT_vr7) :named @p_6)) :rule refl) (step t3.t3 (cl (! (= ?v2 veriT_vr8) :named @p_10)) :rule refl) (step t3.t4 (cl (= @p_2 (! (fun_upd$ veriT_vr6 veriT_vr7 veriT_vr8) :named @p_3))) :rule cong :premises (t3.t1 t3.t2 t3.t3)) (step t3.t5 (cl (! (= ?v3 veriT_vr9) :named @p_7)) :rule refl) (step t3.t6 (cl (= @p_4 (! (fun_app$ @p_3 veriT_vr9) :named @p_5))) :rule cong :premises (t3.t4 t3.t5)) (step t3.t7 (cl @p_6) :rule refl) (step t3.t8 (cl @p_7) :rule refl) (step t3.t9 (cl (= @p_8 (! (= veriT_vr7 veriT_vr9) :named @p_9))) :rule cong :premises (t3.t7 t3.t8)) (step t3.t10 (cl @p_10) :rule refl) (step t3.t11 (cl @p_11) :rule refl) (step t3.t12 (cl @p_7) :rule refl) (step t3.t13 (cl (= @p_12 (! (fun_app$ veriT_vr6 veriT_vr9) :named @p_13))) :rule cong :premises (t3.t11 t3.t12)) (step t3.t14 (cl (= @p_14 (! (ite @p_9 veriT_vr8 @p_13) :named @p_15))) :rule cong :premises (t3.t9 t3.t10 t3.t13)) (step t3.t15 (cl (= @p_16 (! (= @p_5 @p_15) :named @p_17))) :rule cong :premises (t3.t6 t3.t14)) (step t3 (cl (! (= @p_1 (! (forall ((veriT_vr6 A_b_fun$) (veriT_vr7 A$) (veriT_vr8 B$) (veriT_vr9 A$)) @p_17) :named @p_19)) :named @p_18)) :rule bind) (step t4 (cl (not @p_18) (not @p_1) @p_19) :rule equiv_pos2) (step t5 (cl @p_19) :rule th_resolution :premises (axiom1 t3 t4)) (anchor :step t6 :args ((:= (veriT_vr6 A_b_fun$) veriT_vr10) (:= (veriT_vr7 A$) veriT_vr11) (:= (veriT_vr8 B$) veriT_vr12) (:= (veriT_vr9 A$) veriT_vr13))) (step t6.t1 (cl (! (= veriT_vr6 veriT_vr10) :named @p_26)) :rule refl) (step t6.t2 (cl (! (= veriT_vr7 veriT_vr11) :named @p_22)) :rule refl) (step t6.t3 (cl (! (= veriT_vr8 veriT_vr12) :named @p_25)) :rule refl) (step t6.t4 (cl (= @p_3 (! (fun_upd$ veriT_vr10 veriT_vr11 veriT_vr12) :named @p_20))) :rule cong :premises (t6.t1 t6.t2 t6.t3)) (step t6.t5 (cl (! (= veriT_vr9 veriT_vr13) :named @p_23)) :rule refl) (step t6.t6 (cl (= @p_5 (! (fun_app$ @p_20 veriT_vr13) :named @p_21))) :rule cong :premises (t6.t4 t6.t5)) (step t6.t7 (cl @p_22) :rule refl) (step t6.t8 (cl @p_23) :rule refl) (step t6.t9 (cl (= @p_9 (! (= veriT_vr11 veriT_vr13) :named @p_24))) :rule cong :premises (t6.t7 t6.t8)) (step t6.t10 (cl @p_25) :rule refl) (step t6.t11 (cl @p_26) :rule refl) (step t6.t12 (cl @p_23) :rule refl) (step t6.t13 (cl (= @p_13 (! (fun_app$ veriT_vr10 veriT_vr13) :named @p_27))) :rule cong :premises (t6.t11 t6.t12)) (step t6.t14 (cl (= @p_15 (! (ite @p_24 veriT_vr12 @p_27) :named @p_28))) :rule cong :premises (t6.t9 t6.t10 t6.t13)) (step t6.t15 (cl (= @p_17 (! (= @p_21 @p_28) :named @p_29))) :rule cong :premises (t6.t6 t6.t14)) (step t6 (cl (! (= @p_19 (! (forall ((veriT_vr10 A_b_fun$) (veriT_vr11 A$) (veriT_vr12 B$) (veriT_vr13 A$)) @p_29) :named @p_31)) :named @p_30)) :rule bind) (step t7 (cl (not @p_30) (not @p_19) @p_31) :rule equiv_pos2) (step t8 (cl @p_31) :rule th_resolution :premises (t5 t6 t7)) (step t9 (cl (! (= @p_32 (! (and @p_33 (! (not @p_34) :named @p_42)) :named @p_36)) :named @p_35)) :rule bool_simplify) (step t10 (cl (! (not @p_35) :named @p_39) (! (not @p_32) :named @p_37) @p_36) :rule equiv_pos2) (step t11 (cl (not @p_37) @p_38) :rule not_not) (step t12 (cl @p_39 @p_38 @p_36) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_36) :rule th_resolution :premises (axiom2 t9 t12)) (step t14 (cl (! (= @p_36 (! (and @p_40 @p_41 @p_42) :named @p_44)) :named @p_43)) :rule ac_simp) (step t15 (cl (not @p_43) (not @p_36) @p_44) :rule equiv_pos2) (step t16 (cl @p_44) :rule th_resolution :premises (t13 t14 t15)) (step t17 (cl @p_40) :rule and :premises (t16)) (step t18 (cl @p_41) :rule and :premises (t16)) (step t19 (cl @p_42) :rule and :premises (t16)) (step t20 (cl (or (! (not @p_31) :named @p_54) (! (= @p_45 (! (ite @p_46 v2$ (! (fun_app$ @p_47 i$) :named @p_50)) :named @p_49)) :named @p_48))) :rule forall_inst :args ((:= veriT_vr10 @p_47) (:= veriT_vr11 i2$) (:= veriT_vr12 v2$) (:= veriT_vr13 i$))) (anchor :step t21) (assume t21.h1 @p_48) (step t21.t2 (cl (! (= @p_48 (! (and (! (= @p_45 @p_49) :named @p_58) (! (ite @p_46 (= v2$ @p_49) (! (= @p_50 @p_49) :named @p_60)) :named @p_59)) :named @p_51)) :named @p_52)) :rule ite_intro) (step t21.t3 (cl (not @p_52) (! (not @p_48) :named @p_53) @p_51) :rule equiv_pos2) (step t21.t4 (cl @p_51) :rule th_resolution :premises (t21.h1 t21.t2 t21.t3)) (step t21 (cl @p_53 @p_51) :rule subproof :discharge (h1)) (step t22 (cl @p_54 @p_48) :rule or :premises (t20)) (step t23 (cl (! (or @p_54 @p_51) :named @p_56) (! (not @p_54) :named @p_55)) :rule or_neg) (step t24 (cl (not @p_55) @p_31) :rule not_not) (step t25 (cl @p_56 @p_31) :rule th_resolution :premises (t24 t23)) (step t26 (cl @p_56 (! (not @p_51) :named @p_57)) :rule or_neg) (step t27 (cl @p_56) :rule th_resolution :premises (t22 t21 t25 t26)) (step t28 (cl @p_57 @p_58) :rule and_pos) (step t29 (cl (! (not @p_59) :named @p_61) @p_46 @p_60) :rule ite_pos1) (step t30 (cl @p_57 @p_59) :rule and_pos) (step t31 (cl @p_54 @p_51) :rule or :premises (t27)) (step t32 (cl @p_61 @p_60) :rule resolution :premises (t29 t18)) (step t33 (cl @p_51) :rule resolution :premises (t31 t8)) (step t34 (cl @p_58) :rule resolution :premises (t28 t33)) (step t35 (cl @p_59) :rule resolution :premises (t30 t33)) (step t36 (cl @p_60) :rule resolution :premises (t32 t35)) (step t37 (cl (or @p_54 (! (= @p_50 (! (ite @p_62 v1$ @p_63) :named @p_65)) :named @p_64))) :rule forall_inst :args ((:= veriT_vr10 f$) (:= veriT_vr11 i1$) (:= veriT_vr12 v1$) (:= veriT_vr13 i$))) (anchor :step t38) (assume t38.h1 @p_64) (step t38.t2 (cl (! (= @p_64 (! (and (! (= @p_50 @p_65) :named @p_71) (! (ite @p_62 (= v1$ @p_65) (! (= @p_63 @p_65) :named @p_73)) :named @p_72)) :named @p_66)) :named @p_67)) :rule ite_intro) (step t38.t3 (cl (not @p_67) (! (not @p_64) :named @p_68) @p_66) :rule equiv_pos2) (step t38.t4 (cl @p_66) :rule th_resolution :premises (t38.h1 t38.t2 t38.t3)) (step t38 (cl @p_68 @p_66) :rule subproof :discharge (h1)) (step t39 (cl @p_54 @p_64) :rule or :premises (t37)) (step t40 (cl (! (or @p_54 @p_66) :named @p_69) @p_55) :rule or_neg) (step t41 (cl @p_69 @p_31) :rule th_resolution :premises (t24 t40)) (step t42 (cl @p_69 (! (not @p_66) :named @p_70)) :rule or_neg) (step t43 (cl @p_69) :rule th_resolution :premises (t39 t38 t41 t42)) (step t44 (cl @p_70 @p_71) :rule and_pos) (step t45 (cl (! (not @p_72) :named @p_74) @p_62 @p_73) :rule ite_pos1) (step t46 (cl @p_70 @p_72) :rule and_pos) (step t47 (cl @p_54 @p_66) :rule or :premises (t43)) (step t48 (cl @p_74 @p_73) :rule resolution :premises (t45 t17)) (step t49 (cl @p_66) :rule resolution :premises (t47 t8)) (step t50 (cl @p_71) :rule resolution :premises (t44 t49)) (step t51 (cl @p_72) :rule resolution :premises (t46 t49)) (step t52 (cl @p_73) :rule resolution :premises (t48 t51)) (step t53 (cl (not @p_58) (not @p_60) (not @p_71) (not @p_73) @p_34) :rule eq_transitive) (step t54 (cl) :rule resolution :premises (t53 t19 t34 t36 t50 t52)) 6c882ad63eb28539c511430af509d5d6577b2638 38 0 unsat (assume axiom0 (! (forall ((?v0 A$)) (! (= ?v0 (! (id$ ?v0) :named @p_4)) :named @p_6)) :named @p_2)) (assume axiom1 (forall ((?v0 Bool)) (= (id$a ?v0) ?v0))) (assume axiom2 (! (not (! (and (! (= x$ (id$ x$)) :named @p_23) (! (= (! (id$a true) :named @p_17) true) :named @p_1)) :named @p_22)) :named @p_24)) (step t4 (cl (! (and (! (= (! (id$a false) :named @p_16) false) :named @p_15) @p_1) :named @p_18)) :rule bfun_elim :premises (axiom1)) (anchor :step t5 :args ((:= (?v0 A$) veriT_vr0))) (step t5.t1 (cl (! (= ?v0 veriT_vr0) :named @p_3)) :rule refl) (step t5.t2 (cl @p_3) :rule refl) (step t5.t3 (cl (= @p_4 (! (id$ veriT_vr0) :named @p_5))) :rule cong :premises (t5.t2)) (step t5.t4 (cl (= @p_6 (! (= veriT_vr0 @p_5) :named @p_7))) :rule cong :premises (t5.t1 t5.t3)) (step t5 (cl (! (= @p_2 (! (forall ((veriT_vr0 A$)) @p_7) :named @p_9)) :named @p_8)) :rule bind) (step t6 (cl (not @p_8) (not @p_2) @p_9) :rule equiv_pos2) (step t7 (cl @p_9) :rule th_resolution :premises (axiom0 t5 t6)) (anchor :step t8 :args ((:= (veriT_vr0 A$) veriT_vr1))) (step t8.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_10)) :rule refl) (step t8.t2 (cl @p_10) :rule refl) (step t8.t3 (cl (= @p_5 (! (id$ veriT_vr1) :named @p_11))) :rule cong :premises (t8.t2)) (step t8.t4 (cl (= @p_7 (! (= veriT_vr1 @p_11) :named @p_12))) :rule cong :premises (t8.t1 t8.t3)) (step t8 (cl (! (= @p_9 (! (forall ((veriT_vr1 A$)) @p_12) :named @p_14)) :named @p_13)) :rule bind) (step t9 (cl (not @p_13) (not @p_9) @p_14) :rule equiv_pos2) (step t10 (cl @p_14) :rule th_resolution :premises (t7 t8 t9)) (step t11 (cl (= @p_15 (! (not @p_16) :named @p_19))) :rule equiv_simplify) (step t12 (cl (= @p_1 @p_17)) :rule equiv_simplify) (step t13 (cl (! (= @p_18 (! (and @p_19 @p_17) :named @p_21)) :named @p_20)) :rule cong :premises (t11 t12)) (step t14 (cl (not @p_20) (not @p_18) @p_21) :rule equiv_pos2) (step t15 (cl @p_21) :rule th_resolution :premises (t4 t13 t14)) (step t16 (cl (= @p_22 (! (and @p_23 @p_17) :named @p_25))) :rule cong :premises (t12)) (step t17 (cl (! (= @p_24 (! (not @p_25) :named @p_27)) :named @p_26)) :rule cong :premises (t16)) (step t18 (cl (! (not @p_26) :named @p_29) (! (not @p_24) :named @p_28) @p_27) :rule equiv_pos2) (step t19 (cl (not @p_28) @p_22) :rule not_not) (step t20 (cl @p_29 @p_22 @p_27) :rule th_resolution :premises (t19 t18)) (step t21 (cl @p_27) :rule th_resolution :premises (axiom2 t17 t20)) (step t22 (cl @p_17) :rule and :premises (t15)) (step t23 (cl (! (not @p_23) :named @p_30) (not @p_17)) :rule not_and :premises (t21)) (step t24 (cl @p_30) :rule resolution :premises (t23 t22)) (step t25 (cl (or (! (not @p_14) :named @p_31) @p_23)) :rule forall_inst :args ((:= veriT_vr1 x$))) (step t26 (cl @p_31 @p_23) :rule or :premises (t25)) (step t27 (cl) :rule resolution :premises (t26 t10 t24)) 395426c1e77b0b2b43dd37f4fdc5b1794b758773 27 0 unsat (assume axiom0 (not (=> (f$ (! (exists ((?v0 A$)) (! (g$ ?v0) :named @p_2)) :named @p_1)) true))) (step t2 (cl (! (not (! (=> (! (ite @p_1 (! (f$ true) :named @p_6) (! (f$ false) :named @p_7)) :named @p_4) true) :named @p_8)) :named @p_10)) :rule bfun_elim :premises (axiom0)) (anchor :step t3 :args ((:= (?v0 A$) veriT_vr0))) (step t3.t1 (cl (= ?v0 veriT_vr0)) :rule refl) (step t3.t2 (cl (= @p_2 (! (g$ veriT_vr0) :named @p_3))) :rule cong :premises (t3.t1)) (step t3 (cl (= @p_1 (! (exists ((veriT_vr0 A$)) @p_3) :named @p_5))) :rule bind) (step t4 (cl (= @p_4 (! (ite @p_5 @p_6 @p_7) :named @p_9))) :rule cong :premises (t3)) (step t5 (cl (= @p_8 (! (=> @p_9 true) :named @p_11))) :rule cong :premises (t4)) (step t6 (cl (! (= @p_10 (! (not @p_11) :named @p_13)) :named @p_12)) :rule cong :premises (t5)) (step t7 (cl (! (not @p_12) :named @p_15) (! (not @p_10) :named @p_14) @p_13) :rule equiv_pos2) (step t8 (cl (not @p_14) @p_8) :rule not_not) (step t9 (cl @p_15 @p_8 @p_13) :rule th_resolution :premises (t8 t7)) (step t10 (cl @p_13) :rule th_resolution :premises (t2 t6 t9)) (step t11 (cl (! (= @p_13 (! (and @p_9 (! (not true) :named @p_20)) :named @p_17)) :named @p_16)) :rule bool_simplify) (step t12 (cl (! (not @p_16) :named @p_19) (! (not @p_13) :named @p_18) @p_17) :rule equiv_pos2) (step t13 (cl (not @p_18) @p_11) :rule not_not) (step t14 (cl @p_19 @p_11 @p_17) :rule th_resolution :premises (t13 t12)) (step t15 (cl @p_17) :rule th_resolution :premises (t10 t11 t14)) (step t16 (cl (= @p_20 false)) :rule not_simplify) (step t17 (cl (= @p_17 (! (and @p_9 false) :named @p_21))) :rule cong :premises (t16)) (step t18 (cl (= @p_21 false)) :rule and_simplify) (step t19 (cl (! (= @p_17 false) :named @p_22)) :rule trans :premises (t17 t18)) (step t20 (cl (not @p_22) (not @p_17) false) :rule equiv_pos2) (step t21 (cl false) :rule th_resolution :premises (t15 t19 t20)) (step t22 (cl (not false)) :rule false) (step t23 (cl) :rule resolution :premises (t21 t22)) 2bbe312ac0bf24de2abf59ed85b8b3b02d04d4eb 60 0 unsat (assume axiom0 (! (forall ((?v0 Int) (?v1 Int)) (! (= (! (fun_app$ (! (fun_app$a uu$ ?v0) :named @p_2) ?v1) :named @p_4) (! (<= ?v0 ?v1) :named @p_8)) :named @p_10)) :named @p_1)) (assume axiom1 (! (not (! (=> (! (= uu$ le$) :named @p_23) (! (fun_app$ (! (fun_app$a le$ 3) :named @p_40) 42) :named @p_24)) :named @p_28)) :named @p_22)) (anchor :step t3 :args ((:= (?v0 Int) veriT_vr0) (:= (?v1 Int) veriT_vr1))) (step t3.t1 (cl (! (= ?v0 veriT_vr0) :named @p_6)) :rule refl) (step t3.t2 (cl (= @p_2 (! (fun_app$a uu$ veriT_vr0) :named @p_3))) :rule cong :premises (t3.t1)) (step t3.t3 (cl (! (= ?v1 veriT_vr1) :named @p_7)) :rule refl) (step t3.t4 (cl (= @p_4 (! (fun_app$ @p_3 veriT_vr1) :named @p_5))) :rule cong :premises (t3.t2 t3.t3)) (step t3.t5 (cl @p_6) :rule refl) (step t3.t6 (cl @p_7) :rule refl) (step t3.t7 (cl (= @p_8 (! (<= veriT_vr0 veriT_vr1) :named @p_9))) :rule cong :premises (t3.t5 t3.t6)) (step t3.t8 (cl (= @p_10 (! (= @p_5 @p_9) :named @p_11))) :rule cong :premises (t3.t4 t3.t7)) (step t3 (cl (! (= @p_1 (! (forall ((veriT_vr0 Int) (veriT_vr1 Int)) @p_11) :named @p_13)) :named @p_12)) :rule bind) (step t4 (cl (not @p_12) (not @p_1) @p_13) :rule equiv_pos2) (step t5 (cl @p_13) :rule th_resolution :premises (axiom0 t3 t4)) (anchor :step t6 :args ((:= (veriT_vr0 Int) veriT_vr2) (:= (veriT_vr1 Int) veriT_vr3))) (step t6.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_16)) :rule refl) (step t6.t2 (cl (= @p_3 (! (fun_app$a uu$ veriT_vr2) :named @p_14))) :rule cong :premises (t6.t1)) (step t6.t3 (cl (! (= veriT_vr1 veriT_vr3) :named @p_17)) :rule refl) (step t6.t4 (cl (= @p_5 (! (fun_app$ @p_14 veriT_vr3) :named @p_15))) :rule cong :premises (t6.t2 t6.t3)) (step t6.t5 (cl @p_16) :rule refl) (step t6.t6 (cl @p_17) :rule refl) (step t6.t7 (cl (= @p_9 (! (<= veriT_vr2 veriT_vr3) :named @p_18))) :rule cong :premises (t6.t5 t6.t6)) (step t6.t8 (cl (= @p_11 (! (= @p_15 @p_18) :named @p_19))) :rule cong :premises (t6.t4 t6.t7)) (step t6 (cl (! (= @p_13 (! (forall ((veriT_vr2 Int) (veriT_vr3 Int)) @p_19) :named @p_21)) :named @p_20)) :rule bind) (step t7 (cl (not @p_20) (not @p_13) @p_21) :rule equiv_pos2) (step t8 (cl @p_21) :rule th_resolution :premises (t5 t6 t7)) (step t9 (cl (! (= @p_22 (! (and @p_23 (! (not @p_24) :named @p_30)) :named @p_26)) :named @p_25)) :rule bool_simplify) (step t10 (cl (! (not @p_25) :named @p_29) (! (not @p_22) :named @p_27) @p_26) :rule equiv_pos2) (step t11 (cl (not @p_27) @p_28) :rule not_not) (step t12 (cl @p_29 @p_28 @p_26) :rule th_resolution :premises (t11 t10)) (step t13 (cl @p_26) :rule th_resolution :premises (axiom1 t9 t12)) (step t14 (cl @p_23) :rule and :premises (t13)) (step t15 (cl @p_30) :rule and :premises (t13)) (step t16 (cl (or (! (not @p_21) :named @p_37) (! (= (! (fun_app$ (! (fun_app$a uu$ 3) :named @p_41) 42) :named @p_32) (! (<= 3 42) :named @p_33)) :named @p_31))) :rule forall_inst :args ((:= veriT_vr2 3) (:= veriT_vr3 42))) (anchor :step t17) (assume t17.h1 @p_31) (step t17.t2 (cl (= @p_33 true)) :rule comp_simplify) (step t17.t3 (cl (= @p_31 (! (= @p_32 true) :named @p_34))) :rule cong :premises (t17.t2)) (step t17.t4 (cl (= @p_34 @p_32)) :rule equiv_simplify) (step t17.t5 (cl (! (= @p_31 @p_32) :named @p_35)) :rule trans :premises (t17.t3 t17.t4)) (step t17.t6 (cl (not @p_35) (! (not @p_31) :named @p_36) @p_32) :rule equiv_pos2) (step t17.t7 (cl @p_32) :rule th_resolution :premises (t17.h1 t17.t5 t17.t6)) (step t17 (cl @p_36 @p_32) :rule subproof :discharge (h1)) (step t18 (cl @p_37 @p_31) :rule or :premises (t16)) (step t19 (cl (! (or @p_37 @p_32) :named @p_39) (! (not @p_37) :named @p_38)) :rule or_neg) (step t20 (cl (not @p_38) @p_21) :rule not_not) (step t21 (cl @p_39 @p_21) :rule th_resolution :premises (t20 t19)) (step t22 (cl @p_39 (! (not @p_32) :named @p_42)) :rule or_neg) (step t23 (cl @p_39) :rule th_resolution :premises (t18 t17 t21 t22)) (step t24 (cl @p_37 @p_32) :rule or :premises (t23)) (step t25 (cl @p_32) :rule resolution :premises (t24 t8)) (step t26 (cl (not (! (= @p_40 @p_41) :named @p_43)) (! (not (! (= 42 42) :named @p_47)) :named @p_46) @p_42 @p_24) :rule eq_congruent_pred) (step t27 (cl (! (not @p_23) :named @p_45) (not (! (= 3 3) :named @p_44)) @p_43) :rule eq_congruent) (step t28 (cl @p_44) :rule eq_reflexive) (step t29 (cl @p_45 @p_43) :rule th_resolution :premises (t27 t28)) (step t30 (cl @p_46 @p_42 @p_24 @p_45) :rule th_resolution :premises (t26 t29)) (step t31 (cl @p_47) :rule eq_reflexive) (step t32 (cl @p_42 @p_24 @p_45) :rule th_resolution :premises (t30 t31)) (step t33 (cl) :rule resolution :premises (t32 t14 t15 t25)) cff62aa4cee92583faa1651c31b596d6853191d1 27 0 unsat (assume axiom0 (not (=> (f$ (! (forall ((?v0 A$)) (! (g$ ?v0) :named @p_2)) :named @p_1)) true))) (step t2 (cl (! (not (! (=> (! (ite @p_1 (! (f$ true) :named @p_6) (! (f$ false) :named @p_7)) :named @p_4) true) :named @p_8)) :named @p_10)) :rule bfun_elim :premises (axiom0)) (anchor :step t3 :args ((:= (?v0 A$) veriT_vr0))) (step t3.t1 (cl (= ?v0 veriT_vr0)) :rule refl) (step t3.t2 (cl (= @p_2 (! (g$ veriT_vr0) :named @p_3))) :rule cong :premises (t3.t1)) (step t3 (cl (= @p_1 (! (forall ((veriT_vr0 A$)) @p_3) :named @p_5))) :rule bind) (step t4 (cl (= @p_4 (! (ite @p_5 @p_6 @p_7) :named @p_9))) :rule cong :premises (t3)) (step t5 (cl (= @p_8 (! (=> @p_9 true) :named @p_11))) :rule cong :premises (t4)) (step t6 (cl (! (= @p_10 (! (not @p_11) :named @p_13)) :named @p_12)) :rule cong :premises (t5)) (step t7 (cl (! (not @p_12) :named @p_15) (! (not @p_10) :named @p_14) @p_13) :rule equiv_pos2) (step t8 (cl (not @p_14) @p_8) :rule not_not) (step t9 (cl @p_15 @p_8 @p_13) :rule th_resolution :premises (t8 t7)) (step t10 (cl @p_13) :rule th_resolution :premises (t2 t6 t9)) (step t11 (cl (! (= @p_13 (! (and @p_9 (! (not true) :named @p_20)) :named @p_17)) :named @p_16)) :rule bool_simplify) (step t12 (cl (! (not @p_16) :named @p_19) (! (not @p_13) :named @p_18) @p_17) :rule equiv_pos2) (step t13 (cl (not @p_18) @p_11) :rule not_not) (step t14 (cl @p_19 @p_11 @p_17) :rule th_resolution :premises (t13 t12)) (step t15 (cl @p_17) :rule th_resolution :premises (t10 t11 t14)) (step t16 (cl (= @p_20 false)) :rule not_simplify) (step t17 (cl (= @p_17 (! (and @p_9 false) :named @p_21))) :rule cong :premises (t16)) (step t18 (cl (= @p_21 false)) :rule and_simplify) (step t19 (cl (! (= @p_17 false) :named @p_22)) :rule trans :premises (t17 t18)) (step t20 (cl (not @p_22) (not @p_17) false) :rule equiv_pos2) (step t21 (cl false) :rule th_resolution :premises (t15 t19 t20)) (step t22 (cl (not false)) :rule false) (step t23 (cl) :rule resolution :premises (t21 t22)) b98ce8490503b759ff8b992327f85a4088d1bb79 125 0 unsat (assume axiom0 (! (forall ((?v0 Int)) (! (= (! (fun_app$ uu$ ?v0) :named @p_2) (! (+ ?v0 1) :named @p_5)) :named @p_7)) :named @p_1)) (assume axiom1 (! (forall ((?v0 Int_int_fun$)) (! (= nil$ (! (map$ ?v0 nil$) :named @p_22)) :named @p_24)) :named @p_21)) (assume axiom2 (! (forall ((?v0 Int_int_fun$) (?v1 Int) (?v2 Int_list$)) (! (= (! (map$ ?v0 (! (cons$ ?v1 ?v2) :named @p_33)) :named @p_35) (! (cons$ (! (fun_app$ ?v0 ?v1) :named @p_39) (! (map$ ?v0 ?v2) :named @p_42)) :named @p_44)) :named @p_46)) :named @p_32)) (assume axiom3 (not (! (= (! (map$ uu$ (cons$ 0 (! (cons$ 1 nil$) :named @p_62))) :named @p_61) (! (cons$ 1 (! (cons$ 2 nil$) :named @p_90)) :named @p_86)) :named @p_88))) (anchor :step t5 :args ((:= (?v0 Int) veriT_vr0))) (step t5.t1 (cl (! (= ?v0 veriT_vr0) :named @p_4)) :rule refl) (step t5.t2 (cl (= @p_2 (! (fun_app$ uu$ veriT_vr0) :named @p_3))) :rule cong :premises (t5.t1)) (step t5.t3 (cl @p_4) :rule refl) (step t5.t4 (cl (= @p_5 (! (+ veriT_vr0 1) :named @p_6))) :rule cong :premises (t5.t3)) (step t5.t5 (cl (= @p_7 (! (= @p_3 @p_6) :named @p_8))) :rule cong :premises (t5.t2 t5.t4)) (step t5 (cl (! (= @p_1 (! (forall ((veriT_vr0 Int)) @p_8) :named @p_10)) :named @p_9)) :rule bind) (step t6 (cl (not @p_9) (not @p_1) @p_10) :rule equiv_pos2) (step t7 (cl @p_10) :rule th_resolution :premises (axiom0 t5 t6)) (anchor :step t8 :args ((veriT_vr0 Int))) (step t8.t1 (cl (= @p_6 (! (+ 1 veriT_vr0) :named @p_11))) :rule sum_simplify) (step t8.t2 (cl (= @p_8 (! (= @p_3 @p_11) :named @p_12))) :rule cong :premises (t8.t1)) (step t8 (cl (! (= @p_10 (! (forall ((veriT_vr0 Int)) @p_12) :named @p_14)) :named @p_13)) :rule bind) (step t9 (cl (not @p_13) (not @p_10) @p_14) :rule equiv_pos2) (step t10 (cl @p_14) :rule th_resolution :premises (t7 t8 t9)) (anchor :step t11 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t11.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_16)) :rule refl) (step t11.t2 (cl (= @p_3 (! (fun_app$ uu$ veriT_vr1) :named @p_15))) :rule cong :premises (t11.t1)) (step t11.t3 (cl @p_16) :rule refl) (step t11.t4 (cl (= @p_11 (! (+ 1 veriT_vr1) :named @p_17))) :rule cong :premises (t11.t3)) (step t11.t5 (cl (= @p_12 (! (= @p_15 @p_17) :named @p_18))) :rule cong :premises (t11.t2 t11.t4)) (step t11 (cl (! (= @p_14 (! (forall ((veriT_vr1 Int)) @p_18) :named @p_20)) :named @p_19)) :rule bind) (step t12 (cl (not @p_19) (not @p_14) @p_20) :rule equiv_pos2) (step t13 (cl @p_20) :rule th_resolution :premises (t10 t11 t12)) (anchor :step t14 :args ((:= (?v0 Int_int_fun$) veriT_vr2))) (step t14.t1 (cl (= ?v0 veriT_vr2)) :rule refl) (step t14.t2 (cl (= @p_22 (! (map$ veriT_vr2 nil$) :named @p_23))) :rule cong :premises (t14.t1)) (step t14.t3 (cl (= @p_24 (! (= nil$ @p_23) :named @p_25))) :rule cong :premises (t14.t2)) (step t14 (cl (! (= @p_21 (! (forall ((veriT_vr2 Int_int_fun$)) @p_25) :named @p_27)) :named @p_26)) :rule bind) (step t15 (cl (not @p_26) (not @p_21) @p_27) :rule equiv_pos2) (step t16 (cl @p_27) :rule th_resolution :premises (axiom1 t14 t15)) (anchor :step t17 :args ((:= (veriT_vr2 Int_int_fun$) veriT_vr3))) (step t17.t1 (cl (= veriT_vr2 veriT_vr3)) :rule refl) (step t17.t2 (cl (= @p_23 (! (map$ veriT_vr3 nil$) :named @p_28))) :rule cong :premises (t17.t1)) (step t17.t3 (cl (= @p_25 (! (= nil$ @p_28) :named @p_29))) :rule cong :premises (t17.t2)) (step t17 (cl (! (= @p_27 (! (forall ((veriT_vr3 Int_int_fun$)) @p_29) :named @p_31)) :named @p_30)) :rule bind) (step t18 (cl (not @p_30) (not @p_27) @p_31) :rule equiv_pos2) (step t19 (cl @p_31) :rule th_resolution :premises (t16 t17 t18)) (anchor :step t20 :args ((:= (?v0 Int_int_fun$) veriT_vr4) (:= (?v1 Int) veriT_vr5) (:= (?v2 Int_list$) veriT_vr6))) (step t20.t1 (cl (! (= ?v0 veriT_vr4) :named @p_37)) :rule refl) (step t20.t2 (cl (! (= ?v1 veriT_vr5) :named @p_38)) :rule refl) (step t20.t3 (cl (! (= ?v2 veriT_vr6) :named @p_41)) :rule refl) (step t20.t4 (cl (= @p_33 (! (cons$ veriT_vr5 veriT_vr6) :named @p_34))) :rule cong :premises (t20.t2 t20.t3)) (step t20.t5 (cl (= @p_35 (! (map$ veriT_vr4 @p_34) :named @p_36))) :rule cong :premises (t20.t1 t20.t4)) (step t20.t6 (cl @p_37) :rule refl) (step t20.t7 (cl @p_38) :rule refl) (step t20.t8 (cl (= @p_39 (! (fun_app$ veriT_vr4 veriT_vr5) :named @p_40))) :rule cong :premises (t20.t6 t20.t7)) (step t20.t9 (cl @p_37) :rule refl) (step t20.t10 (cl @p_41) :rule refl) (step t20.t11 (cl (= @p_42 (! (map$ veriT_vr4 veriT_vr6) :named @p_43))) :rule cong :premises (t20.t9 t20.t10)) (step t20.t12 (cl (= @p_44 (! (cons$ @p_40 @p_43) :named @p_45))) :rule cong :premises (t20.t8 t20.t11)) (step t20.t13 (cl (= @p_46 (! (= @p_36 @p_45) :named @p_47))) :rule cong :premises (t20.t5 t20.t12)) (step t20 (cl (! (= @p_32 (! (forall ((veriT_vr4 Int_int_fun$) (veriT_vr5 Int) (veriT_vr6 Int_list$)) @p_47) :named @p_49)) :named @p_48)) :rule bind) (step t21 (cl (not @p_48) (not @p_32) @p_49) :rule equiv_pos2) (step t22 (cl @p_49) :rule th_resolution :premises (axiom2 t20 t21)) (anchor :step t23 :args ((:= (veriT_vr4 Int_int_fun$) veriT_vr7) (:= (veriT_vr5 Int) veriT_vr8) (:= (veriT_vr6 Int_list$) veriT_vr9))) (step t23.t1 (cl (! (= veriT_vr4 veriT_vr7) :named @p_52)) :rule refl) (step t23.t2 (cl (! (= veriT_vr5 veriT_vr8) :named @p_53)) :rule refl) (step t23.t3 (cl (! (= veriT_vr6 veriT_vr9) :named @p_55)) :rule refl) (step t23.t4 (cl (= @p_34 (! (cons$ veriT_vr8 veriT_vr9) :named @p_50))) :rule cong :premises (t23.t2 t23.t3)) (step t23.t5 (cl (= @p_36 (! (map$ veriT_vr7 @p_50) :named @p_51))) :rule cong :premises (t23.t1 t23.t4)) (step t23.t6 (cl @p_52) :rule refl) (step t23.t7 (cl @p_53) :rule refl) (step t23.t8 (cl (= @p_40 (! (fun_app$ veriT_vr7 veriT_vr8) :named @p_54))) :rule cong :premises (t23.t6 t23.t7)) (step t23.t9 (cl @p_52) :rule refl) (step t23.t10 (cl @p_55) :rule refl) (step t23.t11 (cl (= @p_43 (! (map$ veriT_vr7 veriT_vr9) :named @p_56))) :rule cong :premises (t23.t9 t23.t10)) (step t23.t12 (cl (= @p_45 (! (cons$ @p_54 @p_56) :named @p_57))) :rule cong :premises (t23.t8 t23.t11)) (step t23.t13 (cl (= @p_47 (! (= @p_51 @p_57) :named @p_58))) :rule cong :premises (t23.t5 t23.t12)) (step t23 (cl (! (= @p_49 (! (forall ((veriT_vr7 Int_int_fun$) (veriT_vr8 Int) (veriT_vr9 Int_list$)) @p_58) :named @p_60)) :named @p_59)) :rule bind) (step t24 (cl (not @p_59) (not @p_49) @p_60) :rule equiv_pos2) (step t25 (cl @p_60) :rule th_resolution :premises (t22 t23 t24)) (step t26 (cl (or (! (not @p_60) :named @p_63) (! (= @p_61 (! (cons$ (! (fun_app$ uu$ 0) :named @p_66) (! (map$ uu$ @p_62) :named @p_65)) :named @p_87)) :named @p_64))) :rule forall_inst :args ((:= veriT_vr7 uu$) (:= veriT_vr8 0) (:= veriT_vr9 @p_62))) (step t27 (cl @p_63 @p_64) :rule or :premises (t26)) (step t28 (cl @p_64) :rule resolution :premises (t27 t25)) (step t29 (cl (or @p_63 (! (= @p_65 (! (cons$ (! (fun_app$ uu$ 1) :named @p_77) (! (map$ uu$ nil$) :named @p_76)) :named @p_92)) :named @p_75))) :rule forall_inst :args ((:= veriT_vr7 uu$) (:= veriT_vr8 1) (:= veriT_vr9 nil$))) (step t30 (cl (or (! (not @p_20) :named @p_72) (! (= @p_66 (! (+ 1 0) :named @p_68)) :named @p_67))) :rule forall_inst :args ((:= veriT_vr1 0))) (anchor :step t31) (assume t31.h1 @p_67) (step t31.t2 (cl (= 1 @p_68)) :rule sum_simplify) (step t31.t3 (cl (! (= @p_67 (! (= 1 @p_66) :named @p_69)) :named @p_70)) :rule cong :premises (t31.t2)) (step t31.t4 (cl (not @p_70) (! (not @p_67) :named @p_71) @p_69) :rule equiv_pos2) (step t31.t5 (cl @p_69) :rule th_resolution :premises (t31.h1 t31.t3 t31.t4)) (step t31 (cl @p_71 @p_69) :rule subproof :discharge (h1)) (step t32 (cl @p_72 @p_67) :rule or :premises (t30)) (step t33 (cl (! (or @p_72 @p_69) :named @p_74) (! (not @p_72) :named @p_73)) :rule or_neg) (step t34 (cl (not @p_73) @p_20) :rule not_not) (step t35 (cl @p_74 @p_20) :rule th_resolution :premises (t34 t33)) (step t36 (cl @p_74 (! (not @p_69) :named @p_89)) :rule or_neg) (step t37 (cl @p_74) :rule th_resolution :premises (t32 t31 t35 t36)) (step t38 (cl @p_63 @p_75) :rule or :premises (t29)) (step t39 (cl @p_75) :rule resolution :premises (t38 t25)) (step t40 (cl @p_72 @p_69) :rule or :premises (t37)) (step t41 (cl @p_69) :rule resolution :premises (t40 t13)) (step t42 (cl (or (! (not @p_31) :named @p_84) (! (= nil$ @p_76) :named @p_85))) :rule forall_inst :args ((:= veriT_vr3 uu$))) (step t43 (cl (or @p_72 (! (= @p_77 (! (+ 1 1) :named @p_79)) :named @p_78))) :rule forall_inst :args ((:= veriT_vr1 1))) (anchor :step t44) (assume t44.h1 @p_78) (step t44.t2 (cl (= 2 @p_79)) :rule sum_simplify) (step t44.t3 (cl (! (= @p_78 (! (= 2 @p_77) :named @p_80)) :named @p_81)) :rule cong :premises (t44.t2)) (step t44.t4 (cl (not @p_81) (! (not @p_78) :named @p_82) @p_80) :rule equiv_pos2) (step t44.t5 (cl @p_80) :rule th_resolution :premises (t44.h1 t44.t3 t44.t4)) (step t44 (cl @p_82 @p_80) :rule subproof :discharge (h1)) (step t45 (cl @p_72 @p_78) :rule or :premises (t43)) (step t46 (cl (! (or @p_72 @p_80) :named @p_83) @p_73) :rule or_neg) (step t47 (cl @p_83 @p_20) :rule th_resolution :premises (t34 t46)) (step t48 (cl @p_83 (! (not @p_80) :named @p_94)) :rule or_neg) (step t49 (cl @p_83) :rule th_resolution :premises (t45 t44 t47 t48)) (step t50 (cl @p_84 @p_85) :rule or :premises (t42)) (step t51 (cl @p_85) :rule resolution :premises (t50 t19)) (step t52 (cl @p_72 @p_80) :rule or :premises (t49)) (step t53 (cl @p_80) :rule resolution :premises (t52 t13)) (step t54 (cl (! (not @p_64) :named @p_98) (not (! (= @p_86 @p_87) :named @p_91)) @p_88) :rule eq_transitive) (step t55 (cl @p_89 (not (! (= @p_90 @p_65) :named @p_93)) @p_91) :rule eq_congruent) (step t56 (cl (not (! (= @p_90 @p_92) :named @p_95)) (! (not @p_75) :named @p_96) @p_93) :rule eq_transitive) (step t57 (cl @p_94 (! (not @p_85) :named @p_97) @p_95) :rule eq_congruent) (step t58 (cl @p_96 @p_93 @p_94 @p_97) :rule th_resolution :premises (t56 t57)) (step t59 (cl @p_89 @p_91 @p_96 @p_94 @p_97) :rule th_resolution :premises (t55 t58)) (step t60 (cl @p_98 @p_88 @p_89 @p_96 @p_94 @p_97) :rule th_resolution :premises (t54 t59)) (step t61 (cl) :rule resolution :premises (t60 axiom3 t28 t39 t41 t51 t53)) 4bfba84466bed875a9d6690636344a2452d6f312 23 0 unsat (assume axiom0 (! (not (! (or (! (forall ((?v0 A$)) (! (p$ ?v0) :named @p_2)) :named @p_1) (! (not @p_1) :named @p_4)) :named @p_6)) :named @p_8)) (anchor :step t2 :args ((:= (?v0 A$) veriT_vr0))) (step t2.t1 (cl (= ?v0 veriT_vr0)) :rule refl) (step t2.t2 (cl (= @p_2 (! (p$ veriT_vr0) :named @p_3))) :rule cong :premises (t2.t1)) (step t2 (cl (= @p_1 (! (forall ((veriT_vr0 A$)) @p_3) :named @p_5))) :rule bind) (step t3 (cl (= @p_4 (! (not @p_5) :named @p_7))) :rule cong :premises (t2)) (step t4 (cl (= @p_6 (! (or @p_5 @p_7) :named @p_9))) :rule cong :premises (t2 t3)) (step t5 (cl (! (= @p_8 (! (not @p_9) :named @p_11)) :named @p_10)) :rule cong :premises (t4)) (step t6 (cl (! (not @p_10) :named @p_13) (! (not @p_8) :named @p_12) @p_11) :rule equiv_pos2) (step t7 (cl (not @p_12) @p_6) :rule not_not) (step t8 (cl @p_13 @p_6 @p_11) :rule th_resolution :premises (t7 t6)) (step t9 (cl @p_11) :rule th_resolution :premises (axiom0 t5 t8)) (step t10 (cl (= @p_9 true)) :rule or_simplify) (step t11 (cl (= @p_11 (! (not true) :named @p_14))) :rule cong :premises (t10)) (step t12 (cl (= @p_14 false)) :rule not_simplify) (step t13 (cl (! (= @p_11 false) :named @p_15)) :rule trans :premises (t11 t12)) (step t14 (cl (! (not @p_15) :named @p_17) (! (not @p_11) :named @p_16) false) :rule equiv_pos2) (step t15 (cl (not @p_16) @p_9) :rule not_not) (step t16 (cl @p_17 @p_9 false) :rule th_resolution :premises (t15 t14)) (step t17 (cl false) :rule th_resolution :premises (t9 t13 t16)) (step t18 (cl (not false)) :rule false) (step t19 (cl) :rule resolution :premises (t17 t18)) 2f429461f6d8832eb5f04c2d4afd47dfd1b769d9 48 0 unsat (assume axiom2 (! (forall ((?v0 A$) (?v1 A$) (?v2 A$)) (! (=> (! (and (! (less_eq$ ?v0 ?v1) :named @p_4) (! (less_eq$ ?v1 ?v2) :named @p_7)) :named @p_9) (! (less_eq$ ?v0 ?v2) :named @p_13)) :named @p_15)) :named @p_3)) (assume axiom3 (! (less_eq$ (! (sup$ (collect$ uu$)) :named @p_2) (! (sup$ (collect$ uua$)) :named @p_1)) :named @p_31)) (assume axiom4 (! (less_eq$ @p_1 @p_2) :named @p_32)) (assume axiom5 (not (! (less_eq$ @p_2 @p_2) :named @p_33))) (anchor :step t5 :args ((:= (?v0 A$) veriT_vr4) (:= (?v1 A$) veriT_vr5) (:= (?v2 A$) veriT_vr6))) (step t5.t1 (cl (! (= ?v0 veriT_vr4) :named @p_11)) :rule refl) (step t5.t2 (cl (! (= ?v1 veriT_vr5) :named @p_6)) :rule refl) (step t5.t3 (cl (= @p_4 (! (less_eq$ veriT_vr4 veriT_vr5) :named @p_5))) :rule cong :premises (t5.t1 t5.t2)) (step t5.t4 (cl @p_6) :rule refl) (step t5.t5 (cl (! (= ?v2 veriT_vr6) :named @p_12)) :rule refl) (step t5.t6 (cl (= @p_7 (! (less_eq$ veriT_vr5 veriT_vr6) :named @p_8))) :rule cong :premises (t5.t4 t5.t5)) (step t5.t7 (cl (= @p_9 (! (and @p_5 @p_8) :named @p_10))) :rule cong :premises (t5.t3 t5.t6)) (step t5.t8 (cl @p_11) :rule refl) (step t5.t9 (cl @p_12) :rule refl) (step t5.t10 (cl (= @p_13 (! (less_eq$ veriT_vr4 veriT_vr6) :named @p_14))) :rule cong :premises (t5.t8 t5.t9)) (step t5.t11 (cl (= @p_15 (! (=> @p_10 @p_14) :named @p_16))) :rule cong :premises (t5.t7 t5.t10)) (step t5 (cl (! (= @p_3 (! (forall ((veriT_vr4 A$) (veriT_vr5 A$) (veriT_vr6 A$)) @p_16) :named @p_18)) :named @p_17)) :rule bind) (step t6 (cl (not @p_17) (not @p_3) @p_18) :rule equiv_pos2) (step t7 (cl @p_18) :rule th_resolution :premises (axiom2 t5 t6)) (anchor :step t8 :args ((:= (veriT_vr4 A$) veriT_vr7) (:= (veriT_vr5 A$) veriT_vr8) (:= (veriT_vr6 A$) veriT_vr9))) (step t8.t1 (cl (! (= veriT_vr4 veriT_vr7) :named @p_23)) :rule refl) (step t8.t2 (cl (! (= veriT_vr5 veriT_vr8) :named @p_20)) :rule refl) (step t8.t3 (cl (= @p_5 (! (less_eq$ veriT_vr7 veriT_vr8) :named @p_19))) :rule cong :premises (t8.t1 t8.t2)) (step t8.t4 (cl @p_20) :rule refl) (step t8.t5 (cl (! (= veriT_vr6 veriT_vr9) :named @p_24)) :rule refl) (step t8.t6 (cl (= @p_8 (! (less_eq$ veriT_vr8 veriT_vr9) :named @p_21))) :rule cong :premises (t8.t4 t8.t5)) (step t8.t7 (cl (= @p_10 (! (and @p_19 @p_21) :named @p_22))) :rule cong :premises (t8.t3 t8.t6)) (step t8.t8 (cl @p_23) :rule refl) (step t8.t9 (cl @p_24) :rule refl) (step t8.t10 (cl (= @p_14 (! (less_eq$ veriT_vr7 veriT_vr9) :named @p_25))) :rule cong :premises (t8.t8 t8.t9)) (step t8.t11 (cl (= @p_16 (! (=> @p_22 @p_25) :named @p_26))) :rule cong :premises (t8.t7 t8.t10)) (step t8 (cl (! (= @p_18 (! (forall ((veriT_vr7 A$) (veriT_vr8 A$) (veriT_vr9 A$)) @p_26) :named @p_28)) :named @p_27)) :rule bind) (step t9 (cl (not @p_27) (not @p_18) @p_28) :rule equiv_pos2) (step t10 (cl @p_28) :rule th_resolution :premises (t7 t8 t9)) (step t11 (cl (or (! (not @p_28) :named @p_29) (! (forall ((veriT_vr7 A$) (veriT_vr8 A$) (veriT_vr9 A$)) (or (not @p_19) (not @p_21) @p_25)) :named @p_30))) :rule qnt_cnf) (step t12 (cl @p_29 @p_30) :rule or :premises (t11)) (step t13 (cl (or (! (not @p_30) :named @p_34) (! (or (! (not @p_31) :named @p_39) (! (not @p_32) :named @p_40) @p_33) :named @p_35))) :rule forall_inst :args ((:= veriT_vr7 @p_2) (:= veriT_vr8 @p_1) (:= veriT_vr9 @p_2))) (step t14 (cl @p_34 @p_35) :rule or :premises (t13)) (step t15 (cl (! (or @p_29 @p_35) :named @p_37) (! (not @p_29) :named @p_36)) :rule or_neg) (step t16 (cl (not @p_36) @p_28) :rule not_not) (step t17 (cl @p_37 @p_28) :rule th_resolution :premises (t16 t15)) (step t18 (cl @p_37 (! (not @p_35) :named @p_38)) :rule or_neg) (step t19 (cl @p_37) :rule th_resolution :premises (t12 t14 t17 t18)) (step t20 (cl @p_38 @p_39 @p_40 @p_33) :rule or_pos) (step t21 (cl @p_29 @p_35) :rule or :premises (t19)) (step t22 (cl @p_38) :rule resolution :premises (t20 axiom3 axiom4 axiom5)) (step t23 (cl) :rule resolution :premises (t21 t10 t22)) 8f4bf1893caa103dd59d38dee5686b501773527d 107 0 unsat (assume axiom0 (! (forall ((?v0 Int)) (! (= (! (dec_10$ ?v0) :named @p_2) (! (ite (! (< ?v0 10) :named @p_5) ?v0 (! (dec_10$ (! (- ?v0 10) :named @p_7)) :named @p_9)) :named @p_11)) :named @p_13)) :named @p_1)) (assume axiom1 (not (! (= (! (dec_10$ (! (* 4 (! (dec_10$ 4) :named @p_38)) :named @p_27)) :named @p_26) 6) :named @p_74))) (anchor :step t3 :args ((:= (?v0 Int) veriT_vr0))) (step t3.t1 (cl (! (= ?v0 veriT_vr0) :named @p_4)) :rule refl) (step t3.t2 (cl (= @p_2 (! (dec_10$ veriT_vr0) :named @p_3))) :rule cong :premises (t3.t1)) (step t3.t3 (cl @p_4) :rule refl) (step t3.t4 (cl (= @p_5 (! (< veriT_vr0 10) :named @p_6))) :rule cong :premises (t3.t3)) (step t3.t5 (cl @p_4) :rule refl) (step t3.t6 (cl @p_4) :rule refl) (step t3.t7 (cl (= @p_7 (! (- veriT_vr0 10) :named @p_8))) :rule cong :premises (t3.t6)) (step t3.t8 (cl (= @p_9 (! (dec_10$ @p_8) :named @p_10))) :rule cong :premises (t3.t7)) (step t3.t9 (cl (= @p_11 (! (ite @p_6 veriT_vr0 @p_10) :named @p_12))) :rule cong :premises (t3.t4 t3.t5 t3.t8)) (step t3.t10 (cl (= @p_13 (! (= @p_3 @p_12) :named @p_14))) :rule cong :premises (t3.t2 t3.t9)) (step t3 (cl (! (= @p_1 (! (forall ((veriT_vr0 Int)) @p_14) :named @p_16)) :named @p_15)) :rule bind) (step t4 (cl (not @p_15) (not @p_1) @p_16) :rule equiv_pos2) (step t5 (cl @p_16) :rule th_resolution :premises (axiom0 t3 t4)) (anchor :step t6 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t6.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_18)) :rule refl) (step t6.t2 (cl (= @p_3 (! (dec_10$ veriT_vr1) :named @p_17))) :rule cong :premises (t6.t1)) (step t6.t3 (cl @p_18) :rule refl) (step t6.t4 (cl (= @p_6 (! (< veriT_vr1 10) :named @p_19))) :rule cong :premises (t6.t3)) (step t6.t5 (cl @p_18) :rule refl) (step t6.t6 (cl @p_18) :rule refl) (step t6.t7 (cl (= @p_8 (! (- veriT_vr1 10) :named @p_20))) :rule cong :premises (t6.t6)) (step t6.t8 (cl (= @p_10 (! (dec_10$ @p_20) :named @p_21))) :rule cong :premises (t6.t7)) (step t6.t9 (cl (= @p_12 (! (ite @p_19 veriT_vr1 @p_21) :named @p_22))) :rule cong :premises (t6.t4 t6.t5 t6.t8)) (step t6.t10 (cl (= @p_14 (! (= @p_17 @p_22) :named @p_23))) :rule cong :premises (t6.t2 t6.t9)) (step t6 (cl (! (= @p_16 (! (forall ((veriT_vr1 Int)) @p_23) :named @p_25)) :named @p_24)) :rule bind) (step t7 (cl (not @p_24) (not @p_16) @p_25) :rule equiv_pos2) (step t8 (cl @p_25) :rule th_resolution :premises (t5 t6 t7)) (step t9 (cl (or (! (not @p_25) :named @p_35) (! (= @p_26 (! (ite (! (< @p_27 10) :named @p_30) @p_27 (! (dec_10$ (! (- @p_27 10) :named @p_54)) :named @p_31)) :named @p_29)) :named @p_28))) :rule forall_inst :args ((:= veriT_vr1 @p_27))) (anchor :step t10) (assume t10.h1 @p_28) (step t10.t2 (cl (! (= @p_28 (! (and (! (= @p_26 @p_29) :named @p_51) (! (ite @p_30 (= @p_27 @p_29) (! (= @p_31 @p_29) :named @p_53)) :named @p_52)) :named @p_32)) :named @p_33)) :rule ite_intro) (step t10.t3 (cl (not @p_33) (! (not @p_28) :named @p_34) @p_32) :rule equiv_pos2) (step t10.t4 (cl @p_32) :rule th_resolution :premises (t10.h1 t10.t2 t10.t3)) (step t10 (cl @p_34 @p_32) :rule subproof :discharge (h1)) (step t11 (cl @p_35 @p_28) :rule or :premises (t9)) (step t12 (cl (! (or @p_35 @p_32) :named @p_37) (! (not @p_35) :named @p_36)) :rule or_neg) (step t13 (cl (not @p_36) @p_25) :rule not_not) (step t14 (cl @p_37 @p_25) :rule th_resolution :premises (t13 t12)) (step t15 (cl @p_37 (! (not @p_32) :named @p_50)) :rule or_neg) (step t16 (cl @p_37) :rule th_resolution :premises (t11 t10 t14 t15)) (step t17 (cl (or @p_35 (! (= @p_38 (! (ite (! (< 4 10) :named @p_40) 4 (! (dec_10$ (! (- 4 10) :named @p_41)) :named @p_42)) :named @p_43)) :named @p_39))) :rule forall_inst :args ((:= veriT_vr1 4))) (anchor :step t18) (assume t18.h1 @p_39) (step t18.t2 (cl (= @p_40 true)) :rule comp_simplify) (step t18.t3 (cl (= @p_41 (- 6))) :rule minus_simplify) (step t18.t4 (cl (= @p_42 (! (dec_10$ (- 6)) :named @p_44))) :rule cong :premises (t18.t3)) (step t18.t5 (cl (= @p_43 (! (ite true 4 @p_44) :named @p_45))) :rule cong :premises (t18.t2 t18.t4)) (step t18.t6 (cl (= 4 @p_45)) :rule ite_simplify) (step t18.t7 (cl (= 4 @p_43)) :rule trans :premises (t18.t5 t18.t6)) (step t18.t8 (cl (! (= @p_39 (! (= 4 @p_38) :named @p_46)) :named @p_47)) :rule cong :premises (t18.t7)) (step t18.t9 (cl (not @p_47) (! (not @p_39) :named @p_48) @p_46) :rule equiv_pos2) (step t18.t10 (cl @p_46) :rule th_resolution :premises (t18.h1 t18.t8 t18.t9)) (step t18 (cl @p_48 @p_46) :rule subproof :discharge (h1)) (step t19 (cl @p_35 @p_39) :rule or :premises (t17)) (step t20 (cl (! (or @p_35 @p_46) :named @p_49) @p_36) :rule or_neg) (step t21 (cl @p_49 @p_25) :rule th_resolution :premises (t13 t20)) (step t22 (cl @p_49 (! (not @p_46) :named @p_67)) :rule or_neg) (step t23 (cl @p_49) :rule th_resolution :premises (t19 t18 t21 t22)) (step t24 (cl @p_50 @p_51) :rule and_pos) (step t25 (cl (not @p_52) @p_30 @p_53) :rule ite_pos1) (step t26 (cl @p_50 @p_52) :rule and_pos) (step t27 (cl @p_35 @p_32) :rule or :premises (t16)) (step t28 (cl @p_32) :rule resolution :premises (t27 t8)) (step t29 (cl @p_51) :rule resolution :premises (t24 t28)) (step t30 (cl @p_52) :rule resolution :premises (t26 t28)) (step t31 (cl @p_35 @p_46) :rule or :premises (t23)) (step t32 (cl @p_46) :rule resolution :premises (t31 t8)) (step t33 (cl (or @p_35 (! (= @p_31 (! (ite (! (< @p_54 10) :named @p_57) @p_54 (! (dec_10$ (- @p_54 10)) :named @p_58)) :named @p_56)) :named @p_55))) :rule forall_inst :args ((:= veriT_vr1 @p_54))) (anchor :step t34) (assume t34.h1 @p_55) (step t34.t2 (cl (! (= @p_55 (! (and (! (= @p_31 @p_56) :named @p_64) (! (ite @p_57 (! (= @p_54 @p_56) :named @p_66) (= @p_58 @p_56)) :named @p_65)) :named @p_59)) :named @p_60)) :rule ite_intro) (step t34.t3 (cl (not @p_60) (! (not @p_55) :named @p_61) @p_59) :rule equiv_pos2) (step t34.t4 (cl @p_59) :rule th_resolution :premises (t34.h1 t34.t2 t34.t3)) (step t34 (cl @p_61 @p_59) :rule subproof :discharge (h1)) (step t35 (cl @p_35 @p_55) :rule or :premises (t33)) (step t36 (cl (! (or @p_35 @p_59) :named @p_62) @p_36) :rule or_neg) (step t37 (cl @p_62 @p_25) :rule th_resolution :premises (t13 t36)) (step t38 (cl @p_62 (! (not @p_59) :named @p_63)) :rule or_neg) (step t39 (cl @p_62) :rule th_resolution :premises (t35 t34 t37 t38)) (step t40 (cl @p_63 @p_64) :rule and_pos) (step t41 (cl (not @p_65) (not @p_57) @p_66) :rule ite_pos2) (step t42 (cl @p_63 @p_65) :rule and_pos) (step t43 (cl @p_35 @p_59) :rule or :premises (t39)) (step t44 (cl @p_59) :rule resolution :premises (t43 t8)) (step t45 (cl @p_64) :rule resolution :premises (t40 t44)) (step t46 (cl @p_65) :rule resolution :premises (t42 t44)) (step t47 (cl @p_57 @p_67) :rule la_generic :args (1 4)) (step t48 (cl @p_57) :rule resolution :premises (t47 t32)) (step t49 (cl @p_66) :rule resolution :premises (t41 t48 t46)) (step t50 (cl (or (! (= 6 @p_54) :named @p_68) (! (not (! (<= 6 @p_54) :named @p_72)) :named @p_69) (! (not (! (<= @p_54 6) :named @p_71)) :named @p_70))) :rule la_disequality) (step t51 (cl @p_68 @p_69 @p_70) :rule or :premises (t50)) (step t52 (cl @p_71 @p_67) :rule la_generic :args (1 4)) (step t53 (cl @p_71) :rule resolution :premises (t52 t32)) (step t54 (cl @p_72 @p_67) :rule la_generic :args (1 (- 4))) (step t55 (cl @p_72) :rule resolution :premises (t54 t32)) (step t56 (cl @p_68) :rule resolution :premises (t51 t55 t53)) (step t57 (cl (! (not @p_64) :named @p_77) (! (not @p_66) :named @p_78) (! (not @p_68) :named @p_79) (! (= 6 @p_31) :named @p_73)) :rule eq_transitive) (step t58 (cl (not @p_73) (! (not @p_53) :named @p_75) (! (not @p_51) :named @p_76) @p_74) :rule eq_transitive) (step t59 (cl @p_75 @p_76 @p_74 @p_77 @p_78 @p_79) :rule th_resolution :premises (t58 t57)) (step t60 (cl @p_75) :rule resolution :premises (t59 axiom1 t29 t45 t49 t56)) (step t61 (cl @p_30) :rule resolution :premises (t25 t60 t30)) (step t62 (cl (not @p_30) @p_67) :rule la_generic :args (1 (- 4))) (step t63 (cl) :rule resolution :premises (t62 t61 t32)) 170ac60e28eef34728b56b5ef83cebdd5889b03f 1434 0 unsat (define-fun veriT_sk0 () A$ (! (choice ((veriT_vr57 A$)) (not (! (=> (! (member$ veriT_vr57 (! (set$ (! (remdups$ xs$) :named @p_9)) :named @p_380)) :named @p_429) (! (not (! (member$ veriT_vr57 (! (set$ (! (remdups$ ys$) :named @p_10)) :named @p_381)) :named @p_431)) :named @p_432)) :named @p_433))) :named @p_449)) (define-fun veriT_sk1 () A$ (! (choice ((veriT_vr58 A$)) (not (! (=> (! (member$ veriT_vr58 @p_381) :named @p_434) (! (not (! (member$ veriT_vr58 @p_380) :named @p_436)) :named @p_437)) :named @p_438))) :named @p_455)) (define-fun veriT_sk4 () A$ (! (choice ((veriT_vr97 A$)) (not (! (=> (! (member$ veriT_vr97 (! (set$ xs$) :named @p_715)) :named @p_825) (! (not (! (member$ veriT_vr97 @p_381) :named @p_827)) :named @p_828)) :named @p_824))) :named @p_799)) (define-fun veriT_sk6 () A$ (! (choice ((veriT_vr103 A$)) (not (! (=> (! (member$ veriT_vr103 @p_380) :named @p_879) (! (not (! (member$ veriT_vr103 @p_381) :named @p_881)) :named @p_882)) :named @p_878))) :named @p_857)) (define-fun veriT_sk11 () A$ (! (choice ((veriT_vr117 A$)) (not (! (=> (! (member$ veriT_vr117 (! (coset$ xs$) :named @p_21)) :named @p_924) (! (member$ veriT_vr117 (! (set$ (! (append$ xs$ @p_10) :named @p_761)) :named @p_760)) :named @p_926)) :named @p_923))) :named @p_909)) (define-fun veriT_sk14 () A$ (! (choice ((veriT_vr123 A$)) (not (! (=> (! (member$ veriT_vr123 @p_715) :named @p_969) (! (not (! (member$ veriT_vr123 @p_381) :named @p_971)) :named @p_972)) :named @p_968))) :named @p_943)) (assume axiom0 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (= (! (fun_app$ (! (fun_app$a uu$ ?v0) :named @p_23) ?v1) :named @p_25) (! (= ?v0 ?v1) :named @p_13)) :named @p_30)) :named @p_22)) (assume axiom1 (! (forall ((?v0 A_list$)) (! (= (! (fun_app$ (! (fun_app$a subset$ @p_21) :named @p_43) (! (set$ ?v0) :named @p_2)) :named @p_46) (! (and (! (less$ zero$ (! (fun_app$b card$ top$) :named @p_1)) :named @p_44) (! (= @p_1 (! (fun_app$b card$ (! (set$ (! (append$ xs$ ?v0) :named @p_49)) :named @p_51)) :named @p_53)) :named @p_55)) :named @p_57)) :named @p_59)) :named @p_42)) (assume axiom2 (! (forall ((?v0 A$) (?v1 A_set$)) (! (=> (! (member$ ?v0 (! (uminus$ ?v1) :named @p_75)) :named @p_77) (! (not (! (member$ ?v0 ?v1) :named @p_81)) :named @p_83)) :named @p_85)) :named @p_74)) (assume axiom3 (! (forall ((?v0 A_list$)) (! (= (! (fun_app$b card$ @p_2) :named @p_100) (! (size$ (! (remdups$ ?v0) :named @p_14)) :named @p_104)) :named @p_106)) :named @p_98)) (assume axiom4 (! (forall ((?v0 A_list$)) (! (finite$ @p_2) :named @p_120)) :named @p_118)) (assume axiom5 (! (forall ((?v0 A$)) (! (member$ ?v0 top$) :named @p_129)) :named @p_128)) (assume axiom6 (not (! (= top$ bot$) :named @p_717))) (assume axiom9 (! (forall ((?v0 Nat$) (?v1 Nat$)) (! (= (! (= ?v0 (! (plus$ ?v0 ?v1) :named @p_138)) :named @p_140) (! (= zero$ ?v1) :named @p_143)) :named @p_145)) :named @p_136)) (assume axiom10 (! (= (! (fun_app$b card$a @p_715) :named @p_1077) (! (size$ @p_9) :named @p_8)) :named @p_1076)) (assume axiom11 (! (= card$ card$a) :named @p_1079)) (assume axiom12 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (=> (! (and (! (finite$ ?v0) :named @p_4) (! (finite$ ?v1) :named @p_159)) :named @p_161) (! (= (! (plus$ (! (fun_app$b card$ ?v0) :named @p_3) (! (fun_app$b card$ ?v1) :named @p_166)) :named @p_168) (! (plus$ (! (fun_app$b card$ (! (sup$ ?v0 ?v1) :named @p_170)) :named @p_172) (! (fun_app$b card$ (! (inf$ ?v0 ?v1) :named @p_6)) :named @p_175)) :named @p_177)) :named @p_179)) :named @p_181)) :named @p_157)) (assume axiom13 (! (forall ((?v0 A_set$)) (! (= (! (= zero$ @p_3) :named @p_204) (! (or (! (= ?v0 bot$) :named @p_207) (! (not @p_4) :named @p_210)) :named @p_212)) :named @p_214)) :named @p_202)) (assume axiom14 (! (forall ((?v0 A_set$)) (! (=> (! (and (! (finite$ top$) :named @p_7) (! (= @p_1 @p_3) :named @p_230)) :named @p_232) (! (= ?v0 top$) :named @p_235)) :named @p_237)) :named @p_228)) (assume axiom15 (! (forall ((?v0 A_list$)) (! (= @p_2 (! (uminus$ (! (coset$ ?v0) :named @p_5)) :named @p_253)) :named @p_255)) :named @p_249)) (assume axiom16 (! (forall ((?v0 A_list$)) (! (= @p_5 (! (uminus$ @p_2) :named @p_270)) :named @p_272)) :named @p_266)) (assume axiom17 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (= (! (= bot$ @p_6) :named @p_285) (! (forall ((?v2 A$)) (! (=> (! (member$ ?v2 ?v0) :named @p_15) (! (not (! (member$ ?v2 ?v1) :named @p_16)) :named @p_294)) :named @p_296)) :named @p_287)) :named @p_298)) :named @p_283)) (assume axiom18 (! (= uu$ eq_set$) :named @p_1100)) (assume axiom19 (! (forall ((?v0 A_set$)) (! (=> @p_4 (! (= (! (finite$ (! (uminus$ ?v0) :named @p_361)) :named @p_363) @p_7) :named @p_365)) :named @p_367)) :named @p_358)) (assume axiom20 (! (= rhs$ (! (ite (! (= zero$ @p_1) :named @p_403) false (! (and (! (= @p_1 (! (plus$ @p_8 (! (size$ @p_10) :named @p_719)) :named @p_1056)) :named @p_400) (! (and (! (forall ((?v0 A$)) (! (=> (! (member$ ?v0 @p_380) :named @p_12) (! (not (! (member$ ?v0 @p_381) :named @p_11)) :named @p_385)) :named @p_387)) :named @p_379) (! (forall ((?v0 A$)) (! (=> @p_11 (! (not @p_12) :named @p_392)) :named @p_394)) :named @p_389)) :named @p_396)) :named @p_399)) :named @p_402)) :named @p_405)) (assume axiom21 (! (forall ((?v0 A_list$) (?v1 A_list$)) (! (= (! (set$ (! (append$ ?v0 ?v1) :named @p_498)) :named @p_500) (! (sup$ @p_2 (! (set$ ?v1) :named @p_20)) :named @p_506)) :named @p_508)) :named @p_497)) (assume axiom22 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (= @p_13 (! (and (! (fun_app$ (! (fun_app$a less_eq$ ?v0) :named @p_19) ?v1) :named @p_17) (! (fun_app$ (! (fun_app$a less_eq$ ?v1) :named @p_528) ?v0) :named @p_530)) :named @p_18)) :named @p_533)) :named @p_522)) (assume axiom23 (! (forall ((?v0 A_list$)) (! (= @p_2 (! (set$ @p_14) :named @p_552)) :named @p_554)) :named @p_548)) (assume axiom24 (! (= subset$ less_eq$) :named @p_1090)) (assume axiom25 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (=> (! (forall ((?v2 A$)) (! (=> @p_15 @p_16) :named @p_571)) :named @p_566) @p_17) :named @p_577)) :named @p_565)) (assume axiom26 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (=> @p_18 @p_13) :named @p_602)) :named @p_593)) (assume axiom27 (! (forall ((?v0 A_set$) (?v1 A_list$)) (! (= (! (fun_app$ @p_19 (! (coset$ ?v1) :named @p_619)) :named @p_621) (! (forall ((?v2 A$)) (! (=> (! (member$ ?v2 @p_20) :named @p_627) (! (not @p_15) :named @p_632)) :named @p_634)) :named @p_623)) :named @p_636)) :named @p_617)) (assume axiom28 (not (= (! (fun_app$ (! (fun_app$a eq_set$ @p_21) :named @p_1097) (! (set$ ys$) :named @p_713)) :named @p_711) rhs$))) (anchor :step t28 :args ((:= (?v0 A_set$) veriT_vr0) (:= (?v1 A_set$) veriT_vr1))) (step t28.t1 (cl (! (= ?v0 veriT_vr0) :named @p_27)) :rule refl) (step t28.t2 (cl (= @p_23 (! (fun_app$a uu$ veriT_vr0) :named @p_24))) :rule cong :premises (t28.t1)) (step t28.t3 (cl (! (= ?v1 veriT_vr1) :named @p_28)) :rule refl) (step t28.t4 (cl (= @p_25 (! (fun_app$ @p_24 veriT_vr1) :named @p_26))) :rule cong :premises (t28.t2 t28.t3)) (step t28.t5 (cl @p_27) :rule refl) (step t28.t6 (cl @p_28) :rule refl) (step t28.t7 (cl (= @p_13 (! (= veriT_vr0 veriT_vr1) :named @p_29))) :rule cong :premises (t28.t5 t28.t6)) (step t28.t8 (cl (= @p_30 (! (= @p_26 @p_29) :named @p_31))) :rule cong :premises (t28.t4 t28.t7)) (step t28 (cl (! (= @p_22 (! (forall ((veriT_vr0 A_set$) (veriT_vr1 A_set$)) @p_31) :named @p_33)) :named @p_32)) :rule bind) (step t29 (cl (not @p_32) (not @p_22) @p_33) :rule equiv_pos2) (step t30 (cl @p_33) :rule th_resolution :premises (axiom0 t28 t29)) (anchor :step t31 :args ((:= (veriT_vr0 A_set$) veriT_vr2) (:= (veriT_vr1 A_set$) veriT_vr3))) (step t31.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_36)) :rule refl) (step t31.t2 (cl (= @p_24 (! (fun_app$a uu$ veriT_vr2) :named @p_34))) :rule cong :premises (t31.t1)) (step t31.t3 (cl (! (= veriT_vr1 veriT_vr3) :named @p_37)) :rule refl) (step t31.t4 (cl (= @p_26 (! (fun_app$ @p_34 veriT_vr3) :named @p_35))) :rule cong :premises (t31.t2 t31.t3)) (step t31.t5 (cl @p_36) :rule refl) (step t31.t6 (cl @p_37) :rule refl) (step t31.t7 (cl (= @p_29 (! (= veriT_vr2 veriT_vr3) :named @p_38))) :rule cong :premises (t31.t5 t31.t6)) (step t31.t8 (cl (= @p_31 (! (= @p_35 @p_38) :named @p_39))) :rule cong :premises (t31.t4 t31.t7)) (step t31 (cl (! (= @p_33 (! (forall ((veriT_vr2 A_set$) (veriT_vr3 A_set$)) @p_39) :named @p_41)) :named @p_40)) :rule bind) (step t32 (cl (not @p_40) (not @p_33) @p_41) :rule equiv_pos2) (step t33 (cl @p_41) :rule th_resolution :premises (t30 t31 t32)) (anchor :step t34 :args ((:= (?v0 A_list$) veriT_vr4))) (step t34.t1 (cl (! (= ?v0 veriT_vr4) :named @p_48)) :rule refl) (step t34.t2 (cl (= @p_2 (! (set$ veriT_vr4) :named @p_45))) :rule cong :premises (t34.t1)) (step t34.t3 (cl (= @p_46 (! (fun_app$ @p_43 @p_45) :named @p_47))) :rule cong :premises (t34.t2)) (step t34.t4 (cl @p_48) :rule refl) (step t34.t5 (cl (= @p_49 (! (append$ xs$ veriT_vr4) :named @p_50))) :rule cong :premises (t34.t4)) (step t34.t6 (cl (= @p_51 (! (set$ @p_50) :named @p_52))) :rule cong :premises (t34.t5)) (step t34.t7 (cl (= @p_53 (! (fun_app$b card$ @p_52) :named @p_54))) :rule cong :premises (t34.t6)) (step t34.t8 (cl (= @p_55 (! (= @p_1 @p_54) :named @p_56))) :rule cong :premises (t34.t7)) (step t34.t9 (cl (= @p_57 (! (and @p_44 @p_56) :named @p_58))) :rule cong :premises (t34.t8)) (step t34.t10 (cl (= @p_59 (! (= @p_47 @p_58) :named @p_60))) :rule cong :premises (t34.t3 t34.t9)) (step t34 (cl (! (= @p_42 (! (forall ((veriT_vr4 A_list$)) @p_60) :named @p_62)) :named @p_61)) :rule bind) (step t35 (cl (not @p_61) (not @p_42) @p_62) :rule equiv_pos2) (step t36 (cl @p_62) :rule th_resolution :premises (axiom1 t34 t35)) (anchor :step t37 :args ((:= (veriT_vr4 A_list$) veriT_vr5))) (step t37.t1 (cl (! (= veriT_vr4 veriT_vr5) :named @p_65)) :rule refl) (step t37.t2 (cl (= @p_45 (! (set$ veriT_vr5) :named @p_63))) :rule cong :premises (t37.t1)) (step t37.t3 (cl (= @p_47 (! (fun_app$ @p_43 @p_63) :named @p_64))) :rule cong :premises (t37.t2)) (step t37.t4 (cl @p_65) :rule refl) (step t37.t5 (cl (= @p_50 (! (append$ xs$ veriT_vr5) :named @p_66))) :rule cong :premises (t37.t4)) (step t37.t6 (cl (= @p_52 (! (set$ @p_66) :named @p_67))) :rule cong :premises (t37.t5)) (step t37.t7 (cl (= @p_54 (! (fun_app$b card$ @p_67) :named @p_68))) :rule cong :premises (t37.t6)) (step t37.t8 (cl (= @p_56 (! (= @p_1 @p_68) :named @p_69))) :rule cong :premises (t37.t7)) (step t37.t9 (cl (= @p_58 (! (and @p_44 @p_69) :named @p_70))) :rule cong :premises (t37.t8)) (step t37.t10 (cl (= @p_60 (! (= @p_64 @p_70) :named @p_71))) :rule cong :premises (t37.t3 t37.t9)) (step t37 (cl (! (= @p_62 (! (forall ((veriT_vr5 A_list$)) @p_71) :named @p_73)) :named @p_72)) :rule bind) (step t38 (cl (not @p_72) (not @p_62) @p_73) :rule equiv_pos2) (step t39 (cl @p_73) :rule th_resolution :premises (t36 t37 t38)) (anchor :step t40 :args ((:= (?v0 A$) veriT_vr6) (:= (?v1 A_set$) veriT_vr7))) (step t40.t1 (cl (! (= ?v0 veriT_vr6) :named @p_79)) :rule refl) (step t40.t2 (cl (! (= ?v1 veriT_vr7) :named @p_80)) :rule refl) (step t40.t3 (cl (= @p_75 (! (uminus$ veriT_vr7) :named @p_76))) :rule cong :premises (t40.t2)) (step t40.t4 (cl (= @p_77 (! (member$ veriT_vr6 @p_76) :named @p_78))) :rule cong :premises (t40.t1 t40.t3)) (step t40.t5 (cl @p_79) :rule refl) (step t40.t6 (cl @p_80) :rule refl) (step t40.t7 (cl (= @p_81 (! (member$ veriT_vr6 veriT_vr7) :named @p_82))) :rule cong :premises (t40.t5 t40.t6)) (step t40.t8 (cl (= @p_83 (! (not @p_82) :named @p_84))) :rule cong :premises (t40.t7)) (step t40.t9 (cl (= @p_85 (! (=> @p_78 @p_84) :named @p_86))) :rule cong :premises (t40.t4 t40.t8)) (step t40 (cl (! (= @p_74 (! (forall ((veriT_vr6 A$) (veriT_vr7 A_set$)) @p_86) :named @p_88)) :named @p_87)) :rule bind) (step t41 (cl (not @p_87) (not @p_74) @p_88) :rule equiv_pos2) (step t42 (cl @p_88) :rule th_resolution :premises (axiom2 t40 t41)) (anchor :step t43 :args ((:= (veriT_vr6 A$) veriT_vr8) (:= (veriT_vr7 A_set$) veriT_vr9))) (step t43.t1 (cl (! (= veriT_vr6 veriT_vr8) :named @p_91)) :rule refl) (step t43.t2 (cl (! (= veriT_vr7 veriT_vr9) :named @p_92)) :rule refl) (step t43.t3 (cl (= @p_76 (! (uminus$ veriT_vr9) :named @p_89))) :rule cong :premises (t43.t2)) (step t43.t4 (cl (= @p_78 (! (member$ veriT_vr8 @p_89) :named @p_90))) :rule cong :premises (t43.t1 t43.t3)) (step t43.t5 (cl @p_91) :rule refl) (step t43.t6 (cl @p_92) :rule refl) (step t43.t7 (cl (= @p_82 (! (member$ veriT_vr8 veriT_vr9) :named @p_93))) :rule cong :premises (t43.t5 t43.t6)) (step t43.t8 (cl (= @p_84 (! (not @p_93) :named @p_94))) :rule cong :premises (t43.t7)) (step t43.t9 (cl (= @p_86 (! (=> @p_90 @p_94) :named @p_95))) :rule cong :premises (t43.t4 t43.t8)) (step t43 (cl (! (= @p_88 (! (forall ((veriT_vr8 A$) (veriT_vr9 A_set$)) @p_95) :named @p_97)) :named @p_96)) :rule bind) (step t44 (cl (not @p_96) (not @p_88) @p_97) :rule equiv_pos2) (step t45 (cl @p_97) :rule th_resolution :premises (t42 t43 t44)) (anchor :step t46 :args ((:= (?v0 A_list$) veriT_vr10))) (step t46.t1 (cl (! (= ?v0 veriT_vr10) :named @p_102)) :rule refl) (step t46.t2 (cl (= @p_2 (! (set$ veriT_vr10) :named @p_99))) :rule cong :premises (t46.t1)) (step t46.t3 (cl (= @p_100 (! (fun_app$b card$ @p_99) :named @p_101))) :rule cong :premises (t46.t2)) (step t46.t4 (cl @p_102) :rule refl) (step t46.t5 (cl (= @p_14 (! (remdups$ veriT_vr10) :named @p_103))) :rule cong :premises (t46.t4)) (step t46.t6 (cl (= @p_104 (! (size$ @p_103) :named @p_105))) :rule cong :premises (t46.t5)) (step t46.t7 (cl (= @p_106 (! (= @p_101 @p_105) :named @p_107))) :rule cong :premises (t46.t3 t46.t6)) (step t46 (cl (! (= @p_98 (! (forall ((veriT_vr10 A_list$)) @p_107) :named @p_109)) :named @p_108)) :rule bind) (step t47 (cl (not @p_108) (not @p_98) @p_109) :rule equiv_pos2) (step t48 (cl @p_109) :rule th_resolution :premises (axiom3 t46 t47)) (anchor :step t49 :args ((:= (veriT_vr10 A_list$) veriT_vr11))) (step t49.t1 (cl (! (= veriT_vr10 veriT_vr11) :named @p_112)) :rule refl) (step t49.t2 (cl (= @p_99 (! (set$ veriT_vr11) :named @p_110))) :rule cong :premises (t49.t1)) (step t49.t3 (cl (= @p_101 (! (fun_app$b card$ @p_110) :named @p_111))) :rule cong :premises (t49.t2)) (step t49.t4 (cl @p_112) :rule refl) (step t49.t5 (cl (= @p_103 (! (remdups$ veriT_vr11) :named @p_113))) :rule cong :premises (t49.t4)) (step t49.t6 (cl (= @p_105 (! (size$ @p_113) :named @p_114))) :rule cong :premises (t49.t5)) (step t49.t7 (cl (= @p_107 (! (= @p_111 @p_114) :named @p_115))) :rule cong :premises (t49.t3 t49.t6)) (step t49 (cl (! (= @p_109 (! (forall ((veriT_vr11 A_list$)) @p_115) :named @p_117)) :named @p_116)) :rule bind) (step t50 (cl (not @p_116) (not @p_109) @p_117) :rule equiv_pos2) (step t51 (cl @p_117) :rule th_resolution :premises (t48 t49 t50)) (anchor :step t52 :args ((:= (?v0 A_list$) veriT_vr12))) (step t52.t1 (cl (= ?v0 veriT_vr12)) :rule refl) (step t52.t2 (cl (= @p_2 (! (set$ veriT_vr12) :named @p_119))) :rule cong :premises (t52.t1)) (step t52.t3 (cl (= @p_120 (! (finite$ @p_119) :named @p_121))) :rule cong :premises (t52.t2)) (step t52 (cl (! (= @p_118 (! (forall ((veriT_vr12 A_list$)) @p_121) :named @p_123)) :named @p_122)) :rule bind) (step t53 (cl (not @p_122) (not @p_118) @p_123) :rule equiv_pos2) (step t54 (cl @p_123) :rule th_resolution :premises (axiom4 t52 t53)) (anchor :step t55 :args ((:= (veriT_vr12 A_list$) veriT_vr13))) (step t55.t1 (cl (= veriT_vr12 veriT_vr13)) :rule refl) (step t55.t2 (cl (= @p_119 (! (set$ veriT_vr13) :named @p_124))) :rule cong :premises (t55.t1)) (step t55.t3 (cl (= @p_121 (! (finite$ @p_124) :named @p_125))) :rule cong :premises (t55.t2)) (step t55 (cl (! (= @p_123 (! (forall ((veriT_vr13 A_list$)) @p_125) :named @p_127)) :named @p_126)) :rule bind) (step t56 (cl (not @p_126) (not @p_123) @p_127) :rule equiv_pos2) (step t57 (cl @p_127) :rule th_resolution :premises (t54 t55 t56)) (anchor :step t58 :args ((:= (?v0 A$) veriT_vr14))) (step t58.t1 (cl (= ?v0 veriT_vr14)) :rule refl) (step t58.t2 (cl (= @p_129 (! (member$ veriT_vr14 top$) :named @p_130))) :rule cong :premises (t58.t1)) (step t58 (cl (! (= @p_128 (! (forall ((veriT_vr14 A$)) @p_130) :named @p_132)) :named @p_131)) :rule bind) (step t59 (cl (not @p_131) (not @p_128) @p_132) :rule equiv_pos2) (step t60 (cl @p_132) :rule th_resolution :premises (axiom5 t58 t59)) (anchor :step t61 :args ((:= (veriT_vr14 A$) veriT_vr15))) (step t61.t1 (cl (= veriT_vr14 veriT_vr15)) :rule refl) (step t61.t2 (cl (= @p_130 (! (member$ veriT_vr15 top$) :named @p_133))) :rule cong :premises (t61.t1)) (step t61 (cl (! (= @p_132 (! (forall ((veriT_vr15 A$)) @p_133) :named @p_135)) :named @p_134)) :rule bind) (step t62 (cl (not @p_134) (not @p_132) @p_135) :rule equiv_pos2) (step t63 (cl @p_135) :rule th_resolution :premises (t60 t61 t62)) (anchor :step t64 :args ((:= (?v0 Nat$) veriT_vr26) (:= (?v1 Nat$) veriT_vr27))) (step t64.t1 (cl (! (= ?v0 veriT_vr26) :named @p_137)) :rule refl) (step t64.t2 (cl @p_137) :rule refl) (step t64.t3 (cl (! (= ?v1 veriT_vr27) :named @p_142)) :rule refl) (step t64.t4 (cl (= @p_138 (! (plus$ veriT_vr26 veriT_vr27) :named @p_139))) :rule cong :premises (t64.t2 t64.t3)) (step t64.t5 (cl (= @p_140 (! (= veriT_vr26 @p_139) :named @p_141))) :rule cong :premises (t64.t1 t64.t4)) (step t64.t6 (cl @p_142) :rule refl) (step t64.t7 (cl (= @p_143 (! (= zero$ veriT_vr27) :named @p_144))) :rule cong :premises (t64.t6)) (step t64.t8 (cl (= @p_145 (! (= @p_141 @p_144) :named @p_146))) :rule cong :premises (t64.t5 t64.t7)) (step t64 (cl (! (= @p_136 (! (forall ((veriT_vr26 Nat$) (veriT_vr27 Nat$)) @p_146) :named @p_148)) :named @p_147)) :rule bind) (step t65 (cl (not @p_147) (not @p_136) @p_148) :rule equiv_pos2) (step t66 (cl @p_148) :rule th_resolution :premises (axiom9 t64 t65)) (anchor :step t67 :args ((:= (veriT_vr26 Nat$) veriT_vr28) (:= (veriT_vr27 Nat$) veriT_vr29))) (step t67.t1 (cl (! (= veriT_vr26 veriT_vr28) :named @p_149)) :rule refl) (step t67.t2 (cl @p_149) :rule refl) (step t67.t3 (cl (! (= veriT_vr27 veriT_vr29) :named @p_152)) :rule refl) (step t67.t4 (cl (= @p_139 (! (plus$ veriT_vr28 veriT_vr29) :named @p_150))) :rule cong :premises (t67.t2 t67.t3)) (step t67.t5 (cl (= @p_141 (! (= veriT_vr28 @p_150) :named @p_151))) :rule cong :premises (t67.t1 t67.t4)) (step t67.t6 (cl @p_152) :rule refl) (step t67.t7 (cl (= @p_144 (! (= zero$ veriT_vr29) :named @p_153))) :rule cong :premises (t67.t6)) (step t67.t8 (cl (= @p_146 (! (= @p_151 @p_153) :named @p_154))) :rule cong :premises (t67.t5 t67.t7)) (step t67 (cl (! (= @p_148 (! (forall ((veriT_vr28 Nat$) (veriT_vr29 Nat$)) @p_154) :named @p_156)) :named @p_155)) :rule bind) (step t68 (cl (not @p_155) (not @p_148) @p_156) :rule equiv_pos2) (step t69 (cl @p_156) :rule th_resolution :premises (t66 t67 t68)) (anchor :step t70 :args ((:= (?v0 A_set$) veriT_vr30) (:= (?v1 A_set$) veriT_vr31))) (step t70.t1 (cl (! (= ?v0 veriT_vr30) :named @p_163)) :rule refl) (step t70.t2 (cl (= @p_4 (! (finite$ veriT_vr30) :named @p_158))) :rule cong :premises (t70.t1)) (step t70.t3 (cl (! (= ?v1 veriT_vr31) :named @p_165)) :rule refl) (step t70.t4 (cl (= @p_159 (! (finite$ veriT_vr31) :named @p_160))) :rule cong :premises (t70.t3)) (step t70.t5 (cl (= @p_161 (! (and @p_158 @p_160) :named @p_162))) :rule cong :premises (t70.t2 t70.t4)) (step t70.t6 (cl @p_163) :rule refl) (step t70.t7 (cl (= @p_3 (! (fun_app$b card$ veriT_vr30) :named @p_164))) :rule cong :premises (t70.t6)) (step t70.t8 (cl @p_165) :rule refl) (step t70.t9 (cl (= @p_166 (! (fun_app$b card$ veriT_vr31) :named @p_167))) :rule cong :premises (t70.t8)) (step t70.t10 (cl (= @p_168 (! (plus$ @p_164 @p_167) :named @p_169))) :rule cong :premises (t70.t7 t70.t9)) (step t70.t11 (cl @p_163) :rule refl) (step t70.t12 (cl @p_165) :rule refl) (step t70.t13 (cl (= @p_170 (! (sup$ veriT_vr30 veriT_vr31) :named @p_171))) :rule cong :premises (t70.t11 t70.t12)) (step t70.t14 (cl (= @p_172 (! (fun_app$b card$ @p_171) :named @p_173))) :rule cong :premises (t70.t13)) (step t70.t15 (cl @p_163) :rule refl) (step t70.t16 (cl @p_165) :rule refl) (step t70.t17 (cl (= @p_6 (! (inf$ veriT_vr30 veriT_vr31) :named @p_174))) :rule cong :premises (t70.t15 t70.t16)) (step t70.t18 (cl (= @p_175 (! (fun_app$b card$ @p_174) :named @p_176))) :rule cong :premises (t70.t17)) (step t70.t19 (cl (= @p_177 (! (plus$ @p_173 @p_176) :named @p_178))) :rule cong :premises (t70.t14 t70.t18)) (step t70.t20 (cl (= @p_179 (! (= @p_169 @p_178) :named @p_180))) :rule cong :premises (t70.t10 t70.t19)) (step t70.t21 (cl (= @p_181 (! (=> @p_162 @p_180) :named @p_182))) :rule cong :premises (t70.t5 t70.t20)) (step t70 (cl (! (= @p_157 (! (forall ((veriT_vr30 A_set$) (veriT_vr31 A_set$)) @p_182) :named @p_184)) :named @p_183)) :rule bind) (step t71 (cl (not @p_183) (not @p_157) @p_184) :rule equiv_pos2) (step t72 (cl @p_184) :rule th_resolution :premises (axiom12 t70 t71)) (anchor :step t73 :args ((:= (veriT_vr30 A_set$) veriT_vr32) (:= (veriT_vr31 A_set$) veriT_vr33))) (step t73.t1 (cl (! (= veriT_vr30 veriT_vr32) :named @p_188)) :rule refl) (step t73.t2 (cl (= @p_158 (! (finite$ veriT_vr32) :named @p_185))) :rule cong :premises (t73.t1)) (step t73.t3 (cl (! (= veriT_vr31 veriT_vr33) :named @p_190)) :rule refl) (step t73.t4 (cl (= @p_160 (! (finite$ veriT_vr33) :named @p_186))) :rule cong :premises (t73.t3)) (step t73.t5 (cl (= @p_162 (! (and @p_185 @p_186) :named @p_187))) :rule cong :premises (t73.t2 t73.t4)) (step t73.t6 (cl @p_188) :rule refl) (step t73.t7 (cl (= @p_164 (! (fun_app$b card$ veriT_vr32) :named @p_189))) :rule cong :premises (t73.t6)) (step t73.t8 (cl @p_190) :rule refl) (step t73.t9 (cl (= @p_167 (! (fun_app$b card$ veriT_vr33) :named @p_191))) :rule cong :premises (t73.t8)) (step t73.t10 (cl (= @p_169 (! (plus$ @p_189 @p_191) :named @p_192))) :rule cong :premises (t73.t7 t73.t9)) (step t73.t11 (cl @p_188) :rule refl) (step t73.t12 (cl @p_190) :rule refl) (step t73.t13 (cl (= @p_171 (! (sup$ veriT_vr32 veriT_vr33) :named @p_193))) :rule cong :premises (t73.t11 t73.t12)) (step t73.t14 (cl (= @p_173 (! (fun_app$b card$ @p_193) :named @p_194))) :rule cong :premises (t73.t13)) (step t73.t15 (cl @p_188) :rule refl) (step t73.t16 (cl @p_190) :rule refl) (step t73.t17 (cl (= @p_174 (! (inf$ veriT_vr32 veriT_vr33) :named @p_195))) :rule cong :premises (t73.t15 t73.t16)) (step t73.t18 (cl (= @p_176 (! (fun_app$b card$ @p_195) :named @p_196))) :rule cong :premises (t73.t17)) (step t73.t19 (cl (= @p_178 (! (plus$ @p_194 @p_196) :named @p_197))) :rule cong :premises (t73.t14 t73.t18)) (step t73.t20 (cl (= @p_180 (! (= @p_192 @p_197) :named @p_198))) :rule cong :premises (t73.t10 t73.t19)) (step t73.t21 (cl (= @p_182 (! (=> @p_187 @p_198) :named @p_199))) :rule cong :premises (t73.t5 t73.t20)) (step t73 (cl (! (= @p_184 (! (forall ((veriT_vr32 A_set$) (veriT_vr33 A_set$)) @p_199) :named @p_201)) :named @p_200)) :rule bind) (step t74 (cl (not @p_200) (not @p_184) @p_201) :rule equiv_pos2) (step t75 (cl @p_201) :rule th_resolution :premises (t72 t73 t74)) (anchor :step t76 :args ((:= (?v0 A_set$) veriT_vr34))) (step t76.t1 (cl (! (= ?v0 veriT_vr34) :named @p_206)) :rule refl) (step t76.t2 (cl (= @p_3 (! (fun_app$b card$ veriT_vr34) :named @p_203))) :rule cong :premises (t76.t1)) (step t76.t3 (cl (= @p_204 (! (= zero$ @p_203) :named @p_205))) :rule cong :premises (t76.t2)) (step t76.t4 (cl @p_206) :rule refl) (step t76.t5 (cl (= @p_207 (! (= bot$ veriT_vr34) :named @p_208))) :rule cong :premises (t76.t4)) (step t76.t6 (cl @p_206) :rule refl) (step t76.t7 (cl (= @p_4 (! (finite$ veriT_vr34) :named @p_209))) :rule cong :premises (t76.t6)) (step t76.t8 (cl (= @p_210 (! (not @p_209) :named @p_211))) :rule cong :premises (t76.t7)) (step t76.t9 (cl (= @p_212 (! (or @p_208 @p_211) :named @p_213))) :rule cong :premises (t76.t5 t76.t8)) (step t76.t10 (cl (= @p_214 (! (= @p_205 @p_213) :named @p_215))) :rule cong :premises (t76.t3 t76.t9)) (step t76 (cl (! (= @p_202 (! (forall ((veriT_vr34 A_set$)) @p_215) :named @p_217)) :named @p_216)) :rule bind) (step t77 (cl (not @p_216) (not @p_202) @p_217) :rule equiv_pos2) (step t78 (cl @p_217) :rule th_resolution :premises (axiom13 t76 t77)) (anchor :step t79 :args ((:= (veriT_vr34 A_set$) veriT_vr35))) (step t79.t1 (cl (! (= veriT_vr34 veriT_vr35) :named @p_220)) :rule refl) (step t79.t2 (cl (= @p_203 (! (fun_app$b card$ veriT_vr35) :named @p_218))) :rule cong :premises (t79.t1)) (step t79.t3 (cl (= @p_205 (! (= zero$ @p_218) :named @p_219))) :rule cong :premises (t79.t2)) (step t79.t4 (cl @p_220) :rule refl) (step t79.t5 (cl (= @p_208 (! (= bot$ veriT_vr35) :named @p_221))) :rule cong :premises (t79.t4)) (step t79.t6 (cl @p_220) :rule refl) (step t79.t7 (cl (= @p_209 (! (finite$ veriT_vr35) :named @p_222))) :rule cong :premises (t79.t6)) (step t79.t8 (cl (= @p_211 (! (not @p_222) :named @p_223))) :rule cong :premises (t79.t7)) (step t79.t9 (cl (= @p_213 (! (or @p_221 @p_223) :named @p_224))) :rule cong :premises (t79.t5 t79.t8)) (step t79.t10 (cl (= @p_215 (! (= @p_219 @p_224) :named @p_225))) :rule cong :premises (t79.t3 t79.t9)) (step t79 (cl (! (= @p_217 (! (forall ((veriT_vr35 A_set$)) @p_225) :named @p_227)) :named @p_226)) :rule bind) (step t80 (cl (not @p_226) (not @p_217) @p_227) :rule equiv_pos2) (step t81 (cl @p_227) :rule th_resolution :premises (t78 t79 t80)) (anchor :step t82 :args ((:= (?v0 A_set$) veriT_vr36))) (step t82.t1 (cl (! (= ?v0 veriT_vr36) :named @p_234)) :rule refl) (step t82.t2 (cl (= @p_3 (! (fun_app$b card$ veriT_vr36) :named @p_229))) :rule cong :premises (t82.t1)) (step t82.t3 (cl (= @p_230 (! (= @p_1 @p_229) :named @p_231))) :rule cong :premises (t82.t2)) (step t82.t4 (cl (= @p_232 (! (and @p_7 @p_231) :named @p_233))) :rule cong :premises (t82.t3)) (step t82.t5 (cl @p_234) :rule refl) (step t82.t6 (cl (= @p_235 (! (= top$ veriT_vr36) :named @p_236))) :rule cong :premises (t82.t5)) (step t82.t7 (cl (= @p_237 (! (=> @p_233 @p_236) :named @p_238))) :rule cong :premises (t82.t4 t82.t6)) (step t82 (cl (! (= @p_228 (! (forall ((veriT_vr36 A_set$)) @p_238) :named @p_240)) :named @p_239)) :rule bind) (step t83 (cl (not @p_239) (not @p_228) @p_240) :rule equiv_pos2) (step t84 (cl @p_240) :rule th_resolution :premises (axiom14 t82 t83)) (anchor :step t85 :args ((:= (veriT_vr36 A_set$) veriT_vr37))) (step t85.t1 (cl (! (= veriT_vr36 veriT_vr37) :named @p_244)) :rule refl) (step t85.t2 (cl (= @p_229 (! (fun_app$b card$ veriT_vr37) :named @p_241))) :rule cong :premises (t85.t1)) (step t85.t3 (cl (= @p_231 (! (= @p_1 @p_241) :named @p_242))) :rule cong :premises (t85.t2)) (step t85.t4 (cl (= @p_233 (! (and @p_7 @p_242) :named @p_243))) :rule cong :premises (t85.t3)) (step t85.t5 (cl @p_244) :rule refl) (step t85.t6 (cl (= @p_236 (! (= top$ veriT_vr37) :named @p_245))) :rule cong :premises (t85.t5)) (step t85.t7 (cl (= @p_238 (! (=> @p_243 @p_245) :named @p_246))) :rule cong :premises (t85.t4 t85.t6)) (step t85 (cl (! (= @p_240 (! (forall ((veriT_vr37 A_set$)) @p_246) :named @p_248)) :named @p_247)) :rule bind) (step t86 (cl (not @p_247) (not @p_240) @p_248) :rule equiv_pos2) (step t87 (cl @p_248) :rule th_resolution :premises (t84 t85 t86)) (anchor :step t88 :args ((:= (?v0 A_list$) veriT_vr38))) (step t88.t1 (cl (! (= ?v0 veriT_vr38) :named @p_251)) :rule refl) (step t88.t2 (cl (= @p_2 (! (set$ veriT_vr38) :named @p_250))) :rule cong :premises (t88.t1)) (step t88.t3 (cl @p_251) :rule refl) (step t88.t4 (cl (= @p_5 (! (coset$ veriT_vr38) :named @p_252))) :rule cong :premises (t88.t3)) (step t88.t5 (cl (= @p_253 (! (uminus$ @p_252) :named @p_254))) :rule cong :premises (t88.t4)) (step t88.t6 (cl (= @p_255 (! (= @p_250 @p_254) :named @p_256))) :rule cong :premises (t88.t2 t88.t5)) (step t88 (cl (! (= @p_249 (! (forall ((veriT_vr38 A_list$)) @p_256) :named @p_258)) :named @p_257)) :rule bind) (step t89 (cl (not @p_257) (not @p_249) @p_258) :rule equiv_pos2) (step t90 (cl @p_258) :rule th_resolution :premises (axiom15 t88 t89)) (anchor :step t91 :args ((:= (veriT_vr38 A_list$) veriT_vr39))) (step t91.t1 (cl (! (= veriT_vr38 veriT_vr39) :named @p_260)) :rule refl) (step t91.t2 (cl (= @p_250 (! (set$ veriT_vr39) :named @p_259))) :rule cong :premises (t91.t1)) (step t91.t3 (cl @p_260) :rule refl) (step t91.t4 (cl (= @p_252 (! (coset$ veriT_vr39) :named @p_261))) :rule cong :premises (t91.t3)) (step t91.t5 (cl (= @p_254 (! (uminus$ @p_261) :named @p_262))) :rule cong :premises (t91.t4)) (step t91.t6 (cl (= @p_256 (! (= @p_259 @p_262) :named @p_263))) :rule cong :premises (t91.t2 t91.t5)) (step t91 (cl (! (= @p_258 (! (forall ((veriT_vr39 A_list$)) @p_263) :named @p_265)) :named @p_264)) :rule bind) (step t92 (cl (not @p_264) (not @p_258) @p_265) :rule equiv_pos2) (step t93 (cl @p_265) :rule th_resolution :premises (t90 t91 t92)) (anchor :step t94 :args ((:= (?v0 A_list$) veriT_vr40))) (step t94.t1 (cl (! (= ?v0 veriT_vr40) :named @p_268)) :rule refl) (step t94.t2 (cl (= @p_5 (! (coset$ veriT_vr40) :named @p_267))) :rule cong :premises (t94.t1)) (step t94.t3 (cl @p_268) :rule refl) (step t94.t4 (cl (= @p_2 (! (set$ veriT_vr40) :named @p_269))) :rule cong :premises (t94.t3)) (step t94.t5 (cl (= @p_270 (! (uminus$ @p_269) :named @p_271))) :rule cong :premises (t94.t4)) (step t94.t6 (cl (= @p_272 (! (= @p_267 @p_271) :named @p_273))) :rule cong :premises (t94.t2 t94.t5)) (step t94 (cl (! (= @p_266 (! (forall ((veriT_vr40 A_list$)) @p_273) :named @p_275)) :named @p_274)) :rule bind) (step t95 (cl (not @p_274) (not @p_266) @p_275) :rule equiv_pos2) (step t96 (cl @p_275) :rule th_resolution :premises (axiom16 t94 t95)) (anchor :step t97 :args ((:= (veriT_vr40 A_list$) veriT_vr41))) (step t97.t1 (cl (! (= veriT_vr40 veriT_vr41) :named @p_277)) :rule refl) (step t97.t2 (cl (= @p_267 (! (coset$ veriT_vr41) :named @p_276))) :rule cong :premises (t97.t1)) (step t97.t3 (cl @p_277) :rule refl) (step t97.t4 (cl (= @p_269 (! (set$ veriT_vr41) :named @p_278))) :rule cong :premises (t97.t3)) (step t97.t5 (cl (= @p_271 (! (uminus$ @p_278) :named @p_279))) :rule cong :premises (t97.t4)) (step t97.t6 (cl (= @p_273 (! (= @p_276 @p_279) :named @p_280))) :rule cong :premises (t97.t2 t97.t5)) (step t97 (cl (! (= @p_275 (! (forall ((veriT_vr41 A_list$)) @p_280) :named @p_282)) :named @p_281)) :rule bind) (step t98 (cl (not @p_281) (not @p_275) @p_282) :rule equiv_pos2) (step t99 (cl @p_282) :rule th_resolution :premises (t96 t97 t98)) (anchor :step t100 :args ((:= (?v0 A_set$) veriT_vr42) (:= (?v1 A_set$) veriT_vr43))) (step t100.t1 (cl (! (= ?v0 veriT_vr42) :named @p_289)) :rule refl) (step t100.t2 (cl (! (= ?v1 veriT_vr43) :named @p_292)) :rule refl) (step t100.t3 (cl (= @p_6 (! (inf$ veriT_vr42 veriT_vr43) :named @p_284))) :rule cong :premises (t100.t1 t100.t2)) (step t100.t4 (cl (= @p_285 (! (= bot$ @p_284) :named @p_286))) :rule cong :premises (t100.t3)) (anchor :step t100.t5 :args ((:= (?v2 A$) veriT_vr44))) (step t100.t5.t1 (cl (! (= ?v2 veriT_vr44) :named @p_291)) :rule refl) (step t100.t5.t2 (cl @p_289) :rule refl) (step t100.t5.t3 (cl (= @p_15 (! (member$ veriT_vr44 veriT_vr42) :named @p_290))) :rule cong :premises (t100.t5.t1 t100.t5.t2)) (step t100.t5.t4 (cl @p_291) :rule refl) (step t100.t5.t5 (cl @p_292) :rule refl) (step t100.t5.t6 (cl (= @p_16 (! (member$ veriT_vr44 veriT_vr43) :named @p_293))) :rule cong :premises (t100.t5.t4 t100.t5.t5)) (step t100.t5.t7 (cl (= @p_294 (! (not @p_293) :named @p_295))) :rule cong :premises (t100.t5.t6)) (step t100.t5.t8 (cl (= @p_296 (! (=> @p_290 @p_295) :named @p_297))) :rule cong :premises (t100.t5.t3 t100.t5.t7)) (step t100.t5 (cl (= @p_287 (! (forall ((veriT_vr44 A$)) @p_297) :named @p_288))) :rule bind) (step t100.t6 (cl (= @p_298 (! (= @p_286 @p_288) :named @p_299))) :rule cong :premises (t100.t4 t100.t5)) (step t100 (cl (! (= @p_283 (! (forall ((veriT_vr42 A_set$) (veriT_vr43 A_set$)) @p_299) :named @p_301)) :named @p_300)) :rule bind) (step t101 (cl (not @p_300) (not @p_283) @p_301) :rule equiv_pos2) (step t102 (cl @p_301) :rule th_resolution :premises (axiom17 t100 t101)) (anchor :step t103 :args ((veriT_vr42 A_set$) (veriT_vr43 A_set$))) (step t103.t1 (cl (= @p_299 (! (and (! (=> @p_286 @p_288) :named @p_315) (! (=> @p_288 @p_286) :named @p_325)) :named @p_302))) :rule connective_def) (step t103 (cl (! (= @p_301 (! (forall ((veriT_vr42 A_set$) (veriT_vr43 A_set$)) @p_302) :named @p_304)) :named @p_303)) :rule bind) (step t104 (cl (not @p_303) (not @p_301) @p_304) :rule equiv_pos2) (step t105 (cl @p_304) :rule th_resolution :premises (t102 t103 t104)) (anchor :step t106 :args ((:= (veriT_vr42 A_set$) veriT_vr45) (:= (veriT_vr43 A_set$) veriT_vr46))) (step t106.t1 (cl (! (= veriT_vr42 veriT_vr45) :named @p_308)) :rule refl) (step t106.t2 (cl (! (= veriT_vr43 veriT_vr46) :named @p_311)) :rule refl) (step t106.t3 (cl (! (= @p_284 (! (inf$ veriT_vr45 veriT_vr46) :named @p_306)) :named @p_323)) :rule cong :premises (t106.t1 t106.t2)) (step t106.t4 (cl (! (= @p_286 (! (= bot$ @p_306) :named @p_305)) :named @p_324)) :rule cong :premises (t106.t3)) (anchor :step t106.t5 :args ((:= (veriT_vr44 A$) veriT_vr47))) (step t106.t5.t1 (cl (! (= veriT_vr44 veriT_vr47) :named @p_310)) :rule refl) (step t106.t5.t2 (cl @p_308) :rule refl) (step t106.t5.t3 (cl (= @p_290 (! (member$ veriT_vr47 veriT_vr45) :named @p_309))) :rule cong :premises (t106.t5.t1 t106.t5.t2)) (step t106.t5.t4 (cl @p_310) :rule refl) (step t106.t5.t5 (cl @p_311) :rule refl) (step t106.t5.t6 (cl (= @p_293 (! (member$ veriT_vr47 veriT_vr46) :named @p_312))) :rule cong :premises (t106.t5.t4 t106.t5.t5)) (step t106.t5.t7 (cl (= @p_295 (! (not @p_312) :named @p_313))) :rule cong :premises (t106.t5.t6)) (step t106.t5.t8 (cl (= @p_297 (! (=> @p_309 @p_313) :named @p_314))) :rule cong :premises (t106.t5.t3 t106.t5.t7)) (step t106.t5 (cl (= @p_288 (! (forall ((veriT_vr47 A$)) @p_314) :named @p_307))) :rule bind) (step t106.t6 (cl (= @p_315 (! (=> @p_305 @p_307) :named @p_316))) :rule cong :premises (t106.t4 t106.t5)) (anchor :step t106.t7 :args ((:= (veriT_vr44 A$) veriT_vr48))) (step t106.t7.t1 (cl (! (= veriT_vr44 veriT_vr48) :named @p_319)) :rule refl) (step t106.t7.t2 (cl @p_308) :rule refl) (step t106.t7.t3 (cl (= @p_290 (! (member$ veriT_vr48 veriT_vr45) :named @p_318))) :rule cong :premises (t106.t7.t1 t106.t7.t2)) (step t106.t7.t4 (cl @p_319) :rule refl) (step t106.t7.t5 (cl @p_311) :rule refl) (step t106.t7.t6 (cl (= @p_293 (! (member$ veriT_vr48 veriT_vr46) :named @p_320))) :rule cong :premises (t106.t7.t4 t106.t7.t5)) (step t106.t7.t7 (cl (= @p_295 (! (not @p_320) :named @p_321))) :rule cong :premises (t106.t7.t6)) (step t106.t7.t8 (cl (= @p_297 (! (=> @p_318 @p_321) :named @p_322))) :rule cong :premises (t106.t7.t3 t106.t7.t7)) (step t106.t7 (cl (= @p_288 (! (forall ((veriT_vr48 A$)) @p_322) :named @p_317))) :rule bind) (step t106.t8 (cl @p_308) :rule refl) (step t106.t9 (cl @p_311) :rule refl) (step t106.t10 (cl @p_323) :rule cong :premises (t106.t8 t106.t9)) (step t106.t11 (cl @p_324) :rule cong :premises (t106.t10)) (step t106.t12 (cl (= @p_325 (! (=> @p_317 @p_305) :named @p_326))) :rule cong :premises (t106.t7 t106.t11)) (step t106.t13 (cl (= @p_302 (! (and @p_316 @p_326) :named @p_327))) :rule cong :premises (t106.t6 t106.t12)) (step t106 (cl (! (= @p_304 (! (forall ((veriT_vr45 A_set$) (veriT_vr46 A_set$)) @p_327) :named @p_329)) :named @p_328)) :rule bind) (step t107 (cl (not @p_328) (not @p_304) @p_329) :rule equiv_pos2) (step t108 (cl @p_329) :rule th_resolution :premises (t105 t106 t107)) (anchor :step t109 :args ((:= (veriT_vr45 A_set$) veriT_vr49) (:= (veriT_vr46 A_set$) veriT_vr50))) (step t109.t1 (cl (! (= veriT_vr45 veriT_vr49) :named @p_333)) :rule refl) (step t109.t2 (cl (! (= veriT_vr46 veriT_vr50) :named @p_336)) :rule refl) (step t109.t3 (cl (! (= @p_306 (! (inf$ veriT_vr49 veriT_vr50) :named @p_332)) :named @p_342)) :rule cong :premises (t109.t1 t109.t2)) (step t109.t4 (cl (! (= @p_305 (! (= bot$ @p_332) :named @p_331)) :named @p_343)) :rule cong :premises (t109.t3)) (anchor :step t109.t5 :args ((:= (veriT_vr47 A$) veriT_vr51))) (step t109.t5.t1 (cl (! (= veriT_vr47 veriT_vr51) :named @p_335)) :rule refl) (step t109.t5.t2 (cl @p_333) :rule refl) (step t109.t5.t3 (cl (= @p_309 (! (member$ veriT_vr51 veriT_vr49) :named @p_334))) :rule cong :premises (t109.t5.t1 t109.t5.t2)) (step t109.t5.t4 (cl @p_335) :rule refl) (step t109.t5.t5 (cl @p_336) :rule refl) (step t109.t5.t6 (cl (= @p_312 (! (member$ veriT_vr51 veriT_vr50) :named @p_337))) :rule cong :premises (t109.t5.t4 t109.t5.t5)) (step t109.t5.t7 (cl (= @p_313 (! (not @p_337) :named @p_338))) :rule cong :premises (t109.t5.t6)) (step t109.t5.t8 (cl (= @p_314 (! (=> @p_334 @p_338) :named @p_339))) :rule cong :premises (t109.t5.t3 t109.t5.t7)) (step t109.t5 (cl (= @p_307 (! (forall ((veriT_vr51 A$)) @p_339) :named @p_330))) :rule bind) (step t109.t6 (cl (= @p_316 (! (=> @p_331 @p_330) :named @p_340))) :rule cong :premises (t109.t4 t109.t5)) (anchor :step t109.t7 :args ((:= (veriT_vr48 A$) veriT_vr51))) (step t109.t7.t1 (cl (! (= veriT_vr48 veriT_vr51) :named @p_341)) :rule refl) (step t109.t7.t2 (cl @p_333) :rule refl) (step t109.t7.t3 (cl (= @p_318 @p_334)) :rule cong :premises (t109.t7.t1 t109.t7.t2)) (step t109.t7.t4 (cl @p_341) :rule refl) (step t109.t7.t5 (cl @p_336) :rule refl) (step t109.t7.t6 (cl (= @p_320 @p_337)) :rule cong :premises (t109.t7.t4 t109.t7.t5)) (step t109.t7.t7 (cl (= @p_321 @p_338)) :rule cong :premises (t109.t7.t6)) (step t109.t7.t8 (cl (= @p_322 @p_339)) :rule cong :premises (t109.t7.t3 t109.t7.t7)) (step t109.t7 (cl (= @p_317 @p_330)) :rule bind) (step t109.t8 (cl @p_333) :rule refl) (step t109.t9 (cl @p_336) :rule refl) (step t109.t10 (cl @p_342) :rule cong :premises (t109.t8 t109.t9)) (step t109.t11 (cl @p_343) :rule cong :premises (t109.t10)) (step t109.t12 (cl (= @p_326 (! (=> @p_330 @p_331) :named @p_344))) :rule cong :premises (t109.t7 t109.t11)) (step t109.t13 (cl (= @p_327 (! (and @p_340 @p_344) :named @p_345))) :rule cong :premises (t109.t6 t109.t12)) (step t109 (cl (! (= @p_329 (! (forall ((veriT_vr49 A_set$) (veriT_vr50 A_set$)) @p_345) :named @p_347)) :named @p_346)) :rule bind) (step t110 (cl (not @p_346) (not @p_329) @p_347) :rule equiv_pos2) (step t111 (cl @p_347) :rule th_resolution :premises (t108 t109 t110)) (anchor :step t112 :args ((:= (veriT_vr49 A_set$) veriT_vr49) (:= (veriT_vr50 A_set$) veriT_vr50))) (anchor :step t112.t1 :args ((:= (veriT_vr51 A$) veriT_vr52))) (step t112.t1.t1 (cl (! (= veriT_vr51 veriT_vr52) :named @p_350)) :rule refl) (step t112.t1.t2 (cl (= @p_334 (! (member$ veriT_vr52 veriT_vr49) :named @p_349))) :rule cong :premises (t112.t1.t1)) (step t112.t1.t3 (cl @p_350) :rule refl) (step t112.t1.t4 (cl (= @p_337 (! (member$ veriT_vr52 veriT_vr50) :named @p_351))) :rule cong :premises (t112.t1.t3)) (step t112.t1.t5 (cl (= @p_338 (! (not @p_351) :named @p_352))) :rule cong :premises (t112.t1.t4)) (step t112.t1.t6 (cl (= @p_339 (! (=> @p_349 @p_352) :named @p_353))) :rule cong :premises (t112.t1.t2 t112.t1.t5)) (step t112.t1 (cl (= @p_330 (! (forall ((veriT_vr52 A$)) @p_353) :named @p_348))) :rule bind) (step t112.t2 (cl (= @p_344 (! (=> @p_348 @p_331) :named @p_354))) :rule cong :premises (t112.t1)) (step t112.t3 (cl (= @p_345 (! (and @p_340 @p_354) :named @p_355))) :rule cong :premises (t112.t2)) (step t112 (cl (! (= @p_347 (! (forall ((veriT_vr49 A_set$) (veriT_vr50 A_set$)) @p_355) :named @p_357)) :named @p_356)) :rule bind) (step t113 (cl (not @p_356) (not @p_347) @p_357) :rule equiv_pos2) (step t114 (cl @p_357) :rule th_resolution :premises (t111 t112 t113)) (anchor :step t115 :args ((:= (?v0 A_set$) veriT_vr53))) (step t115.t1 (cl (! (= ?v0 veriT_vr53) :named @p_360)) :rule refl) (step t115.t2 (cl (= @p_4 (! (finite$ veriT_vr53) :named @p_359))) :rule cong :premises (t115.t1)) (step t115.t3 (cl @p_360) :rule refl) (step t115.t4 (cl (= @p_361 (! (uminus$ veriT_vr53) :named @p_362))) :rule cong :premises (t115.t3)) (step t115.t5 (cl (= @p_363 (! (finite$ @p_362) :named @p_364))) :rule cong :premises (t115.t4)) (step t115.t6 (cl (= @p_365 (! (= @p_364 @p_7) :named @p_366))) :rule cong :premises (t115.t5)) (step t115.t7 (cl (= @p_367 (! (=> @p_359 @p_366) :named @p_368))) :rule cong :premises (t115.t2 t115.t6)) (step t115 (cl (! (= @p_358 (! (forall ((veriT_vr53 A_set$)) @p_368) :named @p_370)) :named @p_369)) :rule bind) (step t116 (cl (not @p_369) (not @p_358) @p_370) :rule equiv_pos2) (step t117 (cl @p_370) :rule th_resolution :premises (axiom19 t115 t116)) (anchor :step t118 :args ((:= (veriT_vr53 A_set$) veriT_vr54))) (step t118.t1 (cl (! (= veriT_vr53 veriT_vr54) :named @p_372)) :rule refl) (step t118.t2 (cl (= @p_359 (! (finite$ veriT_vr54) :named @p_371))) :rule cong :premises (t118.t1)) (step t118.t3 (cl @p_372) :rule refl) (step t118.t4 (cl (= @p_362 (! (uminus$ veriT_vr54) :named @p_373))) :rule cong :premises (t118.t3)) (step t118.t5 (cl (= @p_364 (! (finite$ @p_373) :named @p_374))) :rule cong :premises (t118.t4)) (step t118.t6 (cl (= @p_366 (! (= @p_374 @p_7) :named @p_375))) :rule cong :premises (t118.t5)) (step t118.t7 (cl (= @p_368 (! (=> @p_371 @p_375) :named @p_376))) :rule cong :premises (t118.t2 t118.t6)) (step t118 (cl (! (= @p_370 (! (forall ((veriT_vr54 A_set$)) @p_376) :named @p_378)) :named @p_377)) :rule bind) (step t119 (cl (not @p_377) (not @p_370) @p_378) :rule equiv_pos2) (step t120 (cl @p_378) :rule th_resolution :premises (t117 t118 t119)) (anchor :step t121 :args ((:= (?v0 A$) veriT_vr55))) (step t121.t1 (cl (! (= ?v0 veriT_vr55) :named @p_383)) :rule refl) (step t121.t2 (cl (! (= @p_12 (! (member$ veriT_vr55 @p_380) :named @p_382)) :named @p_391)) :rule cong :premises (t121.t1)) (step t121.t3 (cl @p_383) :rule refl) (step t121.t4 (cl (! (= @p_11 (! (member$ veriT_vr55 @p_381) :named @p_384)) :named @p_390)) :rule cong :premises (t121.t3)) (step t121.t5 (cl (= @p_385 (! (not @p_384) :named @p_386))) :rule cong :premises (t121.t4)) (step t121.t6 (cl (= @p_387 (! (=> @p_382 @p_386) :named @p_388))) :rule cong :premises (t121.t2 t121.t5)) (step t121 (cl (= @p_379 (! (forall ((veriT_vr55 A$)) @p_388) :named @p_397))) :rule bind) (anchor :step t122 :args ((:= (?v0 A$) veriT_vr55))) (step t122.t1 (cl @p_383) :rule refl) (step t122.t2 (cl @p_390) :rule cong :premises (t122.t1)) (step t122.t3 (cl @p_383) :rule refl) (step t122.t4 (cl @p_391) :rule cong :premises (t122.t3)) (step t122.t5 (cl (= @p_392 (! (not @p_382) :named @p_393))) :rule cong :premises (t122.t4)) (step t122.t6 (cl (= @p_394 (! (=> @p_384 @p_393) :named @p_395))) :rule cong :premises (t122.t2 t122.t5)) (step t122 (cl (= @p_389 (! (forall ((veriT_vr55 A$)) @p_395) :named @p_398))) :rule bind) (step t123 (cl (= @p_396 (! (and @p_397 @p_398) :named @p_401))) :rule cong :premises (t121 t122)) (step t124 (cl (= @p_399 (! (and @p_400 @p_401) :named @p_404))) :rule cong :premises (t123)) (step t125 (cl (= @p_402 (! (ite @p_403 false @p_404) :named @p_406))) :rule cong :premises (t124)) (step t126 (cl (! (= @p_405 (! (= rhs$ @p_406) :named @p_408)) :named @p_407)) :rule cong :premises (t125)) (step t127 (cl (not @p_407) (not @p_405) @p_408) :rule equiv_pos2) (step t128 (cl @p_408) :rule th_resolution :premises (axiom20 t126 t127)) (step t129 (cl (= @p_404 (! (and @p_400 @p_397 @p_398) :named @p_409))) :rule ac_simp) (step t130 (cl (= @p_406 (! (ite @p_403 false @p_409) :named @p_410))) :rule cong :premises (t129)) (step t131 (cl (! (= @p_408 (! (= rhs$ @p_410) :named @p_412)) :named @p_411)) :rule cong :premises (t130)) (step t132 (cl (not @p_411) (not @p_408) @p_412) :rule equiv_pos2) (step t133 (cl @p_412) :rule th_resolution :premises (t128 t131 t132)) (step t134 (cl (= @p_410 (! (and (! (not @p_403) :named @p_422) @p_409) :named @p_413))) :rule ite_simplify) (step t135 (cl (! (= @p_412 (! (= rhs$ @p_413) :named @p_415)) :named @p_414)) :rule cong :premises (t134)) (step t136 (cl (not @p_414) (not @p_412) @p_415) :rule equiv_pos2) (step t137 (cl @p_415) :rule th_resolution :premises (t133 t135 t136)) (anchor :step t138 :args ((:= (veriT_vr55 A$) veriT_vr56))) (step t138.t1 (cl (! (= veriT_vr55 veriT_vr56) :named @p_417)) :rule refl) (step t138.t2 (cl (= @p_384 (! (member$ veriT_vr56 @p_381) :named @p_416))) :rule cong :premises (t138.t1)) (step t138.t3 (cl @p_417) :rule refl) (step t138.t4 (cl (= @p_382 (! (member$ veriT_vr56 @p_380) :named @p_418))) :rule cong :premises (t138.t3)) (step t138.t5 (cl (= @p_393 (! (not @p_418) :named @p_419))) :rule cong :premises (t138.t4)) (step t138.t6 (cl (= @p_395 (! (=> @p_416 @p_419) :named @p_420))) :rule cong :premises (t138.t2 t138.t5)) (step t138 (cl (= @p_398 (! (forall ((veriT_vr56 A$)) @p_420) :named @p_421))) :rule bind) (step t139 (cl (= @p_409 (! (and @p_400 @p_397 @p_421) :named @p_423))) :rule cong :premises (t138)) (step t140 (cl (= @p_413 (! (and @p_422 @p_423) :named @p_424))) :rule cong :premises (t139)) (step t141 (cl (! (= @p_415 (! (= rhs$ @p_424) :named @p_426)) :named @p_425)) :rule cong :premises (t140)) (step t142 (cl (not @p_425) (not @p_415) @p_426) :rule equiv_pos2) (step t143 (cl @p_426) :rule th_resolution :premises (t137 t141 t142)) (step t144 (cl (! (= @p_426 (! (and (! (=> rhs$ @p_424) :named @p_442) (! (=> @p_424 rhs$) :named @p_444)) :named @p_428)) :named @p_427)) :rule connective_def) (step t145 (cl (not @p_427) (not @p_426) @p_428) :rule equiv_pos2) (step t146 (cl @p_428) :rule th_resolution :premises (t143 t144 t145)) (anchor :step t147 :args ((:= (veriT_vr55 A$) veriT_vr57))) (step t147.t1 (cl (! (= veriT_vr55 veriT_vr57) :named @p_430)) :rule refl) (step t147.t2 (cl (= @p_382 @p_429)) :rule cong :premises (t147.t1)) (step t147.t3 (cl @p_430) :rule refl) (step t147.t4 (cl (= @p_384 @p_431)) :rule cong :premises (t147.t3)) (step t147.t5 (cl (= @p_386 @p_432)) :rule cong :premises (t147.t4)) (step t147.t6 (cl (= @p_388 @p_433)) :rule cong :premises (t147.t2 t147.t5)) (step t147 (cl (= @p_397 (! (forall ((veriT_vr57 A$)) @p_433) :named @p_439))) :rule bind) (anchor :step t148 :args ((:= (veriT_vr56 A$) veriT_vr58))) (step t148.t1 (cl (! (= veriT_vr56 veriT_vr58) :named @p_435)) :rule refl) (step t148.t2 (cl (= @p_416 @p_434)) :rule cong :premises (t148.t1)) (step t148.t3 (cl @p_435) :rule refl) (step t148.t4 (cl (= @p_418 @p_436)) :rule cong :premises (t148.t3)) (step t148.t5 (cl (= @p_419 @p_437)) :rule cong :premises (t148.t4)) (step t148.t6 (cl (= @p_420 @p_438)) :rule cong :premises (t148.t2 t148.t5)) (step t148 (cl (= @p_421 (! (forall ((veriT_vr58 A$)) @p_438) :named @p_440))) :rule bind) (step t149 (cl (= @p_423 (! (and @p_400 @p_439 @p_440) :named @p_441))) :rule cong :premises (t147 t148)) (step t150 (cl (= @p_424 (! (and @p_422 @p_441) :named @p_443))) :rule cong :premises (t149)) (step t151 (cl (= @p_442 (! (=> rhs$ @p_443) :named @p_445))) :rule cong :premises (t150)) (step t152 (cl (= @p_444 (! (=> @p_443 rhs$) :named @p_446))) :rule cong :premises (t150)) (step t153 (cl (! (= @p_428 (! (and @p_445 @p_446) :named @p_448)) :named @p_447)) :rule cong :premises (t151 t152)) (step t154 (cl (not @p_447) (not @p_428) @p_448) :rule equiv_pos2) (step t155 (cl @p_448) :rule th_resolution :premises (t146 t153 t154)) (anchor :step t156 :args ((:= (veriT_vr57 A$) veriT_sk0))) (step t156.t1 (cl (! (= veriT_vr57 veriT_sk0) :named @p_451)) :rule refl) (step t156.t2 (cl (= @p_429 (! (member$ veriT_sk0 @p_380) :named @p_450))) :rule cong :premises (t156.t1)) (step t156.t3 (cl @p_451) :rule refl) (step t156.t4 (cl (= @p_431 (! (member$ veriT_sk0 @p_381) :named @p_452))) :rule cong :premises (t156.t3)) (step t156.t5 (cl (= @p_432 (! (not @p_452) :named @p_453))) :rule cong :premises (t156.t4)) (step t156.t6 (cl (= @p_433 (! (=> @p_450 @p_453) :named @p_454))) :rule cong :premises (t156.t2 t156.t5)) (step t156 (cl (= @p_439 @p_454)) :rule sko_forall) (anchor :step t157 :args ((:= (veriT_vr58 A$) veriT_sk1))) (step t157.t1 (cl (! (= veriT_vr58 veriT_sk1) :named @p_457)) :rule refl) (step t157.t2 (cl (= @p_434 (! (member$ veriT_sk1 @p_381) :named @p_456))) :rule cong :premises (t157.t1)) (step t157.t3 (cl @p_457) :rule refl) (step t157.t4 (cl (= @p_436 (! (member$ veriT_sk1 @p_380) :named @p_458))) :rule cong :premises (t157.t3)) (step t157.t5 (cl (= @p_437 (! (not @p_458) :named @p_459))) :rule cong :premises (t157.t4)) (step t157.t6 (cl (= @p_438 (! (=> @p_456 @p_459) :named @p_460))) :rule cong :premises (t157.t2 t157.t5)) (step t157 (cl (= @p_440 @p_460)) :rule sko_forall) (step t158 (cl (= @p_441 (! (and @p_400 @p_454 @p_460) :named @p_461))) :rule cong :premises (t156 t157)) (step t159 (cl (= @p_443 (! (and @p_422 @p_461) :named @p_462))) :rule cong :premises (t158)) (step t160 (cl (= @p_446 (! (=> @p_462 rhs$) :named @p_463))) :rule cong :premises (t159)) (step t161 (cl (! (= @p_448 (! (and @p_445 @p_463) :named @p_465)) :named @p_464)) :rule cong :premises (t160)) (step t162 (cl (not @p_464) (not @p_448) @p_465) :rule equiv_pos2) (step t163 (cl @p_465) :rule th_resolution :premises (t155 t161 t162)) (anchor :step t164 :args ((:= (veriT_vr57 A$) veriT_vr59))) (step t164.t1 (cl (! (= veriT_vr57 veriT_vr59) :named @p_467)) :rule refl) (step t164.t2 (cl (= @p_429 (! (member$ veriT_vr59 @p_380) :named @p_466))) :rule cong :premises (t164.t1)) (step t164.t3 (cl @p_467) :rule refl) (step t164.t4 (cl (= @p_431 (! (member$ veriT_vr59 @p_381) :named @p_468))) :rule cong :premises (t164.t3)) (step t164.t5 (cl (= @p_432 (! (not @p_468) :named @p_469))) :rule cong :premises (t164.t4)) (step t164.t6 (cl (= @p_433 (! (=> @p_466 @p_469) :named @p_470))) :rule cong :premises (t164.t2 t164.t5)) (step t164 (cl (= @p_439 (! (forall ((veriT_vr59 A$)) @p_470) :named @p_474))) :rule bind) (anchor :step t165 :args ((:= (veriT_vr58 A$) veriT_vr59))) (step t165.t1 (cl (! (= veriT_vr58 veriT_vr59) :named @p_471)) :rule refl) (step t165.t2 (cl (= @p_434 @p_468)) :rule cong :premises (t165.t1)) (step t165.t3 (cl @p_471) :rule refl) (step t165.t4 (cl (= @p_436 @p_466)) :rule cong :premises (t165.t3)) (step t165.t5 (cl (= @p_437 (! (not @p_466) :named @p_472))) :rule cong :premises (t165.t4)) (step t165.t6 (cl (= @p_438 (! (=> @p_468 @p_472) :named @p_473))) :rule cong :premises (t165.t2 t165.t5)) (step t165 (cl (= @p_440 (! (forall ((veriT_vr59 A$)) @p_473) :named @p_475))) :rule bind) (step t166 (cl (= @p_441 (! (and @p_400 @p_474 @p_475) :named @p_476))) :rule cong :premises (t164 t165)) (step t167 (cl (= @p_443 (! (and @p_422 @p_476) :named @p_477))) :rule cong :premises (t166)) (step t168 (cl (= @p_445 (! (=> rhs$ @p_477) :named @p_478))) :rule cong :premises (t167)) (step t169 (cl (! (= @p_465 (! (and @p_478 @p_463) :named @p_480)) :named @p_479)) :rule cong :premises (t168)) (step t170 (cl (not @p_479) (not @p_465) @p_480) :rule equiv_pos2) (step t171 (cl @p_480) :rule th_resolution :premises (t163 t169 t170)) (step t172 (cl (= @p_477 (! (and @p_422 @p_400 @p_474 @p_475) :named @p_481))) :rule ac_simp) (step t173 (cl (= @p_478 (! (=> rhs$ @p_481) :named @p_483))) :rule cong :premises (t172)) (step t174 (cl (= @p_462 (! (and @p_422 @p_400 @p_454 @p_460) :named @p_482))) :rule ac_simp) (step t175 (cl (= @p_463 (! (=> @p_482 rhs$) :named @p_484))) :rule cong :premises (t174)) (step t176 (cl (! (= @p_480 (! (and @p_483 @p_484) :named @p_486)) :named @p_485)) :rule ac_simp :premises (t173 t175)) (step t177 (cl (not @p_485) (not @p_480) @p_486) :rule equiv_pos2) (step t178 (cl @p_486) :rule th_resolution :premises (t171 t176 t177)) (anchor :step t179 :args ((:= (veriT_vr59 A$) veriT_vr60))) (step t179.t1 (cl (! (= veriT_vr59 veriT_vr60) :named @p_488)) :rule refl) (step t179.t2 (cl (= @p_468 (! (member$ veriT_vr60 @p_381) :named @p_487))) :rule cong :premises (t179.t1)) (step t179.t3 (cl @p_488) :rule refl) (step t179.t4 (cl (= @p_466 (! (member$ veriT_vr60 @p_380) :named @p_489))) :rule cong :premises (t179.t3)) (step t179.t5 (cl (= @p_472 (! (not @p_489) :named @p_490))) :rule cong :premises (t179.t4)) (step t179.t6 (cl (= @p_473 (! (=> @p_487 @p_490) :named @p_491))) :rule cong :premises (t179.t2 t179.t5)) (step t179 (cl (= @p_475 (! (forall ((veriT_vr60 A$)) @p_491) :named @p_492))) :rule bind) (step t180 (cl (= @p_481 (! (and @p_422 @p_400 @p_474 @p_492) :named @p_493))) :rule cong :premises (t179)) (step t181 (cl (= @p_483 (! (=> rhs$ @p_493) :named @p_494))) :rule cong :premises (t180)) (step t182 (cl (! (= @p_486 (! (and @p_494 @p_484) :named @p_496)) :named @p_495)) :rule cong :premises (t181)) (step t183 (cl (not @p_495) (not @p_486) @p_496) :rule equiv_pos2) (step t184 (cl @p_496) :rule th_resolution :premises (t178 t182 t183)) (anchor :step t185 :args ((:= (?v0 A_list$) veriT_vr61) (:= (?v1 A_list$) veriT_vr62))) (step t185.t1 (cl (! (= ?v0 veriT_vr61) :named @p_502)) :rule refl) (step t185.t2 (cl (! (= ?v1 veriT_vr62) :named @p_504)) :rule refl) (step t185.t3 (cl (= @p_498 (! (append$ veriT_vr61 veriT_vr62) :named @p_499))) :rule cong :premises (t185.t1 t185.t2)) (step t185.t4 (cl (= @p_500 (! (set$ @p_499) :named @p_501))) :rule cong :premises (t185.t3)) (step t185.t5 (cl @p_502) :rule refl) (step t185.t6 (cl (= @p_2 (! (set$ veriT_vr61) :named @p_503))) :rule cong :premises (t185.t5)) (step t185.t7 (cl @p_504) :rule refl) (step t185.t8 (cl (= @p_20 (! (set$ veriT_vr62) :named @p_505))) :rule cong :premises (t185.t7)) (step t185.t9 (cl (= @p_506 (! (sup$ @p_503 @p_505) :named @p_507))) :rule cong :premises (t185.t6 t185.t8)) (step t185.t10 (cl (= @p_508 (! (= @p_501 @p_507) :named @p_509))) :rule cong :premises (t185.t4 t185.t9)) (step t185 (cl (! (= @p_497 (! (forall ((veriT_vr61 A_list$) (veriT_vr62 A_list$)) @p_509) :named @p_511)) :named @p_510)) :rule bind) (step t186 (cl (not @p_510) (not @p_497) @p_511) :rule equiv_pos2) (step t187 (cl @p_511) :rule th_resolution :premises (axiom21 t185 t186)) (anchor :step t188 :args ((:= (veriT_vr61 A_list$) veriT_vr63) (:= (veriT_vr62 A_list$) veriT_vr64))) (step t188.t1 (cl (! (= veriT_vr61 veriT_vr63) :named @p_514)) :rule refl) (step t188.t2 (cl (! (= veriT_vr62 veriT_vr64) :named @p_516)) :rule refl) (step t188.t3 (cl (= @p_499 (! (append$ veriT_vr63 veriT_vr64) :named @p_512))) :rule cong :premises (t188.t1 t188.t2)) (step t188.t4 (cl (= @p_501 (! (set$ @p_512) :named @p_513))) :rule cong :premises (t188.t3)) (step t188.t5 (cl @p_514) :rule refl) (step t188.t6 (cl (= @p_503 (! (set$ veriT_vr63) :named @p_515))) :rule cong :premises (t188.t5)) (step t188.t7 (cl @p_516) :rule refl) (step t188.t8 (cl (= @p_505 (! (set$ veriT_vr64) :named @p_517))) :rule cong :premises (t188.t7)) (step t188.t9 (cl (= @p_507 (! (sup$ @p_515 @p_517) :named @p_518))) :rule cong :premises (t188.t6 t188.t8)) (step t188.t10 (cl (= @p_509 (! (= @p_513 @p_518) :named @p_519))) :rule cong :premises (t188.t4 t188.t9)) (step t188 (cl (! (= @p_511 (! (forall ((veriT_vr63 A_list$) (veriT_vr64 A_list$)) @p_519) :named @p_521)) :named @p_520)) :rule bind) (step t189 (cl (not @p_520) (not @p_511) @p_521) :rule equiv_pos2) (step t190 (cl @p_521) :rule th_resolution :premises (t187 t188 t189)) (anchor :step t191 :args ((:= (?v0 A_set$) veriT_vr65) (:= (?v1 A_set$) veriT_vr66))) (step t191.t1 (cl (! (= ?v0 veriT_vr65) :named @p_524)) :rule refl) (step t191.t2 (cl (! (= ?v1 veriT_vr66) :named @p_526)) :rule refl) (step t191.t3 (cl (= @p_13 (! (= veriT_vr65 veriT_vr66) :named @p_523))) :rule cong :premises (t191.t1 t191.t2)) (step t191.t4 (cl @p_524) :rule refl) (step t191.t5 (cl (= @p_19 (! (fun_app$a less_eq$ veriT_vr65) :named @p_525))) :rule cong :premises (t191.t4)) (step t191.t6 (cl @p_526) :rule refl) (step t191.t7 (cl (= @p_17 (! (fun_app$ @p_525 veriT_vr66) :named @p_527))) :rule cong :premises (t191.t5 t191.t6)) (step t191.t8 (cl @p_526) :rule refl) (step t191.t9 (cl (= @p_528 (! (fun_app$a less_eq$ veriT_vr66) :named @p_529))) :rule cong :premises (t191.t8)) (step t191.t10 (cl @p_524) :rule refl) (step t191.t11 (cl (= @p_530 (! (fun_app$ @p_529 veriT_vr65) :named @p_531))) :rule cong :premises (t191.t9 t191.t10)) (step t191.t12 (cl (= @p_18 (! (and @p_527 @p_531) :named @p_532))) :rule cong :premises (t191.t7 t191.t11)) (step t191.t13 (cl (= @p_533 (! (= @p_523 @p_532) :named @p_534))) :rule cong :premises (t191.t3 t191.t12)) (step t191 (cl (! (= @p_522 (! (forall ((veriT_vr65 A_set$) (veriT_vr66 A_set$)) @p_534) :named @p_536)) :named @p_535)) :rule bind) (step t192 (cl (not @p_535) (not @p_522) @p_536) :rule equiv_pos2) (step t193 (cl @p_536) :rule th_resolution :premises (axiom22 t191 t192)) (anchor :step t194 :args ((:= (veriT_vr65 A_set$) veriT_vr67) (:= (veriT_vr66 A_set$) veriT_vr68))) (step t194.t1 (cl (! (= veriT_vr65 veriT_vr67) :named @p_538)) :rule refl) (step t194.t2 (cl (! (= veriT_vr66 veriT_vr68) :named @p_540)) :rule refl) (step t194.t3 (cl (= @p_523 (! (= veriT_vr67 veriT_vr68) :named @p_537))) :rule cong :premises (t194.t1 t194.t2)) (step t194.t4 (cl @p_538) :rule refl) (step t194.t5 (cl (= @p_525 (! (fun_app$a less_eq$ veriT_vr67) :named @p_539))) :rule cong :premises (t194.t4)) (step t194.t6 (cl @p_540) :rule refl) (step t194.t7 (cl (= @p_527 (! (fun_app$ @p_539 veriT_vr68) :named @p_541))) :rule cong :premises (t194.t5 t194.t6)) (step t194.t8 (cl @p_540) :rule refl) (step t194.t9 (cl (= @p_529 (! (fun_app$a less_eq$ veriT_vr68) :named @p_542))) :rule cong :premises (t194.t8)) (step t194.t10 (cl @p_538) :rule refl) (step t194.t11 (cl (= @p_531 (! (fun_app$ @p_542 veriT_vr67) :named @p_543))) :rule cong :premises (t194.t9 t194.t10)) (step t194.t12 (cl (= @p_532 (! (and @p_541 @p_543) :named @p_544))) :rule cong :premises (t194.t7 t194.t11)) (step t194.t13 (cl (= @p_534 (! (= @p_537 @p_544) :named @p_545))) :rule cong :premises (t194.t3 t194.t12)) (step t194 (cl (! (= @p_536 (! (forall ((veriT_vr67 A_set$) (veriT_vr68 A_set$)) @p_545) :named @p_547)) :named @p_546)) :rule bind) (step t195 (cl (not @p_546) (not @p_536) @p_547) :rule equiv_pos2) (step t196 (cl @p_547) :rule th_resolution :premises (t193 t194 t195)) (anchor :step t197 :args ((:= (?v0 A_list$) veriT_vr69))) (step t197.t1 (cl (! (= ?v0 veriT_vr69) :named @p_550)) :rule refl) (step t197.t2 (cl (= @p_2 (! (set$ veriT_vr69) :named @p_549))) :rule cong :premises (t197.t1)) (step t197.t3 (cl @p_550) :rule refl) (step t197.t4 (cl (= @p_14 (! (remdups$ veriT_vr69) :named @p_551))) :rule cong :premises (t197.t3)) (step t197.t5 (cl (= @p_552 (! (set$ @p_551) :named @p_553))) :rule cong :premises (t197.t4)) (step t197.t6 (cl (= @p_554 (! (= @p_549 @p_553) :named @p_555))) :rule cong :premises (t197.t2 t197.t5)) (step t197 (cl (! (= @p_548 (! (forall ((veriT_vr69 A_list$)) @p_555) :named @p_557)) :named @p_556)) :rule bind) (step t198 (cl (not @p_556) (not @p_548) @p_557) :rule equiv_pos2) (step t199 (cl @p_557) :rule th_resolution :premises (axiom23 t197 t198)) (anchor :step t200 :args ((:= (veriT_vr69 A_list$) veriT_vr70))) (step t200.t1 (cl (! (= veriT_vr69 veriT_vr70) :named @p_559)) :rule refl) (step t200.t2 (cl (= @p_549 (! (set$ veriT_vr70) :named @p_558))) :rule cong :premises (t200.t1)) (step t200.t3 (cl @p_559) :rule refl) (step t200.t4 (cl (= @p_551 (! (remdups$ veriT_vr70) :named @p_560))) :rule cong :premises (t200.t3)) (step t200.t5 (cl (= @p_553 (! (set$ @p_560) :named @p_561))) :rule cong :premises (t200.t4)) (step t200.t6 (cl (= @p_555 (! (= @p_558 @p_561) :named @p_562))) :rule cong :premises (t200.t2 t200.t5)) (step t200 (cl (! (= @p_557 (! (forall ((veriT_vr70 A_list$)) @p_562) :named @p_564)) :named @p_563)) :rule bind) (step t201 (cl (not @p_563) (not @p_557) @p_564) :rule equiv_pos2) (step t202 (cl @p_564) :rule th_resolution :premises (t199 t200 t201)) (anchor :step t203 :args ((:= (?v0 A_set$) veriT_vr71) (:= (?v1 A_set$) veriT_vr72))) (anchor :step t203.t1 :args ((:= (?v2 A$) veriT_vr73))) (step t203.t1.t1 (cl (! (= ?v2 veriT_vr73) :named @p_569)) :rule refl) (step t203.t1.t2 (cl (! (= ?v0 veriT_vr71) :named @p_573)) :rule refl) (step t203.t1.t3 (cl (= @p_15 (! (member$ veriT_vr73 veriT_vr71) :named @p_568))) :rule cong :premises (t203.t1.t1 t203.t1.t2)) (step t203.t1.t4 (cl @p_569) :rule refl) (step t203.t1.t5 (cl (! (= ?v1 veriT_vr72) :named @p_575)) :rule refl) (step t203.t1.t6 (cl (= @p_16 (! (member$ veriT_vr73 veriT_vr72) :named @p_570))) :rule cong :premises (t203.t1.t4 t203.t1.t5)) (step t203.t1.t7 (cl (= @p_571 (! (=> @p_568 @p_570) :named @p_572))) :rule cong :premises (t203.t1.t3 t203.t1.t6)) (step t203.t1 (cl (= @p_566 (! (forall ((veriT_vr73 A$)) @p_572) :named @p_567))) :rule bind) (step t203.t2 (cl @p_573) :rule refl) (step t203.t3 (cl (= @p_19 (! (fun_app$a less_eq$ veriT_vr71) :named @p_574))) :rule cong :premises (t203.t2)) (step t203.t4 (cl @p_575) :rule refl) (step t203.t5 (cl (= @p_17 (! (fun_app$ @p_574 veriT_vr72) :named @p_576))) :rule cong :premises (t203.t3 t203.t4)) (step t203.t6 (cl (= @p_577 (! (=> @p_567 @p_576) :named @p_578))) :rule cong :premises (t203.t1 t203.t5)) (step t203 (cl (! (= @p_565 (! (forall ((veriT_vr71 A_set$) (veriT_vr72 A_set$)) @p_578) :named @p_580)) :named @p_579)) :rule bind) (step t204 (cl (not @p_579) (not @p_565) @p_580) :rule equiv_pos2) (step t205 (cl @p_580) :rule th_resolution :premises (axiom25 t203 t204)) (anchor :step t206 :args ((:= (veriT_vr71 A_set$) veriT_vr74) (:= (veriT_vr72 A_set$) veriT_vr75))) (anchor :step t206.t1 :args ((:= (veriT_vr73 A$) veriT_vr76))) (step t206.t1.t1 (cl (! (= veriT_vr73 veriT_vr76) :named @p_583)) :rule refl) (step t206.t1.t2 (cl (! (= veriT_vr71 veriT_vr74) :named @p_586)) :rule refl) (step t206.t1.t3 (cl (= @p_568 (! (member$ veriT_vr76 veriT_vr74) :named @p_582))) :rule cong :premises (t206.t1.t1 t206.t1.t2)) (step t206.t1.t4 (cl @p_583) :rule refl) (step t206.t1.t5 (cl (! (= veriT_vr72 veriT_vr75) :named @p_588)) :rule refl) (step t206.t1.t6 (cl (= @p_570 (! (member$ veriT_vr76 veriT_vr75) :named @p_584))) :rule cong :premises (t206.t1.t4 t206.t1.t5)) (step t206.t1.t7 (cl (= @p_572 (! (=> @p_582 @p_584) :named @p_585))) :rule cong :premises (t206.t1.t3 t206.t1.t6)) (step t206.t1 (cl (= @p_567 (! (forall ((veriT_vr76 A$)) @p_585) :named @p_581))) :rule bind) (step t206.t2 (cl @p_586) :rule refl) (step t206.t3 (cl (= @p_574 (! (fun_app$a less_eq$ veriT_vr74) :named @p_587))) :rule cong :premises (t206.t2)) (step t206.t4 (cl @p_588) :rule refl) (step t206.t5 (cl (= @p_576 (! (fun_app$ @p_587 veriT_vr75) :named @p_589))) :rule cong :premises (t206.t3 t206.t4)) (step t206.t6 (cl (= @p_578 (! (=> @p_581 @p_589) :named @p_590))) :rule cong :premises (t206.t1 t206.t5)) (step t206 (cl (! (= @p_580 (! (forall ((veriT_vr74 A_set$) (veriT_vr75 A_set$)) @p_590) :named @p_592)) :named @p_591)) :rule bind) (step t207 (cl (not @p_591) (not @p_580) @p_592) :rule equiv_pos2) (step t208 (cl @p_592) :rule th_resolution :premises (t205 t206 t207)) (anchor :step t209 :args ((:= (?v0 A_set$) veriT_vr77) (:= (?v1 A_set$) veriT_vr78))) (step t209.t1 (cl (! (= ?v0 veriT_vr77) :named @p_598)) :rule refl) (step t209.t2 (cl (= @p_19 (! (fun_app$a less_eq$ veriT_vr77) :named @p_594))) :rule cong :premises (t209.t1)) (step t209.t3 (cl (! (= ?v1 veriT_vr78) :named @p_596)) :rule refl) (step t209.t4 (cl (= @p_17 (! (fun_app$ @p_594 veriT_vr78) :named @p_595))) :rule cong :premises (t209.t2 t209.t3)) (step t209.t5 (cl @p_596) :rule refl) (step t209.t6 (cl (= @p_528 (! (fun_app$a less_eq$ veriT_vr78) :named @p_597))) :rule cong :premises (t209.t5)) (step t209.t7 (cl @p_598) :rule refl) (step t209.t8 (cl (= @p_530 (! (fun_app$ @p_597 veriT_vr77) :named @p_599))) :rule cong :premises (t209.t6 t209.t7)) (step t209.t9 (cl (= @p_18 (! (and @p_595 @p_599) :named @p_600))) :rule cong :premises (t209.t4 t209.t8)) (step t209.t10 (cl @p_598) :rule refl) (step t209.t11 (cl @p_596) :rule refl) (step t209.t12 (cl (= @p_13 (! (= veriT_vr77 veriT_vr78) :named @p_601))) :rule cong :premises (t209.t10 t209.t11)) (step t209.t13 (cl (= @p_602 (! (=> @p_600 @p_601) :named @p_603))) :rule cong :premises (t209.t9 t209.t12)) (step t209 (cl (! (= @p_593 (! (forall ((veriT_vr77 A_set$) (veriT_vr78 A_set$)) @p_603) :named @p_605)) :named @p_604)) :rule bind) (step t210 (cl (not @p_604) (not @p_593) @p_605) :rule equiv_pos2) (step t211 (cl @p_605) :rule th_resolution :premises (axiom26 t209 t210)) (anchor :step t212 :args ((:= (veriT_vr77 A_set$) veriT_vr79) (:= (veriT_vr78 A_set$) veriT_vr80))) (step t212.t1 (cl (! (= veriT_vr77 veriT_vr79) :named @p_610)) :rule refl) (step t212.t2 (cl (= @p_594 (! (fun_app$a less_eq$ veriT_vr79) :named @p_606))) :rule cong :premises (t212.t1)) (step t212.t3 (cl (! (= veriT_vr78 veriT_vr80) :named @p_608)) :rule refl) (step t212.t4 (cl (= @p_595 (! (fun_app$ @p_606 veriT_vr80) :named @p_607))) :rule cong :premises (t212.t2 t212.t3)) (step t212.t5 (cl @p_608) :rule refl) (step t212.t6 (cl (= @p_597 (! (fun_app$a less_eq$ veriT_vr80) :named @p_609))) :rule cong :premises (t212.t5)) (step t212.t7 (cl @p_610) :rule refl) (step t212.t8 (cl (= @p_599 (! (fun_app$ @p_609 veriT_vr79) :named @p_611))) :rule cong :premises (t212.t6 t212.t7)) (step t212.t9 (cl (= @p_600 (! (and @p_607 @p_611) :named @p_612))) :rule cong :premises (t212.t4 t212.t8)) (step t212.t10 (cl @p_610) :rule refl) (step t212.t11 (cl @p_608) :rule refl) (step t212.t12 (cl (= @p_601 (! (= veriT_vr79 veriT_vr80) :named @p_613))) :rule cong :premises (t212.t10 t212.t11)) (step t212.t13 (cl (= @p_603 (! (=> @p_612 @p_613) :named @p_614))) :rule cong :premises (t212.t9 t212.t12)) (step t212 (cl (! (= @p_605 (! (forall ((veriT_vr79 A_set$) (veriT_vr80 A_set$)) @p_614) :named @p_616)) :named @p_615)) :rule bind) (step t213 (cl (not @p_615) (not @p_605) @p_616) :rule equiv_pos2) (step t214 (cl @p_616) :rule th_resolution :premises (t211 t212 t213)) (anchor :step t215 :args ((:= (?v0 A_set$) veriT_vr81) (:= (?v1 A_list$) veriT_vr82))) (step t215.t1 (cl (! (= ?v0 veriT_vr81) :named @p_630)) :rule refl) (step t215.t2 (cl (= @p_19 (! (fun_app$a less_eq$ veriT_vr81) :named @p_618))) :rule cong :premises (t215.t1)) (step t215.t3 (cl (! (= ?v1 veriT_vr82) :named @p_625)) :rule refl) (step t215.t4 (cl (= @p_619 (! (coset$ veriT_vr82) :named @p_620))) :rule cong :premises (t215.t3)) (step t215.t5 (cl (= @p_621 (! (fun_app$ @p_618 @p_620) :named @p_622))) :rule cong :premises (t215.t2 t215.t4)) (anchor :step t215.t6 :args ((:= (?v2 A$) veriT_vr83))) (step t215.t6.t1 (cl (! (= ?v2 veriT_vr83) :named @p_629)) :rule refl) (step t215.t6.t2 (cl @p_625) :rule refl) (step t215.t6.t3 (cl (= @p_20 (! (set$ veriT_vr82) :named @p_626))) :rule cong :premises (t215.t6.t2)) (step t215.t6.t4 (cl (= @p_627 (! (member$ veriT_vr83 @p_626) :named @p_628))) :rule cong :premises (t215.t6.t1 t215.t6.t3)) (step t215.t6.t5 (cl @p_629) :rule refl) (step t215.t6.t6 (cl @p_630) :rule refl) (step t215.t6.t7 (cl (= @p_15 (! (member$ veriT_vr83 veriT_vr81) :named @p_631))) :rule cong :premises (t215.t6.t5 t215.t6.t6)) (step t215.t6.t8 (cl (= @p_632 (! (not @p_631) :named @p_633))) :rule cong :premises (t215.t6.t7)) (step t215.t6.t9 (cl (= @p_634 (! (=> @p_628 @p_633) :named @p_635))) :rule cong :premises (t215.t6.t4 t215.t6.t8)) (step t215.t6 (cl (= @p_623 (! (forall ((veriT_vr83 A$)) @p_635) :named @p_624))) :rule bind) (step t215.t7 (cl (= @p_636 (! (= @p_622 @p_624) :named @p_637))) :rule cong :premises (t215.t5 t215.t6)) (step t215 (cl (! (= @p_617 (! (forall ((veriT_vr81 A_set$) (veriT_vr82 A_list$)) @p_637) :named @p_639)) :named @p_638)) :rule bind) (step t216 (cl (not @p_638) (not @p_617) @p_639) :rule equiv_pos2) (step t217 (cl @p_639) :rule th_resolution :premises (axiom27 t215 t216)) (anchor :step t218 :args ((veriT_vr81 A_set$) (veriT_vr82 A_list$))) (step t218.t1 (cl (= @p_637 (! (and (! (=> @p_622 @p_624) :named @p_655) (! (=> @p_624 @p_622) :named @p_667)) :named @p_640))) :rule connective_def) (step t218 (cl (! (= @p_639 (! (forall ((veriT_vr81 A_set$) (veriT_vr82 A_list$)) @p_640) :named @p_642)) :named @p_641)) :rule bind) (step t219 (cl (not @p_641) (not @p_639) @p_642) :rule equiv_pos2) (step t220 (cl @p_642) :rule th_resolution :premises (t217 t218 t219)) (anchor :step t221 :args ((:= (veriT_vr81 A_set$) veriT_vr84) (:= (veriT_vr82 A_list$) veriT_vr85))) (step t221.t1 (cl (! (= veriT_vr81 veriT_vr84) :named @p_651)) :rule refl) (step t221.t2 (cl (! (= @p_618 (! (fun_app$a less_eq$ veriT_vr84) :named @p_645)) :named @p_664)) :rule cong :premises (t221.t1)) (step t221.t3 (cl (! (= veriT_vr82 veriT_vr85) :named @p_648)) :rule refl) (step t221.t4 (cl (! (= @p_620 (! (coset$ veriT_vr85) :named @p_646)) :named @p_665)) :rule cong :premises (t221.t3)) (step t221.t5 (cl (! (= @p_622 (! (fun_app$ @p_645 @p_646) :named @p_644)) :named @p_666)) :rule cong :premises (t221.t2 t221.t4)) (anchor :step t221.t6 :args ((:= (veriT_vr83 A$) veriT_vr86))) (step t221.t6.t1 (cl (! (= veriT_vr83 veriT_vr86) :named @p_650)) :rule refl) (step t221.t6.t2 (cl @p_648) :rule refl) (step t221.t6.t3 (cl (! (= @p_626 (! (set$ veriT_vr85) :named @p_643)) :named @p_658)) :rule cong :premises (t221.t6.t2)) (step t221.t6.t4 (cl (= @p_628 (! (member$ veriT_vr86 @p_643) :named @p_649))) :rule cong :premises (t221.t6.t1 t221.t6.t3)) (step t221.t6.t5 (cl @p_650) :rule refl) (step t221.t6.t6 (cl @p_651) :rule refl) (step t221.t6.t7 (cl (= @p_631 (! (member$ veriT_vr86 veriT_vr84) :named @p_652))) :rule cong :premises (t221.t6.t5 t221.t6.t6)) (step t221.t6.t8 (cl (= @p_633 (! (not @p_652) :named @p_653))) :rule cong :premises (t221.t6.t7)) (step t221.t6.t9 (cl (= @p_635 (! (=> @p_649 @p_653) :named @p_654))) :rule cong :premises (t221.t6.t4 t221.t6.t8)) (step t221.t6 (cl (= @p_624 (! (forall ((veriT_vr86 A$)) @p_654) :named @p_647))) :rule bind) (step t221.t7 (cl (= @p_655 (! (=> @p_644 @p_647) :named @p_656))) :rule cong :premises (t221.t5 t221.t6)) (anchor :step t221.t8 :args ((:= (veriT_vr83 A$) veriT_vr87))) (step t221.t8.t1 (cl (! (= veriT_vr83 veriT_vr87) :named @p_660)) :rule refl) (step t221.t8.t2 (cl @p_648) :rule refl) (step t221.t8.t3 (cl @p_658) :rule cong :premises (t221.t8.t2)) (step t221.t8.t4 (cl (= @p_628 (! (member$ veriT_vr87 @p_643) :named @p_659))) :rule cong :premises (t221.t8.t1 t221.t8.t3)) (step t221.t8.t5 (cl @p_660) :rule refl) (step t221.t8.t6 (cl @p_651) :rule refl) (step t221.t8.t7 (cl (= @p_631 (! (member$ veriT_vr87 veriT_vr84) :named @p_661))) :rule cong :premises (t221.t8.t5 t221.t8.t6)) (step t221.t8.t8 (cl (= @p_633 (! (not @p_661) :named @p_662))) :rule cong :premises (t221.t8.t7)) (step t221.t8.t9 (cl (= @p_635 (! (=> @p_659 @p_662) :named @p_663))) :rule cong :premises (t221.t8.t4 t221.t8.t8)) (step t221.t8 (cl (= @p_624 (! (forall ((veriT_vr87 A$)) @p_663) :named @p_657))) :rule bind) (step t221.t9 (cl @p_651) :rule refl) (step t221.t10 (cl @p_664) :rule cong :premises (t221.t9)) (step t221.t11 (cl @p_648) :rule refl) (step t221.t12 (cl @p_665) :rule cong :premises (t221.t11)) (step t221.t13 (cl @p_666) :rule cong :premises (t221.t10 t221.t12)) (step t221.t14 (cl (= @p_667 (! (=> @p_657 @p_644) :named @p_668))) :rule cong :premises (t221.t8 t221.t13)) (step t221.t15 (cl (= @p_640 (! (and @p_656 @p_668) :named @p_669))) :rule cong :premises (t221.t7 t221.t14)) (step t221 (cl (! (= @p_642 (! (forall ((veriT_vr84 A_set$) (veriT_vr85 A_list$)) @p_669) :named @p_671)) :named @p_670)) :rule bind) (step t222 (cl (not @p_670) (not @p_642) @p_671) :rule equiv_pos2) (step t223 (cl @p_671) :rule th_resolution :premises (t220 t221 t222)) (anchor :step t224 :args ((:= (veriT_vr84 A_set$) veriT_vr88) (:= (veriT_vr85 A_list$) veriT_vr89))) (step t224.t1 (cl (! (= veriT_vr84 veriT_vr88) :named @p_680)) :rule refl) (step t224.t2 (cl (! (= @p_645 (! (fun_app$a less_eq$ veriT_vr88) :named @p_674)) :named @p_687)) :rule cong :premises (t224.t1)) (step t224.t3 (cl (! (= veriT_vr85 veriT_vr89) :named @p_676)) :rule refl) (step t224.t4 (cl (! (= @p_646 (! (coset$ veriT_vr89) :named @p_675)) :named @p_688)) :rule cong :premises (t224.t3)) (step t224.t5 (cl (! (= @p_644 (! (fun_app$ @p_674 @p_675) :named @p_673)) :named @p_689)) :rule cong :premises (t224.t2 t224.t4)) (anchor :step t224.t6 :args ((:= (veriT_vr86 A$) veriT_vr90))) (step t224.t6.t1 (cl (! (= veriT_vr86 veriT_vr90) :named @p_679)) :rule refl) (step t224.t6.t2 (cl @p_676) :rule refl) (step t224.t6.t3 (cl (! (= @p_643 (! (set$ veriT_vr89) :named @p_677)) :named @p_685)) :rule cong :premises (t224.t6.t2)) (step t224.t6.t4 (cl (= @p_649 (! (member$ veriT_vr90 @p_677) :named @p_678))) :rule cong :premises (t224.t6.t1 t224.t6.t3)) (step t224.t6.t5 (cl @p_679) :rule refl) (step t224.t6.t6 (cl @p_680) :rule refl) (step t224.t6.t7 (cl (= @p_652 (! (member$ veriT_vr90 veriT_vr88) :named @p_681))) :rule cong :premises (t224.t6.t5 t224.t6.t6)) (step t224.t6.t8 (cl (= @p_653 (! (not @p_681) :named @p_682))) :rule cong :premises (t224.t6.t7)) (step t224.t6.t9 (cl (= @p_654 (! (=> @p_678 @p_682) :named @p_683))) :rule cong :premises (t224.t6.t4 t224.t6.t8)) (step t224.t6 (cl (= @p_647 (! (forall ((veriT_vr90 A$)) @p_683) :named @p_672))) :rule bind) (step t224.t7 (cl (= @p_656 (! (=> @p_673 @p_672) :named @p_684))) :rule cong :premises (t224.t5 t224.t6)) (anchor :step t224.t8 :args ((:= (veriT_vr87 A$) veriT_vr90))) (step t224.t8.t1 (cl (! (= veriT_vr87 veriT_vr90) :named @p_686)) :rule refl) (step t224.t8.t2 (cl @p_676) :rule refl) (step t224.t8.t3 (cl @p_685) :rule cong :premises (t224.t8.t2)) (step t224.t8.t4 (cl (= @p_659 @p_678)) :rule cong :premises (t224.t8.t1 t224.t8.t3)) (step t224.t8.t5 (cl @p_686) :rule refl) (step t224.t8.t6 (cl @p_680) :rule refl) (step t224.t8.t7 (cl (= @p_661 @p_681)) :rule cong :premises (t224.t8.t5 t224.t8.t6)) (step t224.t8.t8 (cl (= @p_662 @p_682)) :rule cong :premises (t224.t8.t7)) (step t224.t8.t9 (cl (= @p_663 @p_683)) :rule cong :premises (t224.t8.t4 t224.t8.t8)) (step t224.t8 (cl (= @p_657 @p_672)) :rule bind) (step t224.t9 (cl @p_680) :rule refl) (step t224.t10 (cl @p_687) :rule cong :premises (t224.t9)) (step t224.t11 (cl @p_676) :rule refl) (step t224.t12 (cl @p_688) :rule cong :premises (t224.t11)) (step t224.t13 (cl @p_689) :rule cong :premises (t224.t10 t224.t12)) (step t224.t14 (cl (= @p_668 (! (=> @p_672 @p_673) :named @p_690))) :rule cong :premises (t224.t8 t224.t13)) (step t224.t15 (cl (= @p_669 (! (and @p_684 @p_690) :named @p_691))) :rule cong :premises (t224.t7 t224.t14)) (step t224 (cl (! (= @p_671 (! (forall ((veriT_vr88 A_set$) (veriT_vr89 A_list$)) @p_691) :named @p_693)) :named @p_692)) :rule bind) (step t225 (cl (not @p_692) (not @p_671) @p_693) :rule equiv_pos2) (step t226 (cl @p_693) :rule th_resolution :premises (t223 t224 t225)) (anchor :step t227 :args ((:= (veriT_vr88 A_set$) veriT_vr88) (:= (veriT_vr89 A_list$) veriT_vr89))) (anchor :step t227.t1 :args ((:= (veriT_vr90 A$) veriT_vr91))) (step t227.t1.t1 (cl (! (= veriT_vr90 veriT_vr91) :named @p_696)) :rule refl) (step t227.t1.t2 (cl (= @p_678 (! (member$ veriT_vr91 @p_677) :named @p_695))) :rule cong :premises (t227.t1.t1)) (step t227.t1.t3 (cl @p_696) :rule refl) (step t227.t1.t4 (cl (= @p_681 (! (member$ veriT_vr91 veriT_vr88) :named @p_697))) :rule cong :premises (t227.t1.t3)) (step t227.t1.t5 (cl (= @p_682 (! (not @p_697) :named @p_698))) :rule cong :premises (t227.t1.t4)) (step t227.t1.t6 (cl (= @p_683 (! (=> @p_695 @p_698) :named @p_699))) :rule cong :premises (t227.t1.t2 t227.t1.t5)) (step t227.t1 (cl (= @p_672 (! (forall ((veriT_vr91 A$)) @p_699) :named @p_694))) :rule bind) (step t227.t2 (cl (= @p_690 (! (=> @p_694 @p_673) :named @p_700))) :rule cong :premises (t227.t1)) (step t227.t3 (cl (= @p_691 (! (and @p_684 @p_700) :named @p_701))) :rule cong :premises (t227.t2)) (step t227 (cl (! (= @p_693 (! (forall ((veriT_vr88 A_set$) (veriT_vr89 A_list$)) @p_701) :named @p_703)) :named @p_702)) :rule bind) (step t228 (cl (not @p_702) (not @p_693) @p_703) :rule equiv_pos2) (step t229 (cl @p_703) :rule th_resolution :premises (t226 t227 t228)) (step t230 (cl @p_494) :rule and :premises (t184)) (step t231 (cl (! (not @p_493) :named @p_704) @p_422) :rule and_pos) (step t232 (cl @p_704 @p_400) :rule and_pos) (step t233 (cl @p_704 @p_474) :rule and_pos) (step t234 (cl (! (not rhs$) :named @p_712) @p_493) :rule implies :premises (t230)) (step t235 (cl @p_484) :rule and :premises (t184)) (step t236 (cl @p_454 @p_450) :rule implies_neg1) (step t237 (cl @p_454 (! (not @p_453) :named @p_705)) :rule implies_neg2) (step t238 (cl (not @p_705) @p_452) :rule not_not) (step t239 (cl @p_454 @p_452) :rule th_resolution :premises (t238 t237)) (step t240 (cl @p_460 @p_456) :rule implies_neg1) (step t241 (cl @p_460 (! (not @p_459) :named @p_706)) :rule implies_neg2) (step t242 (cl (not @p_706) @p_458) :rule not_not) (step t243 (cl @p_460 @p_458) :rule th_resolution :premises (t242 t241)) (step t244 (cl @p_482 (! (not @p_422) :named @p_707) (! (not @p_400) :named @p_708) (! (not @p_454) :named @p_709) (! (not @p_460) :named @p_710)) :rule and_neg) (step t245 (cl (not @p_707) @p_403) :rule not_not) (step t246 (cl @p_482 @p_403 @p_708 @p_709 @p_710) :rule th_resolution :premises (t245 t244)) (step t247 (cl (not @p_482) rhs$) :rule implies :premises (t235)) (step t248 (cl @p_711 rhs$) :rule not_equiv1 :premises (axiom28)) (step t249 (cl (! (not @p_711) :named @p_1099) @p_712) :rule not_equiv2 :premises (axiom28)) (step t250 (cl (or (! (not @p_474) :named @p_1028) (! (forall ((veriT_vr59 A$)) (or @p_472 @p_469)) :named @p_1029))) :rule qnt_cnf) (step t251 (cl (or (! (not @p_564) :named @p_714) (! (= @p_381 @p_713) :named @p_722))) :rule forall_inst :args ((:= veriT_vr70 ys$))) (step t252 (cl (or @p_714 (! (= @p_715 @p_380) :named @p_723))) :rule forall_inst :args ((:= veriT_vr70 xs$))) (step t253 (cl (or @p_714 (! (= @p_381 (! (set$ (! (remdups$ @p_10) :named @p_753)) :named @p_752)) :named @p_724))) :rule forall_inst :args ((:= veriT_vr70 @p_10))) (step t254 (cl (or (! (not @p_282) :named @p_716) (! (= (! (coset$ @p_9) :named @p_792) (! (uminus$ @p_380) :named @p_721)) :named @p_725))) :rule forall_inst :args ((:= veriT_vr41 @p_9))) (step t255 (cl (or @p_716 (! (= @p_21 (! (uminus$ @p_715) :named @p_757)) :named @p_726))) :rule forall_inst :args ((:= veriT_vr41 xs$))) (step t256 (cl (or (! (not @p_265) :named @p_727) (! (= @p_715 (! (uminus$ @p_21) :named @p_1049)) :named @p_728))) :rule forall_inst :args ((:= veriT_vr39 xs$))) (step t257 (cl (or (! (not @p_227) :named @p_735) (! (= @p_403 (! (or @p_717 (! (not @p_7) :named @p_730)) :named @p_729)) :named @p_732))) :rule forall_inst :args ((:= veriT_vr35 top$))) (step t258 (cl (or (! (not @p_127) :named @p_718) (! (finite$ @p_713) :named @p_736))) :rule forall_inst :args ((:= veriT_vr13 ys$))) (step t259 (cl (or @p_718 (! (finite$ @p_381) :named @p_737))) :rule forall_inst :args ((:= veriT_vr13 @p_10))) (step t260 (cl (or @p_718 (! (finite$ @p_715) :named @p_738))) :rule forall_inst :args ((:= veriT_vr13 xs$))) (step t261 (cl (or (! (not @p_117) :named @p_739) (! (= @p_719 (! (fun_app$b card$ @p_713) :named @p_1085)) :named @p_740))) :rule forall_inst :args ((:= veriT_vr11 ys$))) (step t262 (cl (or (! (not @p_97) :named @p_720) (! (=> (! (member$ veriT_sk1 (! (uminus$ @p_381) :named @p_1046)) :named @p_742) (! (not @p_456) :named @p_743)) :named @p_741))) :rule forall_inst :args ((:= veriT_vr8 veriT_sk1) (:= veriT_vr9 @p_381))) (step t263 (cl (or @p_720 (! (=> (! (member$ veriT_sk0 @p_721) :named @p_745) (! (not @p_450) :named @p_746)) :named @p_744))) :rule forall_inst :args ((:= veriT_vr8 veriT_sk0) (:= veriT_vr9 @p_380))) (step t264 (cl (or (! (not @p_41) :named @p_751) (! (= (! (fun_app$ (! (fun_app$a uu$ @p_21) :named @p_1098) @p_713) :named @p_748) (! (= @p_21 @p_713) :named @p_749)) :named @p_747))) :rule forall_inst :args ((:= veriT_vr2 @p_21) (:= veriT_vr3 @p_713))) (step t265 (cl @p_714 @p_722) :rule or :premises (t251)) (step t266 (cl @p_722) :rule resolution :premises (t265 t202)) (step t267 (cl @p_714 @p_723) :rule or :premises (t252)) (step t268 (cl @p_723) :rule resolution :premises (t267 t202)) (step t269 (cl @p_714 @p_724) :rule or :premises (t253)) (step t270 (cl @p_724) :rule resolution :premises (t269 t202)) (step t271 (cl @p_716 @p_725) :rule or :premises (t254)) (step t272 (cl @p_725) :rule resolution :premises (t271 t99)) (step t273 (cl @p_716 @p_726) :rule or :premises (t255)) (step t274 (cl @p_726) :rule resolution :premises (t273 t99)) (step t275 (cl @p_727 @p_728) :rule or :premises (t256)) (step t276 (cl @p_728) :rule resolution :premises (t275 t93)) (step t277 (cl (! (not @p_729) :named @p_733) @p_717 @p_730) :rule or_pos) (step t278 (cl @p_729 (! (not @p_730) :named @p_731)) :rule or_neg) (step t279 (cl (not @p_731) @p_7) :rule not_not) (step t280 (cl @p_729 @p_7) :rule th_resolution :premises (t279 t278)) (step t281 (cl (! (not @p_732) :named @p_734) @p_403 @p_733) :rule equiv_pos1) (step t282 (cl @p_734 @p_422 @p_729) :rule equiv_pos2) (step t283 (cl @p_735 @p_732) :rule or :premises (t257)) (step t284 (cl @p_733 @p_730) :rule resolution :premises (t277 axiom6)) (step t285 (cl @p_732) :rule resolution :premises (t283 t81)) (step t286 (cl @p_718 @p_736) :rule or :premises (t258)) (step t287 (cl @p_736) :rule resolution :premises (t286 t57)) (step t288 (cl @p_718 @p_737) :rule or :premises (t259)) (step t289 (cl @p_737) :rule resolution :premises (t288 t57)) (step t290 (cl @p_718 @p_738) :rule or :premises (t260)) (step t291 (cl @p_738) :rule resolution :premises (t290 t57)) (step t292 (cl @p_739 @p_740) :rule or :premises (t261)) (step t293 (cl @p_740) :rule resolution :premises (t292 t51)) (step t294 (cl (not @p_741) (not @p_742) @p_743) :rule implies_pos) (step t295 (cl @p_720 @p_741) :rule or :premises (t262)) (step t296 (cl @p_741) :rule resolution :premises (t295 t45)) (step t297 (cl (not @p_744) (not @p_745) @p_746) :rule implies_pos) (step t298 (cl @p_720 @p_744) :rule or :premises (t263)) (step t299 (cl @p_744) :rule resolution :premises (t298 t45)) (step t300 (cl (! (not @p_747) :named @p_750) @p_748 (! (not @p_749) :named @p_1040)) :rule equiv_pos1) (step t301 (cl @p_750 (! (not @p_748) :named @p_1107) @p_749) :rule equiv_pos2) (step t302 (cl @p_751 @p_747) :rule or :premises (t264)) (step t303 (cl @p_747) :rule resolution :premises (t302 t33)) (step t304 (cl (or (! (not @p_616) :named @p_765) (! (=> (! (and (! (fun_app$ (! (fun_app$a less_eq$ @p_381) :named @p_854) @p_21) :named @p_762) (! (fun_app$ (! (fun_app$a less_eq$ @p_21) :named @p_907) @p_381) :named @p_763)) :named @p_755) (! (= @p_21 @p_381) :named @p_754)) :named @p_764))) :rule forall_inst :args ((:= veriT_vr79 @p_381) (:= veriT_vr80 @p_21))) (step t305 (cl (not (! (not (! (not @p_592) :named @p_906)) :named @p_936)) @p_592) :rule not_not) (step t306 (cl (or @p_714 (! (= @p_752 (! (set$ (! (remdups$ @p_753) :named @p_938)) :named @p_939)) :named @p_766))) :rule forall_inst :args ((:= veriT_vr70 @p_753))) (step t307 (cl (or (! (not @p_547) :named @p_769) (! (= @p_754 @p_755) :named @p_768))) :rule forall_inst :args ((:= veriT_vr67 @p_381) (:= veriT_vr68 @p_21))) (step t308 (cl (or (! (not @p_521) :named @p_756) (! (= @p_760 (! (sup$ @p_715 @p_381) :named @p_759)) :named @p_770))) :rule forall_inst :args ((:= veriT_vr63 xs$) (:= veriT_vr64 @p_10))) (step t309 (cl (or @p_756 (! (= (! (set$ (append$ xs$ ys$)) :named @p_758) (! (sup$ @p_715 @p_713) :named @p_1115)) :named @p_771))) :rule forall_inst :args ((:= veriT_vr63 xs$) (:= veriT_vr64 ys$))) (step t310 (cl (or (! (not @p_378) :named @p_775) (! (=> @p_738 (! (= (! (finite$ @p_757) :named @p_773) @p_7) :named @p_772)) :named @p_774))) :rule forall_inst :args ((:= veriT_vr54 @p_715))) (step t311 (cl (or (! (not @p_248) :named @p_781) (! (=> (! (and @p_7 (! (= @p_1 (! (fun_app$b card$ @p_758) :named @p_1112)) :named @p_778)) :named @p_777) (! (= top$ @p_758) :named @p_780)) :named @p_779))) :rule forall_inst :args ((:= veriT_vr37 @p_758))) (step t312 (cl (or (! (not @p_201) :named @p_786) (! (=> (! (and @p_738 @p_737) :named @p_782) (! (= (! (plus$ (! (fun_app$b card$ @p_715) :named @p_1073) (! (fun_app$b card$ @p_381) :named @p_1074)) :named @p_1057) (! (plus$ (! (fun_app$b card$ @p_759) :named @p_999) (! (fun_app$b card$ (! (inf$ @p_715 @p_381) :named @p_940)) :named @p_998)) :named @p_1000)) :named @p_785)) :named @p_784))) :rule forall_inst :args ((:= veriT_vr32 @p_715) (:= veriT_vr33 @p_381))) (step t313 (cl (or (! (not @p_73) :named @p_791) (! (= (! (fun_app$ @p_43 @p_760) :named @p_790) (! (and @p_44 (= @p_1 (fun_app$b card$ (set$ (append$ xs$ @p_761))))) :named @p_788)) :named @p_789))) :rule forall_inst :args ((:= veriT_vr5 @p_761))) (step t314 (cl @p_755 (! (not @p_762) :named @p_1003) (! (not @p_763) :named @p_1120)) :rule and_neg) (step t315 (cl (not @p_764) (! (not @p_755) :named @p_767) @p_754) :rule implies_pos) (step t316 (cl @p_765 @p_764) :rule or :premises (t304)) (step t317 (cl @p_764) :rule resolution :premises (t316 t214)) (step t318 (cl @p_714 @p_766) :rule or :premises (t306)) (step t319 (cl @p_766) :rule resolution :premises (t318 t202)) (step t320 (cl @p_767 @p_762) :rule and_pos) (step t321 (cl (not @p_768) (! (not @p_754) :named @p_1051) @p_755) :rule equiv_pos2) (step t322 (cl @p_769 @p_768) :rule or :premises (t307)) (step t323 (cl @p_768) :rule resolution :premises (t322 t196)) (step t324 (cl @p_756 @p_770) :rule or :premises (t308)) (step t325 (cl @p_770) :rule resolution :premises (t324 t190)) (step t326 (cl @p_756 @p_771) :rule or :premises (t309)) (step t327 (cl @p_771) :rule resolution :premises (t326 t190)) (step t328 (cl (not @p_772) (not @p_773) @p_7) :rule equiv_pos2) (step t329 (cl (! (not @p_774) :named @p_776) (! (not @p_738) :named @p_783) @p_772) :rule implies_pos) (step t330 (cl @p_775 @p_774) :rule or :premises (t310)) (step t331 (cl @p_776 @p_772) :rule resolution :premises (t329 t291)) (step t332 (cl @p_774) :rule resolution :premises (t330 t120)) (step t333 (cl @p_772) :rule resolution :premises (t331 t332)) (step t334 (cl @p_777 @p_730 (not @p_778)) :rule and_neg) (step t335 (cl (not @p_779) (not @p_777) @p_780) :rule implies_pos) (step t336 (cl @p_781 @p_779) :rule or :premises (t311)) (step t337 (cl @p_779) :rule resolution :premises (t336 t87)) (step t338 (cl @p_782 @p_783 (not @p_737)) :rule and_neg) (step t339 (cl (! (not @p_784) :named @p_787) (not @p_782) @p_785) :rule implies_pos) (step t340 (cl @p_786 @p_784) :rule or :premises (t312)) (step t341 (cl @p_782) :rule resolution :premises (t338 t291 t289)) (step t342 (cl @p_787 @p_785) :rule resolution :premises (t339 t341)) (step t343 (cl @p_784) :rule resolution :premises (t340 t75)) (step t344 (cl @p_785) :rule resolution :premises (t342 t343)) (step t345 (cl (! (not @p_788) :named @p_1132) @p_44) :rule and_pos) (step t346 (cl (not @p_789) (! (not @p_790) :named @p_1133) @p_788) :rule equiv_pos2) (step t347 (cl @p_791 @p_789) :rule or :premises (t313)) (step t348 (cl @p_789) :rule resolution :premises (t347 t39)) (step t349 (cl (! (not @p_726) :named @p_794) (! (not (! (= @p_721 @p_757) :named @p_793)) :named @p_1044) (! (not @p_725) :named @p_795) (! (= @p_21 @p_792) :named @p_796)) :rule eq_transitive) (step t350 (cl (! (not @p_723) :named @p_797) @p_793) :rule eq_congruent) (step t351 (cl @p_794 @p_795 @p_796 @p_797) :rule th_resolution :premises (t349 t350)) (step t352 (cl (or (! (not @p_703) :named @p_851) (! (and (! (=> @p_762 (! (forall ((veriT_vr90 A$)) (! (=> (! (member$ veriT_vr90 @p_715) :named @p_801) (! (not (! (member$ veriT_vr90 @p_381) :named @p_804)) :named @p_806)) :named @p_808)) :named @p_800)) :named @p_810) (! (=> (! (forall ((veriT_vr91 A$)) (! (=> (! (member$ veriT_vr91 @p_715) :named @p_813) (! (not (! (member$ veriT_vr91 @p_381) :named @p_815)) :named @p_816)) :named @p_817)) :named @p_812) @p_762) :named @p_818)) :named @p_798))) :rule forall_inst :args ((:= veriT_vr88 @p_381) (:= veriT_vr89 xs$))) (anchor :step t353) (assume t353.h1 @p_798) (anchor :step t353.t2 :args ((:= (veriT_vr90 A$) veriT_vr96))) (step t353.t2.t1 (cl (! (= veriT_vr90 veriT_vr96) :named @p_803)) :rule refl) (step t353.t2.t2 (cl (= @p_801 (! (member$ veriT_vr96 @p_715) :named @p_802))) :rule cong :premises (t353.t2.t1)) (step t353.t2.t3 (cl @p_803) :rule refl) (step t353.t2.t4 (cl (= @p_804 (! (member$ veriT_vr96 @p_381) :named @p_805))) :rule cong :premises (t353.t2.t3)) (step t353.t2.t5 (cl (= @p_806 (! (not @p_805) :named @p_807))) :rule cong :premises (t353.t2.t4)) (step t353.t2.t6 (cl (= @p_808 (! (=> @p_802 @p_807) :named @p_809))) :rule cong :premises (t353.t2.t2 t353.t2.t5)) (step t353.t2 (cl (= @p_800 (! (forall ((veriT_vr96 A$)) @p_809) :named @p_811))) :rule bind) (step t353.t3 (cl (= @p_810 (! (=> @p_762 @p_811) :named @p_819))) :rule cong :premises (t353.t2)) (anchor :step t353.t4 :args ((:= (veriT_vr91 A$) veriT_vr96))) (step t353.t4.t1 (cl (! (= veriT_vr91 veriT_vr96) :named @p_814)) :rule refl) (step t353.t4.t2 (cl (= @p_813 @p_802)) :rule cong :premises (t353.t4.t1)) (step t353.t4.t3 (cl @p_814) :rule refl) (step t353.t4.t4 (cl (= @p_815 @p_805)) :rule cong :premises (t353.t4.t3)) (step t353.t4.t5 (cl (= @p_816 @p_807)) :rule cong :premises (t353.t4.t4)) (step t353.t4.t6 (cl (= @p_817 @p_809)) :rule cong :premises (t353.t4.t2 t353.t4.t5)) (step t353.t4 (cl (= @p_812 @p_811)) :rule bind) (step t353.t5 (cl (= @p_818 (! (=> @p_811 @p_762) :named @p_820))) :rule cong :premises (t353.t4)) (step t353.t6 (cl (! (= @p_798 (! (and @p_819 @p_820) :named @p_823)) :named @p_821)) :rule cong :premises (t353.t3 t353.t5)) (step t353.t7 (cl (not @p_821) (! (not @p_798) :named @p_822) @p_823) :rule equiv_pos2) (step t353.t8 (cl @p_823) :rule th_resolution :premises (t353.h1 t353.t6 t353.t7)) (anchor :step t353.t9 :args ((:= (veriT_vr96 A$) veriT_vr97))) (step t353.t9.t1 (cl (! (= veriT_vr96 veriT_vr97) :named @p_826)) :rule refl) (step t353.t9.t2 (cl (= @p_802 @p_825)) :rule cong :premises (t353.t9.t1)) (step t353.t9.t3 (cl @p_826) :rule refl) (step t353.t9.t4 (cl (= @p_805 @p_827)) :rule cong :premises (t353.t9.t3)) (step t353.t9.t5 (cl (= @p_807 @p_828)) :rule cong :premises (t353.t9.t4)) (step t353.t9.t6 (cl (= @p_809 @p_824)) :rule cong :premises (t353.t9.t2 t353.t9.t5)) (step t353.t9 (cl (= @p_811 (! (forall ((veriT_vr97 A$)) @p_824) :named @p_829))) :rule bind) (step t353.t10 (cl (= @p_819 (! (=> @p_762 @p_829) :named @p_830))) :rule cong :premises (t353.t9)) (step t353.t11 (cl (= @p_820 (! (=> @p_829 @p_762) :named @p_831))) :rule cong :premises (t353.t9)) (step t353.t12 (cl (! (= @p_823 (! (and @p_830 @p_831) :named @p_833)) :named @p_832)) :rule cong :premises (t353.t10 t353.t11)) (step t353.t13 (cl (not @p_832) (not @p_823) @p_833) :rule equiv_pos2) (step t353.t14 (cl @p_833) :rule th_resolution :premises (t353.t8 t353.t12 t353.t13)) (anchor :step t353.t15 :args ((:= (veriT_vr97 A$) veriT_sk4))) (step t353.t15.t1 (cl (! (= veriT_vr97 veriT_sk4) :named @p_836)) :rule refl) (step t353.t15.t2 (cl (= @p_825 (! (member$ veriT_sk4 @p_715) :named @p_835))) :rule cong :premises (t353.t15.t1)) (step t353.t15.t3 (cl @p_836) :rule refl) (step t353.t15.t4 (cl (= @p_827 (! (member$ veriT_sk4 @p_381) :named @p_837))) :rule cong :premises (t353.t15.t3)) (step t353.t15.t5 (cl (= @p_828 (! (not @p_837) :named @p_838))) :rule cong :premises (t353.t15.t4)) (step t353.t15.t6 (cl (= @p_824 (! (=> @p_835 @p_838) :named @p_834))) :rule cong :premises (t353.t15.t2 t353.t15.t5)) (step t353.t15 (cl (= @p_829 @p_834)) :rule sko_forall) (step t353.t16 (cl (= @p_831 (! (=> @p_834 @p_762) :named @p_839))) :rule cong :premises (t353.t15)) (step t353.t17 (cl (! (= @p_833 (! (and @p_830 @p_839) :named @p_841)) :named @p_840)) :rule cong :premises (t353.t16)) (step t353.t18 (cl (not @p_840) (not @p_833) @p_841) :rule equiv_pos2) (step t353.t19 (cl @p_841) :rule th_resolution :premises (t353.t14 t353.t17 t353.t18)) (anchor :step t353.t20 :args ((:= (veriT_vr97 A$) veriT_vr98))) (step t353.t20.t1 (cl (! (= veriT_vr97 veriT_vr98) :named @p_844)) :rule refl) (step t353.t20.t2 (cl (= @p_825 (! (member$ veriT_vr98 @p_715) :named @p_843))) :rule cong :premises (t353.t20.t1)) (step t353.t20.t3 (cl @p_844) :rule refl) (step t353.t20.t4 (cl (= @p_827 (! (member$ veriT_vr98 @p_381) :named @p_845))) :rule cong :premises (t353.t20.t3)) (step t353.t20.t5 (cl (= @p_828 (! (not @p_845) :named @p_846))) :rule cong :premises (t353.t20.t4)) (step t353.t20.t6 (cl (= @p_824 (! (=> @p_843 @p_846) :named @p_847))) :rule cong :premises (t353.t20.t2 t353.t20.t5)) (step t353.t20 (cl (= @p_829 (! (forall ((veriT_vr98 A$)) @p_847) :named @p_842))) :rule bind) (step t353.t21 (cl (= @p_830 (! (=> @p_762 @p_842) :named @p_848))) :rule cong :premises (t353.t20)) (step t353.t22 (cl (! (= @p_841 (! (and @p_848 @p_839) :named @p_849)) :named @p_850)) :rule cong :premises (t353.t21)) (step t353.t23 (cl (not @p_850) (not @p_841) @p_849) :rule equiv_pos2) (step t353.t24 (cl @p_849) :rule th_resolution :premises (t353.t19 t353.t22 t353.t23)) (step t353 (cl @p_822 @p_849) :rule subproof :discharge (h1)) (step t354 (cl @p_851 @p_798) :rule or :premises (t352)) (step t355 (cl (! (or @p_851 @p_849) :named @p_853) (! (not @p_851) :named @p_852)) :rule or_neg) (step t356 (cl (not @p_852) @p_703) :rule not_not) (step t357 (cl @p_853 @p_703) :rule th_resolution :premises (t356 t355)) (step t358 (cl @p_853 (! (not @p_849) :named @p_1004)) :rule or_neg) (step t359 (cl @p_853) :rule th_resolution :premises (t354 t353 t357 t358)) (step t360 (cl (or @p_851 (! (and (! (=> (! (fun_app$ @p_854 @p_792) :named @p_855) (! (forall ((veriT_vr90 A$)) (! (=> (! (member$ veriT_vr90 @p_380) :named @p_859) @p_806) :named @p_864)) :named @p_858)) :named @p_866) (! (=> (! (forall ((veriT_vr91 A$)) (! (=> (! (member$ veriT_vr91 @p_380) :named @p_869) @p_816) :named @p_871)) :named @p_868) @p_855) :named @p_872)) :named @p_856))) :rule forall_inst :args ((:= veriT_vr88 @p_381) (:= veriT_vr89 @p_9))) (anchor :step t361) (assume t361.h1 @p_856) (anchor :step t361.t2 :args ((:= (veriT_vr90 A$) veriT_vr102))) (step t361.t2.t1 (cl (! (= veriT_vr90 veriT_vr102) :named @p_861)) :rule refl) (step t361.t2.t2 (cl (= @p_859 (! (member$ veriT_vr102 @p_380) :named @p_860))) :rule cong :premises (t361.t2.t1)) (step t361.t2.t3 (cl @p_861) :rule refl) (step t361.t2.t4 (cl (= @p_804 (! (member$ veriT_vr102 @p_381) :named @p_862))) :rule cong :premises (t361.t2.t3)) (step t361.t2.t5 (cl (= @p_806 (! (not @p_862) :named @p_863))) :rule cong :premises (t361.t2.t4)) (step t361.t2.t6 (cl (= @p_864 (! (=> @p_860 @p_863) :named @p_865))) :rule cong :premises (t361.t2.t2 t361.t2.t5)) (step t361.t2 (cl (= @p_858 (! (forall ((veriT_vr102 A$)) @p_865) :named @p_867))) :rule bind) (step t361.t3 (cl (= @p_866 (! (=> @p_855 @p_867) :named @p_873))) :rule cong :premises (t361.t2)) (anchor :step t361.t4 :args ((:= (veriT_vr91 A$) veriT_vr102))) (step t361.t4.t1 (cl (! (= veriT_vr91 veriT_vr102) :named @p_870)) :rule refl) (step t361.t4.t2 (cl (= @p_869 @p_860)) :rule cong :premises (t361.t4.t1)) (step t361.t4.t3 (cl @p_870) :rule refl) (step t361.t4.t4 (cl (= @p_815 @p_862)) :rule cong :premises (t361.t4.t3)) (step t361.t4.t5 (cl (= @p_816 @p_863)) :rule cong :premises (t361.t4.t4)) (step t361.t4.t6 (cl (= @p_871 @p_865)) :rule cong :premises (t361.t4.t2 t361.t4.t5)) (step t361.t4 (cl (= @p_868 @p_867)) :rule bind) (step t361.t5 (cl (= @p_872 (! (=> @p_867 @p_855) :named @p_874))) :rule cong :premises (t361.t4)) (step t361.t6 (cl (! (= @p_856 (! (and @p_873 @p_874) :named @p_877)) :named @p_875)) :rule cong :premises (t361.t3 t361.t5)) (step t361.t7 (cl (not @p_875) (! (not @p_856) :named @p_876) @p_877) :rule equiv_pos2) (step t361.t8 (cl @p_877) :rule th_resolution :premises (t361.h1 t361.t6 t361.t7)) (anchor :step t361.t9 :args ((:= (veriT_vr102 A$) veriT_vr103))) (step t361.t9.t1 (cl (! (= veriT_vr102 veriT_vr103) :named @p_880)) :rule refl) (step t361.t9.t2 (cl (= @p_860 @p_879)) :rule cong :premises (t361.t9.t1)) (step t361.t9.t3 (cl @p_880) :rule refl) (step t361.t9.t4 (cl (= @p_862 @p_881)) :rule cong :premises (t361.t9.t3)) (step t361.t9.t5 (cl (= @p_863 @p_882)) :rule cong :premises (t361.t9.t4)) (step t361.t9.t6 (cl (= @p_865 @p_878)) :rule cong :premises (t361.t9.t2 t361.t9.t5)) (step t361.t9 (cl (= @p_867 (! (forall ((veriT_vr103 A$)) @p_878) :named @p_883))) :rule bind) (step t361.t10 (cl (= @p_873 (! (=> @p_855 @p_883) :named @p_884))) :rule cong :premises (t361.t9)) (step t361.t11 (cl (= @p_874 (! (=> @p_883 @p_855) :named @p_885))) :rule cong :premises (t361.t9)) (step t361.t12 (cl (! (= @p_877 (! (and @p_884 @p_885) :named @p_887)) :named @p_886)) :rule cong :premises (t361.t10 t361.t11)) (step t361.t13 (cl (not @p_886) (not @p_877) @p_887) :rule equiv_pos2) (step t361.t14 (cl @p_887) :rule th_resolution :premises (t361.t8 t361.t12 t361.t13)) (anchor :step t361.t15 :args ((:= (veriT_vr103 A$) veriT_sk6))) (step t361.t15.t1 (cl (! (= veriT_vr103 veriT_sk6) :named @p_890)) :rule refl) (step t361.t15.t2 (cl (= @p_879 (! (member$ veriT_sk6 @p_380) :named @p_889))) :rule cong :premises (t361.t15.t1)) (step t361.t15.t3 (cl @p_890) :rule refl) (step t361.t15.t4 (cl (= @p_881 (! (member$ veriT_sk6 @p_381) :named @p_891))) :rule cong :premises (t361.t15.t3)) (step t361.t15.t5 (cl (= @p_882 (! (not @p_891) :named @p_892))) :rule cong :premises (t361.t15.t4)) (step t361.t15.t6 (cl (= @p_878 (! (=> @p_889 @p_892) :named @p_888))) :rule cong :premises (t361.t15.t2 t361.t15.t5)) (step t361.t15 (cl (= @p_883 @p_888)) :rule sko_forall) (step t361.t16 (cl (= @p_885 (! (=> @p_888 @p_855) :named @p_893))) :rule cong :premises (t361.t15)) (step t361.t17 (cl (! (= @p_887 (! (and @p_884 @p_893) :named @p_895)) :named @p_894)) :rule cong :premises (t361.t16)) (step t361.t18 (cl (not @p_894) (not @p_887) @p_895) :rule equiv_pos2) (step t361.t19 (cl @p_895) :rule th_resolution :premises (t361.t14 t361.t17 t361.t18)) (anchor :step t361.t20 :args ((:= (veriT_vr103 A$) veriT_vr104))) (step t361.t20.t1 (cl (! (= veriT_vr103 veriT_vr104) :named @p_898)) :rule refl) (step t361.t20.t2 (cl (= @p_879 (! (member$ veriT_vr104 @p_380) :named @p_897))) :rule cong :premises (t361.t20.t1)) (step t361.t20.t3 (cl @p_898) :rule refl) (step t361.t20.t4 (cl (= @p_881 (! (member$ veriT_vr104 @p_381) :named @p_899))) :rule cong :premises (t361.t20.t3)) (step t361.t20.t5 (cl (= @p_882 (! (not @p_899) :named @p_900))) :rule cong :premises (t361.t20.t4)) (step t361.t20.t6 (cl (= @p_878 (! (=> @p_897 @p_900) :named @p_901))) :rule cong :premises (t361.t20.t2 t361.t20.t5)) (step t361.t20 (cl (= @p_883 (! (forall ((veriT_vr104 A$)) @p_901) :named @p_896))) :rule bind) (step t361.t21 (cl (= @p_884 (! (=> @p_855 @p_896) :named @p_902))) :rule cong :premises (t361.t20)) (step t361.t22 (cl (! (= @p_895 (! (and @p_902 @p_893) :named @p_903)) :named @p_904)) :rule cong :premises (t361.t21)) (step t361.t23 (cl (not @p_904) (not @p_895) @p_903) :rule equiv_pos2) (step t361.t24 (cl @p_903) :rule th_resolution :premises (t361.t19 t361.t22 t361.t23)) (step t361 (cl @p_876 @p_903) :rule subproof :discharge (h1)) (step t362 (cl @p_851 @p_856) :rule or :premises (t360)) (step t363 (cl (! (or @p_851 @p_903) :named @p_905) @p_852) :rule or_neg) (step t364 (cl @p_905 @p_703) :rule th_resolution :premises (t356 t363)) (step t365 (cl @p_905 (! (not @p_903) :named @p_1006)) :rule or_neg) (step t366 (cl @p_905) :rule th_resolution :premises (t362 t361 t364 t365)) (step t367 (cl (or @p_906 (! (=> (! (forall ((veriT_vr76 A$)) (! (=> (! (member$ veriT_vr76 @p_21) :named @p_912) (! (member$ veriT_vr76 @p_760) :named @p_915)) :named @p_917)) :named @p_911) (! (fun_app$ @p_907 @p_760) :named @p_910)) :named @p_908))) :rule forall_inst :args ((:= veriT_vr74 @p_21) (:= veriT_vr75 @p_760))) (anchor :step t368) (assume t368.h1 @p_908) (anchor :step t368.t2 :args ((:= (veriT_vr76 A$) veriT_vr116))) (step t368.t2.t1 (cl (! (= veriT_vr76 veriT_vr116) :named @p_914)) :rule refl) (step t368.t2.t2 (cl (= @p_912 (! (member$ veriT_vr116 @p_21) :named @p_913))) :rule cong :premises (t368.t2.t1)) (step t368.t2.t3 (cl @p_914) :rule refl) (step t368.t2.t4 (cl (= @p_915 (! (member$ veriT_vr116 @p_760) :named @p_916))) :rule cong :premises (t368.t2.t3)) (step t368.t2.t5 (cl (= @p_917 (! (=> @p_913 @p_916) :named @p_918))) :rule cong :premises (t368.t2.t2 t368.t2.t4)) (step t368.t2 (cl (= @p_911 (! (forall ((veriT_vr116 A$)) @p_918) :named @p_919))) :rule bind) (step t368.t3 (cl (! (= @p_908 (! (=> @p_919 @p_910) :named @p_922)) :named @p_920)) :rule cong :premises (t368.t2)) (step t368.t4 (cl (not @p_920) (! (not @p_908) :named @p_921) @p_922) :rule equiv_pos2) (step t368.t5 (cl @p_922) :rule th_resolution :premises (t368.h1 t368.t3 t368.t4)) (anchor :step t368.t6 :args ((:= (veriT_vr116 A$) veriT_vr117))) (step t368.t6.t1 (cl (! (= veriT_vr116 veriT_vr117) :named @p_925)) :rule refl) (step t368.t6.t2 (cl (= @p_913 @p_924)) :rule cong :premises (t368.t6.t1)) (step t368.t6.t3 (cl @p_925) :rule refl) (step t368.t6.t4 (cl (= @p_916 @p_926)) :rule cong :premises (t368.t6.t3)) (step t368.t6.t5 (cl (= @p_918 @p_923)) :rule cong :premises (t368.t6.t2 t368.t6.t4)) (step t368.t6 (cl (= @p_919 (! (forall ((veriT_vr117 A$)) @p_923) :named @p_927))) :rule bind) (step t368.t7 (cl (! (= @p_922 (! (=> @p_927 @p_910) :named @p_929)) :named @p_928)) :rule cong :premises (t368.t6)) (step t368.t8 (cl (not @p_928) (not @p_922) @p_929) :rule equiv_pos2) (step t368.t9 (cl @p_929) :rule th_resolution :premises (t368.t5 t368.t7 t368.t8)) (anchor :step t368.t10 :args ((:= (veriT_vr117 A$) veriT_sk11))) (step t368.t10.t1 (cl (! (= veriT_vr117 veriT_sk11) :named @p_932)) :rule refl) (step t368.t10.t2 (cl (= @p_924 (! (member$ veriT_sk11 @p_21) :named @p_931))) :rule cong :premises (t368.t10.t1)) (step t368.t10.t3 (cl @p_932) :rule refl) (step t368.t10.t4 (cl (= @p_926 (! (member$ veriT_sk11 @p_760) :named @p_933))) :rule cong :premises (t368.t10.t3)) (step t368.t10.t5 (cl (= @p_923 (! (=> @p_931 @p_933) :named @p_930))) :rule cong :premises (t368.t10.t2 t368.t10.t4)) (step t368.t10 (cl (= @p_927 @p_930)) :rule sko_forall) (step t368.t11 (cl (! (= @p_929 (! (=> @p_930 @p_910) :named @p_934)) :named @p_935)) :rule cong :premises (t368.t10)) (step t368.t12 (cl (not @p_935) (not @p_929) @p_934) :rule equiv_pos2) (step t368.t13 (cl @p_934) :rule th_resolution :premises (t368.t9 t368.t11 t368.t12)) (step t368 (cl @p_921 @p_934) :rule subproof :discharge (h1)) (step t369 (cl @p_906 @p_908) :rule or :premises (t367)) (step t370 (cl (! (or @p_906 @p_934) :named @p_937) @p_936) :rule or_neg) (step t371 (cl @p_937 @p_592) :rule th_resolution :premises (t305 t370)) (step t372 (cl @p_937 (! (not @p_934) :named @p_1007)) :rule or_neg) (step t373 (cl @p_937) :rule th_resolution :premises (t369 t368 t371 t372)) (step t374 (cl (or @p_756 (! (= (! (set$ (append$ xs$ @p_938)) :named @p_1002) (! (sup$ @p_715 @p_939) :named @p_1061)) :named @p_1008))) :rule forall_inst :args ((:= veriT_vr63 xs$) (:= veriT_vr64 @p_938))) (step t375 (cl (or (! (not @p_357) :named @p_995) (! (and (! (=> (! (= bot$ @p_940) :named @p_941) (! (forall ((veriT_vr51 A$)) (! (=> (! (member$ veriT_vr51 @p_715) :named @p_945) (! (not (! (member$ veriT_vr51 @p_381) :named @p_948)) :named @p_950)) :named @p_952)) :named @p_944)) :named @p_954) (! (=> (! (forall ((veriT_vr52 A$)) (! (=> (! (member$ veriT_vr52 @p_715) :named @p_957) (! (not (! (member$ veriT_vr52 @p_381) :named @p_959)) :named @p_960)) :named @p_961)) :named @p_956) @p_941) :named @p_962)) :named @p_942))) :rule forall_inst :args ((:= veriT_vr49 @p_715) (:= veriT_vr50 @p_381))) (anchor :step t376) (assume t376.h1 @p_942) (anchor :step t376.t2 :args ((:= (veriT_vr51 A$) veriT_vr122))) (step t376.t2.t1 (cl (! (= veriT_vr51 veriT_vr122) :named @p_947)) :rule refl) (step t376.t2.t2 (cl (= @p_945 (! (member$ veriT_vr122 @p_715) :named @p_946))) :rule cong :premises (t376.t2.t1)) (step t376.t2.t3 (cl @p_947) :rule refl) (step t376.t2.t4 (cl (= @p_948 (! (member$ veriT_vr122 @p_381) :named @p_949))) :rule cong :premises (t376.t2.t3)) (step t376.t2.t5 (cl (= @p_950 (! (not @p_949) :named @p_951))) :rule cong :premises (t376.t2.t4)) (step t376.t2.t6 (cl (= @p_952 (! (=> @p_946 @p_951) :named @p_953))) :rule cong :premises (t376.t2.t2 t376.t2.t5)) (step t376.t2 (cl (= @p_944 (! (forall ((veriT_vr122 A$)) @p_953) :named @p_955))) :rule bind) (step t376.t3 (cl (= @p_954 (! (=> @p_941 @p_955) :named @p_963))) :rule cong :premises (t376.t2)) (anchor :step t376.t4 :args ((:= (veriT_vr52 A$) veriT_vr122))) (step t376.t4.t1 (cl (! (= veriT_vr52 veriT_vr122) :named @p_958)) :rule refl) (step t376.t4.t2 (cl (= @p_957 @p_946)) :rule cong :premises (t376.t4.t1)) (step t376.t4.t3 (cl @p_958) :rule refl) (step t376.t4.t4 (cl (= @p_959 @p_949)) :rule cong :premises (t376.t4.t3)) (step t376.t4.t5 (cl (= @p_960 @p_951)) :rule cong :premises (t376.t4.t4)) (step t376.t4.t6 (cl (= @p_961 @p_953)) :rule cong :premises (t376.t4.t2 t376.t4.t5)) (step t376.t4 (cl (= @p_956 @p_955)) :rule bind) (step t376.t5 (cl (= @p_962 (! (=> @p_955 @p_941) :named @p_964))) :rule cong :premises (t376.t4)) (step t376.t6 (cl (! (= @p_942 (! (and @p_963 @p_964) :named @p_967)) :named @p_965)) :rule cong :premises (t376.t3 t376.t5)) (step t376.t7 (cl (not @p_965) (! (not @p_942) :named @p_966) @p_967) :rule equiv_pos2) (step t376.t8 (cl @p_967) :rule th_resolution :premises (t376.h1 t376.t6 t376.t7)) (anchor :step t376.t9 :args ((:= (veriT_vr122 A$) veriT_vr123))) (step t376.t9.t1 (cl (! (= veriT_vr122 veriT_vr123) :named @p_970)) :rule refl) (step t376.t9.t2 (cl (= @p_946 @p_969)) :rule cong :premises (t376.t9.t1)) (step t376.t9.t3 (cl @p_970) :rule refl) (step t376.t9.t4 (cl (= @p_949 @p_971)) :rule cong :premises (t376.t9.t3)) (step t376.t9.t5 (cl (= @p_951 @p_972)) :rule cong :premises (t376.t9.t4)) (step t376.t9.t6 (cl (= @p_953 @p_968)) :rule cong :premises (t376.t9.t2 t376.t9.t5)) (step t376.t9 (cl (= @p_955 (! (forall ((veriT_vr123 A$)) @p_968) :named @p_973))) :rule bind) (step t376.t10 (cl (= @p_963 (! (=> @p_941 @p_973) :named @p_974))) :rule cong :premises (t376.t9)) (step t376.t11 (cl (= @p_964 (! (=> @p_973 @p_941) :named @p_975))) :rule cong :premises (t376.t9)) (step t376.t12 (cl (! (= @p_967 (! (and @p_974 @p_975) :named @p_977)) :named @p_976)) :rule cong :premises (t376.t10 t376.t11)) (step t376.t13 (cl (not @p_976) (not @p_967) @p_977) :rule equiv_pos2) (step t376.t14 (cl @p_977) :rule th_resolution :premises (t376.t8 t376.t12 t376.t13)) (anchor :step t376.t15 :args ((:= (veriT_vr123 A$) veriT_sk14))) (step t376.t15.t1 (cl (! (= veriT_vr123 veriT_sk14) :named @p_980)) :rule refl) (step t376.t15.t2 (cl (= @p_969 (! (member$ veriT_sk14 @p_715) :named @p_979))) :rule cong :premises (t376.t15.t1)) (step t376.t15.t3 (cl @p_980) :rule refl) (step t376.t15.t4 (cl (= @p_971 (! (member$ veriT_sk14 @p_381) :named @p_981))) :rule cong :premises (t376.t15.t3)) (step t376.t15.t5 (cl (= @p_972 (! (not @p_981) :named @p_982))) :rule cong :premises (t376.t15.t4)) (step t376.t15.t6 (cl (= @p_968 (! (=> @p_979 @p_982) :named @p_978))) :rule cong :premises (t376.t15.t2 t376.t15.t5)) (step t376.t15 (cl (= @p_973 @p_978)) :rule sko_forall) (step t376.t16 (cl (= @p_975 (! (=> @p_978 @p_941) :named @p_983))) :rule cong :premises (t376.t15)) (step t376.t17 (cl (! (= @p_977 (! (and @p_974 @p_983) :named @p_985)) :named @p_984)) :rule cong :premises (t376.t16)) (step t376.t18 (cl (not @p_984) (not @p_977) @p_985) :rule equiv_pos2) (step t376.t19 (cl @p_985) :rule th_resolution :premises (t376.t14 t376.t17 t376.t18)) (anchor :step t376.t20 :args ((:= (veriT_vr123 A$) veriT_vr124))) (step t376.t20.t1 (cl (! (= veriT_vr123 veriT_vr124) :named @p_988)) :rule refl) (step t376.t20.t2 (cl (= @p_969 (! (member$ veriT_vr124 @p_715) :named @p_987))) :rule cong :premises (t376.t20.t1)) (step t376.t20.t3 (cl @p_988) :rule refl) (step t376.t20.t4 (cl (= @p_971 (! (member$ veriT_vr124 @p_381) :named @p_989))) :rule cong :premises (t376.t20.t3)) (step t376.t20.t5 (cl (= @p_972 (! (not @p_989) :named @p_990))) :rule cong :premises (t376.t20.t4)) (step t376.t20.t6 (cl (= @p_968 (! (=> @p_987 @p_990) :named @p_991))) :rule cong :premises (t376.t20.t2 t376.t20.t5)) (step t376.t20 (cl (= @p_973 (! (forall ((veriT_vr124 A$)) @p_991) :named @p_986))) :rule bind) (step t376.t21 (cl (= @p_974 (! (=> @p_941 @p_986) :named @p_992))) :rule cong :premises (t376.t20)) (step t376.t22 (cl (! (= @p_985 (! (and @p_992 @p_983) :named @p_993)) :named @p_994)) :rule cong :premises (t376.t21)) (step t376.t23 (cl (not @p_994) (not @p_985) @p_993) :rule equiv_pos2) (step t376.t24 (cl @p_993) :rule th_resolution :premises (t376.t19 t376.t22 t376.t23)) (step t376 (cl @p_966 @p_993) :rule subproof :discharge (h1)) (step t377 (cl @p_995 @p_942) :rule or :premises (t375)) (step t378 (cl (! (or @p_995 @p_993) :named @p_997) (! (not @p_995) :named @p_996)) :rule or_neg) (step t379 (cl (not @p_996) @p_357) :rule not_not) (step t380 (cl @p_997 @p_357) :rule th_resolution :premises (t379 t378)) (step t381 (cl @p_997 (! (not @p_993) :named @p_1010)) :rule or_neg) (step t382 (cl @p_997) :rule th_resolution :premises (t377 t376 t380 t381)) (step t383 (cl (or @p_735 (! (= (! (= zero$ @p_998) :named @p_1001) (! (or @p_941 (not (finite$ @p_940))) :named @p_1011)) :named @p_1012))) :rule forall_inst :args ((:= veriT_vr35 @p_940))) (step t384 (cl (or (! (not @p_156) :named @p_1015) (! (= (! (= @p_999 @p_1000) :named @p_1014) @p_1001) :named @p_1013))) :rule forall_inst :args ((:= veriT_vr28 @p_999) (:= veriT_vr29 @p_998))) (step t385 (cl (or @p_791 (! (= (! (fun_app$ @p_43 @p_939) :named @p_1019) (! (and @p_44 (! (= @p_1 (! (fun_app$b card$ @p_1002) :named @p_1055)) :named @p_1017)) :named @p_1016)) :named @p_1018))) :rule forall_inst :args ((:= veriT_vr5 @p_938))) (step t386 (cl (not @p_848) @p_1003 @p_842) :rule implies_pos) (step t387 (cl @p_1004 @p_848) :rule and_pos) (step t388 (cl @p_851 @p_849) :rule or :premises (t359)) (step t389 (cl @p_849) :rule resolution :premises (t388 t229)) (step t390 (cl @p_848) :rule resolution :premises (t387 t389)) (step t391 (cl @p_888 @p_889) :rule implies_neg1) (step t392 (cl @p_888 (! (not @p_892) :named @p_1005)) :rule implies_neg2) (step t393 (cl (not @p_1005) @p_891) :rule not_not) (step t394 (cl @p_888 @p_891) :rule th_resolution :premises (t393 t392)) (step t395 (cl (not @p_893) (not @p_888) @p_855) :rule implies_pos) (step t396 (cl @p_1006 @p_893) :rule and_pos) (step t397 (cl @p_851 @p_903) :rule or :premises (t366)) (step t398 (cl @p_903) :rule resolution :premises (t397 t229)) (step t399 (cl @p_893) :rule resolution :premises (t396 t398)) (step t400 (cl @p_930 (not @p_933)) :rule implies_neg2) (step t401 (cl @p_1007 (not @p_930) @p_910) :rule implies_pos) (step t402 (cl @p_906 @p_934) :rule or :premises (t373)) (step t403 (cl @p_934) :rule resolution :premises (t402 t208)) (step t404 (cl @p_756 @p_1008) :rule or :premises (t374)) (step t405 (cl @p_1008) :rule resolution :premises (t404 t190)) (step t406 (cl @p_978 @p_979) :rule implies_neg1) (step t407 (cl @p_978 (! (not @p_982) :named @p_1009)) :rule implies_neg2) (step t408 (cl (not @p_1009) @p_981) :rule not_not) (step t409 (cl @p_978 @p_981) :rule th_resolution :premises (t408 t407)) (step t410 (cl (not @p_983) (not @p_978) @p_941) :rule implies_pos) (step t411 (cl @p_1010 @p_983) :rule and_pos) (step t412 (cl @p_995 @p_993) :rule or :premises (t382)) (step t413 (cl @p_993) :rule resolution :premises (t412 t114)) (step t414 (cl @p_983) :rule resolution :premises (t411 t413)) (step t415 (cl @p_1011 (not @p_941)) :rule or_neg) (step t416 (cl (not @p_1012) @p_1001 (not @p_1011)) :rule equiv_pos1) (step t417 (cl @p_735 @p_1012) :rule or :premises (t383)) (step t418 (cl @p_1012) :rule resolution :premises (t417 t81)) (step t419 (cl (not @p_1013) @p_1014 (not @p_1001)) :rule equiv_pos1) (step t420 (cl @p_1015 @p_1013) :rule or :premises (t384)) (step t421 (cl @p_1013) :rule resolution :premises (t420 t69)) (step t422 (cl (! (not @p_1016) :named @p_1020) @p_1017) :rule and_pos) (step t423 (cl @p_1016 (! (not @p_44) :named @p_1131) (! (not @p_1017) :named @p_1054)) :rule and_neg) (step t424 (cl (! (not @p_1018) :named @p_1021) @p_1019 @p_1020) :rule equiv_pos1) (step t425 (cl @p_1021 (! (not @p_1019) :named @p_1121) @p_1016) :rule equiv_pos2) (step t426 (cl @p_791 @p_1018) :rule or :premises (t385)) (step t427 (cl @p_1018) :rule resolution :premises (t426 t39)) (step t428 (cl (or (! (not @p_842) :named @p_1022) (! (forall ((veriT_vr98 A$)) (or (not @p_843) @p_846)) :named @p_1023))) :rule qnt_cnf) (step t429 (cl @p_1022 @p_1023) :rule or :premises (t428)) (step t430 (cl (or (! (not @p_1023) :named @p_1024) (! (or (! (not @p_979) :named @p_1031) @p_982) :named @p_1025))) :rule forall_inst :args ((:= veriT_vr98 veriT_sk14))) (step t431 (cl @p_1024 @p_1025) :rule or :premises (t430)) (step t432 (cl (! (or @p_1022 @p_1025) :named @p_1027) (! (not @p_1022) :named @p_1026)) :rule or_neg) (step t433 (cl (not @p_1026) @p_842) :rule not_not) (step t434 (cl @p_1027 @p_842) :rule th_resolution :premises (t433 t432)) (step t435 (cl @p_1027 (! (not @p_1025) :named @p_1030)) :rule or_neg) (step t436 (cl @p_1027) :rule th_resolution :premises (t429 t431 t434 t435)) (step t437 (cl @p_1028 @p_1029) :rule or :premises (t250)) (step t438 (cl (not (! (not @p_1028) :named @p_1034)) @p_474) :rule not_not) (step t439 (cl @p_1030 @p_1031 @p_982) :rule or_pos) (step t440 (cl @p_1022 @p_1025) :rule or :premises (t436)) (step t441 (cl (or (! (not @p_1029) :named @p_1032) (! (or (! (not @p_889) :named @p_1037) @p_892) :named @p_1033))) :rule forall_inst :args ((:= veriT_vr59 veriT_sk6))) (step t442 (cl @p_1032 @p_1033) :rule or :premises (t441)) (step t443 (cl (! (or @p_1028 @p_1033) :named @p_1035) @p_1034) :rule or_neg) (step t444 (cl @p_1035 @p_474) :rule th_resolution :premises (t438 t443)) (step t445 (cl @p_1035 (! (not @p_1033) :named @p_1036)) :rule or_neg) (step t446 (cl @p_1035) :rule th_resolution :premises (t437 t442 t444 t445)) (step t447 (cl (or (! (not @p_135) :named @p_1038) (! (member$ veriT_sk11 top$) :named @p_1039))) :rule forall_inst :args ((:= veriT_vr15 veriT_sk11))) (step t448 (cl @p_1036 @p_1037 @p_892) :rule or_pos) (step t449 (cl @p_1028 @p_1033) :rule or :premises (t446)) (step t450 (cl @p_1038 @p_1039) :rule or :premises (t447)) (step t451 (cl @p_1039) :rule resolution :premises (t450 t63)) (step t452 (cl @p_1030 @p_978) :rule resolution :premises (t439 t409 t406)) (step t453 (cl @p_1040 (! (not @p_722) :named @p_1043) @p_754) :rule eq_transitive) (step t454 (cl @p_1036 @p_888) :rule resolution :premises (t448 t394 t391)) (step t455 (cl (not (! (= veriT_sk0 veriT_sk0) :named @p_1041)) (! (not (! (= @p_381 @p_721) :named @p_1045)) :named @p_1042) @p_453 @p_745) :rule eq_congruent_pred) (step t456 (cl @p_1041) :rule eq_reflexive) (step t457 (cl @p_1042 @p_453 @p_745) :rule th_resolution :premises (t455 t456)) (step t458 (cl @p_1043 @p_1040 @p_794 @p_1044 @p_1045) :rule eq_transitive) (step t459 (cl @p_1043 @p_1040 @p_794 @p_1045 @p_797) :rule th_resolution :premises (t458 t350)) (step t460 (cl @p_453 @p_745 @p_1043 @p_1040 @p_794 @p_797) :rule th_resolution :premises (t457 t459)) (step t461 (cl @p_453 @p_745 @p_1040) :rule resolution :premises (t460 t266 t268 t274)) (step t462 (cl @p_1040 @p_454) :rule resolution :premises (t461 t297 t239 t236 t299)) (step t463 (cl (not (! (= veriT_sk1 veriT_sk1) :named @p_1047)) (! (not (! (= @p_380 @p_1046) :named @p_1050)) :named @p_1048) @p_459 @p_742) :rule eq_congruent_pred) (step t464 (cl @p_1047) :rule eq_reflexive) (step t465 (cl @p_1048 @p_459 @p_742) :rule th_resolution :premises (t463 t464)) (step t466 (cl @p_797 (! (not @p_728) :named @p_1053) (not (! (= @p_1046 @p_1049) :named @p_1052)) @p_1050) :rule eq_transitive) (step t467 (cl @p_1051 @p_1052) :rule eq_congruent) (step t468 (cl @p_1052 @p_1040 @p_1043) :rule th_resolution :premises (t467 t453)) (step t469 (cl @p_797 @p_1053 @p_1050 @p_1040 @p_1043) :rule th_resolution :premises (t466 t468)) (step t470 (cl @p_459 @p_742 @p_797 @p_1053 @p_1040 @p_1043) :rule th_resolution :premises (t465 t469)) (step t471 (cl @p_459 @p_742 @p_1040) :rule resolution :premises (t470 t266 t268 t276)) (step t472 (cl @p_1054 (! (not (! (= @p_999 @p_1055) :named @p_1058)) :named @p_1119) (! (not @p_1014) :named @p_1070) (! (not @p_785) :named @p_1071) (! (not (! (= @p_1056 @p_1057) :named @p_1075)) :named @p_1072) @p_400) :rule eq_transitive) (step t473 (cl (! (not (! (= card$ card$) :named @p_1059)) :named @p_1087) (! (not (! (= @p_759 @p_1002) :named @p_1062)) :named @p_1060) @p_1058) :rule eq_congruent) (step t474 (cl @p_1059) :rule eq_reflexive) (step t475 (cl @p_1060 @p_1058) :rule th_resolution :premises (t473 t474)) (step t476 (cl (not (! (= @p_759 @p_1061) :named @p_1063)) (! (not @p_1008) :named @p_1069) @p_1062) :rule eq_transitive) (step t477 (cl (! (not (! (= @p_715 @p_715) :named @p_1064)) :named @p_1080) (! (not (! (= @p_381 @p_939) :named @p_1066)) :named @p_1065) @p_1063) :rule eq_congruent) (step t478 (cl @p_1064) :rule eq_reflexive) (step t479 (cl @p_1065 @p_1063) :rule th_resolution :premises (t477 t478)) (step t480 (cl (! (not @p_724) :named @p_1067) (! (not @p_766) :named @p_1068) @p_1066) :rule eq_transitive) (step t481 (cl @p_1063 @p_1067 @p_1068) :rule th_resolution :premises (t479 t480)) (step t482 (cl @p_1069 @p_1062 @p_1067 @p_1068) :rule th_resolution :premises (t476 t481)) (step t483 (cl @p_1058 @p_1069 @p_1067 @p_1068) :rule th_resolution :premises (t475 t482)) (step t484 (cl @p_1054 @p_1070 @p_1071 @p_1072 @p_400 @p_1069 @p_1067 @p_1068) :rule th_resolution :premises (t472 t483)) (step t485 (cl (not (! (= @p_8 @p_1073) :named @p_1078)) (! (not (! (= @p_719 @p_1074) :named @p_1086)) :named @p_1084) @p_1075) :rule eq_congruent) (step t486 (cl (! (not @p_1076) :named @p_1083) (not (! (= @p_1077 @p_1073) :named @p_1081)) @p_1078) :rule eq_transitive) (step t487 (cl (! (not @p_1079) :named @p_1082) @p_1080 @p_1081) :rule eq_congruent) (step t488 (cl @p_1082 @p_1081) :rule th_resolution :premises (t487 t478)) (step t489 (cl @p_1083 @p_1078 @p_1082) :rule th_resolution :premises (t486 t488)) (step t490 (cl @p_1084 @p_1075 @p_1083 @p_1082) :rule th_resolution :premises (t485 t489)) (step t491 (cl (! (not @p_740) :named @p_1089) (not (! (= @p_1085 @p_1074) :named @p_1088)) @p_1086) :rule eq_transitive) (step t492 (cl @p_1087 @p_1043 @p_1088) :rule eq_congruent) (step t493 (cl @p_1043 @p_1088) :rule th_resolution :premises (t492 t474)) (step t494 (cl @p_1089 @p_1086 @p_1043) :rule th_resolution :premises (t491 t493)) (step t495 (cl @p_1075 @p_1083 @p_1082 @p_1089 @p_1043) :rule th_resolution :premises (t490 t494)) (step t496 (cl @p_1054 @p_1070 @p_1071 @p_400 @p_1069 @p_1067 @p_1068 @p_1083 @p_1082 @p_1089 @p_1043) :rule th_resolution :premises (t484 t495)) (step t497 (cl @p_1054 @p_1070 @p_400) :rule resolution :premises (t496 t270 axiom10 t319 t344 t405 t293 axiom11 t266)) (step t498 (cl (not (! (= @p_43 @p_854) :named @p_1091)) (! (not (! (= @p_21 @p_939) :named @p_1094)) :named @p_1093) @p_1019 @p_1003) :rule eq_congruent_pred) (step t499 (cl (! (not @p_1090) :named @p_1092) @p_1051 @p_1091) :rule eq_congruent) (step t500 (cl @p_1092 @p_1091 @p_1040 @p_1043) :rule th_resolution :premises (t499 t453)) (step t501 (cl @p_1093 @p_1019 @p_1003 @p_1092 @p_1040 @p_1043) :rule th_resolution :premises (t498 t500)) (step t502 (cl @p_1040 @p_1043 @p_1067 @p_1068 @p_1094) :rule eq_transitive) (step t503 (cl @p_1019 @p_1003 @p_1092 @p_1040 @p_1043 @p_1040 @p_1043 @p_1067 @p_1068) :rule th_resolution :premises (t501 t502)) (step t504 (cl @p_1019 @p_1003 @p_1092 @p_1040 @p_1043 @p_1067 @p_1068) :rule contraction :premises (t503)) (step t505 (cl @p_1019 @p_1003 @p_1040) :rule resolution :premises (t504 axiom24 t266 t319 t270)) (step t506 (cl (not (! (= @p_713 @p_757) :named @p_1095)) @p_773 (! (not @p_736) :named @p_1096)) :rule eq_congruent_pred) (step t507 (cl @p_1040 @p_794 @p_1095) :rule eq_transitive) (step t508 (cl @p_773 @p_1096 @p_1040 @p_794) :rule th_resolution :premises (t506 t507)) (step t509 (cl @p_773 @p_1040) :rule resolution :premises (t508 t274 t287)) (step t510 (cl @p_1040 @p_754) :rule resolution :premises (t453 t266)) (step t511 (cl (! (not (! (= @p_1097 @p_1098) :named @p_1101)) :named @p_1106) (! (not (! (= @p_713 @p_713) :named @p_1105)) :named @p_1104) @p_748 @p_1099) :rule eq_congruent_pred) (step t512 (cl (! (not @p_1100) :named @p_1103) (! (not (! (= @p_21 @p_21) :named @p_1102)) :named @p_1111) @p_1101) :rule eq_congruent) (step t513 (cl @p_1102) :rule eq_reflexive) (step t514 (cl @p_1103 @p_1101) :rule th_resolution :premises (t512 t513)) (step t515 (cl @p_1104 @p_748 @p_1099 @p_1103) :rule th_resolution :premises (t511 t514)) (step t516 (cl @p_1105) :rule eq_reflexive) (step t517 (cl @p_748 @p_1099 @p_1103) :rule th_resolution :premises (t515 t516)) (step t518 (cl @p_748 @p_1099) :rule resolution :premises (t517 axiom18)) (step t519 (cl rhs$) :rule resolution :premises (t471 t294 t243 t240 t246 t462 t497 t419 t416 t415 t410 t452 t422 t425 t505 t440 t282 t386 t284 t320 t328 t321 t509 t510 t301 t518 t248 t247 t296 t421 t418 t414 t427 t285 t390 t333 t323 t303)) (step t520 (cl @p_493) :rule resolution :premises (t234 t519)) (step t521 (cl @p_1099) :rule resolution :premises (t249 t519)) (step t522 (cl @p_422) :rule resolution :premises (t231 t520)) (step t523 (cl @p_400) :rule resolution :premises (t232 t520)) (step t524 (cl @p_474) :rule resolution :premises (t233 t520)) (step t525 (cl @p_733) :rule resolution :premises (t281 t522 t285)) (step t526 (cl @p_1033) :rule resolution :premises (t449 t524)) (step t527 (cl @p_7) :rule resolution :premises (t280 t525)) (step t528 (cl @p_888) :rule resolution :premises (t454 t526)) (step t529 (cl @p_855) :rule resolution :premises (t395 t528 t399)) (step t530 (cl @p_1106 @p_1104 @p_1107 @p_711) :rule eq_congruent_pred) (step t531 (cl @p_1104 @p_1107 @p_711 @p_1103) :rule th_resolution :premises (t530 t514)) (step t532 (cl @p_1107 @p_711 @p_1103) :rule th_resolution :premises (t531 t516)) (step t533 (cl @p_1107) :rule resolution :premises (t532 axiom18 t521)) (step t534 (cl (! (not (! (= @p_854 @p_854) :named @p_1110)) :named @p_1108) (not @p_796) @p_762 (! (not @p_855) :named @p_1109)) :rule eq_congruent_pred) (step t535 (cl @p_1108 @p_762 @p_1109 @p_794 @p_795 @p_797) :rule th_resolution :premises (t534 t351)) (step t536 (cl @p_1110) :rule eq_reflexive) (step t537 (cl @p_762 @p_1109 @p_794 @p_795 @p_797) :rule th_resolution :premises (t535 t536)) (step t538 (cl @p_762) :rule resolution :premises (t537 t268 t272 t274 t529)) (step t539 (cl @p_1040) :rule resolution :premises (t300 t533 t303)) (step t540 (cl @p_842) :rule resolution :premises (t386 t538 t390)) (step t541 (cl @p_1025) :rule resolution :premises (t440 t540)) (step t542 (cl @p_978) :rule resolution :premises (t452 t541)) (step t543 (cl @p_941) :rule resolution :premises (t410 t542 t414)) (step t544 (cl @p_1011) :rule resolution :premises (t415 t543)) (step t545 (cl @p_1001) :rule resolution :premises (t416 t544 t418)) (step t546 (cl @p_1014) :rule resolution :premises (t419 t545 t421)) (step t547 (cl @p_1111 @p_1051 @p_1043 @p_749) :rule eq_transitive) (step t548 (cl @p_1051 @p_1043 @p_749) :rule th_resolution :premises (t547 t513)) (step t549 (cl @p_1051) :rule resolution :premises (t548 t266 t539)) (step t550 (cl @p_708 @p_1072 @p_1071 @p_1070 (not (! (= @p_1112 @p_999) :named @p_1113)) @p_778) :rule eq_transitive) (step t551 (cl @p_1087 (! (not (! (= @p_758 @p_759) :named @p_1116)) :named @p_1114) @p_1113) :rule eq_congruent) (step t552 (cl @p_1114 @p_1113) :rule th_resolution :premises (t551 t474)) (step t553 (cl (! (not @p_771) :named @p_1118) (! (not (! (= @p_759 @p_1115) :named @p_1117)) :named @p_1127) @p_1116) :rule eq_transitive) (step t554 (cl @p_1080 @p_1043 @p_1117) :rule eq_congruent) (step t555 (cl @p_1043 @p_1117) :rule th_resolution :premises (t554 t478)) (step t556 (cl @p_1118 @p_1116 @p_1043) :rule th_resolution :premises (t553 t555)) (step t557 (cl @p_1113 @p_1118 @p_1043) :rule th_resolution :premises (t552 t556)) (step t558 (cl @p_708 @p_1072 @p_1071 @p_1070 @p_778 @p_1118 @p_1043) :rule th_resolution :premises (t550 t557)) (step t559 (cl @p_708 @p_1071 @p_1070 @p_778 @p_1118 @p_1043 @p_1083 @p_1082 @p_1089 @p_1043) :rule th_resolution :premises (t558 t495)) (step t560 (cl @p_708 @p_1071 @p_1070 @p_778 @p_1118 @p_1043 @p_1083 @p_1082 @p_1089) :rule contraction :premises (t559)) (step t561 (cl @p_778) :rule resolution :premises (t560 axiom11 t523 t266 t293 t327 t344 axiom10 t546)) (step t562 (cl @p_708 @p_1072 @p_1071 @p_1070 @p_1119 @p_1017) :rule eq_transitive) (step t563 (cl @p_708 @p_1072 @p_1071 @p_1070 @p_1017 @p_1069 @p_1067 @p_1068) :rule th_resolution :premises (t562 t483)) (step t564 (cl @p_708 @p_1071 @p_1070 @p_1017 @p_1069 @p_1067 @p_1068 @p_1083 @p_1082 @p_1089 @p_1043) :rule th_resolution :premises (t563 t495)) (step t565 (cl @p_1017) :rule resolution :premises (t564 axiom11 t523 t266 t270 axiom10 t319 t344 t405 t546 t293)) (step t566 (cl @p_767) :rule resolution :premises (t315 t549 t317)) (step t567 (cl @p_777) :rule resolution :premises (t334 t561 t527)) (step t568 (cl @p_1120) :rule resolution :premises (t314 t566 t538)) (step t569 (cl @p_780) :rule resolution :premises (t335 t567 t337)) (step t570 (cl (! (not (! (= @p_43 @p_907) :named @p_1123)) :named @p_1122) @p_1065 @p_1121 @p_763) :rule eq_congruent_pred) (step t571 (cl @p_1122 @p_1121 @p_763 @p_1067 @p_1068) :rule th_resolution :premises (t570 t480)) (step t572 (cl @p_1092 @p_1111 @p_1123) :rule eq_congruent) (step t573 (cl @p_1092 @p_1123) :rule th_resolution :premises (t572 t513)) (step t574 (cl @p_1121 @p_763 @p_1067 @p_1068 @p_1092) :rule th_resolution :premises (t571 t573)) (step t575 (cl @p_1121) :rule resolution :premises (t574 axiom24 t270 t568 t319)) (step t576 (cl (not (! (= veriT_sk11 veriT_sk11) :named @p_1124)) (! (not (! (= top$ @p_760) :named @p_1128)) :named @p_1125) @p_933 (! (not @p_1039) :named @p_1126)) :rule eq_congruent_pred) (step t577 (cl @p_1124) :rule eq_reflexive) (step t578 (cl @p_1125 @p_933 @p_1126) :rule th_resolution :premises (t576 t577)) (step t579 (cl (! (not @p_780) :named @p_1129) @p_1118 @p_1127 (! (not @p_770) :named @p_1130) @p_1128) :rule eq_transitive) (step t580 (cl @p_1129 @p_1118 @p_1130 @p_1128 @p_1043) :rule th_resolution :premises (t579 t555)) (step t581 (cl @p_933 @p_1126 @p_1129 @p_1118 @p_1130 @p_1043) :rule th_resolution :premises (t578 t580)) (step t582 (cl @p_933) :rule resolution :premises (t581 t266 t325 t327 t569 t451)) (step t583 (cl @p_1020) :rule resolution :premises (t424 t575 t427)) (step t584 (cl @p_930) :rule resolution :premises (t400 t582)) (step t585 (cl @p_1131) :rule resolution :premises (t423 t583 t565)) (step t586 (cl @p_910) :rule resolution :premises (t401 t584 t403)) (step t587 (cl @p_1132) :rule resolution :premises (t345 t585)) (step t588 (cl @p_1133) :rule resolution :premises (t346 t587 t348)) (step t589 (cl @p_1122 (! (not (! (= @p_760 @p_760) :named @p_1136)) :named @p_1134) (! (not @p_910) :named @p_1135) @p_790) :rule eq_congruent_pred) (step t590 (cl @p_1134 @p_1135 @p_790 @p_1092) :rule th_resolution :premises (t589 t573)) (step t591 (cl @p_1136) :rule eq_reflexive) (step t592 (cl @p_1135 @p_790 @p_1092) :rule th_resolution :premises (t590 t591)) (step t593 (cl) :rule resolution :premises (t592 axiom24 t588 t586)) d8bde960a6a2cb3d70d1b157d08487440b364301 654 0 unsat (assume axiom0 (! (forall ((?v0 Int)) (! (= (! (fun_app$ uua$ ?v0) :named @p_13) (! (line_integral_exists$ f$ (! (insert$ j$ bot$) :named @p_7)) :named @p_12)) :named @p_15)) :named @p_11)) (assume axiom1 (! (forall ((?v0 Int)) (! (= (! (fun_app$ uu$ ?v0) :named @p_25) (! (line_integral_exists$ f$ (! (insert$ i$ bot$) :named @p_5)) :named @p_24)) :named @p_27)) :named @p_23)) (assume axiom2 (! (forall ((?v0 Int_real_real_real_prod_fun_bool_fun_fun$) (?v1 Int_real_real_real_prod_fun_prod$)) (! (= (! (case_prod$ ?v0 ?v1) :named @p_36) (! (fun_app$a (! (fun_app$ ?v0 (! (fst$ ?v1) :named @p_40)) :named @p_42) (! (snd$ ?v1) :named @p_44)) :named @p_46)) :named @p_48)) :named @p_35)) (assume axiom3 (! (forall ((?v0 Real_real_prod$) (?v1 Real_real_prod$)) (! (=> (! (= (! (insert$ ?v0 bot$) :named @p_3) (! (insert$ ?v1 bot$) :named @p_64)) :named @p_66) (! (= ?v0 ?v1) :named @p_70)) :named @p_72)) :named @p_62)) (assume axiom4 (! (forall ((?v0 Int) (?v1 Real_real_real_prod_fun$)) (! (= ?v1 (! (snd$ (! (pair$ ?v0 ?v1) :named @p_87)) :named @p_89)) :named @p_91)) :named @p_85)) (assume axiom5 (! (forall ((?v0 Real) (?v1 Real)) (! (= ?v1 (! (snd$a (! (fun_app$b (! (pair$a ?v0) :named @p_102) ?v1) :named @p_105)) :named @p_107)) :named @p_109)) :named @p_101)) (assume axiom6 (! (member$ (! (pair$ k$ g$) :named @p_403) one_chain_typeI$) :named @p_402)) (assume axiom7 (! (forall ((?v0 Real_real_prod_set$) (?v1 Real_real_prod$) (?v2 Real_real_prod_set$)) (! (= (! (= bot$ (! (inf$ ?v0 (! (insert$ ?v1 ?v2) :named @p_1)) :named @p_122)) :named @p_124) (! (and (! (not (! (member$a ?v1 ?v0) :named @p_128)) :named @p_130) (! (= bot$ (! (inf$ ?v0 ?v2) :named @p_133)) :named @p_135)) :named @p_137)) :named @p_139)) :named @p_120)) (assume axiom8 (! (finite$ bot$) :named @p_414)) (assume axiom9 (! (forall ((?v0 Real_real_prod_set$) (?v1 Real_real_prod$)) (! (=> (! (finite$ ?v0) :named @p_4) (! (finite$ (! (insert$ ?v1 ?v0) :named @p_160)) :named @p_162)) :named @p_164)) :named @p_157)) (assume axiom10 (! (= i$ (! (fun_app$b (pair$a 1.0) 0.0) :named @p_417)) :named @p_499)) (assume axiom11 (! (forall ((?v0 Real_real_prod$) (?v1 Real_real_prod_set$)) (! (=> (! (member$a ?v0 ?v1) :named @p_176) (! (= ?v1 (! (insert$ ?v0 ?v1) :named @p_2)) :named @p_181)) :named @p_183)) :named @p_175)) (assume axiom12 (! (= j$ (! (fun_app$b (pair$a 0.0) 1.0) :named @p_419)) :named @p_500)) (assume axiom13 (! (forall ((?v0 Real_real_prod_set$)) (! (= bot$ (! (inf$ ?v0 bot$) :named @p_196)) :named @p_198)) :named @p_195)) (assume axiom14 (! (forall ((?v0 Real_real_prod$) (?v1 Real_real_prod$) (?v2 Real_real_prod_set$)) (! (= (! (insert$ ?v0 @p_1) :named @p_208) (! (insert$ ?v1 (! (insert$ ?v0 ?v2) :named @p_213)) :named @p_215)) :named @p_217)) :named @p_206)) (assume axiom15 (! (forall ((?v0 Real_real_prod$) (?v1 Real_real_prod_set$)) (! (= @p_2 (! (sup$ @p_3 ?v1) :named @p_236)) :named @p_238)) :named @p_231)) (assume axiom16 (! (forall ((?v0 Real_real_prod_set$) (?v1 Real_real_prod_real_real_prod_fun$) (?v2 Real_real_prod_set$) (?v3 Real_real_real_prod_fun$) (?v4 Real_real_prod_set$)) (! (=> (! (and @p_4 (! (and (! (fun_app$a (! (line_integral_exists$ ?v1 ?v2) :named @p_252) ?v3) :named @p_254) (! (and (! (fun_app$a (! (line_integral_exists$ ?v1 ?v4) :named @p_257) ?v3) :named @p_260) (! (and (! (= ?v0 (! (sup$ ?v2 ?v4) :named @p_265)) :named @p_267) (! (= bot$ (! (inf$ ?v2 ?v4) :named @p_269)) :named @p_271)) :named @p_273)) :named @p_275)) :named @p_277)) :named @p_279) (! (= (! (line_integral$ ?v1 ?v0 ?v3) :named @p_281) (! (+ (! (line_integral$ ?v1 ?v2 ?v3) :named @p_283) (! (line_integral$ ?v1 ?v4 ?v3) :named @p_285)) :named @p_287)) :named @p_289)) :named @p_291)) :named @p_250)) (assume axiom17 (! (and (! (= (one_chain_line_integral$ f$ @p_5 one_chain_typeI$) (one_chain_line_integral$ f$ @p_5 one_chain_typeII$)) :named @p_337) (! (and (! (forall ((?v0 Int_real_real_real_prod_fun_prod$)) (! (=> (! (member$ ?v0 one_chain_typeI$) :named @p_9) (! (case_prod$ uu$ ?v0) :named @p_6)) :named @p_326)) :named @p_322) (! (forall ((?v0 Int_real_real_real_prod_fun_prod$)) (! (=> (! (member$ ?v0 one_chain_typeII$) :named @p_8) @p_6) :named @p_331)) :named @p_328)) :named @p_333)) :named @p_336)) (assume axiom18 (! (and (! (= (one_chain_line_integral$ f$ @p_7 one_chain_typeII$) (one_chain_line_integral$ f$ @p_7 one_chain_typeI$)) :named @p_377) (! (and (! (forall ((?v0 Int_real_real_real_prod_fun_prod$)) (! (=> @p_8 (! (case_prod$ uua$ ?v0) :named @p_10)) :named @p_366)) :named @p_362) (! (forall ((?v0 Int_real_real_real_prod_fun_prod$)) (! (=> @p_9 @p_10) :named @p_371)) :named @p_368)) :named @p_373)) :named @p_376)) (assume axiom19 (not (! (= (! (line_integral$ f$ (! (insert$ i$ @p_7) :named @p_407) g$) :named @p_462) (! (+ (! (line_integral$ f$ @p_5 g$) :named @p_404) (! (line_integral$ f$ @p_7 g$) :named @p_405)) :named @p_463)) :named @p_410))) (anchor :step t21 :args ((:= (?v0 Int) veriT_vr0))) (step t21.t1 (cl (= ?v0 veriT_vr0)) :rule refl) (step t21.t2 (cl (= @p_13 (! (fun_app$ uua$ veriT_vr0) :named @p_14))) :rule cong :premises (t21.t1)) (step t21.t3 (cl (= @p_15 (! (= @p_12 @p_14) :named @p_16))) :rule cong :premises (t21.t2)) (step t21 (cl (! (= @p_11 (! (forall ((veriT_vr0 Int)) @p_16) :named @p_18)) :named @p_17)) :rule bind) (step t22 (cl (not @p_17) (not @p_11) @p_18) :rule equiv_pos2) (step t23 (cl @p_18) :rule th_resolution :premises (axiom0 t21 t22)) (anchor :step t24 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t24.t1 (cl (= veriT_vr0 veriT_vr1)) :rule refl) (step t24.t2 (cl (= @p_14 (! (fun_app$ uua$ veriT_vr1) :named @p_19))) :rule cong :premises (t24.t1)) (step t24.t3 (cl (= @p_16 (! (= @p_12 @p_19) :named @p_20))) :rule cong :premises (t24.t2)) (step t24 (cl (! (= @p_18 (! (forall ((veriT_vr1 Int)) @p_20) :named @p_22)) :named @p_21)) :rule bind) (step t25 (cl (not @p_21) (not @p_18) @p_22) :rule equiv_pos2) (step t26 (cl @p_22) :rule th_resolution :premises (t23 t24 t25)) (anchor :step t27 :args ((:= (?v0 Int) veriT_vr2))) (step t27.t1 (cl (= ?v0 veriT_vr2)) :rule refl) (step t27.t2 (cl (= @p_25 (! (fun_app$ uu$ veriT_vr2) :named @p_26))) :rule cong :premises (t27.t1)) (step t27.t3 (cl (= @p_27 (! (= @p_24 @p_26) :named @p_28))) :rule cong :premises (t27.t2)) (step t27 (cl (! (= @p_23 (! (forall ((veriT_vr2 Int)) @p_28) :named @p_30)) :named @p_29)) :rule bind) (step t28 (cl (not @p_29) (not @p_23) @p_30) :rule equiv_pos2) (step t29 (cl @p_30) :rule th_resolution :premises (axiom1 t27 t28)) (anchor :step t30 :args ((:= (veriT_vr2 Int) veriT_vr3))) (step t30.t1 (cl (= veriT_vr2 veriT_vr3)) :rule refl) (step t30.t2 (cl (= @p_26 (! (fun_app$ uu$ veriT_vr3) :named @p_31))) :rule cong :premises (t30.t1)) (step t30.t3 (cl (= @p_28 (! (= @p_24 @p_31) :named @p_32))) :rule cong :premises (t30.t2)) (step t30 (cl (! (= @p_30 (! (forall ((veriT_vr3 Int)) @p_32) :named @p_34)) :named @p_33)) :rule bind) (step t31 (cl (not @p_33) (not @p_30) @p_34) :rule equiv_pos2) (step t32 (cl @p_34) :rule th_resolution :premises (t29 t30 t31)) (anchor :step t33 :args ((:= (?v0 Int_real_real_real_prod_fun_bool_fun_fun$) veriT_vr4) (:= (?v1 Int_real_real_real_prod_fun_prod$) veriT_vr5))) (step t33.t1 (cl (! (= ?v0 veriT_vr4) :named @p_38)) :rule refl) (step t33.t2 (cl (! (= ?v1 veriT_vr5) :named @p_39)) :rule refl) (step t33.t3 (cl (= @p_36 (! (case_prod$ veriT_vr4 veriT_vr5) :named @p_37))) :rule cong :premises (t33.t1 t33.t2)) (step t33.t4 (cl @p_38) :rule refl) (step t33.t5 (cl @p_39) :rule refl) (step t33.t6 (cl (= @p_40 (! (fst$ veriT_vr5) :named @p_41))) :rule cong :premises (t33.t5)) (step t33.t7 (cl (= @p_42 (! (fun_app$ veriT_vr4 @p_41) :named @p_43))) :rule cong :premises (t33.t4 t33.t6)) (step t33.t8 (cl @p_39) :rule refl) (step t33.t9 (cl (= @p_44 (! (snd$ veriT_vr5) :named @p_45))) :rule cong :premises (t33.t8)) (step t33.t10 (cl (= @p_46 (! (fun_app$a @p_43 @p_45) :named @p_47))) :rule cong :premises (t33.t7 t33.t9)) (step t33.t11 (cl (= @p_48 (! (= @p_37 @p_47) :named @p_49))) :rule cong :premises (t33.t3 t33.t10)) (step t33 (cl (! (= @p_35 (! (forall ((veriT_vr4 Int_real_real_real_prod_fun_bool_fun_fun$) (veriT_vr5 Int_real_real_real_prod_fun_prod$)) @p_49) :named @p_51)) :named @p_50)) :rule bind) (step t34 (cl (not @p_50) (not @p_35) @p_51) :rule equiv_pos2) (step t35 (cl @p_51) :rule th_resolution :premises (axiom2 t33 t34)) (anchor :step t36 :args ((:= (veriT_vr4 Int_real_real_real_prod_fun_bool_fun_fun$) veriT_vr6) (:= (veriT_vr5 Int_real_real_real_prod_fun_prod$) veriT_vr7))) (step t36.t1 (cl (! (= veriT_vr4 veriT_vr6) :named @p_53)) :rule refl) (step t36.t2 (cl (! (= veriT_vr5 veriT_vr7) :named @p_54)) :rule refl) (step t36.t3 (cl (= @p_37 (! (case_prod$ veriT_vr6 veriT_vr7) :named @p_52))) :rule cong :premises (t36.t1 t36.t2)) (step t36.t4 (cl @p_53) :rule refl) (step t36.t5 (cl @p_54) :rule refl) (step t36.t6 (cl (= @p_41 (! (fst$ veriT_vr7) :named @p_55))) :rule cong :premises (t36.t5)) (step t36.t7 (cl (= @p_43 (! (fun_app$ veriT_vr6 @p_55) :named @p_56))) :rule cong :premises (t36.t4 t36.t6)) (step t36.t8 (cl @p_54) :rule refl) (step t36.t9 (cl (= @p_45 (! (snd$ veriT_vr7) :named @p_57))) :rule cong :premises (t36.t8)) (step t36.t10 (cl (= @p_47 (! (fun_app$a @p_56 @p_57) :named @p_58))) :rule cong :premises (t36.t7 t36.t9)) (step t36.t11 (cl (= @p_49 (! (= @p_52 @p_58) :named @p_59))) :rule cong :premises (t36.t3 t36.t10)) (step t36 (cl (! (= @p_51 (! (forall ((veriT_vr6 Int_real_real_real_prod_fun_bool_fun_fun$) (veriT_vr7 Int_real_real_real_prod_fun_prod$)) @p_59) :named @p_61)) :named @p_60)) :rule bind) (step t37 (cl (not @p_60) (not @p_51) @p_61) :rule equiv_pos2) (step t38 (cl @p_61) :rule th_resolution :premises (t35 t36 t37)) (anchor :step t39 :args ((:= (?v0 Real_real_prod$) veriT_vr8) (:= (?v1 Real_real_prod$) veriT_vr9))) (step t39.t1 (cl (! (= ?v0 veriT_vr8) :named @p_68)) :rule refl) (step t39.t2 (cl (= @p_3 (! (insert$ veriT_vr8 bot$) :named @p_63))) :rule cong :premises (t39.t1)) (step t39.t3 (cl (! (= ?v1 veriT_vr9) :named @p_69)) :rule refl) (step t39.t4 (cl (= @p_64 (! (insert$ veriT_vr9 bot$) :named @p_65))) :rule cong :premises (t39.t3)) (step t39.t5 (cl (= @p_66 (! (= @p_63 @p_65) :named @p_67))) :rule cong :premises (t39.t2 t39.t4)) (step t39.t6 (cl @p_68) :rule refl) (step t39.t7 (cl @p_69) :rule refl) (step t39.t8 (cl (= @p_70 (! (= veriT_vr8 veriT_vr9) :named @p_71))) :rule cong :premises (t39.t6 t39.t7)) (step t39.t9 (cl (= @p_72 (! (=> @p_67 @p_71) :named @p_73))) :rule cong :premises (t39.t5 t39.t8)) (step t39 (cl (! (= @p_62 (! (forall ((veriT_vr8 Real_real_prod$) (veriT_vr9 Real_real_prod$)) @p_73) :named @p_75)) :named @p_74)) :rule bind) (step t40 (cl (not @p_74) (not @p_62) @p_75) :rule equiv_pos2) (step t41 (cl @p_75) :rule th_resolution :premises (axiom3 t39 t40)) (anchor :step t42 :args ((:= (veriT_vr8 Real_real_prod$) veriT_vr10) (:= (veriT_vr9 Real_real_prod$) veriT_vr11))) (step t42.t1 (cl (! (= veriT_vr8 veriT_vr10) :named @p_79)) :rule refl) (step t42.t2 (cl (= @p_63 (! (insert$ veriT_vr10 bot$) :named @p_76))) :rule cong :premises (t42.t1)) (step t42.t3 (cl (! (= veriT_vr9 veriT_vr11) :named @p_80)) :rule refl) (step t42.t4 (cl (= @p_65 (! (insert$ veriT_vr11 bot$) :named @p_77))) :rule cong :premises (t42.t3)) (step t42.t5 (cl (= @p_67 (! (= @p_76 @p_77) :named @p_78))) :rule cong :premises (t42.t2 t42.t4)) (step t42.t6 (cl @p_79) :rule refl) (step t42.t7 (cl @p_80) :rule refl) (step t42.t8 (cl (= @p_71 (! (= veriT_vr10 veriT_vr11) :named @p_81))) :rule cong :premises (t42.t6 t42.t7)) (step t42.t9 (cl (= @p_73 (! (=> @p_78 @p_81) :named @p_82))) :rule cong :premises (t42.t5 t42.t8)) (step t42 (cl (! (= @p_75 (! (forall ((veriT_vr10 Real_real_prod$) (veriT_vr11 Real_real_prod$)) @p_82) :named @p_84)) :named @p_83)) :rule bind) (step t43 (cl (not @p_83) (not @p_75) @p_84) :rule equiv_pos2) (step t44 (cl @p_84) :rule th_resolution :premises (t41 t42 t43)) (anchor :step t45 :args ((:= (?v0 Int) veriT_vr12) (:= (?v1 Real_real_real_prod_fun$) veriT_vr13))) (step t45.t1 (cl (! (= ?v1 veriT_vr13) :named @p_86)) :rule refl) (step t45.t2 (cl (= ?v0 veriT_vr12)) :rule refl) (step t45.t3 (cl @p_86) :rule refl) (step t45.t4 (cl (= @p_87 (! (pair$ veriT_vr12 veriT_vr13) :named @p_88))) :rule cong :premises (t45.t2 t45.t3)) (step t45.t5 (cl (= @p_89 (! (snd$ @p_88) :named @p_90))) :rule cong :premises (t45.t4)) (step t45.t6 (cl (= @p_91 (! (= veriT_vr13 @p_90) :named @p_92))) :rule cong :premises (t45.t1 t45.t5)) (step t45 (cl (! (= @p_85 (! (forall ((veriT_vr12 Int) (veriT_vr13 Real_real_real_prod_fun$)) @p_92) :named @p_94)) :named @p_93)) :rule bind) (step t46 (cl (not @p_93) (not @p_85) @p_94) :rule equiv_pos2) (step t47 (cl @p_94) :rule th_resolution :premises (axiom4 t45 t46)) (anchor :step t48 :args ((:= (veriT_vr12 Int) veriT_vr14) (:= (veriT_vr13 Real_real_real_prod_fun$) veriT_vr15))) (step t48.t1 (cl (! (= veriT_vr13 veriT_vr15) :named @p_95)) :rule refl) (step t48.t2 (cl (= veriT_vr12 veriT_vr14)) :rule refl) (step t48.t3 (cl @p_95) :rule refl) (step t48.t4 (cl (= @p_88 (! (pair$ veriT_vr14 veriT_vr15) :named @p_96))) :rule cong :premises (t48.t2 t48.t3)) (step t48.t5 (cl (= @p_90 (! (snd$ @p_96) :named @p_97))) :rule cong :premises (t48.t4)) (step t48.t6 (cl (= @p_92 (! (= veriT_vr15 @p_97) :named @p_98))) :rule cong :premises (t48.t1 t48.t5)) (step t48 (cl (! (= @p_94 (! (forall ((veriT_vr14 Int) (veriT_vr15 Real_real_real_prod_fun$)) @p_98) :named @p_100)) :named @p_99)) :rule bind) (step t49 (cl (not @p_99) (not @p_94) @p_100) :rule equiv_pos2) (step t50 (cl @p_100) :rule th_resolution :premises (t47 t48 t49)) (anchor :step t51 :args ((:= (?v0 Real) veriT_vr16) (:= (?v1 Real) veriT_vr17))) (step t51.t1 (cl (! (= ?v1 veriT_vr17) :named @p_104)) :rule refl) (step t51.t2 (cl (= ?v0 veriT_vr16)) :rule refl) (step t51.t3 (cl (= @p_102 (! (pair$a veriT_vr16) :named @p_103))) :rule cong :premises (t51.t2)) (step t51.t4 (cl @p_104) :rule refl) (step t51.t5 (cl (= @p_105 (! (fun_app$b @p_103 veriT_vr17) :named @p_106))) :rule cong :premises (t51.t3 t51.t4)) (step t51.t6 (cl (= @p_107 (! (snd$a @p_106) :named @p_108))) :rule cong :premises (t51.t5)) (step t51.t7 (cl (= @p_109 (! (= veriT_vr17 @p_108) :named @p_110))) :rule cong :premises (t51.t1 t51.t6)) (step t51 (cl (! (= @p_101 (! (forall ((veriT_vr16 Real) (veriT_vr17 Real)) @p_110) :named @p_112)) :named @p_111)) :rule bind) (step t52 (cl (not @p_111) (not @p_101) @p_112) :rule equiv_pos2) (step t53 (cl @p_112) :rule th_resolution :premises (axiom5 t51 t52)) (anchor :step t54 :args ((:= (veriT_vr16 Real) veriT_vr18) (:= (veriT_vr17 Real) veriT_vr19))) (step t54.t1 (cl (! (= veriT_vr17 veriT_vr19) :named @p_114)) :rule refl) (step t54.t2 (cl (= veriT_vr16 veriT_vr18)) :rule refl) (step t54.t3 (cl (= @p_103 (! (pair$a veriT_vr18) :named @p_113))) :rule cong :premises (t54.t2)) (step t54.t4 (cl @p_114) :rule refl) (step t54.t5 (cl (= @p_106 (! (fun_app$b @p_113 veriT_vr19) :named @p_115))) :rule cong :premises (t54.t3 t54.t4)) (step t54.t6 (cl (= @p_108 (! (snd$a @p_115) :named @p_116))) :rule cong :premises (t54.t5)) (step t54.t7 (cl (= @p_110 (! (= veriT_vr19 @p_116) :named @p_117))) :rule cong :premises (t54.t1 t54.t6)) (step t54 (cl (! (= @p_112 (! (forall ((veriT_vr18 Real) (veriT_vr19 Real)) @p_117) :named @p_119)) :named @p_118)) :rule bind) (step t55 (cl (not @p_118) (not @p_112) @p_119) :rule equiv_pos2) (step t56 (cl @p_119) :rule th_resolution :premises (t53 t54 t55)) (anchor :step t57 :args ((:= (?v0 Real_real_prod_set$) veriT_vr20) (:= (?v1 Real_real_prod$) veriT_vr21) (:= (?v2 Real_real_prod_set$) veriT_vr22))) (step t57.t1 (cl (! (= ?v0 veriT_vr20) :named @p_127)) :rule refl) (step t57.t2 (cl (! (= ?v1 veriT_vr21) :named @p_126)) :rule refl) (step t57.t3 (cl (! (= ?v2 veriT_vr22) :named @p_132)) :rule refl) (step t57.t4 (cl (= @p_1 (! (insert$ veriT_vr21 veriT_vr22) :named @p_121))) :rule cong :premises (t57.t2 t57.t3)) (step t57.t5 (cl (= @p_122 (! (inf$ veriT_vr20 @p_121) :named @p_123))) :rule cong :premises (t57.t1 t57.t4)) (step t57.t6 (cl (= @p_124 (! (= bot$ @p_123) :named @p_125))) :rule cong :premises (t57.t5)) (step t57.t7 (cl @p_126) :rule refl) (step t57.t8 (cl @p_127) :rule refl) (step t57.t9 (cl (= @p_128 (! (member$a veriT_vr21 veriT_vr20) :named @p_129))) :rule cong :premises (t57.t7 t57.t8)) (step t57.t10 (cl (= @p_130 (! (not @p_129) :named @p_131))) :rule cong :premises (t57.t9)) (step t57.t11 (cl @p_127) :rule refl) (step t57.t12 (cl @p_132) :rule refl) (step t57.t13 (cl (= @p_133 (! (inf$ veriT_vr20 veriT_vr22) :named @p_134))) :rule cong :premises (t57.t11 t57.t12)) (step t57.t14 (cl (= @p_135 (! (= bot$ @p_134) :named @p_136))) :rule cong :premises (t57.t13)) (step t57.t15 (cl (= @p_137 (! (and @p_131 @p_136) :named @p_138))) :rule cong :premises (t57.t10 t57.t14)) (step t57.t16 (cl (= @p_139 (! (= @p_125 @p_138) :named @p_140))) :rule cong :premises (t57.t6 t57.t15)) (step t57 (cl (! (= @p_120 (! (forall ((veriT_vr20 Real_real_prod_set$) (veriT_vr21 Real_real_prod$) (veriT_vr22 Real_real_prod_set$)) @p_140) :named @p_142)) :named @p_141)) :rule bind) (step t58 (cl (not @p_141) (not @p_120) @p_142) :rule equiv_pos2) (step t59 (cl @p_142) :rule th_resolution :premises (axiom7 t57 t58)) (anchor :step t60 :args ((:= (veriT_vr20 Real_real_prod_set$) veriT_vr23) (:= (veriT_vr21 Real_real_prod$) veriT_vr24) (:= (veriT_vr22 Real_real_prod_set$) veriT_vr25))) (step t60.t1 (cl (! (= veriT_vr20 veriT_vr23) :named @p_147)) :rule refl) (step t60.t2 (cl (! (= veriT_vr21 veriT_vr24) :named @p_146)) :rule refl) (step t60.t3 (cl (! (= veriT_vr22 veriT_vr25) :named @p_150)) :rule refl) (step t60.t4 (cl (= @p_121 (! (insert$ veriT_vr24 veriT_vr25) :named @p_143))) :rule cong :premises (t60.t2 t60.t3)) (step t60.t5 (cl (= @p_123 (! (inf$ veriT_vr23 @p_143) :named @p_144))) :rule cong :premises (t60.t1 t60.t4)) (step t60.t6 (cl (= @p_125 (! (= bot$ @p_144) :named @p_145))) :rule cong :premises (t60.t5)) (step t60.t7 (cl @p_146) :rule refl) (step t60.t8 (cl @p_147) :rule refl) (step t60.t9 (cl (= @p_129 (! (member$a veriT_vr24 veriT_vr23) :named @p_148))) :rule cong :premises (t60.t7 t60.t8)) (step t60.t10 (cl (= @p_131 (! (not @p_148) :named @p_149))) :rule cong :premises (t60.t9)) (step t60.t11 (cl @p_147) :rule refl) (step t60.t12 (cl @p_150) :rule refl) (step t60.t13 (cl (= @p_134 (! (inf$ veriT_vr23 veriT_vr25) :named @p_151))) :rule cong :premises (t60.t11 t60.t12)) (step t60.t14 (cl (= @p_136 (! (= bot$ @p_151) :named @p_152))) :rule cong :premises (t60.t13)) (step t60.t15 (cl (= @p_138 (! (and @p_149 @p_152) :named @p_153))) :rule cong :premises (t60.t10 t60.t14)) (step t60.t16 (cl (= @p_140 (! (= @p_145 @p_153) :named @p_154))) :rule cong :premises (t60.t6 t60.t15)) (step t60 (cl (! (= @p_142 (! (forall ((veriT_vr23 Real_real_prod_set$) (veriT_vr24 Real_real_prod$) (veriT_vr25 Real_real_prod_set$)) @p_154) :named @p_156)) :named @p_155)) :rule bind) (step t61 (cl (not @p_155) (not @p_142) @p_156) :rule equiv_pos2) (step t62 (cl @p_156) :rule th_resolution :premises (t59 t60 t61)) (anchor :step t63 :args ((:= (?v0 Real_real_prod_set$) veriT_vr26) (:= (?v1 Real_real_prod$) veriT_vr27))) (step t63.t1 (cl (! (= ?v0 veriT_vr26) :named @p_159)) :rule refl) (step t63.t2 (cl (= @p_4 (! (finite$ veriT_vr26) :named @p_158))) :rule cong :premises (t63.t1)) (step t63.t3 (cl (= ?v1 veriT_vr27)) :rule refl) (step t63.t4 (cl @p_159) :rule refl) (step t63.t5 (cl (= @p_160 (! (insert$ veriT_vr27 veriT_vr26) :named @p_161))) :rule cong :premises (t63.t3 t63.t4)) (step t63.t6 (cl (= @p_162 (! (finite$ @p_161) :named @p_163))) :rule cong :premises (t63.t5)) (step t63.t7 (cl (= @p_164 (! (=> @p_158 @p_163) :named @p_165))) :rule cong :premises (t63.t2 t63.t6)) (step t63 (cl (! (= @p_157 (! (forall ((veriT_vr26 Real_real_prod_set$) (veriT_vr27 Real_real_prod$)) @p_165) :named @p_167)) :named @p_166)) :rule bind) (step t64 (cl (not @p_166) (not @p_157) @p_167) :rule equiv_pos2) (step t65 (cl @p_167) :rule th_resolution :premises (axiom9 t63 t64)) (anchor :step t66 :args ((:= (veriT_vr26 Real_real_prod_set$) veriT_vr28) (:= (veriT_vr27 Real_real_prod$) veriT_vr29))) (step t66.t1 (cl (! (= veriT_vr26 veriT_vr28) :named @p_169)) :rule refl) (step t66.t2 (cl (= @p_158 (! (finite$ veriT_vr28) :named @p_168))) :rule cong :premises (t66.t1)) (step t66.t3 (cl (= veriT_vr27 veriT_vr29)) :rule refl) (step t66.t4 (cl @p_169) :rule refl) (step t66.t5 (cl (= @p_161 (! (insert$ veriT_vr29 veriT_vr28) :named @p_170))) :rule cong :premises (t66.t3 t66.t4)) (step t66.t6 (cl (= @p_163 (! (finite$ @p_170) :named @p_171))) :rule cong :premises (t66.t5)) (step t66.t7 (cl (= @p_165 (! (=> @p_168 @p_171) :named @p_172))) :rule cong :premises (t66.t2 t66.t6)) (step t66 (cl (! (= @p_167 (! (forall ((veriT_vr28 Real_real_prod_set$) (veriT_vr29 Real_real_prod$)) @p_172) :named @p_174)) :named @p_173)) :rule bind) (step t67 (cl (not @p_173) (not @p_167) @p_174) :rule equiv_pos2) (step t68 (cl @p_174) :rule th_resolution :premises (t65 t66 t67)) (anchor :step t69 :args ((:= (?v0 Real_real_prod$) veriT_vr30) (:= (?v1 Real_real_prod_set$) veriT_vr31))) (step t69.t1 (cl (! (= ?v0 veriT_vr30) :named @p_179)) :rule refl) (step t69.t2 (cl (! (= ?v1 veriT_vr31) :named @p_178)) :rule refl) (step t69.t3 (cl (= @p_176 (! (member$a veriT_vr30 veriT_vr31) :named @p_177))) :rule cong :premises (t69.t1 t69.t2)) (step t69.t4 (cl @p_178) :rule refl) (step t69.t5 (cl @p_179) :rule refl) (step t69.t6 (cl @p_178) :rule refl) (step t69.t7 (cl (= @p_2 (! (insert$ veriT_vr30 veriT_vr31) :named @p_180))) :rule cong :premises (t69.t5 t69.t6)) (step t69.t8 (cl (= @p_181 (! (= veriT_vr31 @p_180) :named @p_182))) :rule cong :premises (t69.t4 t69.t7)) (step t69.t9 (cl (= @p_183 (! (=> @p_177 @p_182) :named @p_184))) :rule cong :premises (t69.t3 t69.t8)) (step t69 (cl (! (= @p_175 (! (forall ((veriT_vr30 Real_real_prod$) (veriT_vr31 Real_real_prod_set$)) @p_184) :named @p_186)) :named @p_185)) :rule bind) (step t70 (cl (not @p_185) (not @p_175) @p_186) :rule equiv_pos2) (step t71 (cl @p_186) :rule th_resolution :premises (axiom11 t69 t70)) (anchor :step t72 :args ((:= (veriT_vr30 Real_real_prod$) veriT_vr32) (:= (veriT_vr31 Real_real_prod_set$) veriT_vr33))) (step t72.t1 (cl (! (= veriT_vr30 veriT_vr32) :named @p_189)) :rule refl) (step t72.t2 (cl (! (= veriT_vr31 veriT_vr33) :named @p_188)) :rule refl) (step t72.t3 (cl (= @p_177 (! (member$a veriT_vr32 veriT_vr33) :named @p_187))) :rule cong :premises (t72.t1 t72.t2)) (step t72.t4 (cl @p_188) :rule refl) (step t72.t5 (cl @p_189) :rule refl) (step t72.t6 (cl @p_188) :rule refl) (step t72.t7 (cl (= @p_180 (! (insert$ veriT_vr32 veriT_vr33) :named @p_190))) :rule cong :premises (t72.t5 t72.t6)) (step t72.t8 (cl (= @p_182 (! (= veriT_vr33 @p_190) :named @p_191))) :rule cong :premises (t72.t4 t72.t7)) (step t72.t9 (cl (= @p_184 (! (=> @p_187 @p_191) :named @p_192))) :rule cong :premises (t72.t3 t72.t8)) (step t72 (cl (! (= @p_186 (! (forall ((veriT_vr32 Real_real_prod$) (veriT_vr33 Real_real_prod_set$)) @p_192) :named @p_194)) :named @p_193)) :rule bind) (step t73 (cl (not @p_193) (not @p_186) @p_194) :rule equiv_pos2) (step t74 (cl @p_194) :rule th_resolution :premises (t71 t72 t73)) (anchor :step t75 :args ((:= (?v0 Real_real_prod_set$) veriT_vr34))) (step t75.t1 (cl (= ?v0 veriT_vr34)) :rule refl) (step t75.t2 (cl (= @p_196 (! (inf$ veriT_vr34 bot$) :named @p_197))) :rule cong :premises (t75.t1)) (step t75.t3 (cl (= @p_198 (! (= bot$ @p_197) :named @p_199))) :rule cong :premises (t75.t2)) (step t75 (cl (! (= @p_195 (! (forall ((veriT_vr34 Real_real_prod_set$)) @p_199) :named @p_201)) :named @p_200)) :rule bind) (step t76 (cl (not @p_200) (not @p_195) @p_201) :rule equiv_pos2) (step t77 (cl @p_201) :rule th_resolution :premises (axiom13 t75 t76)) (anchor :step t78 :args ((:= (veriT_vr34 Real_real_prod_set$) veriT_vr35))) (step t78.t1 (cl (= veriT_vr34 veriT_vr35)) :rule refl) (step t78.t2 (cl (= @p_197 (! (inf$ veriT_vr35 bot$) :named @p_202))) :rule cong :premises (t78.t1)) (step t78.t3 (cl (= @p_199 (! (= bot$ @p_202) :named @p_203))) :rule cong :premises (t78.t2)) (step t78 (cl (! (= @p_201 (! (forall ((veriT_vr35 Real_real_prod_set$)) @p_203) :named @p_205)) :named @p_204)) :rule bind) (step t79 (cl (not @p_204) (not @p_201) @p_205) :rule equiv_pos2) (step t80 (cl @p_205) :rule th_resolution :premises (t77 t78 t79)) (anchor :step t81 :args ((:= (?v0 Real_real_prod$) veriT_vr36) (:= (?v1 Real_real_prod$) veriT_vr37) (:= (?v2 Real_real_prod_set$) veriT_vr38))) (step t81.t1 (cl (! (= ?v0 veriT_vr36) :named @p_211)) :rule refl) (step t81.t2 (cl (! (= ?v1 veriT_vr37) :named @p_210)) :rule refl) (step t81.t3 (cl (! (= ?v2 veriT_vr38) :named @p_212)) :rule refl) (step t81.t4 (cl (= @p_1 (! (insert$ veriT_vr37 veriT_vr38) :named @p_207))) :rule cong :premises (t81.t2 t81.t3)) (step t81.t5 (cl (= @p_208 (! (insert$ veriT_vr36 @p_207) :named @p_209))) :rule cong :premises (t81.t1 t81.t4)) (step t81.t6 (cl @p_210) :rule refl) (step t81.t7 (cl @p_211) :rule refl) (step t81.t8 (cl @p_212) :rule refl) (step t81.t9 (cl (= @p_213 (! (insert$ veriT_vr36 veriT_vr38) :named @p_214))) :rule cong :premises (t81.t7 t81.t8)) (step t81.t10 (cl (= @p_215 (! (insert$ veriT_vr37 @p_214) :named @p_216))) :rule cong :premises (t81.t6 t81.t9)) (step t81.t11 (cl (= @p_217 (! (= @p_209 @p_216) :named @p_218))) :rule cong :premises (t81.t5 t81.t10)) (step t81 (cl (! (= @p_206 (! (forall ((veriT_vr36 Real_real_prod$) (veriT_vr37 Real_real_prod$) (veriT_vr38 Real_real_prod_set$)) @p_218) :named @p_220)) :named @p_219)) :rule bind) (step t82 (cl (not @p_219) (not @p_206) @p_220) :rule equiv_pos2) (step t83 (cl @p_220) :rule th_resolution :premises (axiom14 t81 t82)) (anchor :step t84 :args ((:= (veriT_vr36 Real_real_prod$) veriT_vr39) (:= (veriT_vr37 Real_real_prod$) veriT_vr40) (:= (veriT_vr38 Real_real_prod_set$) veriT_vr41))) (step t84.t1 (cl (! (= veriT_vr36 veriT_vr39) :named @p_224)) :rule refl) (step t84.t2 (cl (! (= veriT_vr37 veriT_vr40) :named @p_223)) :rule refl) (step t84.t3 (cl (! (= veriT_vr38 veriT_vr41) :named @p_225)) :rule refl) (step t84.t4 (cl (= @p_207 (! (insert$ veriT_vr40 veriT_vr41) :named @p_221))) :rule cong :premises (t84.t2 t84.t3)) (step t84.t5 (cl (= @p_209 (! (insert$ veriT_vr39 @p_221) :named @p_222))) :rule cong :premises (t84.t1 t84.t4)) (step t84.t6 (cl @p_223) :rule refl) (step t84.t7 (cl @p_224) :rule refl) (step t84.t8 (cl @p_225) :rule refl) (step t84.t9 (cl (= @p_214 (! (insert$ veriT_vr39 veriT_vr41) :named @p_226))) :rule cong :premises (t84.t7 t84.t8)) (step t84.t10 (cl (= @p_216 (! (insert$ veriT_vr40 @p_226) :named @p_227))) :rule cong :premises (t84.t6 t84.t9)) (step t84.t11 (cl (= @p_218 (! (= @p_222 @p_227) :named @p_228))) :rule cong :premises (t84.t5 t84.t10)) (step t84 (cl (! (= @p_220 (! (forall ((veriT_vr39 Real_real_prod$) (veriT_vr40 Real_real_prod$) (veriT_vr41 Real_real_prod_set$)) @p_228) :named @p_230)) :named @p_229)) :rule bind) (step t85 (cl (not @p_229) (not @p_220) @p_230) :rule equiv_pos2) (step t86 (cl @p_230) :rule th_resolution :premises (t83 t84 t85)) (anchor :step t87 :args ((:= (?v0 Real_real_prod$) veriT_vr42) (:= (?v1 Real_real_prod_set$) veriT_vr43))) (step t87.t1 (cl (! (= ?v0 veriT_vr42) :named @p_233)) :rule refl) (step t87.t2 (cl (! (= ?v1 veriT_vr43) :named @p_235)) :rule refl) (step t87.t3 (cl (= @p_2 (! (insert$ veriT_vr42 veriT_vr43) :named @p_232))) :rule cong :premises (t87.t1 t87.t2)) (step t87.t4 (cl @p_233) :rule refl) (step t87.t5 (cl (= @p_3 (! (insert$ veriT_vr42 bot$) :named @p_234))) :rule cong :premises (t87.t4)) (step t87.t6 (cl @p_235) :rule refl) (step t87.t7 (cl (= @p_236 (! (sup$ @p_234 veriT_vr43) :named @p_237))) :rule cong :premises (t87.t5 t87.t6)) (step t87.t8 (cl (= @p_238 (! (= @p_232 @p_237) :named @p_239))) :rule cong :premises (t87.t3 t87.t7)) (step t87 (cl (! (= @p_231 (! (forall ((veriT_vr42 Real_real_prod$) (veriT_vr43 Real_real_prod_set$)) @p_239) :named @p_241)) :named @p_240)) :rule bind) (step t88 (cl (not @p_240) (not @p_231) @p_241) :rule equiv_pos2) (step t89 (cl @p_241) :rule th_resolution :premises (axiom15 t87 t88)) (anchor :step t90 :args ((:= (veriT_vr42 Real_real_prod$) veriT_vr44) (:= (veriT_vr43 Real_real_prod_set$) veriT_vr45))) (step t90.t1 (cl (! (= veriT_vr42 veriT_vr44) :named @p_243)) :rule refl) (step t90.t2 (cl (! (= veriT_vr43 veriT_vr45) :named @p_245)) :rule refl) (step t90.t3 (cl (= @p_232 (! (insert$ veriT_vr44 veriT_vr45) :named @p_242))) :rule cong :premises (t90.t1 t90.t2)) (step t90.t4 (cl @p_243) :rule refl) (step t90.t5 (cl (= @p_234 (! (insert$ veriT_vr44 bot$) :named @p_244))) :rule cong :premises (t90.t4)) (step t90.t6 (cl @p_245) :rule refl) (step t90.t7 (cl (= @p_237 (! (sup$ @p_244 veriT_vr45) :named @p_246))) :rule cong :premises (t90.t5 t90.t6)) (step t90.t8 (cl (= @p_239 (! (= @p_242 @p_246) :named @p_247))) :rule cong :premises (t90.t3 t90.t7)) (step t90 (cl (! (= @p_241 (! (forall ((veriT_vr44 Real_real_prod$) (veriT_vr45 Real_real_prod_set$)) @p_247) :named @p_249)) :named @p_248)) :rule bind) (step t91 (cl (not @p_248) (not @p_241) @p_249) :rule equiv_pos2) (step t92 (cl @p_249) :rule th_resolution :premises (t89 t90 t91)) (anchor :step t93 :args ((:= (?v0 Real_real_prod_set$) veriT_vr46) (:= (?v1 Real_real_prod_real_real_prod_fun$) veriT_vr47) (:= (?v2 Real_real_prod_set$) veriT_vr48) (:= (?v3 Real_real_real_prod_fun$) veriT_vr49) (:= (?v4 Real_real_prod_set$) veriT_vr50))) (step t93.t1 (cl (! (= ?v0 veriT_vr46) :named @p_262)) :rule refl) (step t93.t2 (cl (= @p_4 (! (finite$ veriT_vr46) :named @p_251))) :rule cong :premises (t93.t1)) (step t93.t3 (cl (! (= ?v1 veriT_vr47) :named @p_256)) :rule refl) (step t93.t4 (cl (! (= ?v2 veriT_vr48) :named @p_263)) :rule refl) (step t93.t5 (cl (= @p_252 (! (line_integral_exists$ veriT_vr47 veriT_vr48) :named @p_253))) :rule cong :premises (t93.t3 t93.t4)) (step t93.t6 (cl (! (= ?v3 veriT_vr49) :named @p_259)) :rule refl) (step t93.t7 (cl (= @p_254 (! (fun_app$a @p_253 veriT_vr49) :named @p_255))) :rule cong :premises (t93.t5 t93.t6)) (step t93.t8 (cl @p_256) :rule refl) (step t93.t9 (cl (! (= ?v4 veriT_vr50) :named @p_264)) :rule refl) (step t93.t10 (cl (= @p_257 (! (line_integral_exists$ veriT_vr47 veriT_vr50) :named @p_258))) :rule cong :premises (t93.t8 t93.t9)) (step t93.t11 (cl @p_259) :rule refl) (step t93.t12 (cl (= @p_260 (! (fun_app$a @p_258 veriT_vr49) :named @p_261))) :rule cong :premises (t93.t10 t93.t11)) (step t93.t13 (cl @p_262) :rule refl) (step t93.t14 (cl @p_263) :rule refl) (step t93.t15 (cl @p_264) :rule refl) (step t93.t16 (cl (= @p_265 (! (sup$ veriT_vr48 veriT_vr50) :named @p_266))) :rule cong :premises (t93.t14 t93.t15)) (step t93.t17 (cl (= @p_267 (! (= veriT_vr46 @p_266) :named @p_268))) :rule cong :premises (t93.t13 t93.t16)) (step t93.t18 (cl @p_263) :rule refl) (step t93.t19 (cl @p_264) :rule refl) (step t93.t20 (cl (= @p_269 (! (inf$ veriT_vr48 veriT_vr50) :named @p_270))) :rule cong :premises (t93.t18 t93.t19)) (step t93.t21 (cl (= @p_271 (! (= bot$ @p_270) :named @p_272))) :rule cong :premises (t93.t20)) (step t93.t22 (cl (= @p_273 (! (and @p_268 @p_272) :named @p_274))) :rule cong :premises (t93.t17 t93.t21)) (step t93.t23 (cl (= @p_275 (! (and @p_261 @p_274) :named @p_276))) :rule cong :premises (t93.t12 t93.t22)) (step t93.t24 (cl (= @p_277 (! (and @p_255 @p_276) :named @p_278))) :rule cong :premises (t93.t7 t93.t23)) (step t93.t25 (cl (= @p_279 (! (and @p_251 @p_278) :named @p_280))) :rule cong :premises (t93.t2 t93.t24)) (step t93.t26 (cl @p_256) :rule refl) (step t93.t27 (cl @p_262) :rule refl) (step t93.t28 (cl @p_259) :rule refl) (step t93.t29 (cl (= @p_281 (! (line_integral$ veriT_vr47 veriT_vr46 veriT_vr49) :named @p_282))) :rule cong :premises (t93.t26 t93.t27 t93.t28)) (step t93.t30 (cl @p_256) :rule refl) (step t93.t31 (cl @p_263) :rule refl) (step t93.t32 (cl @p_259) :rule refl) (step t93.t33 (cl (= @p_283 (! (line_integral$ veriT_vr47 veriT_vr48 veriT_vr49) :named @p_284))) :rule cong :premises (t93.t30 t93.t31 t93.t32)) (step t93.t34 (cl @p_256) :rule refl) (step t93.t35 (cl @p_264) :rule refl) (step t93.t36 (cl @p_259) :rule refl) (step t93.t37 (cl (= @p_285 (! (line_integral$ veriT_vr47 veriT_vr50 veriT_vr49) :named @p_286))) :rule cong :premises (t93.t34 t93.t35 t93.t36)) (step t93.t38 (cl (= @p_287 (! (+ @p_284 @p_286) :named @p_288))) :rule cong :premises (t93.t33 t93.t37)) (step t93.t39 (cl (= @p_289 (! (= @p_282 @p_288) :named @p_290))) :rule cong :premises (t93.t29 t93.t38)) (step t93.t40 (cl (= @p_291 (! (=> @p_280 @p_290) :named @p_292))) :rule cong :premises (t93.t25 t93.t39)) (step t93 (cl (! (= @p_250 (! (forall ((veriT_vr46 Real_real_prod_set$) (veriT_vr47 Real_real_prod_real_real_prod_fun$) (veriT_vr48 Real_real_prod_set$) (veriT_vr49 Real_real_real_prod_fun$) (veriT_vr50 Real_real_prod_set$)) @p_292) :named @p_294)) :named @p_293)) :rule bind) (step t94 (cl (not @p_293) (not @p_250) @p_294) :rule equiv_pos2) (step t95 (cl @p_294) :rule th_resolution :premises (axiom16 t93 t94)) (anchor :step t96 :args ((veriT_vr46 Real_real_prod_set$) (veriT_vr47 Real_real_prod_real_real_prod_fun$) (veriT_vr48 Real_real_prod_set$) (veriT_vr49 Real_real_real_prod_fun$) (veriT_vr50 Real_real_prod_set$))) (step t96.t1 (cl (= @p_280 (! (and @p_251 @p_255 @p_261 @p_268 @p_272) :named @p_295))) :rule ac_simp) (step t96.t2 (cl (= @p_292 (! (=> @p_295 @p_290) :named @p_296))) :rule cong :premises (t96.t1)) (step t96 (cl (! (= @p_294 (! (forall ((veriT_vr46 Real_real_prod_set$) (veriT_vr47 Real_real_prod_real_real_prod_fun$) (veriT_vr48 Real_real_prod_set$) (veriT_vr49 Real_real_real_prod_fun$) (veriT_vr50 Real_real_prod_set$)) @p_296) :named @p_298)) :named @p_297)) :rule bind) (step t97 (cl (not @p_297) (not @p_294) @p_298) :rule equiv_pos2) (step t98 (cl @p_298) :rule th_resolution :premises (t95 t96 t97)) (anchor :step t99 :args ((:= (veriT_vr46 Real_real_prod_set$) veriT_vr51) (:= (veriT_vr47 Real_real_prod_real_real_prod_fun$) veriT_vr52) (:= (veriT_vr48 Real_real_prod_set$) veriT_vr53) (:= (veriT_vr49 Real_real_real_prod_fun$) veriT_vr54) (:= (veriT_vr50 Real_real_prod_set$) veriT_vr55))) (step t99.t1 (cl (! (= veriT_vr46 veriT_vr51) :named @p_306)) :rule refl) (step t99.t2 (cl (= @p_251 (! (finite$ veriT_vr51) :named @p_299))) :rule cong :premises (t99.t1)) (step t99.t3 (cl (! (= veriT_vr47 veriT_vr52) :named @p_302)) :rule refl) (step t99.t4 (cl (! (= veriT_vr48 veriT_vr53) :named @p_307)) :rule refl) (step t99.t5 (cl (= @p_253 (! (line_integral_exists$ veriT_vr52 veriT_vr53) :named @p_300))) :rule cong :premises (t99.t3 t99.t4)) (step t99.t6 (cl (! (= veriT_vr49 veriT_vr54) :named @p_304)) :rule refl) (step t99.t7 (cl (= @p_255 (! (fun_app$a @p_300 veriT_vr54) :named @p_301))) :rule cong :premises (t99.t5 t99.t6)) (step t99.t8 (cl @p_302) :rule refl) (step t99.t9 (cl (! (= veriT_vr50 veriT_vr55) :named @p_308)) :rule refl) (step t99.t10 (cl (= @p_258 (! (line_integral_exists$ veriT_vr52 veriT_vr55) :named @p_303))) :rule cong :premises (t99.t8 t99.t9)) (step t99.t11 (cl @p_304) :rule refl) (step t99.t12 (cl (= @p_261 (! (fun_app$a @p_303 veriT_vr54) :named @p_305))) :rule cong :premises (t99.t10 t99.t11)) (step t99.t13 (cl @p_306) :rule refl) (step t99.t14 (cl @p_307) :rule refl) (step t99.t15 (cl @p_308) :rule refl) (step t99.t16 (cl (= @p_266 (! (sup$ veriT_vr53 veriT_vr55) :named @p_309))) :rule cong :premises (t99.t14 t99.t15)) (step t99.t17 (cl (= @p_268 (! (= veriT_vr51 @p_309) :named @p_310))) :rule cong :premises (t99.t13 t99.t16)) (step t99.t18 (cl @p_307) :rule refl) (step t99.t19 (cl @p_308) :rule refl) (step t99.t20 (cl (= @p_270 (! (inf$ veriT_vr53 veriT_vr55) :named @p_311))) :rule cong :premises (t99.t18 t99.t19)) (step t99.t21 (cl (= @p_272 (! (= bot$ @p_311) :named @p_312))) :rule cong :premises (t99.t20)) (step t99.t22 (cl (= @p_295 (! (and @p_299 @p_301 @p_305 @p_310 @p_312) :named @p_313))) :rule cong :premises (t99.t2 t99.t7 t99.t12 t99.t17 t99.t21)) (step t99.t23 (cl @p_302) :rule refl) (step t99.t24 (cl @p_306) :rule refl) (step t99.t25 (cl @p_304) :rule refl) (step t99.t26 (cl (= @p_282 (! (line_integral$ veriT_vr52 veriT_vr51 veriT_vr54) :named @p_314))) :rule cong :premises (t99.t23 t99.t24 t99.t25)) (step t99.t27 (cl @p_302) :rule refl) (step t99.t28 (cl @p_307) :rule refl) (step t99.t29 (cl @p_304) :rule refl) (step t99.t30 (cl (= @p_284 (! (line_integral$ veriT_vr52 veriT_vr53 veriT_vr54) :named @p_315))) :rule cong :premises (t99.t27 t99.t28 t99.t29)) (step t99.t31 (cl @p_302) :rule refl) (step t99.t32 (cl @p_308) :rule refl) (step t99.t33 (cl @p_304) :rule refl) (step t99.t34 (cl (= @p_286 (! (line_integral$ veriT_vr52 veriT_vr55 veriT_vr54) :named @p_316))) :rule cong :premises (t99.t31 t99.t32 t99.t33)) (step t99.t35 (cl (= @p_288 (! (+ @p_315 @p_316) :named @p_317))) :rule cong :premises (t99.t30 t99.t34)) (step t99.t36 (cl (= @p_290 (! (= @p_314 @p_317) :named @p_318))) :rule cong :premises (t99.t26 t99.t35)) (step t99.t37 (cl (= @p_296 (! (=> @p_313 @p_318) :named @p_319))) :rule cong :premises (t99.t22 t99.t36)) (step t99 (cl (! (= @p_298 (! (forall ((veriT_vr51 Real_real_prod_set$) (veriT_vr52 Real_real_prod_real_real_prod_fun$) (veriT_vr53 Real_real_prod_set$) (veriT_vr54 Real_real_real_prod_fun$) (veriT_vr55 Real_real_prod_set$)) @p_319) :named @p_321)) :named @p_320)) :rule bind) (step t100 (cl (not @p_320) (not @p_298) @p_321) :rule equiv_pos2) (step t101 (cl @p_321) :rule th_resolution :premises (t98 t99 t100)) (anchor :step t102 :args ((:= (?v0 Int_real_real_real_prod_fun_prod$) veriT_vr56))) (step t102.t1 (cl (! (= ?v0 veriT_vr56) :named @p_324)) :rule refl) (step t102.t2 (cl (= @p_9 (! (member$ veriT_vr56 one_chain_typeI$) :named @p_323))) :rule cong :premises (t102.t1)) (step t102.t3 (cl @p_324) :rule refl) (step t102.t4 (cl (! (= @p_6 (! (case_prod$ uu$ veriT_vr56) :named @p_325)) :named @p_330)) :rule cong :premises (t102.t3)) (step t102.t5 (cl (= @p_326 (! (=> @p_323 @p_325) :named @p_327))) :rule cong :premises (t102.t2 t102.t4)) (step t102 (cl (= @p_322 (! (forall ((veriT_vr56 Int_real_real_real_prod_fun_prod$)) @p_327) :named @p_334))) :rule bind) (anchor :step t103 :args ((:= (?v0 Int_real_real_real_prod_fun_prod$) veriT_vr56))) (step t103.t1 (cl @p_324) :rule refl) (step t103.t2 (cl (= @p_8 (! (member$ veriT_vr56 one_chain_typeII$) :named @p_329))) :rule cong :premises (t103.t1)) (step t103.t3 (cl @p_324) :rule refl) (step t103.t4 (cl @p_330) :rule cong :premises (t103.t3)) (step t103.t5 (cl (= @p_331 (! (=> @p_329 @p_325) :named @p_332))) :rule cong :premises (t103.t2 t103.t4)) (step t103 (cl (= @p_328 (! (forall ((veriT_vr56 Int_real_real_real_prod_fun_prod$)) @p_332) :named @p_335))) :rule bind) (step t104 (cl (= @p_333 (! (and @p_334 @p_335) :named @p_338))) :rule cong :premises (t102 t103)) (step t105 (cl (! (= @p_336 (! (and @p_337 @p_338) :named @p_340)) :named @p_339)) :rule cong :premises (t104)) (step t106 (cl (not @p_339) (not @p_336) @p_340) :rule equiv_pos2) (step t107 (cl @p_340) :rule th_resolution :premises (axiom17 t105 t106)) (step t108 (cl (! (= @p_340 (! (and @p_337 @p_334 @p_335) :named @p_342)) :named @p_341)) :rule ac_simp) (step t109 (cl (not @p_341) (not @p_340) @p_342) :rule equiv_pos2) (step t110 (cl @p_342) :rule th_resolution :premises (t107 t108 t109)) (anchor :step t111 :args ((:= (veriT_vr56 Int_real_real_real_prod_fun_prod$) veriT_vr57))) (step t111.t1 (cl (! (= veriT_vr56 veriT_vr57) :named @p_344)) :rule refl) (step t111.t2 (cl (= @p_329 (! (member$ veriT_vr57 one_chain_typeII$) :named @p_343))) :rule cong :premises (t111.t1)) (step t111.t3 (cl @p_344) :rule refl) (step t111.t4 (cl (= @p_325 (! (case_prod$ uu$ veriT_vr57) :named @p_345))) :rule cong :premises (t111.t3)) (step t111.t5 (cl (= @p_332 (! (=> @p_343 @p_345) :named @p_346))) :rule cong :premises (t111.t2 t111.t4)) (step t111 (cl (= @p_335 (! (forall ((veriT_vr57 Int_real_real_real_prod_fun_prod$)) @p_346) :named @p_347))) :rule bind) (step t112 (cl (! (= @p_342 (! (and @p_337 @p_334 @p_347) :named @p_349)) :named @p_348)) :rule cong :premises (t111)) (step t113 (cl (not @p_348) (not @p_342) @p_349) :rule equiv_pos2) (step t114 (cl @p_349) :rule th_resolution :premises (t110 t112 t113)) (anchor :step t115 :args ((:= (veriT_vr56 Int_real_real_real_prod_fun_prod$) veriT_vr58))) (step t115.t1 (cl (! (= veriT_vr56 veriT_vr58) :named @p_351)) :rule refl) (step t115.t2 (cl (= @p_323 (! (member$ veriT_vr58 one_chain_typeI$) :named @p_350))) :rule cong :premises (t115.t1)) (step t115.t3 (cl @p_351) :rule refl) (step t115.t4 (cl (= @p_325 (! (case_prod$ uu$ veriT_vr58) :named @p_352))) :rule cong :premises (t115.t3)) (step t115.t5 (cl (= @p_327 (! (=> @p_350 @p_352) :named @p_353))) :rule cong :premises (t115.t2 t115.t4)) (step t115 (cl (= @p_334 (! (forall ((veriT_vr58 Int_real_real_real_prod_fun_prod$)) @p_353) :named @p_358))) :rule bind) (anchor :step t116 :args ((:= (veriT_vr57 Int_real_real_real_prod_fun_prod$) veriT_vr59))) (step t116.t1 (cl (! (= veriT_vr57 veriT_vr59) :named @p_355)) :rule refl) (step t116.t2 (cl (= @p_343 (! (member$ veriT_vr59 one_chain_typeII$) :named @p_354))) :rule cong :premises (t116.t1)) (step t116.t3 (cl @p_355) :rule refl) (step t116.t4 (cl (= @p_345 (! (case_prod$ uu$ veriT_vr59) :named @p_356))) :rule cong :premises (t116.t3)) (step t116.t5 (cl (= @p_346 (! (=> @p_354 @p_356) :named @p_357))) :rule cong :premises (t116.t2 t116.t4)) (step t116 (cl (= @p_347 (! (forall ((veriT_vr59 Int_real_real_real_prod_fun_prod$)) @p_357) :named @p_359))) :rule bind) (step t117 (cl (! (= @p_349 (! (and @p_337 @p_358 @p_359) :named @p_361)) :named @p_360)) :rule cong :premises (t115 t116)) (step t118 (cl (not @p_360) (not @p_349) @p_361) :rule equiv_pos2) (step t119 (cl @p_361) :rule th_resolution :premises (t114 t117 t118)) (anchor :step t120 :args ((:= (?v0 Int_real_real_real_prod_fun_prod$) veriT_vr60))) (step t120.t1 (cl (! (= ?v0 veriT_vr60) :named @p_364)) :rule refl) (step t120.t2 (cl (= @p_8 (! (member$ veriT_vr60 one_chain_typeII$) :named @p_363))) :rule cong :premises (t120.t1)) (step t120.t3 (cl @p_364) :rule refl) (step t120.t4 (cl (! (= @p_10 (! (case_prod$ uua$ veriT_vr60) :named @p_365)) :named @p_370)) :rule cong :premises (t120.t3)) (step t120.t5 (cl (= @p_366 (! (=> @p_363 @p_365) :named @p_367))) :rule cong :premises (t120.t2 t120.t4)) (step t120 (cl (= @p_362 (! (forall ((veriT_vr60 Int_real_real_real_prod_fun_prod$)) @p_367) :named @p_374))) :rule bind) (anchor :step t121 :args ((:= (?v0 Int_real_real_real_prod_fun_prod$) veriT_vr60))) (step t121.t1 (cl @p_364) :rule refl) (step t121.t2 (cl (= @p_9 (! (member$ veriT_vr60 one_chain_typeI$) :named @p_369))) :rule cong :premises (t121.t1)) (step t121.t3 (cl @p_364) :rule refl) (step t121.t4 (cl @p_370) :rule cong :premises (t121.t3)) (step t121.t5 (cl (= @p_371 (! (=> @p_369 @p_365) :named @p_372))) :rule cong :premises (t121.t2 t121.t4)) (step t121 (cl (= @p_368 (! (forall ((veriT_vr60 Int_real_real_real_prod_fun_prod$)) @p_372) :named @p_375))) :rule bind) (step t122 (cl (= @p_373 (! (and @p_374 @p_375) :named @p_378))) :rule cong :premises (t120 t121)) (step t123 (cl (! (= @p_376 (! (and @p_377 @p_378) :named @p_380)) :named @p_379)) :rule cong :premises (t122)) (step t124 (cl (not @p_379) (not @p_376) @p_380) :rule equiv_pos2) (step t125 (cl @p_380) :rule th_resolution :premises (axiom18 t123 t124)) (step t126 (cl (! (= @p_380 (! (and @p_377 @p_374 @p_375) :named @p_382)) :named @p_381)) :rule ac_simp) (step t127 (cl (not @p_381) (not @p_380) @p_382) :rule equiv_pos2) (step t128 (cl @p_382) :rule th_resolution :premises (t125 t126 t127)) (anchor :step t129 :args ((:= (veriT_vr60 Int_real_real_real_prod_fun_prod$) veriT_vr61))) (step t129.t1 (cl (! (= veriT_vr60 veriT_vr61) :named @p_384)) :rule refl) (step t129.t2 (cl (= @p_369 (! (member$ veriT_vr61 one_chain_typeI$) :named @p_383))) :rule cong :premises (t129.t1)) (step t129.t3 (cl @p_384) :rule refl) (step t129.t4 (cl (= @p_365 (! (case_prod$ uua$ veriT_vr61) :named @p_385))) :rule cong :premises (t129.t3)) (step t129.t5 (cl (= @p_372 (! (=> @p_383 @p_385) :named @p_386))) :rule cong :premises (t129.t2 t129.t4)) (step t129 (cl (= @p_375 (! (forall ((veriT_vr61 Int_real_real_real_prod_fun_prod$)) @p_386) :named @p_387))) :rule bind) (step t130 (cl (! (= @p_382 (! (and @p_377 @p_374 @p_387) :named @p_389)) :named @p_388)) :rule cong :premises (t129)) (step t131 (cl (not @p_388) (not @p_382) @p_389) :rule equiv_pos2) (step t132 (cl @p_389) :rule th_resolution :premises (t128 t130 t131)) (anchor :step t133 :args ((:= (veriT_vr60 Int_real_real_real_prod_fun_prod$) veriT_vr62))) (step t133.t1 (cl (! (= veriT_vr60 veriT_vr62) :named @p_391)) :rule refl) (step t133.t2 (cl (= @p_363 (! (member$ veriT_vr62 one_chain_typeII$) :named @p_390))) :rule cong :premises (t133.t1)) (step t133.t3 (cl @p_391) :rule refl) (step t133.t4 (cl (= @p_365 (! (case_prod$ uua$ veriT_vr62) :named @p_392))) :rule cong :premises (t133.t3)) (step t133.t5 (cl (= @p_367 (! (=> @p_390 @p_392) :named @p_393))) :rule cong :premises (t133.t2 t133.t4)) (step t133 (cl (= @p_374 (! (forall ((veriT_vr62 Int_real_real_real_prod_fun_prod$)) @p_393) :named @p_398))) :rule bind) (anchor :step t134 :args ((:= (veriT_vr61 Int_real_real_real_prod_fun_prod$) veriT_vr63))) (step t134.t1 (cl (! (= veriT_vr61 veriT_vr63) :named @p_395)) :rule refl) (step t134.t2 (cl (= @p_383 (! (member$ veriT_vr63 one_chain_typeI$) :named @p_394))) :rule cong :premises (t134.t1)) (step t134.t3 (cl @p_395) :rule refl) (step t134.t4 (cl (= @p_385 (! (case_prod$ uua$ veriT_vr63) :named @p_396))) :rule cong :premises (t134.t3)) (step t134.t5 (cl (= @p_386 (! (=> @p_394 @p_396) :named @p_397))) :rule cong :premises (t134.t2 t134.t4)) (step t134 (cl (= @p_387 (! (forall ((veriT_vr63 Int_real_real_real_prod_fun_prod$)) @p_397) :named @p_399))) :rule bind) (step t135 (cl (! (= @p_389 (! (and @p_377 @p_398 @p_399) :named @p_401)) :named @p_400)) :rule cong :premises (t133 t134)) (step t136 (cl (not @p_400) (not @p_389) @p_401) :rule equiv_pos2) (step t137 (cl @p_401) :rule th_resolution :premises (t132 t135 t136)) (step t138 (cl @p_358) :rule and :premises (t119)) (step t139 (cl @p_399) :rule and :premises (t137)) (step t140 (cl (or (! (not @p_399) :named @p_422) (! (=> @p_402 (! (case_prod$ uua$ @p_403) :named @p_421)) :named @p_420))) :rule forall_inst :args ((:= veriT_vr63 @p_403))) (step t141 (cl (or (! (not @p_358) :named @p_427) (! (=> @p_402 (! (case_prod$ uu$ @p_403) :named @p_426)) :named @p_424))) :rule forall_inst :args ((:= veriT_vr58 @p_403))) (step t142 (cl (or (! (not @p_321) :named @p_406) (! (=> (! (and (! (finite$ @p_5) :named @p_415) (! (fun_app$a @p_12 g$) :named @p_409) (! (fun_app$a @p_24 g$) :named @p_408) (! (= @p_5 (! (sup$ @p_7 @p_5) :named @p_466)) :named @p_430) (! (= bot$ (inf$ @p_7 @p_5)) :named @p_431)) :named @p_429) (! (= @p_404 (! (+ @p_405 @p_404) :named @p_528)) :named @p_433)) :named @p_432))) :rule forall_inst :args ((:= veriT_vr51 @p_5) (:= veriT_vr52 f$) (:= veriT_vr53 @p_7) (:= veriT_vr54 g$) (:= veriT_vr55 @p_5))) (step t143 (cl (or @p_406 (! (=> (! (and (! (finite$ @p_407) :named @p_412) @p_408 @p_409 (! (= @p_407 (sup$ @p_5 @p_7)) :named @p_411) (! (= bot$ (inf$ @p_5 @p_7)) :named @p_437)) :named @p_434) @p_410) :named @p_438))) :rule forall_inst :args ((:= veriT_vr51 @p_407) (:= veriT_vr52 f$) (:= veriT_vr53 @p_5) (:= veriT_vr54 g$) (:= veriT_vr55 @p_7))) (step t144 (cl (or (! (not @p_249) :named @p_441) @p_411)) :rule forall_inst :args ((:= veriT_vr44 i$) (:= veriT_vr45 @p_7))) (step t145 (cl (or (! (not @p_230) :named @p_442) (! (= @p_407 (! (insert$ j$ @p_5) :named @p_467)) :named @p_443))) :rule forall_inst :args ((:= veriT_vr39 j$) (:= veriT_vr40 i$) (:= veriT_vr41 bot$))) (step t146 (cl (or (! (not @p_194) :named @p_447) (! (=> (! (member$a i$ @p_7) :named @p_445) (! (= @p_7 @p_407) :named @p_446)) :named @p_444))) :rule forall_inst :args ((:= veriT_vr32 i$) (:= veriT_vr33 @p_7))) (step t147 (cl (or (! (not @p_174) :named @p_413) (! (=> (! (finite$ @p_7) :named @p_416) @p_412) :named @p_448))) :rule forall_inst :args ((:= veriT_vr28 @p_7) (:= veriT_vr29 i$))) (step t148 (cl (or @p_413 (! (=> @p_414 @p_415) :named @p_449))) :rule forall_inst :args ((:= veriT_vr28 bot$) (:= veriT_vr29 i$))) (step t149 (cl (or @p_413 (! (=> @p_414 @p_416) :named @p_451))) :rule forall_inst :args ((:= veriT_vr28 bot$) (:= veriT_vr29 j$))) (step t150 (cl (or (! (not @p_119) :named @p_418) (! (= 0.0 (! (snd$a @p_417) :named @p_495)) :named @p_454))) :rule forall_inst :args ((:= veriT_vr18 1.0) (:= veriT_vr19 0.0))) (step t151 (cl (or @p_418 (! (= 1.0 (! (snd$a @p_419) :named @p_496)) :named @p_455))) :rule forall_inst :args ((:= veriT_vr18 0.0) (:= veriT_vr19 1.0))) (step t152 (cl (or (! (not @p_100) :named @p_456) (! (= g$ (! (snd$ @p_403) :named @p_471)) :named @p_457))) :rule forall_inst :args ((:= veriT_vr14 k$) (:= veriT_vr15 g$))) (step t153 (cl (or (! (not @p_84) :named @p_461) (! (=> (! (= @p_7 @p_5) :named @p_459) (! (= j$ i$) :named @p_460)) :named @p_458))) :rule forall_inst :args ((:= veriT_vr10 i$) (:= veriT_vr11 j$))) (step t154 (cl (! (not @p_420) :named @p_423) (! (not @p_402) :named @p_425) @p_421) :rule implies_pos) (step t155 (cl @p_422 @p_420) :rule or :premises (t140)) (step t156 (cl @p_423 @p_421) :rule resolution :premises (t154 axiom6)) (step t157 (cl @p_420) :rule resolution :premises (t155 t139)) (step t158 (cl @p_421) :rule resolution :premises (t156 t157)) (step t159 (cl (! (not @p_424) :named @p_428) @p_425 @p_426) :rule implies_pos) (step t160 (cl @p_427 @p_424) :rule or :premises (t141)) (step t161 (cl @p_428 @p_426) :rule resolution :premises (t159 axiom6)) (step t162 (cl @p_424) :rule resolution :premises (t160 t138)) (step t163 (cl @p_426) :rule resolution :premises (t161 t162)) (step t164 (cl @p_429 (not @p_415) (! (not @p_409) :named @p_436) (! (not @p_408) :named @p_435) (not @p_430) (not @p_431)) :rule and_neg) (step t165 (cl (not @p_432) (not @p_429) @p_433) :rule implies_pos) (step t166 (cl @p_406 @p_432) :rule or :premises (t142)) (step t167 (cl @p_432) :rule resolution :premises (t166 t101)) (step t168 (cl @p_434 (not @p_412) @p_435 @p_436 (not @p_411) (! (not @p_437) :named @p_515)) :rule and_neg) (step t169 (cl (! (not @p_438) :named @p_439) (! (not @p_434) :named @p_440) @p_410) :rule implies_pos) (step t170 (cl @p_406 @p_438) :rule or :premises (t143)) (step t171 (cl @p_439 @p_440) :rule resolution :premises (t169 axiom19)) (step t172 (cl @p_438) :rule resolution :premises (t170 t101)) (step t173 (cl @p_440) :rule resolution :premises (t171 t172)) (step t174 (cl @p_441 @p_411) :rule or :premises (t144)) (step t175 (cl @p_411) :rule resolution :premises (t174 t92)) (step t176 (cl @p_442 @p_443) :rule or :premises (t145)) (step t177 (cl @p_443) :rule resolution :premises (t176 t86)) (step t178 (cl (not @p_444) (! (not @p_445) :named @p_470) @p_446) :rule implies_pos) (step t179 (cl @p_447 @p_444) :rule or :premises (t146)) (step t180 (cl @p_444) :rule resolution :premises (t179 t74)) (step t181 (cl (not @p_448) (not @p_416) @p_412) :rule implies_pos) (step t182 (cl @p_413 @p_448) :rule or :premises (t147)) (step t183 (cl @p_448) :rule resolution :premises (t182 t68)) (step t184 (cl (! (not @p_449) :named @p_450) (! (not @p_414) :named @p_452) @p_415) :rule implies_pos) (step t185 (cl @p_413 @p_449) :rule or :premises (t148)) (step t186 (cl @p_450 @p_415) :rule resolution :premises (t184 axiom8)) (step t187 (cl @p_449) :rule resolution :premises (t185 t68)) (step t188 (cl @p_415) :rule resolution :premises (t186 t187)) (step t189 (cl (! (not @p_451) :named @p_453) @p_452 @p_416) :rule implies_pos) (step t190 (cl @p_413 @p_451) :rule or :premises (t149)) (step t191 (cl @p_453 @p_416) :rule resolution :premises (t189 axiom8)) (step t192 (cl @p_451) :rule resolution :premises (t190 t68)) (step t193 (cl @p_416) :rule resolution :premises (t191 t192)) (step t194 (cl @p_412) :rule resolution :premises (t181 t193 t183)) (step t195 (cl @p_418 @p_454) :rule or :premises (t150)) (step t196 (cl @p_454) :rule resolution :premises (t195 t56)) (step t197 (cl @p_418 @p_455) :rule or :premises (t151)) (step t198 (cl @p_455) :rule resolution :premises (t197 t56)) (step t199 (cl @p_456 @p_457) :rule or :premises (t152)) (step t200 (cl @p_457) :rule resolution :premises (t199 t50)) (step t201 (cl (not @p_458) (! (not @p_459) :named @p_507) @p_460) :rule implies_pos) (step t202 (cl @p_461 @p_458) :rule or :premises (t153)) (step t203 (cl @p_458) :rule resolution :premises (t202 t44)) (step t204 (cl (or @p_410 (! (not (! (<= @p_462 @p_463) :named @p_534)) :named @p_464) (! (not (! (<= @p_463 @p_462) :named @p_535)) :named @p_465))) :rule la_disequality) (step t205 (cl @p_410 @p_464 @p_465) :rule or :premises (t204)) (step t206 (cl @p_464 @p_465) :rule resolution :premises (t205 axiom19)) (step t207 (cl (or @p_441 (! (= @p_466 @p_467) :named @p_474))) :rule forall_inst :args ((:= veriT_vr44 j$) (:= veriT_vr45 @p_5))) (step t208 (cl (or @p_447 (! (=> (! (member$a j$ @p_5) :named @p_468) (! (= @p_5 @p_467) :named @p_477)) :named @p_475))) :rule forall_inst :args ((:= veriT_vr32 j$) (:= veriT_vr33 @p_5))) (step t209 (cl (or (! (not @p_156) :named @p_469) (! (= @p_437 (! (and (! (not @p_468) :named @p_476) (! (= bot$ (inf$ @p_5 bot$)) :named @p_479)) :named @p_478)) :named @p_482))) :rule forall_inst :args ((:= veriT_vr23 @p_5) (:= veriT_vr24 j$) (:= veriT_vr25 bot$))) (step t210 (cl (or @p_469 (! (= @p_431 (! (and @p_470 (! (= bot$ (inf$ @p_7 bot$)) :named @p_485)) :named @p_483)) :named @p_487))) :rule forall_inst :args ((:= veriT_vr23 @p_7) (:= veriT_vr24 i$) (:= veriT_vr25 bot$))) (step t211 (cl (or (! (not @p_61) :named @p_472) (! (= @p_426 (! (fun_app$a (! (fun_app$ uu$ (! (fst$ @p_403) :named @p_473)) :named @p_508) @p_471) :named @p_489)) :named @p_488))) :rule forall_inst :args ((:= veriT_vr6 uu$) (:= veriT_vr7 @p_403))) (step t212 (cl (or @p_472 (! (= @p_421 (! (fun_app$a (! (fun_app$ uua$ @p_473) :named @p_509) @p_471) :named @p_492)) :named @p_491))) :rule forall_inst :args ((:= veriT_vr6 uua$) (:= veriT_vr7 @p_403))) (step t213 (cl @p_441 @p_474) :rule or :premises (t207)) (step t214 (cl @p_474) :rule resolution :premises (t213 t92)) (step t215 (cl (not @p_475) @p_476 @p_477) :rule implies_pos) (step t216 (cl @p_447 @p_475) :rule or :premises (t208)) (step t217 (cl @p_475) :rule resolution :premises (t216 t74)) (step t218 (cl (not (! (not @p_470) :named @p_484)) @p_445) :rule not_not) (step t219 (cl @p_478 (! (not @p_476) :named @p_480) (! (not @p_479) :named @p_481)) :rule and_neg) (step t220 (cl (not @p_480) @p_468) :rule not_not) (step t221 (cl @p_478 @p_468 @p_481) :rule th_resolution :premises (t220 t219)) (step t222 (cl (not @p_482) @p_437 (! (not @p_478) :named @p_516)) :rule equiv_pos1) (step t223 (cl @p_469 @p_482) :rule or :premises (t209)) (step t224 (cl @p_482) :rule resolution :premises (t223 t62)) (step t225 (cl @p_483 @p_484 (! (not @p_485) :named @p_486)) :rule and_neg) (step t226 (cl @p_483 @p_445 @p_486) :rule th_resolution :premises (t218 t225)) (step t227 (cl (not @p_487) @p_431 (not @p_483)) :rule equiv_pos1) (step t228 (cl @p_469 @p_487) :rule or :premises (t210)) (step t229 (cl @p_487) :rule resolution :premises (t228 t62)) (step t230 (cl (! (not @p_488) :named @p_490) (not @p_426) @p_489) :rule equiv_pos2) (step t231 (cl @p_472 @p_488) :rule or :premises (t211)) (step t232 (cl @p_490 @p_489) :rule resolution :premises (t230 t163)) (step t233 (cl @p_488) :rule resolution :premises (t231 t38)) (step t234 (cl @p_489) :rule resolution :premises (t232 t233)) (step t235 (cl (! (not @p_491) :named @p_493) (not @p_421) @p_492) :rule equiv_pos2) (step t236 (cl @p_472 @p_491) :rule or :premises (t212)) (step t237 (cl @p_493 @p_492) :rule resolution :premises (t235 t158)) (step t238 (cl @p_491) :rule resolution :premises (t236 t38)) (step t239 (cl @p_492) :rule resolution :premises (t237 t238)) (step t240 (cl (or (! (not @p_205) :named @p_494) @p_479)) :rule forall_inst :args ((:= veriT_vr35 @p_5))) (step t241 (cl (or @p_494 @p_485)) :rule forall_inst :args ((:= veriT_vr35 @p_7))) (step t242 (cl @p_494 @p_479) :rule or :premises (t240)) (step t243 (cl @p_479) :rule resolution :premises (t242 t80)) (step t244 (cl @p_494 @p_485) :rule or :premises (t241)) (step t245 (cl @p_485) :rule resolution :premises (t244 t80)) (step t246 (cl (! (= @p_7 @p_7) :named @p_520)) :rule eq_reflexive) (step t247 (cl (not (! (= 1.0 0.0) :named @p_497))) :rule la_generic :args ((- 1))) (step t248 (cl (! (not @p_455) :named @p_505) (not (! (= @p_495 @p_496) :named @p_498)) (! (not @p_454) :named @p_506) @p_497) :rule eq_transitive) (step t249 (cl (not (! (= @p_417 @p_419) :named @p_501)) @p_498) :rule eq_congruent) (step t250 (cl (! (not @p_499) :named @p_502) (! (not @p_460) :named @p_503) (! (not @p_500) :named @p_504) @p_501) :rule eq_transitive) (step t251 (cl @p_498 @p_502 @p_503 @p_504) :rule th_resolution :premises (t249 t250)) (step t252 (cl @p_505 @p_506 @p_497 @p_502 @p_503 @p_504) :rule th_resolution :premises (t248 t251)) (step t253 (cl @p_505 @p_506 @p_502 @p_503 @p_504) :rule th_resolution :premises (t247 t252)) (step t254 (cl @p_503) :rule resolution :premises (t253 axiom10 axiom12 t196 t198)) (step t255 (cl @p_507) :rule resolution :premises (t201 t254 t203)) (step t256 (cl (! (= f$ f$) :named @p_523)) :rule eq_reflexive) (step t257 (cl (! (= g$ g$) :named @p_524)) :rule eq_reflexive) (step t258 (cl (! (= @p_405 @p_405) :named @p_527)) :rule eq_reflexive) (step t259 (cl (or (! (not @p_34) :named @p_510) (! (= @p_24 @p_508) :named @p_511))) :rule forall_inst :args ((:= veriT_vr3 @p_473))) (step t260 (cl (or (! (not @p_22) :named @p_512) (! (= @p_12 @p_509) :named @p_513))) :rule forall_inst :args ((:= veriT_vr1 @p_473))) (step t261 (cl @p_510 @p_511) :rule or :premises (t259)) (step t262 (cl @p_511) :rule resolution :premises (t261 t32)) (step t263 (cl @p_512 @p_513) :rule or :premises (t260)) (step t264 (cl @p_513) :rule resolution :premises (t263 t26)) (step t265 (cl (not @p_511) (! (not @p_457) :named @p_514) @p_408 (not @p_489)) :rule eq_congruent_pred) (step t266 (cl @p_408) :rule resolution :premises (t265 t200 t234 t262)) (step t267 (cl (not @p_513) @p_514 @p_409 (not @p_492)) :rule eq_congruent_pred) (step t268 (cl @p_409) :rule resolution :premises (t267 t200 t239 t264)) (step t269 (cl @p_515) :rule resolution :premises (t168 t268 t194 t173 t175 t266)) (step t270 (cl @p_516) :rule resolution :premises (t222 t269 t224)) (step t271 (cl @p_468) :rule resolution :premises (t221 t270 t243)) (step t272 (cl @p_477) :rule resolution :premises (t215 t271 t217)) (step t273 (cl (! (not @p_477) :named @p_517) (not @p_474) @p_430) :rule eq_transitive) (step t274 (cl @p_430) :rule resolution :premises (t273 t214 t272)) (step t275 (cl @p_517 (! (not @p_443) :named @p_518) (! (= @p_5 @p_407) :named @p_519)) :rule eq_transitive) (step t276 (cl @p_518 @p_517 @p_519) :rule eq_transitive) (step t277 (cl (not @p_520) (! (not @p_446) :named @p_521) (! (not @p_519) :named @p_522) @p_459) :rule eq_transitive) (step t278 (cl @p_521 @p_522 @p_459) :rule th_resolution :premises (t277 t246)) (step t279 (cl @p_521 @p_459 @p_518 @p_517) :rule th_resolution :premises (t278 t276)) (step t280 (cl @p_521) :rule resolution :premises (t279 t177 t255 t272)) (step t281 (cl (not @p_523) @p_522 (! (not @p_524) :named @p_525) (! (= @p_462 @p_404) :named @p_526)) :rule eq_congruent) (step t282 (cl @p_522 @p_525 @p_526) :rule th_resolution :premises (t281 t256)) (step t283 (cl @p_522 @p_526) :rule th_resolution :premises (t282 t257)) (step t284 (cl @p_526 @p_517 @p_518) :rule th_resolution :premises (t283 t275)) (step t285 (cl @p_470) :rule resolution :premises (t178 t280 t180)) (step t286 (cl @p_483) :rule resolution :premises (t226 t285 t245)) (step t287 (cl @p_431) :rule resolution :premises (t227 t286 t229)) (step t288 (cl @p_429) :rule resolution :premises (t164 t287 t188 t266 t268 t274)) (step t289 (cl @p_433) :rule resolution :premises (t165 t288 t167)) (step t290 (cl (not @p_527) (! (not @p_526) :named @p_529) (! (= (! (+ @p_405 @p_462) :named @p_531) @p_528) :named @p_530)) :rule eq_congruent) (step t291 (cl @p_529 @p_530) :rule th_resolution :premises (t290 t258)) (step t292 (cl @p_530 @p_517 @p_518) :rule th_resolution :premises (t291 t284)) (step t293 (cl (! (not @p_433) :named @p_532) (not @p_530) (! (= @p_404 @p_531) :named @p_533)) :rule eq_transitive) (step t294 (cl @p_532 @p_533 @p_517 @p_518) :rule th_resolution :premises (t293 t292)) (step t295 (cl @p_534 @p_529 (! (not @p_533) :named @p_536)) :rule la_generic :args (1.0 (- 2) (- 1))) (step t296 (cl @p_534 @p_517 @p_518 @p_532) :rule th_resolution :premises (t295 t284 t294)) (step t297 (cl @p_534) :rule resolution :premises (t296 t289 t177 t272)) (step t298 (cl @p_465) :rule resolution :premises (t206 t297)) (step t299 (cl @p_535 @p_529 @p_536) :rule la_generic :args (1.0 2 1)) (step t300 (cl @p_535 @p_517 @p_518 @p_532) :rule th_resolution :premises (t299 t284 t294)) (step t301 (cl) :rule resolution :premises (t300 t289 t177 t298 t272)) 5fc0f54f0190d5b6a4967f69daa36736ef3d3325 6 0 unsat (assume axiom0 (! (< 0.0 (+ x$ (! (* 2.0 y$) :named @p_1))) :named @p_2)) (assume axiom1 (! (< 0.0 (- x$ @p_1)) :named @p_3)) (assume axiom2 (! (< x$ 0.0) :named @p_4)) (step t4 (cl (not @p_2) (not @p_3) (not @p_4)) :rule la_generic :args (1.0 1.0 2.0)) (step t5 (cl) :rule resolution :premises (t4 axiom0 axiom1 axiom2)) 809cc2252f6f1da6c3a4347a531733952ab0b99f 467 0 unsat (assume axiom0 (! (forall ((?v0 Real)) (! (= (! (fun_app$ uuc$ ?v0) :named @p_9) (! (pair$ (! (times$ (! (- ?v0 (! (divide$ 1.0 2.0) :named @p_7)) :named @p_12) d$) :named @p_1) (! (diamond_y$ @p_1) :named @p_16)) :named @p_18)) :named @p_20)) :named @p_6)) (assume axiom3 (! (forall ((?v0 Real)) (! (= (! (fun_app$ uub$ ?v0) :named @p_37) (! (pair$ (! (- (! (divide$ d$ 2.0) :named @p_3)) :named @p_2) (! (times$ (! (- (! (* 2.0 ?v0) :named @p_40) 1.0) :named @p_42) (! (diamond_y$ @p_2) :named @p_36)) :named @p_44)) :named @p_46)) :named @p_48)) :named @p_35)) (assume axiom4 (! (< 0.0 d$) :named @p_257)) (assume axiom5 (! (forall ((?v0 Real)) (! (= (! (diamond_y$ ?v0) :named @p_62) (! (- @p_3 (! (ite (! (< ?v0 0.0) :named @p_65) (! (- ?v0) :named @p_4) ?v0) :named @p_68)) :named @p_70)) :named @p_72)) :named @p_61)) (assume axiom7 (! (forall ((?v0 Real) (?v1 Real) (?v2 Real)) (! (= (! (< (! (divide$ ?v0 ?v1) :named @p_5) (! (divide$ ?v2 ?v1) :named @p_88)) :named @p_90) (! (and (! (=> (! (< 0.0 ?v1) :named @p_92) (! (< ?v0 ?v2) :named @p_96)) :named @p_98) (! (and (! (=> (! (< ?v1 0.0) :named @p_100) (! (< ?v2 ?v0) :named @p_102)) :named @p_104) (! (not (! (= 0.0 ?v1) :named @p_106)) :named @p_108)) :named @p_110)) :named @p_112)) :named @p_114)) :named @p_85)) (assume axiom8 (! (forall ((?v0 Real) (?v1 Real)) (! (= (! (divide$ @p_4 ?v1) :named @p_142) (! (- @p_5) :named @p_147)) :named @p_149)) :named @p_140)) (assume axiom9 (! (forall ((?v0 Real) (?v1 Real)) (! (= (! (times$ @p_4 ?v1) :named @p_164) (! (- (! (times$ ?v0 ?v1) :named @p_168)) :named @p_170)) :named @p_172)) :named @p_162)) (assume axiom10 (! (forall ((?v0 Real) (?v1 Real) (?v2 Real) (?v3 Real)) (! (= (! (= (! (pair$ ?v0 ?v1) :named @p_186) (! (pair$ ?v2 ?v3) :named @p_188)) :named @p_190) (! (and (! (= ?v0 ?v2) :named @p_194) (! (= ?v1 ?v3) :named @p_198)) :named @p_200)) :named @p_202)) :named @p_185)) (assume axiom11 (! (not (! (=> (! (and (! (not (= uua$ uu$)) :named @p_226) (! (= uuc$ uub$) :named @p_227)) :named @p_220) false) :named @p_224)) :named @p_219)) (anchor :step t10 :args ((:= (?v0 Real) veriT_vr0))) (step t10.t1 (cl (! (= ?v0 veriT_vr0) :named @p_11)) :rule refl) (step t10.t2 (cl (= @p_9 (! (fun_app$ uuc$ veriT_vr0) :named @p_10))) :rule cong :premises (t10.t1)) (step t10.t3 (cl @p_11) :rule refl) (step t10.t4 (cl (! (= @p_12 (! (- veriT_vr0 @p_7) :named @p_13)) :named @p_14)) :rule cong :premises (t10.t3)) (step t10.t5 (cl (! (= @p_1 (! (times$ @p_13 d$) :named @p_8)) :named @p_15)) :rule cong :premises (t10.t4)) (step t10.t6 (cl @p_11) :rule refl) (step t10.t7 (cl @p_14) :rule cong :premises (t10.t6)) (step t10.t8 (cl @p_15) :rule cong :premises (t10.t7)) (step t10.t9 (cl (= @p_16 (! (diamond_y$ @p_8) :named @p_17))) :rule cong :premises (t10.t8)) (step t10.t10 (cl (= @p_18 (! (pair$ @p_8 @p_17) :named @p_19))) :rule cong :premises (t10.t5 t10.t9)) (step t10.t11 (cl (= @p_20 (! (= @p_10 @p_19) :named @p_21))) :rule cong :premises (t10.t2 t10.t10)) (step t10 (cl (! (= @p_6 (! (forall ((veriT_vr0 Real)) @p_21) :named @p_23)) :named @p_22)) :rule bind) (step t11 (cl (not @p_22) (not @p_6) @p_23) :rule equiv_pos2) (step t12 (cl @p_23) :rule th_resolution :premises (axiom0 t10 t11)) (anchor :step t13 :args ((:= (veriT_vr0 Real) veriT_vr1))) (step t13.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_26)) :rule refl) (step t13.t2 (cl (= @p_10 (! (fun_app$ uuc$ veriT_vr1) :named @p_25))) :rule cong :premises (t13.t1)) (step t13.t3 (cl @p_26) :rule refl) (step t13.t4 (cl (! (= @p_13 (! (- veriT_vr1 @p_7) :named @p_27)) :named @p_28)) :rule cong :premises (t13.t3)) (step t13.t5 (cl (! (= @p_8 (! (times$ @p_27 d$) :named @p_24)) :named @p_29)) :rule cong :premises (t13.t4)) (step t13.t6 (cl @p_26) :rule refl) (step t13.t7 (cl @p_28) :rule cong :premises (t13.t6)) (step t13.t8 (cl @p_29) :rule cong :premises (t13.t7)) (step t13.t9 (cl (= @p_17 (! (diamond_y$ @p_24) :named @p_30))) :rule cong :premises (t13.t8)) (step t13.t10 (cl (= @p_19 (! (pair$ @p_24 @p_30) :named @p_31))) :rule cong :premises (t13.t5 t13.t9)) (step t13.t11 (cl (= @p_21 (! (= @p_25 @p_31) :named @p_32))) :rule cong :premises (t13.t2 t13.t10)) (step t13 (cl (! (= @p_23 (! (forall ((veriT_vr1 Real)) @p_32) :named @p_34)) :named @p_33)) :rule bind) (step t14 (cl (not @p_33) (not @p_23) @p_34) :rule equiv_pos2) (step t15 (cl @p_34) :rule th_resolution :premises (t12 t13 t14)) (anchor :step t16 :args ((:= (?v0 Real) veriT_vr6))) (step t16.t1 (cl (! (= ?v0 veriT_vr6) :named @p_39)) :rule refl) (step t16.t2 (cl (= @p_37 (! (fun_app$ uub$ veriT_vr6) :named @p_38))) :rule cong :premises (t16.t1)) (step t16.t3 (cl @p_39) :rule refl) (step t16.t4 (cl (= @p_40 (! (* 2.0 veriT_vr6) :named @p_41))) :rule cong :premises (t16.t3)) (step t16.t5 (cl (= @p_42 (! (- @p_41 1.0) :named @p_43))) :rule cong :premises (t16.t4)) (step t16.t6 (cl (= @p_44 (! (times$ @p_43 @p_36) :named @p_45))) :rule cong :premises (t16.t5)) (step t16.t7 (cl (= @p_46 (! (pair$ @p_2 @p_45) :named @p_47))) :rule cong :premises (t16.t6)) (step t16.t8 (cl (= @p_48 (! (= @p_38 @p_47) :named @p_49))) :rule cong :premises (t16.t2 t16.t7)) (step t16 (cl (! (= @p_35 (! (forall ((veriT_vr6 Real)) @p_49) :named @p_51)) :named @p_50)) :rule bind) (step t17 (cl (not @p_50) (not @p_35) @p_51) :rule equiv_pos2) (step t18 (cl @p_51) :rule th_resolution :premises (axiom3 t16 t17)) (anchor :step t19 :args ((:= (veriT_vr6 Real) veriT_vr7))) (step t19.t1 (cl (! (= veriT_vr6 veriT_vr7) :named @p_53)) :rule refl) (step t19.t2 (cl (= @p_38 (! (fun_app$ uub$ veriT_vr7) :named @p_52))) :rule cong :premises (t19.t1)) (step t19.t3 (cl @p_53) :rule refl) (step t19.t4 (cl (= @p_41 (! (* 2.0 veriT_vr7) :named @p_54))) :rule cong :premises (t19.t3)) (step t19.t5 (cl (= @p_43 (! (- @p_54 1.0) :named @p_55))) :rule cong :premises (t19.t4)) (step t19.t6 (cl (= @p_45 (! (times$ @p_55 @p_36) :named @p_56))) :rule cong :premises (t19.t5)) (step t19.t7 (cl (= @p_47 (! (pair$ @p_2 @p_56) :named @p_57))) :rule cong :premises (t19.t6)) (step t19.t8 (cl (= @p_49 (! (= @p_52 @p_57) :named @p_58))) :rule cong :premises (t19.t2 t19.t7)) (step t19 (cl (! (= @p_51 (! (forall ((veriT_vr7 Real)) @p_58) :named @p_60)) :named @p_59)) :rule bind) (step t20 (cl (not @p_59) (not @p_51) @p_60) :rule equiv_pos2) (step t21 (cl @p_60) :rule th_resolution :premises (t18 t19 t20)) (anchor :step t22 :args ((:= (?v0 Real) veriT_vr8))) (step t22.t1 (cl (! (= ?v0 veriT_vr8) :named @p_64)) :rule refl) (step t22.t2 (cl (= @p_62 (! (diamond_y$ veriT_vr8) :named @p_63))) :rule cong :premises (t22.t1)) (step t22.t3 (cl @p_64) :rule refl) (step t22.t4 (cl (= @p_65 (! (< veriT_vr8 0.0) :named @p_66))) :rule cong :premises (t22.t3)) (step t22.t5 (cl @p_64) :rule refl) (step t22.t6 (cl (= @p_4 (! (- veriT_vr8) :named @p_67))) :rule cong :premises (t22.t5)) (step t22.t7 (cl @p_64) :rule refl) (step t22.t8 (cl (= @p_68 (! (ite @p_66 @p_67 veriT_vr8) :named @p_69))) :rule cong :premises (t22.t4 t22.t6 t22.t7)) (step t22.t9 (cl (= @p_70 (! (- @p_3 @p_69) :named @p_71))) :rule cong :premises (t22.t8)) (step t22.t10 (cl (= @p_72 (! (= @p_63 @p_71) :named @p_73))) :rule cong :premises (t22.t2 t22.t9)) (step t22 (cl (! (= @p_61 (! (forall ((veriT_vr8 Real)) @p_73) :named @p_75)) :named @p_74)) :rule bind) (step t23 (cl (not @p_74) (not @p_61) @p_75) :rule equiv_pos2) (step t24 (cl @p_75) :rule th_resolution :premises (axiom5 t22 t23)) (anchor :step t25 :args ((:= (veriT_vr8 Real) veriT_vr9))) (step t25.t1 (cl (! (= veriT_vr8 veriT_vr9) :named @p_77)) :rule refl) (step t25.t2 (cl (= @p_63 (! (diamond_y$ veriT_vr9) :named @p_76))) :rule cong :premises (t25.t1)) (step t25.t3 (cl @p_77) :rule refl) (step t25.t4 (cl (= @p_66 (! (< veriT_vr9 0.0) :named @p_78))) :rule cong :premises (t25.t3)) (step t25.t5 (cl @p_77) :rule refl) (step t25.t6 (cl (= @p_67 (! (- veriT_vr9) :named @p_79))) :rule cong :premises (t25.t5)) (step t25.t7 (cl @p_77) :rule refl) (step t25.t8 (cl (= @p_69 (! (ite @p_78 @p_79 veriT_vr9) :named @p_80))) :rule cong :premises (t25.t4 t25.t6 t25.t7)) (step t25.t9 (cl (= @p_71 (! (- @p_3 @p_80) :named @p_81))) :rule cong :premises (t25.t8)) (step t25.t10 (cl (= @p_73 (! (= @p_76 @p_81) :named @p_82))) :rule cong :premises (t25.t2 t25.t9)) (step t25 (cl (! (= @p_75 (! (forall ((veriT_vr9 Real)) @p_82) :named @p_84)) :named @p_83)) :rule bind) (step t26 (cl (not @p_83) (not @p_75) @p_84) :rule equiv_pos2) (step t27 (cl @p_84) :rule th_resolution :premises (t24 t25 t26)) (anchor :step t28 :args ((:= (?v0 Real) veriT_vr10) (:= (?v1 Real) veriT_vr11) (:= (?v2 Real) veriT_vr12))) (step t28.t1 (cl (! (= ?v0 veriT_vr10) :named @p_94)) :rule refl) (step t28.t2 (cl (! (= ?v1 veriT_vr11) :named @p_87)) :rule refl) (step t28.t3 (cl (= @p_5 (! (divide$ veriT_vr10 veriT_vr11) :named @p_86))) :rule cong :premises (t28.t1 t28.t2)) (step t28.t4 (cl (! (= ?v2 veriT_vr12) :named @p_95)) :rule refl) (step t28.t5 (cl @p_87) :rule refl) (step t28.t6 (cl (= @p_88 (! (divide$ veriT_vr12 veriT_vr11) :named @p_89))) :rule cong :premises (t28.t4 t28.t5)) (step t28.t7 (cl (= @p_90 (! (< @p_86 @p_89) :named @p_91))) :rule cong :premises (t28.t3 t28.t6)) (step t28.t8 (cl @p_87) :rule refl) (step t28.t9 (cl (= @p_92 (! (< 0.0 veriT_vr11) :named @p_93))) :rule cong :premises (t28.t8)) (step t28.t10 (cl @p_94) :rule refl) (step t28.t11 (cl @p_95) :rule refl) (step t28.t12 (cl (= @p_96 (! (< veriT_vr10 veriT_vr12) :named @p_97))) :rule cong :premises (t28.t10 t28.t11)) (step t28.t13 (cl (= @p_98 (! (=> @p_93 @p_97) :named @p_99))) :rule cong :premises (t28.t9 t28.t12)) (step t28.t14 (cl @p_87) :rule refl) (step t28.t15 (cl (= @p_100 (! (< veriT_vr11 0.0) :named @p_101))) :rule cong :premises (t28.t14)) (step t28.t16 (cl @p_95) :rule refl) (step t28.t17 (cl @p_94) :rule refl) (step t28.t18 (cl (= @p_102 (! (< veriT_vr12 veriT_vr10) :named @p_103))) :rule cong :premises (t28.t16 t28.t17)) (step t28.t19 (cl (= @p_104 (! (=> @p_101 @p_103) :named @p_105))) :rule cong :premises (t28.t15 t28.t18)) (step t28.t20 (cl @p_87) :rule refl) (step t28.t21 (cl (= @p_106 (! (= 0.0 veriT_vr11) :named @p_107))) :rule cong :premises (t28.t20)) (step t28.t22 (cl (= @p_108 (! (not @p_107) :named @p_109))) :rule cong :premises (t28.t21)) (step t28.t23 (cl (= @p_110 (! (and @p_105 @p_109) :named @p_111))) :rule cong :premises (t28.t19 t28.t22)) (step t28.t24 (cl (= @p_112 (! (and @p_99 @p_111) :named @p_113))) :rule cong :premises (t28.t13 t28.t23)) (step t28.t25 (cl (= @p_114 (! (= @p_91 @p_113) :named @p_115))) :rule cong :premises (t28.t7 t28.t24)) (step t28 (cl (! (= @p_85 (! (forall ((veriT_vr10 Real) (veriT_vr11 Real) (veriT_vr12 Real)) @p_115) :named @p_117)) :named @p_116)) :rule bind) (step t29 (cl (not @p_116) (not @p_85) @p_117) :rule equiv_pos2) (step t30 (cl @p_117) :rule th_resolution :premises (axiom7 t28 t29)) (anchor :step t31 :args ((veriT_vr10 Real) (veriT_vr11 Real) (veriT_vr12 Real))) (step t31.t1 (cl (= @p_113 (! (and @p_99 @p_105 @p_109) :named @p_118))) :rule ac_simp) (step t31.t2 (cl (= @p_115 (! (= @p_91 @p_118) :named @p_119))) :rule cong :premises (t31.t1)) (step t31 (cl (! (= @p_117 (! (forall ((veriT_vr10 Real) (veriT_vr11 Real) (veriT_vr12 Real)) @p_119) :named @p_121)) :named @p_120)) :rule bind) (step t32 (cl (not @p_120) (not @p_117) @p_121) :rule equiv_pos2) (step t33 (cl @p_121) :rule th_resolution :premises (t30 t31 t32)) (anchor :step t34 :args ((:= (veriT_vr10 Real) veriT_vr13) (:= (veriT_vr11 Real) veriT_vr14) (:= (veriT_vr12 Real) veriT_vr15))) (step t34.t1 (cl (! (= veriT_vr10 veriT_vr13) :named @p_127)) :rule refl) (step t34.t2 (cl (! (= veriT_vr11 veriT_vr14) :named @p_123)) :rule refl) (step t34.t3 (cl (= @p_86 (! (divide$ veriT_vr13 veriT_vr14) :named @p_122))) :rule cong :premises (t34.t1 t34.t2)) (step t34.t4 (cl (! (= veriT_vr12 veriT_vr15) :named @p_128)) :rule refl) (step t34.t5 (cl @p_123) :rule refl) (step t34.t6 (cl (= @p_89 (! (divide$ veriT_vr15 veriT_vr14) :named @p_124))) :rule cong :premises (t34.t4 t34.t5)) (step t34.t7 (cl (= @p_91 (! (< @p_122 @p_124) :named @p_125))) :rule cong :premises (t34.t3 t34.t6)) (step t34.t8 (cl @p_123) :rule refl) (step t34.t9 (cl (= @p_93 (! (< 0.0 veriT_vr14) :named @p_126))) :rule cong :premises (t34.t8)) (step t34.t10 (cl @p_127) :rule refl) (step t34.t11 (cl @p_128) :rule refl) (step t34.t12 (cl (= @p_97 (! (< veriT_vr13 veriT_vr15) :named @p_129))) :rule cong :premises (t34.t10 t34.t11)) (step t34.t13 (cl (= @p_99 (! (=> @p_126 @p_129) :named @p_130))) :rule cong :premises (t34.t9 t34.t12)) (step t34.t14 (cl @p_123) :rule refl) (step t34.t15 (cl (= @p_101 (! (< veriT_vr14 0.0) :named @p_131))) :rule cong :premises (t34.t14)) (step t34.t16 (cl @p_128) :rule refl) (step t34.t17 (cl @p_127) :rule refl) (step t34.t18 (cl (= @p_103 (! (< veriT_vr15 veriT_vr13) :named @p_132))) :rule cong :premises (t34.t16 t34.t17)) (step t34.t19 (cl (= @p_105 (! (=> @p_131 @p_132) :named @p_133))) :rule cong :premises (t34.t15 t34.t18)) (step t34.t20 (cl @p_123) :rule refl) (step t34.t21 (cl (= @p_107 (! (= 0.0 veriT_vr14) :named @p_134))) :rule cong :premises (t34.t20)) (step t34.t22 (cl (= @p_109 (! (not @p_134) :named @p_135))) :rule cong :premises (t34.t21)) (step t34.t23 (cl (= @p_118 (! (and @p_130 @p_133 @p_135) :named @p_136))) :rule cong :premises (t34.t13 t34.t19 t34.t22)) (step t34.t24 (cl (= @p_119 (! (= @p_125 @p_136) :named @p_137))) :rule cong :premises (t34.t7 t34.t23)) (step t34 (cl (! (= @p_121 (! (forall ((veriT_vr13 Real) (veriT_vr14 Real) (veriT_vr15 Real)) @p_137) :named @p_139)) :named @p_138)) :rule bind) (step t35 (cl (not @p_138) (not @p_121) @p_139) :rule equiv_pos2) (step t36 (cl @p_139) :rule th_resolution :premises (t33 t34 t35)) (anchor :step t37 :args ((:= (?v0 Real) veriT_vr16) (:= (?v1 Real) veriT_vr17))) (step t37.t1 (cl (! (= ?v0 veriT_vr16) :named @p_144)) :rule refl) (step t37.t2 (cl (= @p_4 (! (- veriT_vr16) :named @p_141))) :rule cong :premises (t37.t1)) (step t37.t3 (cl (! (= ?v1 veriT_vr17) :named @p_145)) :rule refl) (step t37.t4 (cl (= @p_142 (! (divide$ @p_141 veriT_vr17) :named @p_143))) :rule cong :premises (t37.t2 t37.t3)) (step t37.t5 (cl @p_144) :rule refl) (step t37.t6 (cl @p_145) :rule refl) (step t37.t7 (cl (= @p_5 (! (divide$ veriT_vr16 veriT_vr17) :named @p_146))) :rule cong :premises (t37.t5 t37.t6)) (step t37.t8 (cl (= @p_147 (! (- @p_146) :named @p_148))) :rule cong :premises (t37.t7)) (step t37.t9 (cl (= @p_149 (! (= @p_143 @p_148) :named @p_150))) :rule cong :premises (t37.t4 t37.t8)) (step t37 (cl (! (= @p_140 (! (forall ((veriT_vr16 Real) (veriT_vr17 Real)) @p_150) :named @p_152)) :named @p_151)) :rule bind) (step t38 (cl (not @p_151) (not @p_140) @p_152) :rule equiv_pos2) (step t39 (cl @p_152) :rule th_resolution :premises (axiom8 t37 t38)) (anchor :step t40 :args ((:= (veriT_vr16 Real) veriT_vr18) (:= (veriT_vr17 Real) veriT_vr19))) (step t40.t1 (cl (! (= veriT_vr16 veriT_vr18) :named @p_155)) :rule refl) (step t40.t2 (cl (= @p_141 (! (- veriT_vr18) :named @p_153))) :rule cong :premises (t40.t1)) (step t40.t3 (cl (! (= veriT_vr17 veriT_vr19) :named @p_156)) :rule refl) (step t40.t4 (cl (= @p_143 (! (divide$ @p_153 veriT_vr19) :named @p_154))) :rule cong :premises (t40.t2 t40.t3)) (step t40.t5 (cl @p_155) :rule refl) (step t40.t6 (cl @p_156) :rule refl) (step t40.t7 (cl (= @p_146 (! (divide$ veriT_vr18 veriT_vr19) :named @p_157))) :rule cong :premises (t40.t5 t40.t6)) (step t40.t8 (cl (= @p_148 (! (- @p_157) :named @p_158))) :rule cong :premises (t40.t7)) (step t40.t9 (cl (= @p_150 (! (= @p_154 @p_158) :named @p_159))) :rule cong :premises (t40.t4 t40.t8)) (step t40 (cl (! (= @p_152 (! (forall ((veriT_vr18 Real) (veriT_vr19 Real)) @p_159) :named @p_161)) :named @p_160)) :rule bind) (step t41 (cl (not @p_160) (not @p_152) @p_161) :rule equiv_pos2) (step t42 (cl @p_161) :rule th_resolution :premises (t39 t40 t41)) (anchor :step t43 :args ((:= (?v0 Real) veriT_vr20) (:= (?v1 Real) veriT_vr21))) (step t43.t1 (cl (! (= ?v0 veriT_vr20) :named @p_166)) :rule refl) (step t43.t2 (cl (= @p_4 (! (- veriT_vr20) :named @p_163))) :rule cong :premises (t43.t1)) (step t43.t3 (cl (! (= ?v1 veriT_vr21) :named @p_167)) :rule refl) (step t43.t4 (cl (= @p_164 (! (times$ @p_163 veriT_vr21) :named @p_165))) :rule cong :premises (t43.t2 t43.t3)) (step t43.t5 (cl @p_166) :rule refl) (step t43.t6 (cl @p_167) :rule refl) (step t43.t7 (cl (= @p_168 (! (times$ veriT_vr20 veriT_vr21) :named @p_169))) :rule cong :premises (t43.t5 t43.t6)) (step t43.t8 (cl (= @p_170 (! (- @p_169) :named @p_171))) :rule cong :premises (t43.t7)) (step t43.t9 (cl (= @p_172 (! (= @p_165 @p_171) :named @p_173))) :rule cong :premises (t43.t4 t43.t8)) (step t43 (cl (! (= @p_162 (! (forall ((veriT_vr20 Real) (veriT_vr21 Real)) @p_173) :named @p_175)) :named @p_174)) :rule bind) (step t44 (cl (not @p_174) (not @p_162) @p_175) :rule equiv_pos2) (step t45 (cl @p_175) :rule th_resolution :premises (axiom9 t43 t44)) (anchor :step t46 :args ((:= (veriT_vr20 Real) veriT_vr22) (:= (veriT_vr21 Real) veriT_vr23))) (step t46.t1 (cl (! (= veriT_vr20 veriT_vr22) :named @p_178)) :rule refl) (step t46.t2 (cl (= @p_163 (! (- veriT_vr22) :named @p_176))) :rule cong :premises (t46.t1)) (step t46.t3 (cl (! (= veriT_vr21 veriT_vr23) :named @p_179)) :rule refl) (step t46.t4 (cl (= @p_165 (! (times$ @p_176 veriT_vr23) :named @p_177))) :rule cong :premises (t46.t2 t46.t3)) (step t46.t5 (cl @p_178) :rule refl) (step t46.t6 (cl @p_179) :rule refl) (step t46.t7 (cl (= @p_169 (! (times$ veriT_vr22 veriT_vr23) :named @p_180))) :rule cong :premises (t46.t5 t46.t6)) (step t46.t8 (cl (= @p_171 (! (- @p_180) :named @p_181))) :rule cong :premises (t46.t7)) (step t46.t9 (cl (= @p_173 (! (= @p_177 @p_181) :named @p_182))) :rule cong :premises (t46.t4 t46.t8)) (step t46 (cl (! (= @p_175 (! (forall ((veriT_vr22 Real) (veriT_vr23 Real)) @p_182) :named @p_184)) :named @p_183)) :rule bind) (step t47 (cl (not @p_183) (not @p_175) @p_184) :rule equiv_pos2) (step t48 (cl @p_184) :rule th_resolution :premises (t45 t46 t47)) (anchor :step t49 :args ((:= (?v0 Real) veriT_vr24) (:= (?v1 Real) veriT_vr25) (:= (?v2 Real) veriT_vr26) (:= (?v3 Real) veriT_vr27))) (step t49.t1 (cl (! (= ?v0 veriT_vr24) :named @p_192)) :rule refl) (step t49.t2 (cl (! (= ?v1 veriT_vr25) :named @p_196)) :rule refl) (step t49.t3 (cl (= @p_186 (! (pair$ veriT_vr24 veriT_vr25) :named @p_187))) :rule cong :premises (t49.t1 t49.t2)) (step t49.t4 (cl (! (= ?v2 veriT_vr26) :named @p_193)) :rule refl) (step t49.t5 (cl (! (= ?v3 veriT_vr27) :named @p_197)) :rule refl) (step t49.t6 (cl (= @p_188 (! (pair$ veriT_vr26 veriT_vr27) :named @p_189))) :rule cong :premises (t49.t4 t49.t5)) (step t49.t7 (cl (= @p_190 (! (= @p_187 @p_189) :named @p_191))) :rule cong :premises (t49.t3 t49.t6)) (step t49.t8 (cl @p_192) :rule refl) (step t49.t9 (cl @p_193) :rule refl) (step t49.t10 (cl (= @p_194 (! (= veriT_vr24 veriT_vr26) :named @p_195))) :rule cong :premises (t49.t8 t49.t9)) (step t49.t11 (cl @p_196) :rule refl) (step t49.t12 (cl @p_197) :rule refl) (step t49.t13 (cl (= @p_198 (! (= veriT_vr25 veriT_vr27) :named @p_199))) :rule cong :premises (t49.t11 t49.t12)) (step t49.t14 (cl (= @p_200 (! (and @p_195 @p_199) :named @p_201))) :rule cong :premises (t49.t10 t49.t13)) (step t49.t15 (cl (= @p_202 (! (= @p_191 @p_201) :named @p_203))) :rule cong :premises (t49.t7 t49.t14)) (step t49 (cl (! (= @p_185 (! (forall ((veriT_vr24 Real) (veriT_vr25 Real) (veriT_vr26 Real) (veriT_vr27 Real)) @p_203) :named @p_205)) :named @p_204)) :rule bind) (step t50 (cl (not @p_204) (not @p_185) @p_205) :rule equiv_pos2) (step t51 (cl @p_205) :rule th_resolution :premises (axiom10 t49 t50)) (anchor :step t52 :args ((:= (veriT_vr24 Real) veriT_vr28) (:= (veriT_vr25 Real) veriT_vr29) (:= (veriT_vr26 Real) veriT_vr30) (:= (veriT_vr27 Real) veriT_vr31))) (step t52.t1 (cl (! (= veriT_vr24 veriT_vr28) :named @p_209)) :rule refl) (step t52.t2 (cl (! (= veriT_vr25 veriT_vr29) :named @p_212)) :rule refl) (step t52.t3 (cl (= @p_187 (! (pair$ veriT_vr28 veriT_vr29) :named @p_206))) :rule cong :premises (t52.t1 t52.t2)) (step t52.t4 (cl (! (= veriT_vr26 veriT_vr30) :named @p_210)) :rule refl) (step t52.t5 (cl (! (= veriT_vr27 veriT_vr31) :named @p_213)) :rule refl) (step t52.t6 (cl (= @p_189 (! (pair$ veriT_vr30 veriT_vr31) :named @p_207))) :rule cong :premises (t52.t4 t52.t5)) (step t52.t7 (cl (= @p_191 (! (= @p_206 @p_207) :named @p_208))) :rule cong :premises (t52.t3 t52.t6)) (step t52.t8 (cl @p_209) :rule refl) (step t52.t9 (cl @p_210) :rule refl) (step t52.t10 (cl (= @p_195 (! (= veriT_vr28 veriT_vr30) :named @p_211))) :rule cong :premises (t52.t8 t52.t9)) (step t52.t11 (cl @p_212) :rule refl) (step t52.t12 (cl @p_213) :rule refl) (step t52.t13 (cl (= @p_199 (! (= veriT_vr29 veriT_vr31) :named @p_214))) :rule cong :premises (t52.t11 t52.t12)) (step t52.t14 (cl (= @p_201 (! (and @p_211 @p_214) :named @p_215))) :rule cong :premises (t52.t10 t52.t13)) (step t52.t15 (cl (= @p_203 (! (= @p_208 @p_215) :named @p_216))) :rule cong :premises (t52.t7 t52.t14)) (step t52 (cl (! (= @p_205 (! (forall ((veriT_vr28 Real) (veriT_vr29 Real) (veriT_vr30 Real) (veriT_vr31 Real)) @p_216) :named @p_218)) :named @p_217)) :rule bind) (step t53 (cl (not @p_217) (not @p_205) @p_218) :rule equiv_pos2) (step t54 (cl @p_218) :rule th_resolution :premises (t51 t52 t53)) (step t55 (cl (! (= @p_219 (! (and @p_220 (! (not false) :named @p_228)) :named @p_222)) :named @p_221)) :rule bool_simplify) (step t56 (cl (! (not @p_221) :named @p_225) (! (not @p_219) :named @p_223) @p_222) :rule equiv_pos2) (step t57 (cl (not @p_223) @p_224) :rule not_not) (step t58 (cl @p_225 @p_224 @p_222) :rule th_resolution :premises (t57 t56)) (step t59 (cl @p_222) :rule th_resolution :premises (axiom11 t55 t58)) (step t60 (cl (! (= @p_222 (! (and @p_226 @p_227 @p_228) :named @p_230)) :named @p_229)) :rule ac_simp) (step t61 (cl (not @p_229) (not @p_222) @p_230) :rule equiv_pos2) (step t62 (cl @p_230) :rule th_resolution :premises (t59 t60 t61)) (step t63 (cl (! (= @p_228 true) :named @p_269)) :rule not_simplify) (step t64 (cl (= @p_230 (! (and @p_226 @p_227 true) :named @p_231))) :rule cong :premises (t63)) (step t65 (cl (= @p_231 @p_220)) :rule and_simplify) (step t66 (cl (! (= @p_230 @p_220) :named @p_232)) :rule trans :premises (t64 t65)) (step t67 (cl (not @p_232) (not @p_230) @p_220) :rule equiv_pos2) (step t68 (cl @p_220) :rule th_resolution :premises (t62 t66 t67)) (step t69 (cl @p_227) :rule and :premises (t68)) (step t70 (cl (or (! (not @p_218) :named @p_233) (! (forall ((veriT_vr28 Real) (veriT_vr29 Real) (veriT_vr30 Real) (veriT_vr31 Real)) (or (not @p_208) @p_211)) :named @p_359))) :rule qnt_cnf) (step t71 (cl (not (! (not @p_233) :named @p_364)) @p_218) :rule not_not) (step t72 (cl (not (! (not (! (not @p_184) :named @p_234)) :named @p_243)) @p_184) :rule not_not) (step t73 (cl (or @p_234 (! (= (! (times$ (! (- 0.0) :named @p_238) d$) :named @p_239) (! (- (! (times$ 0.0 d$) :named @p_236)) :named @p_237)) :named @p_235))) :rule forall_inst :args ((:= veriT_vr22 0.0) (:= veriT_vr23 d$))) (anchor :step t74) (assume t74.h1 @p_235) (step t74.t2 (cl (! (= 0.0 @p_238) :named @p_249)) :rule minus_simplify) (step t74.t3 (cl (= @p_239 @p_236)) :rule cong :premises (t74.t2)) (step t74.t4 (cl (! (= @p_235 (! (= @p_236 @p_237) :named @p_240)) :named @p_241)) :rule cong :premises (t74.t3)) (step t74.t5 (cl (not @p_241) (! (not @p_235) :named @p_242) @p_240) :rule equiv_pos2) (step t74.t6 (cl @p_240) :rule th_resolution :premises (t74.h1 t74.t4 t74.t5)) (step t74 (cl @p_242 @p_240) :rule subproof :discharge (h1)) (step t75 (cl @p_234 @p_235) :rule or :premises (t73)) (step t76 (cl (! (or @p_234 @p_240) :named @p_244) @p_243) :rule or_neg) (step t77 (cl @p_244 @p_184) :rule th_resolution :premises (t72 t76)) (step t78 (cl @p_244 (! (not @p_240) :named @p_350)) :rule or_neg) (step t79 (cl @p_244) :rule th_resolution :premises (t75 t74 t77 t78)) (step t80 (cl (not (! (not (! (not @p_161) :named @p_245)) :named @p_254)) @p_161) :rule not_not) (step t81 (cl (or @p_245 (! (= (! (divide$ @p_238 2.0) :named @p_250) (! (- (! (divide$ 0.0 2.0) :named @p_247)) :named @p_248)) :named @p_246))) :rule forall_inst :args ((:= veriT_vr18 0.0) (:= veriT_vr19 2.0))) (anchor :step t82) (assume t82.h1 @p_246) (step t82.t2 (cl @p_249) :rule minus_simplify) (step t82.t3 (cl (= @p_250 @p_247)) :rule cong :premises (t82.t2)) (step t82.t4 (cl (! (= @p_246 (! (= @p_247 @p_248) :named @p_251)) :named @p_252)) :rule cong :premises (t82.t3)) (step t82.t5 (cl (not @p_252) (! (not @p_246) :named @p_253) @p_251) :rule equiv_pos2) (step t82.t6 (cl @p_251) :rule th_resolution :premises (t82.h1 t82.t4 t82.t5)) (step t82 (cl @p_253 @p_251) :rule subproof :discharge (h1)) (step t83 (cl @p_245 @p_246) :rule or :premises (t81)) (step t84 (cl (! (or @p_245 @p_251) :named @p_255) @p_254) :rule or_neg) (step t85 (cl @p_255 @p_161) :rule th_resolution :premises (t80 t84)) (step t86 (cl @p_255 (! (not @p_251) :named @p_336)) :rule or_neg) (step t87 (cl @p_255) :rule th_resolution :premises (t83 t82 t85 t86)) (step t88 (cl (not (! (not (! (not @p_139) :named @p_256)) :named @p_276)) @p_139) :rule not_not) (step t89 (cl (or @p_256 (! (= (! (< @p_247 @p_3) :named @p_259) (! (and (! (=> (! (< 0.0 2.0) :named @p_260) @p_257) :named @p_261) (! (=> (! (< 2.0 0.0) :named @p_263) (! (< d$ 0.0) :named @p_265)) :named @p_264) (! (not (! (= 2.0 0.0) :named @p_267)) :named @p_268)) :named @p_270)) :named @p_258))) :rule forall_inst :args ((:= veriT_vr13 0.0) (:= veriT_vr14 2.0) (:= veriT_vr15 d$))) (anchor :step t90) (assume t90.h1 @p_258) (step t90.t2 (cl (= @p_260 true)) :rule comp_simplify) (step t90.t3 (cl (= @p_261 (! (=> true @p_257) :named @p_262))) :rule cong :premises (t90.t2)) (step t90.t4 (cl (= @p_262 @p_257)) :rule implies_simplify) (step t90.t5 (cl (= @p_261 @p_257)) :rule trans :premises (t90.t3 t90.t4)) (step t90.t6 (cl (= @p_263 false)) :rule comp_simplify) (step t90.t7 (cl (= @p_264 (! (=> false @p_265) :named @p_266))) :rule cong :premises (t90.t6)) (step t90.t8 (cl (= @p_266 true)) :rule implies_simplify) (step t90.t9 (cl (= @p_264 true)) :rule trans :premises (t90.t7 t90.t8)) (step t90.t10 (cl (= @p_267 false)) :rule eq_simplify) (step t90.t11 (cl (= @p_268 @p_228)) :rule cong :premises (t90.t10)) (step t90.t12 (cl @p_269) :rule not_simplify) (step t90.t13 (cl (= @p_268 true)) :rule trans :premises (t90.t11 t90.t12)) (step t90.t14 (cl (= @p_270 (! (and @p_257 true true) :named @p_271))) :rule cong :premises (t90.t5 t90.t9 t90.t13)) (step t90.t15 (cl (= @p_271 (! (and @p_257) :named @p_272))) :rule and_simplify) (step t90.t16 (cl (= @p_272 @p_257)) :rule and_simplify) (step t90.t17 (cl (= @p_270 @p_257)) :rule trans :premises (t90.t14 t90.t15 t90.t16)) (step t90.t18 (cl (! (= @p_258 (! (= @p_259 @p_257) :named @p_273)) :named @p_274)) :rule cong :premises (t90.t17)) (step t90.t19 (cl (not @p_274) (! (not @p_258) :named @p_275) @p_273) :rule equiv_pos2) (step t90.t20 (cl @p_273) :rule th_resolution :premises (t90.h1 t90.t18 t90.t19)) (step t90 (cl @p_275 @p_273) :rule subproof :discharge (h1)) (step t91 (cl @p_256 @p_258) :rule or :premises (t89)) (step t92 (cl (! (or @p_256 @p_273) :named @p_277) @p_276) :rule or_neg) (step t93 (cl @p_277 @p_139) :rule th_resolution :premises (t88 t92)) (step t94 (cl @p_277 (! (not @p_273) :named @p_326)) :rule or_neg) (step t95 (cl @p_277) :rule th_resolution :premises (t91 t90 t93 t94)) (step t96 (cl (not (! (not (! (not @p_84) :named @p_278)) :named @p_291)) @p_84) :rule not_not) (step t97 (cl (or @p_278 (! (= @p_36 (! (- @p_3 (! (ite (! (< @p_2 0.0) :named @p_280) (! (- @p_2) :named @p_282) @p_2) :named @p_283)) :named @p_284)) :named @p_279))) :rule forall_inst :args ((:= veriT_vr9 @p_2))) (anchor :step t98) (assume t98.h1 @p_279) (step t98.t2 (cl (= @p_3 @p_282)) :rule minus_simplify) (step t98.t3 (cl (= @p_283 (! (ite @p_280 @p_3 @p_2) :named @p_281))) :rule cong :premises (t98.t2)) (step t98.t4 (cl (= @p_284 (! (- @p_3 @p_281) :named @p_285))) :rule cong :premises (t98.t3)) (step t98.t5 (cl (! (= @p_279 (! (= @p_36 @p_285) :named @p_288)) :named @p_286)) :rule cong :premises (t98.t4)) (step t98.t6 (cl (not @p_286) (! (not @p_279) :named @p_287) @p_288) :rule equiv_pos2) (step t98.t7 (cl @p_288) :rule th_resolution :premises (t98.h1 t98.t5 t98.t6)) (step t98.t8 (cl (! (= @p_288 (! (and (! (= @p_36 (- @p_3 @p_281)) :named @p_328) (! (ite @p_280 (! (= @p_3 @p_281) :named @p_330) (= @p_2 @p_281)) :named @p_329)) :named @p_289)) :named @p_290)) :rule ite_intro) (step t98.t9 (cl (not @p_290) (not @p_288) @p_289) :rule equiv_pos2) (step t98.t10 (cl @p_289) :rule th_resolution :premises (t98.t7 t98.t8 t98.t9)) (step t98 (cl @p_287 @p_289) :rule subproof :discharge (h1)) (step t99 (cl @p_278 @p_279) :rule or :premises (t97)) (step t100 (cl (! (or @p_278 @p_289) :named @p_292) @p_291) :rule or_neg) (step t101 (cl @p_292 @p_84) :rule th_resolution :premises (t96 t100)) (step t102 (cl @p_292 (! (not @p_289) :named @p_327)) :rule or_neg) (step t103 (cl @p_292) :rule th_resolution :premises (t99 t98 t101 t102)) (step t104 (cl (or @p_278 (! (= (! (diamond_y$ @p_3) :named @p_294) (- @p_3 (! (ite (! (< @p_3 0.0) :named @p_296) @p_2 @p_3) :named @p_295))) :named @p_293))) :rule forall_inst :args ((:= veriT_vr9 @p_3))) (anchor :step t105) (assume t105.h1 @p_293) (step t105.t2 (cl (! (= @p_293 (! (and (= @p_294 (- @p_3 @p_295)) (! (ite @p_296 (! (= @p_2 @p_295) :named @p_342) (! (= @p_3 @p_295) :named @p_332)) :named @p_331)) :named @p_297)) :named @p_298)) :rule ite_intro) (step t105.t3 (cl (not @p_298) (! (not @p_293) :named @p_299) @p_297) :rule equiv_pos2) (step t105.t4 (cl @p_297) :rule th_resolution :premises (t105.h1 t105.t2 t105.t3)) (step t105 (cl @p_299 @p_297) :rule subproof :discharge (h1)) (step t106 (cl @p_278 @p_293) :rule or :premises (t104)) (step t107 (cl (! (or @p_278 @p_297) :named @p_300) @p_291) :rule or_neg) (step t108 (cl @p_300 @p_84) :rule th_resolution :premises (t96 t107)) (step t109 (cl @p_300 (! (not @p_297) :named @p_333)) :rule or_neg) (step t110 (cl @p_300) :rule th_resolution :premises (t106 t105 t108 t109)) (step t111 (cl (or @p_278 (! (= (! (diamond_y$ 0.0) :named @p_302) (! (- @p_3 (! (ite (! (< 0.0 0.0) :named @p_303) @p_238 0.0) :named @p_304)) :named @p_306)) :named @p_301))) :rule forall_inst :args ((:= veriT_vr9 0.0))) (anchor :step t112) (assume t112.h1 @p_301) (step t112.t2 (cl (= @p_303 false)) :rule comp_simplify) (step t112.t3 (cl @p_249) :rule minus_simplify) (step t112.t4 (cl (= @p_304 (! (ite false 0.0 0.0) :named @p_305))) :rule cong :premises (t112.t2 t112.t3)) (step t112.t5 (cl (= 0.0 @p_305)) :rule ite_simplify) (step t112.t6 (cl (= 0.0 @p_304)) :rule trans :premises (t112.t4 t112.t5)) (step t112.t7 (cl (= @p_306 (! (- @p_3 0.0) :named @p_307))) :rule cong :premises (t112.t6)) (step t112.t8 (cl (= @p_3 @p_307)) :rule minus_simplify) (step t112.t9 (cl (= @p_3 @p_306)) :rule trans :premises (t112.t7 t112.t8)) (step t112.t10 (cl (! (= @p_301 (! (= @p_3 @p_302) :named @p_308)) :named @p_309)) :rule cong :premises (t112.t9)) (step t112.t11 (cl (not @p_309) (! (not @p_301) :named @p_310) @p_308) :rule equiv_pos2) (step t112.t12 (cl @p_308) :rule th_resolution :premises (t112.h1 t112.t10 t112.t11)) (step t112 (cl @p_310 @p_308) :rule subproof :discharge (h1)) (step t113 (cl @p_278 @p_301) :rule or :premises (t111)) (step t114 (cl (! (or @p_278 @p_308) :named @p_311) @p_291) :rule or_neg) (step t115 (cl @p_311 @p_84) :rule th_resolution :premises (t96 t114)) (step t116 (cl @p_311 (! (not @p_308) :named @p_376)) :rule or_neg) (step t117 (cl @p_311) :rule th_resolution :premises (t113 t112 t115 t116)) (step t118 (cl (or (! (not @p_60) :named @p_334) (! (= (! (fun_app$ uub$ @p_7) :named @p_369) (! (pair$ @p_2 (! (times$ (- (* 2.0 @p_7) 1.0) @p_36) :named @p_361)) :named @p_360)) :named @p_335))) :rule forall_inst :args ((:= veriT_vr7 @p_7))) (step t119 (cl (or (! (not @p_34) :named @p_323) (! (= (! (fun_app$ uuc$ @p_7) :named @p_314) (! (pair$ (! (times$ (! (- @p_7 @p_7) :named @p_315) d$) :named @p_312) (! (diamond_y$ @p_312) :named @p_316)) :named @p_318)) :named @p_313))) :rule forall_inst :args ((:= veriT_vr1 @p_7))) (anchor :step t120) (assume t120.h1 @p_313) (step t120.t2 (cl (= 0.0 @p_315)) :rule minus_simplify) (step t120.t3 (cl (= @p_236 @p_312)) :rule cong :premises (t120.t2)) (step t120.t4 (cl (= @p_316 (! (diamond_y$ @p_236) :named @p_317))) :rule cong :premises (t120.t3)) (step t120.t5 (cl (= @p_318 (! (pair$ @p_236 @p_317) :named @p_319))) :rule cong :premises (t120.t3 t120.t4)) (step t120.t6 (cl (! (= @p_313 (! (= @p_314 @p_319) :named @p_320)) :named @p_321)) :rule cong :premises (t120.t5)) (step t120.t7 (cl (not @p_321) (! (not @p_313) :named @p_322) @p_320) :rule equiv_pos2) (step t120.t8 (cl @p_320) :rule th_resolution :premises (t120.h1 t120.t6 t120.t7)) (step t120 (cl @p_322 @p_320) :rule subproof :discharge (h1)) (step t121 (cl @p_323 @p_313) :rule or :premises (t119)) (step t122 (cl (! (or @p_323 @p_320) :named @p_325) (! (not @p_323) :named @p_324)) :rule or_neg) (step t123 (cl (not @p_324) @p_34) :rule not_not) (step t124 (cl @p_325 @p_34) :rule th_resolution :premises (t123 t122)) (step t125 (cl @p_325 (! (not @p_320) :named @p_370)) :rule or_neg) (step t126 (cl @p_325) :rule th_resolution :premises (t121 t120 t124 t125)) (step t127 (cl @p_234 @p_240) :rule or :premises (t79)) (step t128 (cl @p_240) :rule resolution :premises (t127 t48)) (step t129 (cl @p_245 @p_251) :rule or :premises (t87)) (step t130 (cl @p_251) :rule resolution :premises (t129 t42)) (step t131 (cl @p_326 @p_259 (not @p_257)) :rule equiv_pos1) (step t132 (cl @p_256 @p_273) :rule or :premises (t95)) (step t133 (cl @p_326 @p_259) :rule resolution :premises (t131 axiom4)) (step t134 (cl @p_273) :rule resolution :premises (t132 t36)) (step t135 (cl @p_259) :rule resolution :premises (t133 t134)) (step t136 (cl @p_327 @p_328) :rule and_pos) (step t137 (cl (not @p_329) (! (not @p_280) :named @p_339) @p_330) :rule ite_pos2) (step t138 (cl @p_327 @p_329) :rule and_pos) (step t139 (cl @p_278 @p_289) :rule or :premises (t103)) (step t140 (cl @p_289) :rule resolution :premises (t139 t27)) (step t141 (cl @p_328) :rule resolution :premises (t136 t140)) (step t142 (cl @p_329) :rule resolution :premises (t138 t140)) (step t143 (cl (not @p_331) @p_296 @p_332) :rule ite_pos1) (step t144 (cl @p_333 @p_331) :rule and_pos) (step t145 (cl @p_278 @p_297) :rule or :premises (t110)) (step t146 (cl @p_297) :rule resolution :premises (t145 t27)) (step t147 (cl @p_331) :rule resolution :premises (t144 t146)) (step t148 (cl @p_278 @p_308) :rule or :premises (t117)) (step t149 (cl @p_308) :rule resolution :premises (t148 t27)) (step t150 (cl @p_334 @p_335) :rule or :premises (t118)) (step t151 (cl @p_335) :rule resolution :premises (t150 t21)) (step t152 (cl @p_323 @p_320) :rule or :premises (t126)) (step t153 (cl @p_320) :rule resolution :premises (t152 t15)) (step t154 (cl (! (not @p_259) :named @p_338) (! (not @p_296) :named @p_337) @p_336) :rule la_generic :args (2.0 2.0 1)) (step t155 (cl @p_337) :rule resolution :premises (t154 t130 t135)) (step t156 (cl @p_332) :rule resolution :premises (t143 t155 t147)) (step t157 (cl @p_338 @p_280 @p_336) :rule la_generic :args (2.0 2.0 1)) (step t158 (cl @p_280) :rule resolution :premises (t157 t130 t135)) (step t159 (cl @p_330) :rule resolution :premises (t137 t158 t142)) (step t160 (cl (! (not (! (= @p_3 @p_2) :named @p_343)) :named @p_341) (not (! (= 0.0 0.0) :named @p_340)) @p_339 @p_296) :rule eq_congruent_pred) (step t161 (cl @p_340) :rule eq_reflexive) (step t162 (cl @p_341 @p_339 @p_296) :rule th_resolution :premises (t160 t161)) (step t163 (cl (! (not @p_332) :named @p_344) (! (not @p_342) :named @p_345) @p_343) :rule eq_transitive) (step t164 (cl @p_339 @p_296 @p_344 @p_345) :rule th_resolution :premises (t162 t163)) (step t165 (cl @p_345) :rule resolution :premises (t164 t155 t158 t156)) (step t166 (cl (! (= @p_7 @p_7) :named @p_372)) :rule eq_reflexive) (step t167 (cl (or (! (= 0.0 @p_236) :named @p_346) (! (not (! (<= 0.0 @p_236) :named @p_351)) :named @p_347) (! (not (! (<= @p_236 0.0) :named @p_349)) :named @p_348))) :rule la_disequality) (step t168 (cl @p_346 @p_347 @p_348) :rule or :premises (t167)) (step t169 (cl @p_349 @p_350) :rule la_generic :args (2.0 (- 1))) (step t170 (cl @p_349) :rule resolution :premises (t169 t128)) (step t171 (cl @p_351 @p_350) :rule la_generic :args (2.0 1)) (step t172 (cl @p_351) :rule resolution :premises (t171 t128)) (step t173 (cl @p_346) :rule resolution :premises (t168 t172 t170)) (step t174 (cl (or (! (= @p_36 @p_236) :named @p_352) (! (not (! (<= @p_36 @p_236) :named @p_356)) :named @p_353) (! (not (! (<= @p_236 @p_36) :named @p_355)) :named @p_354))) :rule la_disequality) (step t175 (cl @p_352 @p_353 @p_354) :rule or :premises (t174)) (step t176 (cl @p_355 @p_350 (! (not @p_328) :named @p_357) (! (not @p_330) :named @p_358)) :rule la_generic :args (2.0 (- 1) 2 2)) (step t177 (cl @p_355) :rule resolution :premises (t176 t128 t141 t159)) (step t178 (cl @p_356 @p_350 @p_357 @p_358) :rule la_generic :args (2.0 1 (- 2) (- 2))) (step t179 (cl @p_356) :rule resolution :premises (t178 t128 t141 t159)) (step t180 (cl @p_352) :rule resolution :premises (t175 t179 t177)) (step t181 (cl @p_233 @p_359) :rule or :premises (t70)) (step t182 (cl (or (! (not @p_359) :named @p_362) (! (or (! (not (! (= @p_360 @p_319) :named @p_371)) :named @p_367) (! (= @p_2 @p_236) :named @p_368)) :named @p_363))) :rule forall_inst :args ((:= veriT_vr28 @p_2) (:= veriT_vr29 @p_361) (:= veriT_vr30 @p_236) (:= veriT_vr31 @p_317))) (step t183 (cl @p_362 @p_363) :rule or :premises (t182)) (step t184 (cl (! (or @p_233 @p_363) :named @p_365) @p_364) :rule or_neg) (step t185 (cl @p_365 @p_218) :rule th_resolution :premises (t71 t184)) (step t186 (cl @p_365 (! (not @p_363) :named @p_366)) :rule or_neg) (step t187 (cl @p_365) :rule th_resolution :premises (t181 t183 t185 t186)) (step t188 (cl @p_366 @p_367 @p_368) :rule or_pos) (step t189 (cl @p_233 @p_363) :rule or :premises (t187)) (step t190 (cl @p_363) :rule resolution :premises (t189 t54)) (step t191 (cl (! (not @p_335) :named @p_375) (not (! (= @p_369 @p_314) :named @p_373)) @p_370 @p_371) :rule eq_transitive) (step t192 (cl (! (not @p_227) :named @p_374) (not @p_372) @p_373) :rule eq_congruent) (step t193 (cl @p_374 @p_373) :rule th_resolution :premises (t192 t166)) (step t194 (cl @p_375 @p_370 @p_371 @p_374) :rule th_resolution :premises (t191 t193)) (step t195 (cl @p_371) :rule resolution :premises (t194 t69 t151 t153)) (step t196 (cl @p_368) :rule resolution :premises (t188 t195 t190)) (step t197 (cl (! (not @p_368) :named @p_378) (! (not @p_352) :named @p_381) (not (! (= @p_36 @p_302) :named @p_377)) @p_376 @p_344 @p_342) :rule eq_transitive) (step t198 (cl (not (! (= @p_2 0.0) :named @p_379)) @p_377) :rule eq_congruent) (step t199 (cl @p_378 (! (not @p_346) :named @p_380) @p_379) :rule eq_transitive) (step t200 (cl @p_377 @p_378 @p_380) :rule th_resolution :premises (t198 t199)) (step t201 (cl @p_378 @p_381 @p_376 @p_344 @p_342 @p_378 @p_380) :rule th_resolution :premises (t197 t200)) (step t202 (cl @p_378 @p_381 @p_376 @p_344 @p_342 @p_380) :rule contraction :premises (t201)) (step t203 (cl) :rule resolution :premises (t202 t156 t165 t149 t173 t180 t196)) 42dfad143cfae67cfed01ebeb5997e53c8d08b98 26 0 unsat (assume axiom0 (! (forall ((?v0 A$) (?v1 B$) (?v2 C$)) (! (p$ ?v0 ?v1) :named @p_2)) :named @p_1)) (assume axiom1 (! (=> (! (p$ z$ y$) :named @p_12) false) :named @p_11)) (step t3 (cl (! (= @p_1 (! (forall ((?v0 A$) (?v1 B$)) @p_2) :named @p_4)) :named @p_3)) :rule qnt_rm_unused) (step t4 (cl (not @p_3) (not @p_1) @p_4) :rule equiv_pos2) (step t5 (cl @p_4) :rule th_resolution :premises (axiom0 t3 t4)) (anchor :step t6 :args ((:= (?v0 A$) veriT_vr0) (:= (?v1 B$) veriT_vr1))) (step t6.t1 (cl (= ?v0 veriT_vr0)) :rule refl) (step t6.t2 (cl (= ?v1 veriT_vr1)) :rule refl) (step t6.t3 (cl (= @p_2 (! (p$ veriT_vr0 veriT_vr1) :named @p_5))) :rule cong :premises (t6.t1 t6.t2)) (step t6 (cl (! (= @p_4 (! (forall ((veriT_vr0 A$) (veriT_vr1 B$)) @p_5) :named @p_7)) :named @p_6)) :rule bind) (step t7 (cl (not @p_6) (not @p_4) @p_7) :rule equiv_pos2) (step t8 (cl @p_7) :rule th_resolution :premises (t5 t6 t7)) (anchor :step t9 :args ((:= (veriT_vr0 A$) veriT_vr2) (:= (veriT_vr1 B$) veriT_vr3))) (step t9.t1 (cl (= veriT_vr0 veriT_vr2)) :rule refl) (step t9.t2 (cl (= veriT_vr1 veriT_vr3)) :rule refl) (step t9.t3 (cl (= @p_5 (! (p$ veriT_vr2 veriT_vr3) :named @p_8))) :rule cong :premises (t9.t1 t9.t2)) (step t9 (cl (! (= @p_7 (! (forall ((veriT_vr2 A$) (veriT_vr3 B$)) @p_8) :named @p_10)) :named @p_9)) :rule bind) (step t10 (cl (not @p_9) (not @p_7) @p_10) :rule equiv_pos2) (step t11 (cl @p_10) :rule th_resolution :premises (t8 t9 t10)) (step t12 (cl (! (= @p_11 (! (not @p_12) :named @p_14)) :named @p_13)) :rule implies_simplify) (step t13 (cl (not @p_13) (not @p_11) @p_14) :rule equiv_pos2) (step t14 (cl @p_14) :rule th_resolution :premises (axiom1 t12 t13)) (step t15 (cl (or (! (not @p_10) :named @p_15) @p_12)) :rule forall_inst :args ((:= veriT_vr2 z$) (:= veriT_vr3 y$))) (step t16 (cl @p_15 @p_12) :rule or :premises (t15)) (step t17 (cl) :rule resolution :premises (t16 t11 t14)) d91a4d59e816a47672957ce0be20acd9aa3eef3e 23 0 unsat (assume axiom0 (! (not (! (<= y$ (! (ite (! (<= x$ y$) :named @p_3) y$ x$) :named @p_2)) :named @p_7)) :named @p_1)) (step t2 (cl (! (= @p_1 (! (and (! (not (! (<= y$ @p_2) :named @p_13)) :named @p_9) (! (ite @p_3 (! (= y$ @p_2) :named @p_12) (! (= x$ @p_2) :named @p_11)) :named @p_10)) :named @p_5)) :named @p_4)) :rule ite_intro) (step t3 (cl (! (not @p_4) :named @p_8) (! (not @p_1) :named @p_6) @p_5) :rule equiv_pos2) (step t4 (cl (not @p_6) @p_7) :rule not_not) (step t5 (cl @p_8 @p_7 @p_5) :rule th_resolution :premises (t4 t3)) (step t6 (cl @p_5) :rule th_resolution :premises (axiom0 t2 t5)) (step t7 (cl @p_9) :rule and :premises (t6)) (step t8 (cl @p_10) :rule and :premises (t6)) (step t9 (cl @p_3 @p_11) :rule ite1 :premises (t8)) (step t10 (cl (! (not @p_3) :named @p_15) @p_12) :rule ite2 :premises (t8)) (step t11 (cl @p_13 @p_3 (! (not @p_11) :named @p_14)) :rule la_generic :args (1 1 (- 1))) (step t12 (cl @p_3 @p_14) :rule resolution :premises (t11 t7)) (step t13 (cl (not (! (= y$ x$) :named @p_17)) (! (not @p_12) :named @p_16) @p_15 @p_13) :rule eq_congruent_pred) (step t14 (cl @p_16 @p_14 @p_17) :rule eq_transitive) (step t15 (cl @p_16 @p_15 @p_13 @p_16 @p_14) :rule th_resolution :premises (t13 t14)) (step t16 (cl @p_16 @p_15 @p_13 @p_14) :rule contraction :premises (t15)) (step t17 (cl @p_16 @p_15 @p_14) :rule resolution :premises (t16 t7)) (step t18 (cl @p_14) :rule resolution :premises (t17 t10 t12)) (step t19 (cl @p_3) :rule resolution :premises (t9 t18)) (step t20 (cl @p_12) :rule resolution :premises (t10 t19)) (step t21 (cl @p_13 @p_16) :rule la_generic :args (1 (- 1))) (step t22 (cl) :rule resolution :premises (t21 t7 t20)) 47bc3239fb0fd8c5f8f4969f0c6c1996f0a21574 567 0 unsat (define-fun veriT_sk0 () A$ (! (choice ((veriT_vr145 A$)) (not (! (not (! (and (! (= (! (arg_min_on$ f$ (! (image$b g$ b$) :named @p_6)) :named @p_336) (! (fun_app$b g$ veriT_vr145) :named @p_378)) :named @p_379) (! (member$a veriT_vr145 b$) :named @p_381)) :named @p_382)) :named @p_377))) :named @p_357)) (assume axiom29 (! (forall ((?v0 B_set$) (?v1 B_c_fun$)) (! (=> (! (and (! (finite$ ?v0) :named @p_1) (! (not (! (= ?v0 bot$) :named @p_10)) :named @p_2)) :named @p_13) (! (member$ (! (arg_min_on$ ?v1 ?v0) :named @p_15) ?v0) :named @p_17)) :named @p_19)) :named @p_7)) (assume axiom31 (! (forall ((?v0 B_set$) (?v1 B$) (?v2 B_c_fun$)) (! (=> (! (and @p_1 (! (and @p_2 (! (member$ ?v1 ?v0) :named @p_38)) :named @p_40)) :named @p_42) (! (less_eq$ (! (fun_app$ ?v2 (! (arg_min_on$ ?v2 ?v0) :named @p_45)) :named @p_47) (! (fun_app$ ?v2 ?v1) :named @p_50)) :named @p_52)) :named @p_54)) :named @p_33)) (assume axiom33 (! (forall ((?v0 B_c_fun$) (?v1 A_b_fun$) (?v2 A$)) (! (= (! (fun_app$a (! (comp$ ?v0 ?v1) :named @p_78) ?v2) :named @p_80) (! (fun_app$ ?v0 (! (fun_app$b ?v1 ?v2) :named @p_3)) :named @p_86)) :named @p_88)) :named @p_77)) (assume axiom36 (! (forall ((?v0 A_set$) (?v1 A_b_fun$)) (! (=> (! (finite$a ?v0) :named @p_103) (! (finite$ (! (image$b ?v1 ?v0) :named @p_106)) :named @p_108)) :named @p_110)) :named @p_102)) (assume axiom40 (! (forall ((?v0 B$) (?v1 A_b_fun$) (?v2 A_set$)) (! (=> (! (and (! (member$ ?v0 (! (image$b ?v1 ?v2) :named @p_122)) :named @p_124) (! (forall ((?v3 A$)) (! (=> (! (and (! (= ?v0 (! (fun_app$b ?v1 ?v3) :named @p_130)) :named @p_132) (! (member$a ?v3 ?v2) :named @p_136)) :named @p_138) false) :named @p_140)) :named @p_126)) :named @p_142) false) :named @p_144)) :named @p_121)) (assume axiom44 (! (forall ((?v0 B$) (?v1 A_b_fun$) (?v2 A$) (?v3 A_set$)) (! (=> (! (and (! (= @p_3 ?v0) :named @p_173) (! (member$a ?v2 ?v3) :named @p_176)) :named @p_178) (! (member$ ?v0 (! (image$b ?v1 ?v3) :named @p_183)) :named @p_185)) :named @p_187)) :named @p_171)) (assume axiom48 (! (forall ((?v0 A_b_fun$) (?v1 A_set$)) (! (= (! (= bot$ (! (image$b ?v0 ?v1) :named @p_205)) :named @p_207) (! (= bot$a ?v1) :named @p_210)) :named @p_212)) :named @p_204)) (assume axiom50 (! (forall ((?v0 B_c_fun$) (?v1 B_set$) (?v2 B$) (?v3 B$)) (! (=> (! (and (! (inj_on$ ?v0 ?v1) :named @p_224) (! (and (! (= (! (fun_app$ ?v0 ?v2) :named @p_227) (! (fun_app$ ?v0 ?v3) :named @p_229)) :named @p_231) (! (and (! (member$ ?v2 ?v1) :named @p_235) (! (member$ ?v3 ?v1) :named @p_238)) :named @p_240)) :named @p_242)) :named @p_244) (! (= ?v3 ?v2) :named @p_246)) :named @p_248)) :named @p_223)) (assume axiom51 (! (forall ((?v0 C$) (?v1 C$)) (! (= (! (less$ ?v0 ?v1) :named @p_272) (! (and (! (less_eq$ ?v0 ?v1) :named @p_276) (! (not (! (= ?v0 ?v1) :named @p_278)) :named @p_280)) :named @p_282)) :named @p_284)) :named @p_271)) (assume axiom23 (! (inj_on$ f$ @p_6) :named @p_353)) (assume axiom24 (! (finite$a b$) :named @p_332)) (assume axiom25 (not (! (= bot$a b$) :named @p_331))) (assume axiom26 (! (member$a (! (arg_min_on$a (! (comp$ f$ g$) :named @p_4) b$) :named @p_5) b$) :named @p_423)) (assume axiom27 (! (not (! (exists ((?v0 A$)) (! (and (! (member$a ?v0 b$) :named @p_300) (! (less$ (! (fun_app$a @p_4 ?v0) :named @p_303) (! (fun_app$a @p_4 @p_5) :named @p_299)) :named @p_305)) :named @p_307)) :named @p_298)) :named @p_309)) (assume axiom52 (not (! (= @p_336 (! (fun_app$b g$ @p_5) :named @p_333)) :named @p_355))) (anchor :step t16 :args ((:= (?v0 B_set$) veriT_vr0) (:= (?v1 B_c_fun$) veriT_vr1))) (step t16.t1 (cl (! (= ?v0 veriT_vr0) :named @p_9)) :rule refl) (step t16.t2 (cl (= @p_1 (! (finite$ veriT_vr0) :named @p_8))) :rule cong :premises (t16.t1)) (step t16.t3 (cl @p_9) :rule refl) (step t16.t4 (cl (= @p_10 (! (= bot$ veriT_vr0) :named @p_11))) :rule cong :premises (t16.t3)) (step t16.t5 (cl (= @p_2 (! (not @p_11) :named @p_12))) :rule cong :premises (t16.t4)) (step t16.t6 (cl (= @p_13 (! (and @p_8 @p_12) :named @p_14))) :rule cong :premises (t16.t2 t16.t5)) (step t16.t7 (cl (= ?v1 veriT_vr1)) :rule refl) (step t16.t8 (cl @p_9) :rule refl) (step t16.t9 (cl (= @p_15 (! (arg_min_on$ veriT_vr1 veriT_vr0) :named @p_16))) :rule cong :premises (t16.t7 t16.t8)) (step t16.t10 (cl @p_9) :rule refl) (step t16.t11 (cl (= @p_17 (! (member$ @p_16 veriT_vr0) :named @p_18))) :rule cong :premises (t16.t9 t16.t10)) (step t16.t12 (cl (= @p_19 (! (=> @p_14 @p_18) :named @p_20))) :rule cong :premises (t16.t6 t16.t11)) (step t16 (cl (! (= @p_7 (! (forall ((veriT_vr0 B_set$) (veriT_vr1 B_c_fun$)) @p_20) :named @p_22)) :named @p_21)) :rule bind) (step t17 (cl (not @p_21) (not @p_7) @p_22) :rule equiv_pos2) (step t18 (cl @p_22) :rule th_resolution :premises (axiom29 t16 t17)) (anchor :step t19 :args ((:= (veriT_vr0 B_set$) veriT_vr2) (:= (veriT_vr1 B_c_fun$) veriT_vr3))) (step t19.t1 (cl (! (= veriT_vr0 veriT_vr2) :named @p_24)) :rule refl) (step t19.t2 (cl (= @p_8 (! (finite$ veriT_vr2) :named @p_23))) :rule cong :premises (t19.t1)) (step t19.t3 (cl @p_24) :rule refl) (step t19.t4 (cl (= @p_11 (! (= bot$ veriT_vr2) :named @p_25))) :rule cong :premises (t19.t3)) (step t19.t5 (cl (= @p_12 (! (not @p_25) :named @p_26))) :rule cong :premises (t19.t4)) (step t19.t6 (cl (= @p_14 (! (and @p_23 @p_26) :named @p_27))) :rule cong :premises (t19.t2 t19.t5)) (step t19.t7 (cl (= veriT_vr1 veriT_vr3)) :rule refl) (step t19.t8 (cl @p_24) :rule refl) (step t19.t9 (cl (= @p_16 (! (arg_min_on$ veriT_vr3 veriT_vr2) :named @p_28))) :rule cong :premises (t19.t7 t19.t8)) (step t19.t10 (cl @p_24) :rule refl) (step t19.t11 (cl (= @p_18 (! (member$ @p_28 veriT_vr2) :named @p_29))) :rule cong :premises (t19.t9 t19.t10)) (step t19.t12 (cl (= @p_20 (! (=> @p_27 @p_29) :named @p_30))) :rule cong :premises (t19.t6 t19.t11)) (step t19 (cl (! (= @p_22 (! (forall ((veriT_vr2 B_set$) (veriT_vr3 B_c_fun$)) @p_30) :named @p_32)) :named @p_31)) :rule bind) (step t20 (cl (not @p_31) (not @p_22) @p_32) :rule equiv_pos2) (step t21 (cl @p_32) :rule th_resolution :premises (t18 t19 t20)) (anchor :step t22 :args ((:= (?v0 B_set$) veriT_vr8) (:= (?v1 B$) veriT_vr9) (:= (?v2 B_c_fun$) veriT_vr10))) (step t22.t1 (cl (! (= ?v0 veriT_vr8) :named @p_35)) :rule refl) (step t22.t2 (cl (= @p_1 (! (finite$ veriT_vr8) :named @p_34))) :rule cong :premises (t22.t1)) (step t22.t3 (cl @p_35) :rule refl) (step t22.t4 (cl (= @p_10 (! (= bot$ veriT_vr8) :named @p_36))) :rule cong :premises (t22.t3)) (step t22.t5 (cl (= @p_2 (! (not @p_36) :named @p_37))) :rule cong :premises (t22.t4)) (step t22.t6 (cl (! (= ?v1 veriT_vr9) :named @p_49)) :rule refl) (step t22.t7 (cl @p_35) :rule refl) (step t22.t8 (cl (= @p_38 (! (member$ veriT_vr9 veriT_vr8) :named @p_39))) :rule cong :premises (t22.t6 t22.t7)) (step t22.t9 (cl (= @p_40 (! (and @p_37 @p_39) :named @p_41))) :rule cong :premises (t22.t5 t22.t8)) (step t22.t10 (cl (= @p_42 (! (and @p_34 @p_41) :named @p_43))) :rule cong :premises (t22.t2 t22.t9)) (step t22.t11 (cl (! (= ?v2 veriT_vr10) :named @p_44)) :rule refl) (step t22.t12 (cl @p_44) :rule refl) (step t22.t13 (cl @p_35) :rule refl) (step t22.t14 (cl (= @p_45 (! (arg_min_on$ veriT_vr10 veriT_vr8) :named @p_46))) :rule cong :premises (t22.t12 t22.t13)) (step t22.t15 (cl (= @p_47 (! (fun_app$ veriT_vr10 @p_46) :named @p_48))) :rule cong :premises (t22.t11 t22.t14)) (step t22.t16 (cl @p_44) :rule refl) (step t22.t17 (cl @p_49) :rule refl) (step t22.t18 (cl (= @p_50 (! (fun_app$ veriT_vr10 veriT_vr9) :named @p_51))) :rule cong :premises (t22.t16 t22.t17)) (step t22.t19 (cl (= @p_52 (! (less_eq$ @p_48 @p_51) :named @p_53))) :rule cong :premises (t22.t15 t22.t18)) (step t22.t20 (cl (= @p_54 (! (=> @p_43 @p_53) :named @p_55))) :rule cong :premises (t22.t10 t22.t19)) (step t22 (cl (! (= @p_33 (! (forall ((veriT_vr8 B_set$) (veriT_vr9 B$) (veriT_vr10 B_c_fun$)) @p_55) :named @p_57)) :named @p_56)) :rule bind) (step t23 (cl (not @p_56) (not @p_33) @p_57) :rule equiv_pos2) (step t24 (cl @p_57) :rule th_resolution :premises (axiom31 t22 t23)) (anchor :step t25 :args ((veriT_vr8 B_set$) (veriT_vr9 B$) (veriT_vr10 B_c_fun$))) (step t25.t1 (cl (= @p_43 (! (and @p_34 @p_37 @p_39) :named @p_58))) :rule ac_simp) (step t25.t2 (cl (= @p_55 (! (=> @p_58 @p_53) :named @p_59))) :rule cong :premises (t25.t1)) (step t25 (cl (! (= @p_57 (! (forall ((veriT_vr8 B_set$) (veriT_vr9 B$) (veriT_vr10 B_c_fun$)) @p_59) :named @p_61)) :named @p_60)) :rule bind) (step t26 (cl (not @p_60) (not @p_57) @p_61) :rule equiv_pos2) (step t27 (cl @p_61) :rule th_resolution :premises (t24 t25 t26)) (anchor :step t28 :args ((:= (veriT_vr8 B_set$) veriT_vr11) (:= (veriT_vr9 B$) veriT_vr12) (:= (veriT_vr10 B_c_fun$) veriT_vr13))) (step t28.t1 (cl (! (= veriT_vr8 veriT_vr11) :named @p_63)) :rule refl) (step t28.t2 (cl (= @p_34 (! (finite$ veriT_vr11) :named @p_62))) :rule cong :premises (t28.t1)) (step t28.t3 (cl @p_63) :rule refl) (step t28.t4 (cl (= @p_36 (! (= bot$ veriT_vr11) :named @p_64))) :rule cong :premises (t28.t3)) (step t28.t5 (cl (= @p_37 (! (not @p_64) :named @p_65))) :rule cong :premises (t28.t4)) (step t28.t6 (cl (! (= veriT_vr9 veriT_vr12) :named @p_71)) :rule refl) (step t28.t7 (cl @p_63) :rule refl) (step t28.t8 (cl (= @p_39 (! (member$ veriT_vr12 veriT_vr11) :named @p_66))) :rule cong :premises (t28.t6 t28.t7)) (step t28.t9 (cl (= @p_58 (! (and @p_62 @p_65 @p_66) :named @p_67))) :rule cong :premises (t28.t2 t28.t5 t28.t8)) (step t28.t10 (cl (! (= veriT_vr10 veriT_vr13) :named @p_68)) :rule refl) (step t28.t11 (cl @p_68) :rule refl) (step t28.t12 (cl @p_63) :rule refl) (step t28.t13 (cl (= @p_46 (! (arg_min_on$ veriT_vr13 veriT_vr11) :named @p_69))) :rule cong :premises (t28.t11 t28.t12)) (step t28.t14 (cl (= @p_48 (! (fun_app$ veriT_vr13 @p_69) :named @p_70))) :rule cong :premises (t28.t10 t28.t13)) (step t28.t15 (cl @p_68) :rule refl) (step t28.t16 (cl @p_71) :rule refl) (step t28.t17 (cl (= @p_51 (! (fun_app$ veriT_vr13 veriT_vr12) :named @p_72))) :rule cong :premises (t28.t15 t28.t16)) (step t28.t18 (cl (= @p_53 (! (less_eq$ @p_70 @p_72) :named @p_73))) :rule cong :premises (t28.t14 t28.t17)) (step t28.t19 (cl (= @p_59 (! (=> @p_67 @p_73) :named @p_74))) :rule cong :premises (t28.t9 t28.t18)) (step t28 (cl (! (= @p_61 (! (forall ((veriT_vr11 B_set$) (veriT_vr12 B$) (veriT_vr13 B_c_fun$)) @p_74) :named @p_76)) :named @p_75)) :rule bind) (step t29 (cl (not @p_75) (not @p_61) @p_76) :rule equiv_pos2) (step t30 (cl @p_76) :rule th_resolution :premises (t27 t28 t29)) (anchor :step t31 :args ((:= (?v0 B_c_fun$) veriT_vr20) (:= (?v1 A_b_fun$) veriT_vr21) (:= (?v2 A$) veriT_vr22))) (step t31.t1 (cl (! (= ?v0 veriT_vr20) :named @p_82)) :rule refl) (step t31.t2 (cl (! (= ?v1 veriT_vr21) :named @p_83)) :rule refl) (step t31.t3 (cl (= @p_78 (! (comp$ veriT_vr20 veriT_vr21) :named @p_79))) :rule cong :premises (t31.t1 t31.t2)) (step t31.t4 (cl (! (= ?v2 veriT_vr22) :named @p_84)) :rule refl) (step t31.t5 (cl (= @p_80 (! (fun_app$a @p_79 veriT_vr22) :named @p_81))) :rule cong :premises (t31.t3 t31.t4)) (step t31.t6 (cl @p_82) :rule refl) (step t31.t7 (cl @p_83) :rule refl) (step t31.t8 (cl @p_84) :rule refl) (step t31.t9 (cl (= @p_3 (! (fun_app$b veriT_vr21 veriT_vr22) :named @p_85))) :rule cong :premises (t31.t7 t31.t8)) (step t31.t10 (cl (= @p_86 (! (fun_app$ veriT_vr20 @p_85) :named @p_87))) :rule cong :premises (t31.t6 t31.t9)) (step t31.t11 (cl (= @p_88 (! (= @p_81 @p_87) :named @p_89))) :rule cong :premises (t31.t5 t31.t10)) (step t31 (cl (! (= @p_77 (! (forall ((veriT_vr20 B_c_fun$) (veriT_vr21 A_b_fun$) (veriT_vr22 A$)) @p_89) :named @p_91)) :named @p_90)) :rule bind) (step t32 (cl (not @p_90) (not @p_77) @p_91) :rule equiv_pos2) (step t33 (cl @p_91) :rule th_resolution :premises (axiom33 t31 t32)) (anchor :step t34 :args ((:= (veriT_vr20 B_c_fun$) veriT_vr23) (:= (veriT_vr21 A_b_fun$) veriT_vr24) (:= (veriT_vr22 A$) veriT_vr25))) (step t34.t1 (cl (! (= veriT_vr20 veriT_vr23) :named @p_94)) :rule refl) (step t34.t2 (cl (! (= veriT_vr21 veriT_vr24) :named @p_95)) :rule refl) (step t34.t3 (cl (= @p_79 (! (comp$ veriT_vr23 veriT_vr24) :named @p_92))) :rule cong :premises (t34.t1 t34.t2)) (step t34.t4 (cl (! (= veriT_vr22 veriT_vr25) :named @p_96)) :rule refl) (step t34.t5 (cl (= @p_81 (! (fun_app$a @p_92 veriT_vr25) :named @p_93))) :rule cong :premises (t34.t3 t34.t4)) (step t34.t6 (cl @p_94) :rule refl) (step t34.t7 (cl @p_95) :rule refl) (step t34.t8 (cl @p_96) :rule refl) (step t34.t9 (cl (= @p_85 (! (fun_app$b veriT_vr24 veriT_vr25) :named @p_97))) :rule cong :premises (t34.t7 t34.t8)) (step t34.t10 (cl (= @p_87 (! (fun_app$ veriT_vr23 @p_97) :named @p_98))) :rule cong :premises (t34.t6 t34.t9)) (step t34.t11 (cl (= @p_89 (! (= @p_93 @p_98) :named @p_99))) :rule cong :premises (t34.t5 t34.t10)) (step t34 (cl (! (= @p_91 (! (forall ((veriT_vr23 B_c_fun$) (veriT_vr24 A_b_fun$) (veriT_vr25 A$)) @p_99) :named @p_101)) :named @p_100)) :rule bind) (step t35 (cl (not @p_100) (not @p_91) @p_101) :rule equiv_pos2) (step t36 (cl @p_101) :rule th_resolution :premises (t33 t34 t35)) (anchor :step t37 :args ((:= (?v0 A_set$) veriT_vr34) (:= (?v1 A_b_fun$) veriT_vr35))) (step t37.t1 (cl (! (= ?v0 veriT_vr34) :named @p_105)) :rule refl) (step t37.t2 (cl (= @p_103 (! (finite$a veriT_vr34) :named @p_104))) :rule cong :premises (t37.t1)) (step t37.t3 (cl (= ?v1 veriT_vr35)) :rule refl) (step t37.t4 (cl @p_105) :rule refl) (step t37.t5 (cl (= @p_106 (! (image$b veriT_vr35 veriT_vr34) :named @p_107))) :rule cong :premises (t37.t3 t37.t4)) (step t37.t6 (cl (= @p_108 (! (finite$ @p_107) :named @p_109))) :rule cong :premises (t37.t5)) (step t37.t7 (cl (= @p_110 (! (=> @p_104 @p_109) :named @p_111))) :rule cong :premises (t37.t2 t37.t6)) (step t37 (cl (! (= @p_102 (! (forall ((veriT_vr34 A_set$) (veriT_vr35 A_b_fun$)) @p_111) :named @p_113)) :named @p_112)) :rule bind) (step t38 (cl (not @p_112) (not @p_102) @p_113) :rule equiv_pos2) (step t39 (cl @p_113) :rule th_resolution :premises (axiom36 t37 t38)) (anchor :step t40 :args ((:= (veriT_vr34 A_set$) veriT_vr36) (:= (veriT_vr35 A_b_fun$) veriT_vr37))) (step t40.t1 (cl (! (= veriT_vr34 veriT_vr36) :named @p_115)) :rule refl) (step t40.t2 (cl (= @p_104 (! (finite$a veriT_vr36) :named @p_114))) :rule cong :premises (t40.t1)) (step t40.t3 (cl (= veriT_vr35 veriT_vr37)) :rule refl) (step t40.t4 (cl @p_115) :rule refl) (step t40.t5 (cl (= @p_107 (! (image$b veriT_vr37 veriT_vr36) :named @p_116))) :rule cong :premises (t40.t3 t40.t4)) (step t40.t6 (cl (= @p_109 (! (finite$ @p_116) :named @p_117))) :rule cong :premises (t40.t5)) (step t40.t7 (cl (= @p_111 (! (=> @p_114 @p_117) :named @p_118))) :rule cong :premises (t40.t2 t40.t6)) (step t40 (cl (! (= @p_113 (! (forall ((veriT_vr36 A_set$) (veriT_vr37 A_b_fun$)) @p_118) :named @p_120)) :named @p_119)) :rule bind) (step t41 (cl (not @p_119) (not @p_113) @p_120) :rule equiv_pos2) (step t42 (cl @p_120) :rule th_resolution :premises (t39 t40 t41)) (anchor :step t43 :args ((:= (?v0 B$) veriT_vr58) (:= (?v1 A_b_fun$) veriT_vr59) (:= (?v2 A_set$) veriT_vr60))) (step t43.t1 (cl (! (= ?v0 veriT_vr58) :named @p_128)) :rule refl) (step t43.t2 (cl (! (= ?v1 veriT_vr59) :named @p_129)) :rule refl) (step t43.t3 (cl (! (= ?v2 veriT_vr60) :named @p_135)) :rule refl) (step t43.t4 (cl (= @p_122 (! (image$b veriT_vr59 veriT_vr60) :named @p_123))) :rule cong :premises (t43.t2 t43.t3)) (step t43.t5 (cl (= @p_124 (! (member$ veriT_vr58 @p_123) :named @p_125))) :rule cong :premises (t43.t1 t43.t4)) (anchor :step t43.t6 :args ((:= (?v3 A$) veriT_vr61))) (step t43.t6.t1 (cl @p_128) :rule refl) (step t43.t6.t2 (cl @p_129) :rule refl) (step t43.t6.t3 (cl (! (= ?v3 veriT_vr61) :named @p_134)) :rule refl) (step t43.t6.t4 (cl (= @p_130 (! (fun_app$b veriT_vr59 veriT_vr61) :named @p_131))) :rule cong :premises (t43.t6.t2 t43.t6.t3)) (step t43.t6.t5 (cl (= @p_132 (! (= veriT_vr58 @p_131) :named @p_133))) :rule cong :premises (t43.t6.t1 t43.t6.t4)) (step t43.t6.t6 (cl @p_134) :rule refl) (step t43.t6.t7 (cl @p_135) :rule refl) (step t43.t6.t8 (cl (= @p_136 (! (member$a veriT_vr61 veriT_vr60) :named @p_137))) :rule cong :premises (t43.t6.t6 t43.t6.t7)) (step t43.t6.t9 (cl (= @p_138 (! (and @p_133 @p_137) :named @p_139))) :rule cong :premises (t43.t6.t5 t43.t6.t8)) (step t43.t6.t10 (cl (= @p_140 (! (=> @p_139 false) :named @p_141))) :rule cong :premises (t43.t6.t9)) (step t43.t6 (cl (= @p_126 (! (forall ((veriT_vr61 A$)) @p_141) :named @p_127))) :rule bind) (step t43.t7 (cl (= @p_142 (! (and @p_125 @p_127) :named @p_143))) :rule cong :premises (t43.t5 t43.t6)) (step t43.t8 (cl (= @p_144 (! (=> @p_143 false) :named @p_145))) :rule cong :premises (t43.t7)) (step t43 (cl (! (= @p_121 (! (forall ((veriT_vr58 B$) (veriT_vr59 A_b_fun$) (veriT_vr60 A_set$)) @p_145) :named @p_147)) :named @p_146)) :rule bind) (step t44 (cl (not @p_146) (not @p_121) @p_147) :rule equiv_pos2) (step t45 (cl @p_147) :rule th_resolution :premises (axiom40 t43 t44)) (anchor :step t46 :args ((veriT_vr58 B$) (veriT_vr59 A_b_fun$) (veriT_vr60 A_set$))) (anchor :step t46.t1 :args ((veriT_vr61 A$))) (step t46.t1.t1 (cl (= @p_141 (! (not @p_139) :named @p_149))) :rule implies_simplify) (step t46.t1 (cl (= @p_127 (! (forall ((veriT_vr61 A$)) @p_149) :named @p_148))) :rule bind) (step t46.t2 (cl (= @p_143 (! (and @p_125 @p_148) :named @p_150))) :rule cong :premises (t46.t1)) (step t46.t3 (cl (= @p_145 (! (=> @p_150 false) :named @p_151))) :rule cong :premises (t46.t2)) (step t46.t4 (cl (= @p_151 (! (not @p_150) :named @p_152))) :rule implies_simplify) (step t46.t5 (cl (= @p_145 @p_152)) :rule trans :premises (t46.t3 t46.t4)) (step t46 (cl (! (= @p_147 (! (forall ((veriT_vr58 B$) (veriT_vr59 A_b_fun$) (veriT_vr60 A_set$)) @p_152) :named @p_154)) :named @p_153)) :rule bind) (step t47 (cl (not @p_153) (not @p_147) @p_154) :rule equiv_pos2) (step t48 (cl @p_154) :rule th_resolution :premises (t45 t46 t47)) (anchor :step t49 :args ((:= (veriT_vr58 B$) veriT_vr62) (:= (veriT_vr59 A_b_fun$) veriT_vr63) (:= (veriT_vr60 A_set$) veriT_vr64))) (step t49.t1 (cl (! (= veriT_vr58 veriT_vr62) :named @p_158)) :rule refl) (step t49.t2 (cl (! (= veriT_vr59 veriT_vr63) :named @p_159)) :rule refl) (step t49.t3 (cl (! (= veriT_vr60 veriT_vr64) :named @p_163)) :rule refl) (step t49.t4 (cl (= @p_123 (! (image$b veriT_vr63 veriT_vr64) :named @p_155))) :rule cong :premises (t49.t2 t49.t3)) (step t49.t5 (cl (= @p_125 (! (member$ veriT_vr62 @p_155) :named @p_156))) :rule cong :premises (t49.t1 t49.t4)) (anchor :step t49.t6 :args ((:= (veriT_vr61 A$) veriT_vr65))) (step t49.t6.t1 (cl @p_158) :rule refl) (step t49.t6.t2 (cl @p_159) :rule refl) (step t49.t6.t3 (cl (! (= veriT_vr61 veriT_vr65) :named @p_162)) :rule refl) (step t49.t6.t4 (cl (= @p_131 (! (fun_app$b veriT_vr63 veriT_vr65) :named @p_160))) :rule cong :premises (t49.t6.t2 t49.t6.t3)) (step t49.t6.t5 (cl (= @p_133 (! (= veriT_vr62 @p_160) :named @p_161))) :rule cong :premises (t49.t6.t1 t49.t6.t4)) (step t49.t6.t6 (cl @p_162) :rule refl) (step t49.t6.t7 (cl @p_163) :rule refl) (step t49.t6.t8 (cl (= @p_137 (! (member$a veriT_vr65 veriT_vr64) :named @p_164))) :rule cong :premises (t49.t6.t6 t49.t6.t7)) (step t49.t6.t9 (cl (= @p_139 (! (and @p_161 @p_164) :named @p_165))) :rule cong :premises (t49.t6.t5 t49.t6.t8)) (step t49.t6.t10 (cl (= @p_149 (! (not @p_165) :named @p_166))) :rule cong :premises (t49.t6.t9)) (step t49.t6 (cl (= @p_148 (! (forall ((veriT_vr65 A$)) @p_166) :named @p_157))) :rule bind) (step t49.t7 (cl (= @p_150 (! (and @p_156 @p_157) :named @p_167))) :rule cong :premises (t49.t5 t49.t6)) (step t49.t8 (cl (= @p_152 (! (not @p_167) :named @p_168))) :rule cong :premises (t49.t7)) (step t49 (cl (! (= @p_154 (! (forall ((veriT_vr62 B$) (veriT_vr63 A_b_fun$) (veriT_vr64 A_set$)) @p_168) :named @p_170)) :named @p_169)) :rule bind) (step t50 (cl (not @p_169) (not @p_154) @p_170) :rule equiv_pos2) (step t51 (cl @p_170) :rule th_resolution :premises (t48 t49 t50)) (anchor :step t52 :args ((:= (?v0 B$) veriT_vr90) (:= (?v1 A_b_fun$) veriT_vr91) (:= (?v2 A$) veriT_vr92) (:= (?v3 A_set$) veriT_vr93))) (step t52.t1 (cl (! (= ?v1 veriT_vr91) :named @p_181)) :rule refl) (step t52.t2 (cl (! (= ?v2 veriT_vr92) :named @p_175)) :rule refl) (step t52.t3 (cl (= @p_3 (! (fun_app$b veriT_vr91 veriT_vr92) :named @p_172))) :rule cong :premises (t52.t1 t52.t2)) (step t52.t4 (cl (! (= ?v0 veriT_vr90) :named @p_180)) :rule refl) (step t52.t5 (cl (= @p_173 (! (= veriT_vr90 @p_172) :named @p_174))) :rule cong :premises (t52.t3 t52.t4)) (step t52.t6 (cl @p_175) :rule refl) (step t52.t7 (cl (! (= ?v3 veriT_vr93) :named @p_182)) :rule refl) (step t52.t8 (cl (= @p_176 (! (member$a veriT_vr92 veriT_vr93) :named @p_177))) :rule cong :premises (t52.t6 t52.t7)) (step t52.t9 (cl (= @p_178 (! (and @p_174 @p_177) :named @p_179))) :rule cong :premises (t52.t5 t52.t8)) (step t52.t10 (cl @p_180) :rule refl) (step t52.t11 (cl @p_181) :rule refl) (step t52.t12 (cl @p_182) :rule refl) (step t52.t13 (cl (= @p_183 (! (image$b veriT_vr91 veriT_vr93) :named @p_184))) :rule cong :premises (t52.t11 t52.t12)) (step t52.t14 (cl (= @p_185 (! (member$ veriT_vr90 @p_184) :named @p_186))) :rule cong :premises (t52.t10 t52.t13)) (step t52.t15 (cl (= @p_187 (! (=> @p_179 @p_186) :named @p_188))) :rule cong :premises (t52.t9 t52.t14)) (step t52 (cl (! (= @p_171 (! (forall ((veriT_vr90 B$) (veriT_vr91 A_b_fun$) (veriT_vr92 A$) (veriT_vr93 A_set$)) @p_188) :named @p_190)) :named @p_189)) :rule bind) (step t53 (cl (not @p_189) (not @p_171) @p_190) :rule equiv_pos2) (step t54 (cl @p_190) :rule th_resolution :premises (axiom44 t52 t53)) (anchor :step t55 :args ((:= (veriT_vr90 B$) veriT_vr94) (:= (veriT_vr91 A_b_fun$) veriT_vr95) (:= (veriT_vr92 A$) veriT_vr96) (:= (veriT_vr93 A_set$) veriT_vr97))) (step t55.t1 (cl (! (= veriT_vr90 veriT_vr94) :named @p_196)) :rule refl) (step t55.t2 (cl (! (= veriT_vr91 veriT_vr95) :named @p_197)) :rule refl) (step t55.t3 (cl (! (= veriT_vr92 veriT_vr96) :named @p_193)) :rule refl) (step t55.t4 (cl (= @p_172 (! (fun_app$b veriT_vr95 veriT_vr96) :named @p_191))) :rule cong :premises (t55.t2 t55.t3)) (step t55.t5 (cl (= @p_174 (! (= veriT_vr94 @p_191) :named @p_192))) :rule cong :premises (t55.t1 t55.t4)) (step t55.t6 (cl @p_193) :rule refl) (step t55.t7 (cl (! (= veriT_vr93 veriT_vr97) :named @p_198)) :rule refl) (step t55.t8 (cl (= @p_177 (! (member$a veriT_vr96 veriT_vr97) :named @p_194))) :rule cong :premises (t55.t6 t55.t7)) (step t55.t9 (cl (= @p_179 (! (and @p_192 @p_194) :named @p_195))) :rule cong :premises (t55.t5 t55.t8)) (step t55.t10 (cl @p_196) :rule refl) (step t55.t11 (cl @p_197) :rule refl) (step t55.t12 (cl @p_198) :rule refl) (step t55.t13 (cl (= @p_184 (! (image$b veriT_vr95 veriT_vr97) :named @p_199))) :rule cong :premises (t55.t11 t55.t12)) (step t55.t14 (cl (= @p_186 (! (member$ veriT_vr94 @p_199) :named @p_200))) :rule cong :premises (t55.t10 t55.t13)) (step t55.t15 (cl (= @p_188 (! (=> @p_195 @p_200) :named @p_201))) :rule cong :premises (t55.t9 t55.t14)) (step t55 (cl (! (= @p_190 (! (forall ((veriT_vr94 B$) (veriT_vr95 A_b_fun$) (veriT_vr96 A$) (veriT_vr97 A_set$)) @p_201) :named @p_203)) :named @p_202)) :rule bind) (step t56 (cl (not @p_202) (not @p_190) @p_203) :rule equiv_pos2) (step t57 (cl @p_203) :rule th_resolution :premises (t54 t55 t56)) (anchor :step t58 :args ((:= (?v0 A_b_fun$) veriT_vr114) (:= (?v1 A_set$) veriT_vr115))) (step t58.t1 (cl (= ?v0 veriT_vr114)) :rule refl) (step t58.t2 (cl (! (= ?v1 veriT_vr115) :named @p_209)) :rule refl) (step t58.t3 (cl (= @p_205 (! (image$b veriT_vr114 veriT_vr115) :named @p_206))) :rule cong :premises (t58.t1 t58.t2)) (step t58.t4 (cl (= @p_207 (! (= bot$ @p_206) :named @p_208))) :rule cong :premises (t58.t3)) (step t58.t5 (cl @p_209) :rule refl) (step t58.t6 (cl (= @p_210 (! (= bot$a veriT_vr115) :named @p_211))) :rule cong :premises (t58.t5)) (step t58.t7 (cl (= @p_212 (! (= @p_208 @p_211) :named @p_213))) :rule cong :premises (t58.t4 t58.t6)) (step t58 (cl (! (= @p_204 (! (forall ((veriT_vr114 A_b_fun$) (veriT_vr115 A_set$)) @p_213) :named @p_215)) :named @p_214)) :rule bind) (step t59 (cl (not @p_214) (not @p_204) @p_215) :rule equiv_pos2) (step t60 (cl @p_215) :rule th_resolution :premises (axiom48 t58 t59)) (anchor :step t61 :args ((:= (veriT_vr114 A_b_fun$) veriT_vr116) (:= (veriT_vr115 A_set$) veriT_vr117))) (step t61.t1 (cl (= veriT_vr114 veriT_vr116)) :rule refl) (step t61.t2 (cl (! (= veriT_vr115 veriT_vr117) :named @p_218)) :rule refl) (step t61.t3 (cl (= @p_206 (! (image$b veriT_vr116 veriT_vr117) :named @p_216))) :rule cong :premises (t61.t1 t61.t2)) (step t61.t4 (cl (= @p_208 (! (= bot$ @p_216) :named @p_217))) :rule cong :premises (t61.t3)) (step t61.t5 (cl @p_218) :rule refl) (step t61.t6 (cl (= @p_211 (! (= bot$a veriT_vr117) :named @p_219))) :rule cong :premises (t61.t5)) (step t61.t7 (cl (= @p_213 (! (= @p_217 @p_219) :named @p_220))) :rule cong :premises (t61.t4 t61.t6)) (step t61 (cl (! (= @p_215 (! (forall ((veriT_vr116 A_b_fun$) (veriT_vr117 A_set$)) @p_220) :named @p_222)) :named @p_221)) :rule bind) (step t62 (cl (not @p_221) (not @p_215) @p_222) :rule equiv_pos2) (step t63 (cl @p_222) :rule th_resolution :premises (t60 t61 t62)) (anchor :step t64 :args ((:= (?v0 B_c_fun$) veriT_vr122) (:= (?v1 B_set$) veriT_vr123) (:= (?v2 B$) veriT_vr124) (:= (?v3 B$) veriT_vr125))) (step t64.t1 (cl (! (= ?v0 veriT_vr122) :named @p_226)) :rule refl) (step t64.t2 (cl (! (= ?v1 veriT_vr123) :named @p_234)) :rule refl) (step t64.t3 (cl (= @p_224 (! (inj_on$ veriT_vr122 veriT_vr123) :named @p_225))) :rule cong :premises (t64.t1 t64.t2)) (step t64.t4 (cl @p_226) :rule refl) (step t64.t5 (cl (! (= ?v2 veriT_vr124) :named @p_233)) :rule refl) (step t64.t6 (cl (= @p_227 (! (fun_app$ veriT_vr122 veriT_vr124) :named @p_228))) :rule cong :premises (t64.t4 t64.t5)) (step t64.t7 (cl @p_226) :rule refl) (step t64.t8 (cl (! (= ?v3 veriT_vr125) :named @p_237)) :rule refl) (step t64.t9 (cl (= @p_229 (! (fun_app$ veriT_vr122 veriT_vr125) :named @p_230))) :rule cong :premises (t64.t7 t64.t8)) (step t64.t10 (cl (= @p_231 (! (= @p_228 @p_230) :named @p_232))) :rule cong :premises (t64.t6 t64.t9)) (step t64.t11 (cl @p_233) :rule refl) (step t64.t12 (cl @p_234) :rule refl) (step t64.t13 (cl (= @p_235 (! (member$ veriT_vr124 veriT_vr123) :named @p_236))) :rule cong :premises (t64.t11 t64.t12)) (step t64.t14 (cl @p_237) :rule refl) (step t64.t15 (cl @p_234) :rule refl) (step t64.t16 (cl (= @p_238 (! (member$ veriT_vr125 veriT_vr123) :named @p_239))) :rule cong :premises (t64.t14 t64.t15)) (step t64.t17 (cl (= @p_240 (! (and @p_236 @p_239) :named @p_241))) :rule cong :premises (t64.t13 t64.t16)) (step t64.t18 (cl (= @p_242 (! (and @p_232 @p_241) :named @p_243))) :rule cong :premises (t64.t10 t64.t17)) (step t64.t19 (cl (= @p_244 (! (and @p_225 @p_243) :named @p_245))) :rule cong :premises (t64.t3 t64.t18)) (step t64.t20 (cl @p_237) :rule refl) (step t64.t21 (cl @p_233) :rule refl) (step t64.t22 (cl (= @p_246 (! (= veriT_vr124 veriT_vr125) :named @p_247))) :rule cong :premises (t64.t20 t64.t21)) (step t64.t23 (cl (= @p_248 (! (=> @p_245 @p_247) :named @p_249))) :rule cong :premises (t64.t19 t64.t22)) (step t64 (cl (! (= @p_223 (! (forall ((veriT_vr122 B_c_fun$) (veriT_vr123 B_set$) (veriT_vr124 B$) (veriT_vr125 B$)) @p_249) :named @p_251)) :named @p_250)) :rule bind) (step t65 (cl (not @p_250) (not @p_223) @p_251) :rule equiv_pos2) (step t66 (cl @p_251) :rule th_resolution :premises (axiom50 t64 t65)) (anchor :step t67 :args ((veriT_vr122 B_c_fun$) (veriT_vr123 B_set$) (veriT_vr124 B$) (veriT_vr125 B$))) (step t67.t1 (cl (= @p_245 (! (and @p_225 @p_232 @p_236 @p_239) :named @p_252))) :rule ac_simp) (step t67.t2 (cl (= @p_249 (! (=> @p_252 @p_247) :named @p_253))) :rule cong :premises (t67.t1)) (step t67 (cl (! (= @p_251 (! (forall ((veriT_vr122 B_c_fun$) (veriT_vr123 B_set$) (veriT_vr124 B$) (veriT_vr125 B$)) @p_253) :named @p_255)) :named @p_254)) :rule bind) (step t68 (cl (not @p_254) (not @p_251) @p_255) :rule equiv_pos2) (step t69 (cl @p_255) :rule th_resolution :premises (t66 t67 t68)) (anchor :step t70 :args ((:= (veriT_vr122 B_c_fun$) veriT_vr126) (:= (veriT_vr123 B_set$) veriT_vr127) (:= (veriT_vr124 B$) veriT_vr128) (:= (veriT_vr125 B$) veriT_vr129))) (step t70.t1 (cl (! (= veriT_vr122 veriT_vr126) :named @p_257)) :rule refl) (step t70.t2 (cl (! (= veriT_vr123 veriT_vr127) :named @p_262)) :rule refl) (step t70.t3 (cl (= @p_225 (! (inj_on$ veriT_vr126 veriT_vr127) :named @p_256))) :rule cong :premises (t70.t1 t70.t2)) (step t70.t4 (cl @p_257) :rule refl) (step t70.t5 (cl (! (= veriT_vr124 veriT_vr128) :named @p_261)) :rule refl) (step t70.t6 (cl (= @p_228 (! (fun_app$ veriT_vr126 veriT_vr128) :named @p_258))) :rule cong :premises (t70.t4 t70.t5)) (step t70.t7 (cl @p_257) :rule refl) (step t70.t8 (cl (! (= veriT_vr125 veriT_vr129) :named @p_264)) :rule refl) (step t70.t9 (cl (= @p_230 (! (fun_app$ veriT_vr126 veriT_vr129) :named @p_259))) :rule cong :premises (t70.t7 t70.t8)) (step t70.t10 (cl (= @p_232 (! (= @p_258 @p_259) :named @p_260))) :rule cong :premises (t70.t6 t70.t9)) (step t70.t11 (cl @p_261) :rule refl) (step t70.t12 (cl @p_262) :rule refl) (step t70.t13 (cl (= @p_236 (! (member$ veriT_vr128 veriT_vr127) :named @p_263))) :rule cong :premises (t70.t11 t70.t12)) (step t70.t14 (cl @p_264) :rule refl) (step t70.t15 (cl @p_262) :rule refl) (step t70.t16 (cl (= @p_239 (! (member$ veriT_vr129 veriT_vr127) :named @p_265))) :rule cong :premises (t70.t14 t70.t15)) (step t70.t17 (cl (= @p_252 (! (and @p_256 @p_260 @p_263 @p_265) :named @p_266))) :rule cong :premises (t70.t3 t70.t10 t70.t13 t70.t16)) (step t70.t18 (cl @p_261) :rule refl) (step t70.t19 (cl @p_264) :rule refl) (step t70.t20 (cl (= @p_247 (! (= veriT_vr128 veriT_vr129) :named @p_267))) :rule cong :premises (t70.t18 t70.t19)) (step t70.t21 (cl (= @p_253 (! (=> @p_266 @p_267) :named @p_268))) :rule cong :premises (t70.t17 t70.t20)) (step t70 (cl (! (= @p_255 (! (forall ((veriT_vr126 B_c_fun$) (veriT_vr127 B_set$) (veriT_vr128 B$) (veriT_vr129 B$)) @p_268) :named @p_270)) :named @p_269)) :rule bind) (step t71 (cl (not @p_269) (not @p_255) @p_270) :rule equiv_pos2) (step t72 (cl @p_270) :rule th_resolution :premises (t69 t70 t71)) (anchor :step t73 :args ((:= (?v0 C$) veriT_vr130) (:= (?v1 C$) veriT_vr131))) (step t73.t1 (cl (! (= ?v0 veriT_vr130) :named @p_274)) :rule refl) (step t73.t2 (cl (! (= ?v1 veriT_vr131) :named @p_275)) :rule refl) (step t73.t3 (cl (= @p_272 (! (less$ veriT_vr130 veriT_vr131) :named @p_273))) :rule cong :premises (t73.t1 t73.t2)) (step t73.t4 (cl @p_274) :rule refl) (step t73.t5 (cl @p_275) :rule refl) (step t73.t6 (cl (= @p_276 (! (less_eq$ veriT_vr130 veriT_vr131) :named @p_277))) :rule cong :premises (t73.t4 t73.t5)) (step t73.t7 (cl @p_274) :rule refl) (step t73.t8 (cl @p_275) :rule refl) (step t73.t9 (cl (= @p_278 (! (= veriT_vr130 veriT_vr131) :named @p_279))) :rule cong :premises (t73.t7 t73.t8)) (step t73.t10 (cl (= @p_280 (! (not @p_279) :named @p_281))) :rule cong :premises (t73.t9)) (step t73.t11 (cl (= @p_282 (! (and @p_277 @p_281) :named @p_283))) :rule cong :premises (t73.t6 t73.t10)) (step t73.t12 (cl (= @p_284 (! (= @p_273 @p_283) :named @p_285))) :rule cong :premises (t73.t3 t73.t11)) (step t73 (cl (! (= @p_271 (! (forall ((veriT_vr130 C$) (veriT_vr131 C$)) @p_285) :named @p_287)) :named @p_286)) :rule bind) (step t74 (cl (not @p_286) (not @p_271) @p_287) :rule equiv_pos2) (step t75 (cl @p_287) :rule th_resolution :premises (axiom51 t73 t74)) (anchor :step t76 :args ((:= (veriT_vr130 C$) veriT_vr132) (:= (veriT_vr131 C$) veriT_vr133))) (step t76.t1 (cl (! (= veriT_vr130 veriT_vr132) :named @p_289)) :rule refl) (step t76.t2 (cl (! (= veriT_vr131 veriT_vr133) :named @p_290)) :rule refl) (step t76.t3 (cl (= @p_273 (! (less$ veriT_vr132 veriT_vr133) :named @p_288))) :rule cong :premises (t76.t1 t76.t2)) (step t76.t4 (cl @p_289) :rule refl) (step t76.t5 (cl @p_290) :rule refl) (step t76.t6 (cl (= @p_277 (! (less_eq$ veriT_vr132 veriT_vr133) :named @p_291))) :rule cong :premises (t76.t4 t76.t5)) (step t76.t7 (cl @p_289) :rule refl) (step t76.t8 (cl @p_290) :rule refl) (step t76.t9 (cl (= @p_279 (! (= veriT_vr132 veriT_vr133) :named @p_292))) :rule cong :premises (t76.t7 t76.t8)) (step t76.t10 (cl (= @p_281 (! (not @p_292) :named @p_293))) :rule cong :premises (t76.t9)) (step t76.t11 (cl (= @p_283 (! (and @p_291 @p_293) :named @p_294))) :rule cong :premises (t76.t6 t76.t10)) (step t76.t12 (cl (= @p_285 (! (= @p_288 @p_294) :named @p_295))) :rule cong :premises (t76.t3 t76.t11)) (step t76 (cl (! (= @p_287 (! (forall ((veriT_vr132 C$) (veriT_vr133 C$)) @p_295) :named @p_297)) :named @p_296)) :rule bind) (step t77 (cl (not @p_296) (not @p_287) @p_297) :rule equiv_pos2) (step t78 (cl @p_297) :rule th_resolution :premises (t75 t76 t77)) (anchor :step t79 :args ((:= (?v0 A$) veriT_vr134))) (step t79.t1 (cl (! (= ?v0 veriT_vr134) :named @p_302)) :rule refl) (step t79.t2 (cl (= @p_300 (! (member$a veriT_vr134 b$) :named @p_301))) :rule cong :premises (t79.t1)) (step t79.t3 (cl @p_302) :rule refl) (step t79.t4 (cl (= @p_303 (! (fun_app$a @p_4 veriT_vr134) :named @p_304))) :rule cong :premises (t79.t3)) (step t79.t5 (cl (= @p_305 (! (less$ @p_304 @p_299) :named @p_306))) :rule cong :premises (t79.t4)) (step t79.t6 (cl (= @p_307 (! (and @p_301 @p_306) :named @p_308))) :rule cong :premises (t79.t2 t79.t5)) (step t79 (cl (= @p_298 (! (exists ((veriT_vr134 A$)) @p_308) :named @p_310))) :rule bind) (step t80 (cl (! (= @p_309 (! (not @p_310) :named @p_312)) :named @p_311)) :rule cong :premises (t79)) (step t81 (cl (! (not @p_311) :named @p_314) (! (not @p_309) :named @p_313) @p_312) :rule equiv_pos2) (step t82 (cl (not @p_313) @p_298) :rule not_not) (step t83 (cl @p_314 @p_298 @p_312) :rule th_resolution :premises (t82 t81)) (step t84 (cl @p_312) :rule th_resolution :premises (axiom27 t80 t83)) (anchor :step t85 :args ((:= (veriT_vr134 A$) veriT_vr135))) (step t85.t1 (cl (! (= veriT_vr134 veriT_vr135) :named @p_316)) :rule refl) (step t85.t2 (cl (= @p_301 (! (member$a veriT_vr135 b$) :named @p_315))) :rule cong :premises (t85.t1)) (step t85.t3 (cl @p_316) :rule refl) (step t85.t4 (cl (= @p_304 (! (fun_app$a @p_4 veriT_vr135) :named @p_317))) :rule cong :premises (t85.t3)) (step t85.t5 (cl (= @p_306 (! (less$ @p_317 @p_299) :named @p_318))) :rule cong :premises (t85.t4)) (step t85.t6 (cl (= @p_308 (! (and @p_315 @p_318) :named @p_319))) :rule cong :premises (t85.t2 t85.t5)) (step t85 (cl (= @p_310 (! (exists ((veriT_vr135 A$)) @p_319) :named @p_320))) :rule bind) (step t86 (cl (! (= @p_312 (! (not @p_320) :named @p_322)) :named @p_321)) :rule cong :premises (t85)) (step t87 (cl (! (not @p_321) :named @p_324) (! (not @p_312) :named @p_323) @p_322) :rule equiv_pos2) (step t88 (cl (not @p_323) @p_310) :rule not_not) (step t89 (cl @p_324 @p_310 @p_322) :rule th_resolution :premises (t88 t87)) (step t90 (cl @p_322) :rule th_resolution :premises (t84 t86 t89)) (step t91 (cl (= @p_320 (! (not (! (forall ((veriT_vr135 A$)) (not @p_319)) :named @p_330)) :named @p_325))) :rule connective_def) (step t92 (cl (! (= @p_322 (! (not @p_325) :named @p_327)) :named @p_326)) :rule cong :premises (t91)) (step t93 (cl (! (not @p_326) :named @p_329) (! (not @p_322) :named @p_328) @p_327) :rule equiv_pos2) (step t94 (cl (not @p_328) @p_320) :rule not_not) (step t95 (cl @p_329 @p_320 @p_327) :rule th_resolution :premises (t94 t93)) (step t96 (cl (not @p_327) @p_330) :rule not_not) (step t97 (cl @p_329 @p_320 @p_330) :rule th_resolution :premises (t96 t95)) (step t98 (cl @p_327) :rule th_resolution :premises (t90 t92 t97)) (step t99 (cl @p_330) :rule th_resolution :premises (t96 t98)) (step t100 (cl (or (! (not @p_203) :named @p_421) (! (forall ((veriT_vr94 B$) (veriT_vr95 A_b_fun$) (veriT_vr96 A$) (veriT_vr97 A_set$)) (or (not @p_192) (not @p_194) @p_200)) :named @p_422))) :rule qnt_cnf) (step t101 (cl (or (! (not @p_222) :named @p_339) (! (= (! (= bot$ @p_6) :named @p_335) @p_331) :named @p_337))) :rule forall_inst :args ((:= veriT_vr116 g$) (:= veriT_vr117 b$))) (step t102 (cl (or (! (not @p_120) :named @p_342) (! (=> @p_332 (! (finite$ @p_6) :named @p_334)) :named @p_341))) :rule forall_inst :args ((:= veriT_vr36 b$) (:= veriT_vr37 g$))) (step t103 (cl (or (! (not @p_101) :named @p_344) (! (= @p_299 (! (fun_app$ f$ @p_333) :named @p_354)) :named @p_345))) :rule forall_inst :args ((:= veriT_vr23 f$) (:= veriT_vr24 g$) (:= veriT_vr25 @p_5))) (step t104 (cl (or (! (not @p_32) :named @p_351) (! (=> (! (and @p_334 (! (not @p_335) :named @p_338)) :named @p_346) (! (member$ @p_336 @p_6) :named @p_350)) :named @p_349))) :rule forall_inst :args ((:= veriT_vr2 @p_6) (:= veriT_vr3 f$))) (step t105 (cl (! (not @p_337) :named @p_340) @p_338 @p_331) :rule equiv_pos2) (step t106 (cl @p_339 @p_337) :rule or :premises (t101)) (step t107 (cl @p_340 @p_338) :rule resolution :premises (t105 axiom25)) (step t108 (cl @p_337) :rule resolution :premises (t106 t63)) (step t109 (cl @p_338) :rule resolution :premises (t107 t108)) (step t110 (cl (! (not @p_341) :named @p_343) (not @p_332) @p_334) :rule implies_pos) (step t111 (cl @p_342 @p_341) :rule or :premises (t102)) (step t112 (cl @p_343 @p_334) :rule resolution :premises (t110 axiom24)) (step t113 (cl @p_341) :rule resolution :premises (t111 t42)) (step t114 (cl @p_334) :rule resolution :premises (t112 t113)) (step t115 (cl @p_344 @p_345) :rule or :premises (t103)) (step t116 (cl @p_345) :rule resolution :premises (t115 t36)) (step t117 (cl @p_346 (! (not @p_334) :named @p_348) (! (not @p_338) :named @p_347)) :rule and_neg) (step t118 (cl (not @p_347) @p_335) :rule not_not) (step t119 (cl @p_346 @p_348 @p_335) :rule th_resolution :premises (t118 t117)) (step t120 (cl (! (not @p_349) :named @p_352) (not @p_346) @p_350) :rule implies_pos) (step t121 (cl @p_351 @p_349) :rule or :premises (t104)) (step t122 (cl @p_346) :rule resolution :premises (t119 t109 t114)) (step t123 (cl @p_352 @p_350) :rule resolution :premises (t120 t122)) (step t124 (cl @p_349) :rule resolution :premises (t121 t21)) (step t125 (cl @p_350) :rule resolution :premises (t123 t124)) (step t126 (cl (or (! (not @p_270) :named @p_410) (! (=> (! (and @p_353 (! (= @p_354 (! (fun_app$ f$ @p_336) :named @p_406)) :named @p_408) @p_350 (! (member$ @p_333 @p_6) :named @p_405)) :named @p_407) @p_355) :named @p_409))) :rule forall_inst :args ((:= veriT_vr126 f$) (:= veriT_vr127 @p_6) (:= veriT_vr128 @p_336) (:= veriT_vr129 @p_333))) (step t127 (cl (or (! (not @p_170) :named @p_401) (! (not (! (and @p_350 (! (forall ((veriT_vr65 A$)) (! (not (! (and (! (= @p_336 (! (fun_app$b g$ veriT_vr65) :named @p_359)) :named @p_361) (! (member$a veriT_vr65 b$) :named @p_364)) :named @p_366)) :named @p_368)) :named @p_358)) :named @p_370)) :named @p_356))) :rule forall_inst :args ((:= veriT_vr62 @p_336) (:= veriT_vr63 g$) (:= veriT_vr64 b$))) (anchor :step t128) (assume t128.h1 @p_356) (anchor :step t128.t2 :args ((:= (veriT_vr65 A$) veriT_vr144))) (step t128.t2.t1 (cl (! (= veriT_vr65 veriT_vr144) :named @p_363)) :rule refl) (step t128.t2.t2 (cl (= @p_359 (! (fun_app$b g$ veriT_vr144) :named @p_360))) :rule cong :premises (t128.t2.t1)) (step t128.t2.t3 (cl (= @p_361 (! (= @p_336 @p_360) :named @p_362))) :rule cong :premises (t128.t2.t2)) (step t128.t2.t4 (cl @p_363) :rule refl) (step t128.t2.t5 (cl (= @p_364 (! (member$a veriT_vr144 b$) :named @p_365))) :rule cong :premises (t128.t2.t4)) (step t128.t2.t6 (cl (= @p_366 (! (and @p_362 @p_365) :named @p_367))) :rule cong :premises (t128.t2.t3 t128.t2.t5)) (step t128.t2.t7 (cl (= @p_368 (! (not @p_367) :named @p_369))) :rule cong :premises (t128.t2.t6)) (step t128.t2 (cl (= @p_358 (! (forall ((veriT_vr144 A$)) @p_369) :named @p_371))) :rule bind) (step t128.t3 (cl (= @p_370 (! (and @p_350 @p_371) :named @p_372))) :rule cong :premises (t128.t2)) (step t128.t4 (cl (! (= @p_356 (! (not @p_372) :named @p_375)) :named @p_373)) :rule cong :premises (t128.t3)) (step t128.t5 (cl (! (not @p_373) :named @p_376) (! (not @p_356) :named @p_374) @p_375) :rule equiv_pos2) (step t128.t6 (cl (! (not @p_374) :named @p_400) @p_370) :rule not_not) (step t128.t7 (cl @p_376 @p_370 @p_375) :rule th_resolution :premises (t128.t6 t128.t5)) (step t128.t8 (cl @p_375) :rule th_resolution :premises (t128.h1 t128.t4 t128.t7)) (anchor :step t128.t9 :args ((:= (veriT_vr144 A$) veriT_vr145))) (step t128.t9.t1 (cl (! (= veriT_vr144 veriT_vr145) :named @p_380)) :rule refl) (step t128.t9.t2 (cl (= @p_360 @p_378)) :rule cong :premises (t128.t9.t1)) (step t128.t9.t3 (cl (= @p_362 @p_379)) :rule cong :premises (t128.t9.t2)) (step t128.t9.t4 (cl @p_380) :rule refl) (step t128.t9.t5 (cl (= @p_365 @p_381)) :rule cong :premises (t128.t9.t4)) (step t128.t9.t6 (cl (= @p_367 @p_382)) :rule cong :premises (t128.t9.t3 t128.t9.t5)) (step t128.t9.t7 (cl (= @p_369 @p_377)) :rule cong :premises (t128.t9.t6)) (step t128.t9 (cl (= @p_371 (! (forall ((veriT_vr145 A$)) @p_377) :named @p_383))) :rule bind) (step t128.t10 (cl (= @p_372 (! (and @p_350 @p_383) :named @p_384))) :rule cong :premises (t128.t9)) (step t128.t11 (cl (! (= @p_375 (! (not @p_384) :named @p_386)) :named @p_385)) :rule cong :premises (t128.t10)) (step t128.t12 (cl (! (not @p_385) :named @p_388) (! (not @p_375) :named @p_387) @p_386) :rule equiv_pos2) (step t128.t13 (cl (not @p_387) @p_372) :rule not_not) (step t128.t14 (cl @p_388 @p_372 @p_386) :rule th_resolution :premises (t128.t13 t128.t12)) (step t128.t15 (cl @p_386) :rule th_resolution :premises (t128.t8 t128.t11 t128.t14)) (anchor :step t128.t16 :args ((:= (veriT_vr145 A$) veriT_sk0))) (step t128.t16.t1 (cl (! (= veriT_vr145 veriT_sk0) :named @p_392)) :rule refl) (step t128.t16.t2 (cl (= @p_378 (! (fun_app$b g$ veriT_sk0) :named @p_390))) :rule cong :premises (t128.t16.t1)) (step t128.t16.t3 (cl (= @p_379 (! (= @p_336 @p_390) :named @p_391))) :rule cong :premises (t128.t16.t2)) (step t128.t16.t4 (cl @p_392) :rule refl) (step t128.t16.t5 (cl (= @p_381 (! (member$a veriT_sk0 b$) :named @p_393))) :rule cong :premises (t128.t16.t4)) (step t128.t16.t6 (cl (= @p_382 (! (and @p_391 @p_393) :named @p_394))) :rule cong :premises (t128.t16.t3 t128.t16.t5)) (step t128.t16.t7 (cl (= @p_377 (! (not @p_394) :named @p_389))) :rule cong :premises (t128.t16.t6)) (step t128.t16 (cl (= @p_383 @p_389)) :rule sko_forall) (step t128.t17 (cl (= @p_384 (! (and @p_350 @p_389) :named @p_395))) :rule cong :premises (t128.t16)) (step t128.t18 (cl (! (= @p_386 (! (not @p_395) :named @p_396)) :named @p_397)) :rule cong :premises (t128.t17)) (step t128.t19 (cl (! (not @p_397) :named @p_399) (! (not @p_386) :named @p_398) @p_396) :rule equiv_pos2) (step t128.t20 (cl (not @p_398) @p_384) :rule not_not) (step t128.t21 (cl @p_399 @p_384 @p_396) :rule th_resolution :premises (t128.t20 t128.t19)) (step t128.t22 (cl @p_396) :rule th_resolution :premises (t128.t15 t128.t18 t128.t21)) (step t128 (cl @p_374 @p_396) :rule subproof :discharge (h1)) (step t129 (cl @p_400 @p_370) :rule not_not) (step t130 (cl @p_370 @p_396) :rule th_resolution :premises (t129 t128)) (step t131 (cl @p_401 @p_356) :rule or :premises (t127)) (step t132 (cl (! (or @p_401 @p_396) :named @p_403) (! (not @p_401) :named @p_402)) :rule or_neg) (step t133 (cl (not @p_402) @p_170) :rule not_not) (step t134 (cl @p_403 @p_170) :rule th_resolution :premises (t133 t132)) (step t135 (cl @p_403 (! (not @p_396) :named @p_404)) :rule or_neg) (step t136 (cl (not @p_404) @p_395) :rule not_not) (step t137 (cl @p_403 @p_395) :rule th_resolution :premises (t136 t135)) (step t138 (cl @p_403) :rule th_resolution :premises (t131 t130 t134 t137)) (step t139 (cl (or (! (not @p_76) :named @p_420) (! (=> (! (and @p_334 @p_338 @p_405) :named @p_417) (! (less_eq$ @p_406 @p_354) :named @p_419)) :named @p_418))) :rule forall_inst :args ((:= veriT_vr11 @p_6) (:= veriT_vr12 @p_333) (:= veriT_vr13 f$))) (step t140 (cl @p_407 (not @p_353) (! (not @p_408) :named @p_411) (! (not @p_350) :named @p_415) (! (not @p_405) :named @p_412)) :rule and_neg) (step t141 (cl (! (not @p_409) :named @p_413) (! (not @p_407) :named @p_414) @p_355) :rule implies_pos) (step t142 (cl @p_410 @p_409) :rule or :premises (t126)) (step t143 (cl @p_407 @p_411 @p_412) :rule resolution :premises (t140 axiom23 t125)) (step t144 (cl @p_413 @p_414) :rule resolution :premises (t141 axiom52)) (step t145 (cl @p_409) :rule resolution :premises (t142 t72)) (step t146 (cl @p_414) :rule resolution :premises (t144 t145)) (step t147 (cl @p_389 @p_391) :rule and_pos) (step t148 (cl @p_389 @p_393) :rule and_pos) (step t149 (cl @p_395 @p_415 (! (not @p_389) :named @p_416)) :rule and_neg) (step t150 (cl (not @p_416) @p_394) :rule not_not) (step t151 (cl @p_395 @p_415 @p_394) :rule th_resolution :premises (t150 t149)) (step t152 (cl @p_401 @p_396) :rule or :premises (t138)) (step t153 (cl @p_395 @p_394) :rule resolution :premises (t151 t125)) (step t154 (cl @p_396) :rule resolution :premises (t152 t51)) (step t155 (cl @p_394) :rule resolution :premises (t153 t154)) (step t156 (cl @p_391) :rule resolution :premises (t147 t155)) (step t157 (cl @p_393) :rule resolution :premises (t148 t155)) (step t158 (cl @p_417 @p_348 @p_347 @p_412) :rule and_neg) (step t159 (cl @p_417 @p_348 @p_335 @p_412) :rule th_resolution :premises (t118 t158)) (step t160 (cl (not @p_418) (not @p_417) @p_419) :rule implies_pos) (step t161 (cl @p_420 @p_418) :rule or :premises (t139)) (step t162 (cl @p_417 @p_412) :rule resolution :premises (t159 t109 t114)) (step t163 (cl @p_418) :rule resolution :premises (t161 t30)) (step t164 (cl @p_421 @p_422) :rule or :premises (t100)) (step t165 (cl (or (! (not @p_422) :named @p_424) (! (or (! (not (! (= @p_333 @p_333) :named @p_430)) :named @p_431) (! (not @p_423) :named @p_429) @p_405) :named @p_425))) :rule forall_inst :args ((:= veriT_vr94 @p_333) (:= veriT_vr95 g$) (:= veriT_vr96 @p_5) (:= veriT_vr97 b$))) (step t166 (cl @p_424 @p_425) :rule or :premises (t165)) (step t167 (cl (! (or @p_421 @p_425) :named @p_427) (! (not @p_421) :named @p_426)) :rule or_neg) (step t168 (cl (not @p_426) @p_203) :rule not_not) (step t169 (cl @p_427 @p_203) :rule th_resolution :premises (t168 t167)) (step t170 (cl @p_427 (! (not @p_425) :named @p_428)) :rule or_neg) (step t171 (cl @p_427) :rule th_resolution :premises (t164 t166 t169 t170)) (anchor :step t172) (assume t172.h1 @p_425) (step t172.t2 (cl (= @p_430 true)) :rule eq_simplify) (step t172.t3 (cl (= @p_431 (! (not true) :named @p_432))) :rule cong :premises (t172.t2)) (step t172.t4 (cl (= @p_432 false)) :rule not_simplify) (step t172.t5 (cl (= @p_431 false)) :rule trans :premises (t172.t3 t172.t4)) (step t172.t6 (cl (= @p_425 (! (or false @p_429 @p_405) :named @p_433))) :rule cong :premises (t172.t5)) (step t172.t7 (cl (= @p_433 (! (or @p_429 @p_405) :named @p_434))) :rule or_simplify) (step t172.t8 (cl (! (= @p_425 @p_434) :named @p_435)) :rule trans :premises (t172.t6 t172.t7)) (step t172.t9 (cl (not @p_435) @p_428 @p_434) :rule equiv_pos2) (step t172.t10 (cl @p_434) :rule th_resolution :premises (t172.h1 t172.t8 t172.t9)) (step t172 (cl @p_428 @p_434) :rule subproof :discharge (h1)) (step t173 (cl @p_421 @p_425) :rule or :premises (t171)) (step t174 (cl (! (or @p_421 @p_434) :named @p_436) @p_426) :rule or_neg) (step t175 (cl @p_436 @p_203) :rule th_resolution :premises (t168 t174)) (step t176 (cl @p_436 (! (not @p_434) :named @p_437)) :rule or_neg) (step t177 (cl @p_436) :rule th_resolution :premises (t173 t172 t175 t176)) (step t178 (cl @p_437 @p_429 @p_405) :rule or_pos) (step t179 (cl @p_421 @p_434) :rule or :premises (t177)) (step t180 (cl @p_437 @p_405) :rule resolution :premises (t178 axiom26)) (step t181 (cl @p_434) :rule resolution :premises (t179 t57)) (step t182 (cl @p_405) :rule resolution :premises (t180 t181)) (step t183 (cl @p_411) :rule resolution :premises (t143 t182 t146)) (step t184 (cl @p_417) :rule resolution :premises (t162 t182)) (step t185 (cl @p_419) :rule resolution :premises (t160 t184 t163)) (step t186 (cl (or @p_325 (! (not (! (and @p_393 (! (less$ (! (fun_app$a @p_4 veriT_sk0) :named @p_438) @p_299) :named @p_440)) :named @p_439)) :named @p_441))) :rule forall_inst :args ((:= veriT_vr135 veriT_sk0))) (step t187 (cl (or (! (not @p_297) :named @p_448) (! (= (! (less$ @p_406 @p_354) :named @p_447) (! (and @p_419 @p_411) :named @p_443)) :named @p_446))) :rule forall_inst :args ((:= veriT_vr132 @p_406) (:= veriT_vr133 @p_354))) (step t188 (cl (or @p_344 (! (= @p_438 (! (fun_app$ f$ @p_390) :named @p_451)) :named @p_450))) :rule forall_inst :args ((:= veriT_vr23 f$) (:= veriT_vr24 g$) (:= veriT_vr25 veriT_sk0))) (step t189 (cl @p_439 (not @p_393) (! (not @p_440) :named @p_442)) :rule and_neg) (step t190 (cl @p_325 @p_441) :rule or :premises (t186)) (step t191 (cl @p_439 @p_442) :rule resolution :premises (t189 t157)) (step t192 (cl @p_441) :rule resolution :premises (t190 t99)) (step t193 (cl @p_442) :rule resolution :premises (t191 t192)) (step t194 (cl @p_443 (! (not @p_419) :named @p_445) (! (not @p_411) :named @p_444)) :rule and_neg) (step t195 (cl (not @p_444) @p_408) :rule not_not) (step t196 (cl @p_443 @p_445 @p_408) :rule th_resolution :premises (t195 t194)) (step t197 (cl (! (not @p_446) :named @p_449) @p_447 (not @p_443)) :rule equiv_pos1) (step t198 (cl @p_448 @p_446) :rule or :premises (t187)) (step t199 (cl @p_443) :rule resolution :premises (t196 t183 t185)) (step t200 (cl @p_449 @p_447) :rule resolution :premises (t197 t199)) (step t201 (cl @p_446) :rule resolution :premises (t198 t78)) (step t202 (cl @p_447) :rule resolution :premises (t200 t201)) (step t203 (cl @p_344 @p_450) :rule or :premises (t188)) (step t204 (cl @p_450) :rule resolution :premises (t203 t36)) (step t205 (cl (not (! (= @p_406 @p_438) :named @p_452)) (! (not @p_345) :named @p_457) (! (not @p_447) :named @p_458) @p_440) :rule eq_congruent_pred) (step t206 (cl (not (! (= @p_406 @p_451) :named @p_453)) (! (not @p_450) :named @p_456) @p_452) :rule eq_transitive) (step t207 (cl (not (! (= f$ f$) :named @p_454)) (! (not @p_391) :named @p_455) @p_453) :rule eq_congruent) (step t208 (cl @p_454) :rule eq_reflexive) (step t209 (cl @p_455 @p_453) :rule th_resolution :premises (t207 t208)) (step t210 (cl @p_456 @p_452 @p_455) :rule th_resolution :premises (t206 t209)) (step t211 (cl @p_457 @p_458 @p_440 @p_456 @p_455) :rule th_resolution :premises (t205 t210)) (step t212 (cl) :rule resolution :premises (t211 t116 t156 t193 t202 t204)) 2c004ebfd8457fdbede51bb75b1997f1f1e2bc6d 791 0 unsat (assume axiom0 (! (forall ((?v0 Real)) (! (= (! (fun_app$ uuc$ ?v0) :named @p_9) (! (pair$ (! (times$ (! (- ?v0 (! (divide$ 1.0 2.0) :named @p_7)) :named @p_12) d$) :named @p_1) (! (diamond_y$ @p_1) :named @p_16)) :named @p_18)) :named @p_20)) :named @p_6)) (assume axiom3 (! (forall ((?v0 Real)) (! (= (! (fun_app$ uub$ ?v0) :named @p_37) (! (pair$ (! (- (! (divide$ d$ 2.0) :named @p_3)) :named @p_2) (! (times$ (! (- (! (* 2.0 ?v0) :named @p_40) 1.0) :named @p_42) (! (diamond_y$ @p_2) :named @p_36)) :named @p_44)) :named @p_46)) :named @p_48)) :named @p_35)) (assume axiom4 (! (< 0.0 d$) :named @p_453)) (assume axiom5 (! (forall ((?v0 Real)) (! (= (! (diamond_y$ ?v0) :named @p_62) (! (- @p_3 (! (ite (! (< ?v0 0.0) :named @p_65) (! (- ?v0) :named @p_4) ?v0) :named @p_68)) :named @p_70)) :named @p_72)) :named @p_61)) (assume axiom7 (! (forall ((?v0 Real) (?v1 Real) (?v2 Real)) (! (= (! (< (! (divide$ ?v0 ?v1) :named @p_5) (! (divide$ ?v2 ?v1) :named @p_88)) :named @p_90) (! (and (! (=> (! (< 0.0 ?v1) :named @p_92) (! (< ?v0 ?v2) :named @p_96)) :named @p_98) (! (and (! (=> (! (< ?v1 0.0) :named @p_100) (! (< ?v2 ?v0) :named @p_102)) :named @p_104) (! (not (! (= 0.0 ?v1) :named @p_106)) :named @p_108)) :named @p_110)) :named @p_112)) :named @p_114)) :named @p_85)) (assume axiom8 (! (forall ((?v0 Real) (?v1 Real)) (! (= (! (divide$ @p_4 ?v1) :named @p_142) (! (- @p_5) :named @p_147)) :named @p_149)) :named @p_140)) (assume axiom9 (! (forall ((?v0 Real) (?v1 Real)) (! (= (! (times$ @p_4 ?v1) :named @p_164) (! (- (! (times$ ?v0 ?v1) :named @p_168)) :named @p_170)) :named @p_172)) :named @p_162)) (assume axiom10 (! (forall ((?v0 Real) (?v1 Real) (?v2 Real) (?v3 Real)) (! (= (! (= (! (pair$ ?v0 ?v1) :named @p_186) (! (pair$ ?v2 ?v3) :named @p_188)) :named @p_190) (! (and (! (= ?v0 ?v2) :named @p_194) (! (= ?v1 ?v3) :named @p_198)) :named @p_200)) :named @p_202)) :named @p_185)) (assume axiom11 (! (not (! (=> (! (and (! (not (= uua$ uu$)) :named @p_226) (! (= uuc$ uub$) :named @p_227)) :named @p_220) false) :named @p_224)) :named @p_219)) (anchor :step t10 :args ((:= (?v0 Real) veriT_vr0))) (step t10.t1 (cl (! (= ?v0 veriT_vr0) :named @p_11)) :rule refl) (step t10.t2 (cl (= @p_9 (! (fun_app$ uuc$ veriT_vr0) :named @p_10))) :rule cong :premises (t10.t1)) (step t10.t3 (cl @p_11) :rule refl) (step t10.t4 (cl (! (= @p_12 (! (- veriT_vr0 @p_7) :named @p_13)) :named @p_14)) :rule cong :premises (t10.t3)) (step t10.t5 (cl (! (= @p_1 (! (times$ @p_13 d$) :named @p_8)) :named @p_15)) :rule cong :premises (t10.t4)) (step t10.t6 (cl @p_11) :rule refl) (step t10.t7 (cl @p_14) :rule cong :premises (t10.t6)) (step t10.t8 (cl @p_15) :rule cong :premises (t10.t7)) (step t10.t9 (cl (= @p_16 (! (diamond_y$ @p_8) :named @p_17))) :rule cong :premises (t10.t8)) (step t10.t10 (cl (= @p_18 (! (pair$ @p_8 @p_17) :named @p_19))) :rule cong :premises (t10.t5 t10.t9)) (step t10.t11 (cl (= @p_20 (! (= @p_10 @p_19) :named @p_21))) :rule cong :premises (t10.t2 t10.t10)) (step t10 (cl (! (= @p_6 (! (forall ((veriT_vr0 Real)) @p_21) :named @p_23)) :named @p_22)) :rule bind) (step t11 (cl (not @p_22) (not @p_6) @p_23) :rule equiv_pos2) (step t12 (cl @p_23) :rule th_resolution :premises (axiom0 t10 t11)) (anchor :step t13 :args ((:= (veriT_vr0 Real) veriT_vr1))) (step t13.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_26)) :rule refl) (step t13.t2 (cl (= @p_10 (! (fun_app$ uuc$ veriT_vr1) :named @p_25))) :rule cong :premises (t13.t1)) (step t13.t3 (cl @p_26) :rule refl) (step t13.t4 (cl (! (= @p_13 (! (- veriT_vr1 @p_7) :named @p_27)) :named @p_28)) :rule cong :premises (t13.t3)) (step t13.t5 (cl (! (= @p_8 (! (times$ @p_27 d$) :named @p_24)) :named @p_29)) :rule cong :premises (t13.t4)) (step t13.t6 (cl @p_26) :rule refl) (step t13.t7 (cl @p_28) :rule cong :premises (t13.t6)) (step t13.t8 (cl @p_29) :rule cong :premises (t13.t7)) (step t13.t9 (cl (= @p_17 (! (diamond_y$ @p_24) :named @p_30))) :rule cong :premises (t13.t8)) (step t13.t10 (cl (= @p_19 (! (pair$ @p_24 @p_30) :named @p_31))) :rule cong :premises (t13.t5 t13.t9)) (step t13.t11 (cl (= @p_21 (! (= @p_25 @p_31) :named @p_32))) :rule cong :premises (t13.t2 t13.t10)) (step t13 (cl (! (= @p_23 (! (forall ((veriT_vr1 Real)) @p_32) :named @p_34)) :named @p_33)) :rule bind) (step t14 (cl (not @p_33) (not @p_23) @p_34) :rule equiv_pos2) (step t15 (cl @p_34) :rule th_resolution :premises (t12 t13 t14)) (anchor :step t16 :args ((:= (?v0 Real) veriT_vr6))) (step t16.t1 (cl (! (= ?v0 veriT_vr6) :named @p_39)) :rule refl) (step t16.t2 (cl (= @p_37 (! (fun_app$ uub$ veriT_vr6) :named @p_38))) :rule cong :premises (t16.t1)) (step t16.t3 (cl @p_39) :rule refl) (step t16.t4 (cl (= @p_40 (! (* 2.0 veriT_vr6) :named @p_41))) :rule cong :premises (t16.t3)) (step t16.t5 (cl (= @p_42 (! (- @p_41 1.0) :named @p_43))) :rule cong :premises (t16.t4)) (step t16.t6 (cl (= @p_44 (! (times$ @p_43 @p_36) :named @p_45))) :rule cong :premises (t16.t5)) (step t16.t7 (cl (= @p_46 (! (pair$ @p_2 @p_45) :named @p_47))) :rule cong :premises (t16.t6)) (step t16.t8 (cl (= @p_48 (! (= @p_38 @p_47) :named @p_49))) :rule cong :premises (t16.t2 t16.t7)) (step t16 (cl (! (= @p_35 (! (forall ((veriT_vr6 Real)) @p_49) :named @p_51)) :named @p_50)) :rule bind) (step t17 (cl (not @p_50) (not @p_35) @p_51) :rule equiv_pos2) (step t18 (cl @p_51) :rule th_resolution :premises (axiom3 t16 t17)) (anchor :step t19 :args ((:= (veriT_vr6 Real) veriT_vr7))) (step t19.t1 (cl (! (= veriT_vr6 veriT_vr7) :named @p_53)) :rule refl) (step t19.t2 (cl (= @p_38 (! (fun_app$ uub$ veriT_vr7) :named @p_52))) :rule cong :premises (t19.t1)) (step t19.t3 (cl @p_53) :rule refl) (step t19.t4 (cl (= @p_41 (! (* 2.0 veriT_vr7) :named @p_54))) :rule cong :premises (t19.t3)) (step t19.t5 (cl (= @p_43 (! (- @p_54 1.0) :named @p_55))) :rule cong :premises (t19.t4)) (step t19.t6 (cl (= @p_45 (! (times$ @p_55 @p_36) :named @p_56))) :rule cong :premises (t19.t5)) (step t19.t7 (cl (= @p_47 (! (pair$ @p_2 @p_56) :named @p_57))) :rule cong :premises (t19.t6)) (step t19.t8 (cl (= @p_49 (! (= @p_52 @p_57) :named @p_58))) :rule cong :premises (t19.t2 t19.t7)) (step t19 (cl (! (= @p_51 (! (forall ((veriT_vr7 Real)) @p_58) :named @p_60)) :named @p_59)) :rule bind) (step t20 (cl (not @p_59) (not @p_51) @p_60) :rule equiv_pos2) (step t21 (cl @p_60) :rule th_resolution :premises (t18 t19 t20)) (anchor :step t22 :args ((:= (?v0 Real) veriT_vr8))) (step t22.t1 (cl (! (= ?v0 veriT_vr8) :named @p_64)) :rule refl) (step t22.t2 (cl (= @p_62 (! (diamond_y$ veriT_vr8) :named @p_63))) :rule cong :premises (t22.t1)) (step t22.t3 (cl @p_64) :rule refl) (step t22.t4 (cl (= @p_65 (! (< veriT_vr8 0.0) :named @p_66))) :rule cong :premises (t22.t3)) (step t22.t5 (cl @p_64) :rule refl) (step t22.t6 (cl (= @p_4 (! (- veriT_vr8) :named @p_67))) :rule cong :premises (t22.t5)) (step t22.t7 (cl @p_64) :rule refl) (step t22.t8 (cl (= @p_68 (! (ite @p_66 @p_67 veriT_vr8) :named @p_69))) :rule cong :premises (t22.t4 t22.t6 t22.t7)) (step t22.t9 (cl (= @p_70 (! (- @p_3 @p_69) :named @p_71))) :rule cong :premises (t22.t8)) (step t22.t10 (cl (= @p_72 (! (= @p_63 @p_71) :named @p_73))) :rule cong :premises (t22.t2 t22.t9)) (step t22 (cl (! (= @p_61 (! (forall ((veriT_vr8 Real)) @p_73) :named @p_75)) :named @p_74)) :rule bind) (step t23 (cl (not @p_74) (not @p_61) @p_75) :rule equiv_pos2) (step t24 (cl @p_75) :rule th_resolution :premises (axiom5 t22 t23)) (anchor :step t25 :args ((:= (veriT_vr8 Real) veriT_vr9))) (step t25.t1 (cl (! (= veriT_vr8 veriT_vr9) :named @p_77)) :rule refl) (step t25.t2 (cl (= @p_63 (! (diamond_y$ veriT_vr9) :named @p_76))) :rule cong :premises (t25.t1)) (step t25.t3 (cl @p_77) :rule refl) (step t25.t4 (cl (= @p_66 (! (< veriT_vr9 0.0) :named @p_78))) :rule cong :premises (t25.t3)) (step t25.t5 (cl @p_77) :rule refl) (step t25.t6 (cl (= @p_67 (! (- veriT_vr9) :named @p_79))) :rule cong :premises (t25.t5)) (step t25.t7 (cl @p_77) :rule refl) (step t25.t8 (cl (= @p_69 (! (ite @p_78 @p_79 veriT_vr9) :named @p_80))) :rule cong :premises (t25.t4 t25.t6 t25.t7)) (step t25.t9 (cl (= @p_71 (! (- @p_3 @p_80) :named @p_81))) :rule cong :premises (t25.t8)) (step t25.t10 (cl (= @p_73 (! (= @p_76 @p_81) :named @p_82))) :rule cong :premises (t25.t2 t25.t9)) (step t25 (cl (! (= @p_75 (! (forall ((veriT_vr9 Real)) @p_82) :named @p_84)) :named @p_83)) :rule bind) (step t26 (cl (not @p_83) (not @p_75) @p_84) :rule equiv_pos2) (step t27 (cl @p_84) :rule th_resolution :premises (t24 t25 t26)) (anchor :step t28 :args ((:= (?v0 Real) veriT_vr10) (:= (?v1 Real) veriT_vr11) (:= (?v2 Real) veriT_vr12))) (step t28.t1 (cl (! (= ?v0 veriT_vr10) :named @p_94)) :rule refl) (step t28.t2 (cl (! (= ?v1 veriT_vr11) :named @p_87)) :rule refl) (step t28.t3 (cl (= @p_5 (! (divide$ veriT_vr10 veriT_vr11) :named @p_86))) :rule cong :premises (t28.t1 t28.t2)) (step t28.t4 (cl (! (= ?v2 veriT_vr12) :named @p_95)) :rule refl) (step t28.t5 (cl @p_87) :rule refl) (step t28.t6 (cl (= @p_88 (! (divide$ veriT_vr12 veriT_vr11) :named @p_89))) :rule cong :premises (t28.t4 t28.t5)) (step t28.t7 (cl (= @p_90 (! (< @p_86 @p_89) :named @p_91))) :rule cong :premises (t28.t3 t28.t6)) (step t28.t8 (cl @p_87) :rule refl) (step t28.t9 (cl (= @p_92 (! (< 0.0 veriT_vr11) :named @p_93))) :rule cong :premises (t28.t8)) (step t28.t10 (cl @p_94) :rule refl) (step t28.t11 (cl @p_95) :rule refl) (step t28.t12 (cl (= @p_96 (! (< veriT_vr10 veriT_vr12) :named @p_97))) :rule cong :premises (t28.t10 t28.t11)) (step t28.t13 (cl (= @p_98 (! (=> @p_93 @p_97) :named @p_99))) :rule cong :premises (t28.t9 t28.t12)) (step t28.t14 (cl @p_87) :rule refl) (step t28.t15 (cl (= @p_100 (! (< veriT_vr11 0.0) :named @p_101))) :rule cong :premises (t28.t14)) (step t28.t16 (cl @p_95) :rule refl) (step t28.t17 (cl @p_94) :rule refl) (step t28.t18 (cl (= @p_102 (! (< veriT_vr12 veriT_vr10) :named @p_103))) :rule cong :premises (t28.t16 t28.t17)) (step t28.t19 (cl (= @p_104 (! (=> @p_101 @p_103) :named @p_105))) :rule cong :premises (t28.t15 t28.t18)) (step t28.t20 (cl @p_87) :rule refl) (step t28.t21 (cl (= @p_106 (! (= 0.0 veriT_vr11) :named @p_107))) :rule cong :premises (t28.t20)) (step t28.t22 (cl (= @p_108 (! (not @p_107) :named @p_109))) :rule cong :premises (t28.t21)) (step t28.t23 (cl (= @p_110 (! (and @p_105 @p_109) :named @p_111))) :rule cong :premises (t28.t19 t28.t22)) (step t28.t24 (cl (= @p_112 (! (and @p_99 @p_111) :named @p_113))) :rule cong :premises (t28.t13 t28.t23)) (step t28.t25 (cl (= @p_114 (! (= @p_91 @p_113) :named @p_115))) :rule cong :premises (t28.t7 t28.t24)) (step t28 (cl (! (= @p_85 (! (forall ((veriT_vr10 Real) (veriT_vr11 Real) (veriT_vr12 Real)) @p_115) :named @p_117)) :named @p_116)) :rule bind) (step t29 (cl (not @p_116) (not @p_85) @p_117) :rule equiv_pos2) (step t30 (cl @p_117) :rule th_resolution :premises (axiom7 t28 t29)) (anchor :step t31 :args ((veriT_vr10 Real) (veriT_vr11 Real) (veriT_vr12 Real))) (step t31.t1 (cl (= @p_113 (! (and @p_99 @p_105 @p_109) :named @p_118))) :rule ac_simp) (step t31.t2 (cl (= @p_115 (! (= @p_91 @p_118) :named @p_119))) :rule cong :premises (t31.t1)) (step t31 (cl (! (= @p_117 (! (forall ((veriT_vr10 Real) (veriT_vr11 Real) (veriT_vr12 Real)) @p_119) :named @p_121)) :named @p_120)) :rule bind) (step t32 (cl (not @p_120) (not @p_117) @p_121) :rule equiv_pos2) (step t33 (cl @p_121) :rule th_resolution :premises (t30 t31 t32)) (anchor :step t34 :args ((:= (veriT_vr10 Real) veriT_vr13) (:= (veriT_vr11 Real) veriT_vr14) (:= (veriT_vr12 Real) veriT_vr15))) (step t34.t1 (cl (! (= veriT_vr10 veriT_vr13) :named @p_127)) :rule refl) (step t34.t2 (cl (! (= veriT_vr11 veriT_vr14) :named @p_123)) :rule refl) (step t34.t3 (cl (= @p_86 (! (divide$ veriT_vr13 veriT_vr14) :named @p_122))) :rule cong :premises (t34.t1 t34.t2)) (step t34.t4 (cl (! (= veriT_vr12 veriT_vr15) :named @p_128)) :rule refl) (step t34.t5 (cl @p_123) :rule refl) (step t34.t6 (cl (= @p_89 (! (divide$ veriT_vr15 veriT_vr14) :named @p_124))) :rule cong :premises (t34.t4 t34.t5)) (step t34.t7 (cl (= @p_91 (! (< @p_122 @p_124) :named @p_125))) :rule cong :premises (t34.t3 t34.t6)) (step t34.t8 (cl @p_123) :rule refl) (step t34.t9 (cl (= @p_93 (! (< 0.0 veriT_vr14) :named @p_126))) :rule cong :premises (t34.t8)) (step t34.t10 (cl @p_127) :rule refl) (step t34.t11 (cl @p_128) :rule refl) (step t34.t12 (cl (= @p_97 (! (< veriT_vr13 veriT_vr15) :named @p_129))) :rule cong :premises (t34.t10 t34.t11)) (step t34.t13 (cl (= @p_99 (! (=> @p_126 @p_129) :named @p_130))) :rule cong :premises (t34.t9 t34.t12)) (step t34.t14 (cl @p_123) :rule refl) (step t34.t15 (cl (= @p_101 (! (< veriT_vr14 0.0) :named @p_131))) :rule cong :premises (t34.t14)) (step t34.t16 (cl @p_128) :rule refl) (step t34.t17 (cl @p_127) :rule refl) (step t34.t18 (cl (= @p_103 (! (< veriT_vr15 veriT_vr13) :named @p_132))) :rule cong :premises (t34.t16 t34.t17)) (step t34.t19 (cl (= @p_105 (! (=> @p_131 @p_132) :named @p_133))) :rule cong :premises (t34.t15 t34.t18)) (step t34.t20 (cl @p_123) :rule refl) (step t34.t21 (cl (= @p_107 (! (= 0.0 veriT_vr14) :named @p_134))) :rule cong :premises (t34.t20)) (step t34.t22 (cl (= @p_109 (! (not @p_134) :named @p_135))) :rule cong :premises (t34.t21)) (step t34.t23 (cl (= @p_118 (! (and @p_130 @p_133 @p_135) :named @p_136))) :rule cong :premises (t34.t13 t34.t19 t34.t22)) (step t34.t24 (cl (= @p_119 (! (= @p_125 @p_136) :named @p_137))) :rule cong :premises (t34.t7 t34.t23)) (step t34 (cl (! (= @p_121 (! (forall ((veriT_vr13 Real) (veriT_vr14 Real) (veriT_vr15 Real)) @p_137) :named @p_139)) :named @p_138)) :rule bind) (step t35 (cl (not @p_138) (not @p_121) @p_139) :rule equiv_pos2) (step t36 (cl @p_139) :rule th_resolution :premises (t33 t34 t35)) (anchor :step t37 :args ((:= (?v0 Real) veriT_vr16) (:= (?v1 Real) veriT_vr17))) (step t37.t1 (cl (! (= ?v0 veriT_vr16) :named @p_144)) :rule refl) (step t37.t2 (cl (= @p_4 (! (- veriT_vr16) :named @p_141))) :rule cong :premises (t37.t1)) (step t37.t3 (cl (! (= ?v1 veriT_vr17) :named @p_145)) :rule refl) (step t37.t4 (cl (= @p_142 (! (divide$ @p_141 veriT_vr17) :named @p_143))) :rule cong :premises (t37.t2 t37.t3)) (step t37.t5 (cl @p_144) :rule refl) (step t37.t6 (cl @p_145) :rule refl) (step t37.t7 (cl (= @p_5 (! (divide$ veriT_vr16 veriT_vr17) :named @p_146))) :rule cong :premises (t37.t5 t37.t6)) (step t37.t8 (cl (= @p_147 (! (- @p_146) :named @p_148))) :rule cong :premises (t37.t7)) (step t37.t9 (cl (= @p_149 (! (= @p_143 @p_148) :named @p_150))) :rule cong :premises (t37.t4 t37.t8)) (step t37 (cl (! (= @p_140 (! (forall ((veriT_vr16 Real) (veriT_vr17 Real)) @p_150) :named @p_152)) :named @p_151)) :rule bind) (step t38 (cl (not @p_151) (not @p_140) @p_152) :rule equiv_pos2) (step t39 (cl @p_152) :rule th_resolution :premises (axiom8 t37 t38)) (anchor :step t40 :args ((:= (veriT_vr16 Real) veriT_vr18) (:= (veriT_vr17 Real) veriT_vr19))) (step t40.t1 (cl (! (= veriT_vr16 veriT_vr18) :named @p_155)) :rule refl) (step t40.t2 (cl (= @p_141 (! (- veriT_vr18) :named @p_153))) :rule cong :premises (t40.t1)) (step t40.t3 (cl (! (= veriT_vr17 veriT_vr19) :named @p_156)) :rule refl) (step t40.t4 (cl (= @p_143 (! (divide$ @p_153 veriT_vr19) :named @p_154))) :rule cong :premises (t40.t2 t40.t3)) (step t40.t5 (cl @p_155) :rule refl) (step t40.t6 (cl @p_156) :rule refl) (step t40.t7 (cl (= @p_146 (! (divide$ veriT_vr18 veriT_vr19) :named @p_157))) :rule cong :premises (t40.t5 t40.t6)) (step t40.t8 (cl (= @p_148 (! (- @p_157) :named @p_158))) :rule cong :premises (t40.t7)) (step t40.t9 (cl (= @p_150 (! (= @p_154 @p_158) :named @p_159))) :rule cong :premises (t40.t4 t40.t8)) (step t40 (cl (! (= @p_152 (! (forall ((veriT_vr18 Real) (veriT_vr19 Real)) @p_159) :named @p_161)) :named @p_160)) :rule bind) (step t41 (cl (not @p_160) (not @p_152) @p_161) :rule equiv_pos2) (step t42 (cl @p_161) :rule th_resolution :premises (t39 t40 t41)) (anchor :step t43 :args ((:= (?v0 Real) veriT_vr20) (:= (?v1 Real) veriT_vr21))) (step t43.t1 (cl (! (= ?v0 veriT_vr20) :named @p_166)) :rule refl) (step t43.t2 (cl (= @p_4 (! (- veriT_vr20) :named @p_163))) :rule cong :premises (t43.t1)) (step t43.t3 (cl (! (= ?v1 veriT_vr21) :named @p_167)) :rule refl) (step t43.t4 (cl (= @p_164 (! (times$ @p_163 veriT_vr21) :named @p_165))) :rule cong :premises (t43.t2 t43.t3)) (step t43.t5 (cl @p_166) :rule refl) (step t43.t6 (cl @p_167) :rule refl) (step t43.t7 (cl (= @p_168 (! (times$ veriT_vr20 veriT_vr21) :named @p_169))) :rule cong :premises (t43.t5 t43.t6)) (step t43.t8 (cl (= @p_170 (! (- @p_169) :named @p_171))) :rule cong :premises (t43.t7)) (step t43.t9 (cl (= @p_172 (! (= @p_165 @p_171) :named @p_173))) :rule cong :premises (t43.t4 t43.t8)) (step t43 (cl (! (= @p_162 (! (forall ((veriT_vr20 Real) (veriT_vr21 Real)) @p_173) :named @p_175)) :named @p_174)) :rule bind) (step t44 (cl (not @p_174) (not @p_162) @p_175) :rule equiv_pos2) (step t45 (cl @p_175) :rule th_resolution :premises (axiom9 t43 t44)) (anchor :step t46 :args ((:= (veriT_vr20 Real) veriT_vr22) (:= (veriT_vr21 Real) veriT_vr23))) (step t46.t1 (cl (! (= veriT_vr20 veriT_vr22) :named @p_178)) :rule refl) (step t46.t2 (cl (= @p_163 (! (- veriT_vr22) :named @p_176))) :rule cong :premises (t46.t1)) (step t46.t3 (cl (! (= veriT_vr21 veriT_vr23) :named @p_179)) :rule refl) (step t46.t4 (cl (= @p_165 (! (times$ @p_176 veriT_vr23) :named @p_177))) :rule cong :premises (t46.t2 t46.t3)) (step t46.t5 (cl @p_178) :rule refl) (step t46.t6 (cl @p_179) :rule refl) (step t46.t7 (cl (= @p_169 (! (times$ veriT_vr22 veriT_vr23) :named @p_180))) :rule cong :premises (t46.t5 t46.t6)) (step t46.t8 (cl (= @p_171 (! (- @p_180) :named @p_181))) :rule cong :premises (t46.t7)) (step t46.t9 (cl (= @p_173 (! (= @p_177 @p_181) :named @p_182))) :rule cong :premises (t46.t4 t46.t8)) (step t46 (cl (! (= @p_175 (! (forall ((veriT_vr22 Real) (veriT_vr23 Real)) @p_182) :named @p_184)) :named @p_183)) :rule bind) (step t47 (cl (not @p_183) (not @p_175) @p_184) :rule equiv_pos2) (step t48 (cl @p_184) :rule th_resolution :premises (t45 t46 t47)) (anchor :step t49 :args ((:= (?v0 Real) veriT_vr24) (:= (?v1 Real) veriT_vr25) (:= (?v2 Real) veriT_vr26) (:= (?v3 Real) veriT_vr27))) (step t49.t1 (cl (! (= ?v0 veriT_vr24) :named @p_192)) :rule refl) (step t49.t2 (cl (! (= ?v1 veriT_vr25) :named @p_196)) :rule refl) (step t49.t3 (cl (= @p_186 (! (pair$ veriT_vr24 veriT_vr25) :named @p_187))) :rule cong :premises (t49.t1 t49.t2)) (step t49.t4 (cl (! (= ?v2 veriT_vr26) :named @p_193)) :rule refl) (step t49.t5 (cl (! (= ?v3 veriT_vr27) :named @p_197)) :rule refl) (step t49.t6 (cl (= @p_188 (! (pair$ veriT_vr26 veriT_vr27) :named @p_189))) :rule cong :premises (t49.t4 t49.t5)) (step t49.t7 (cl (= @p_190 (! (= @p_187 @p_189) :named @p_191))) :rule cong :premises (t49.t3 t49.t6)) (step t49.t8 (cl @p_192) :rule refl) (step t49.t9 (cl @p_193) :rule refl) (step t49.t10 (cl (= @p_194 (! (= veriT_vr24 veriT_vr26) :named @p_195))) :rule cong :premises (t49.t8 t49.t9)) (step t49.t11 (cl @p_196) :rule refl) (step t49.t12 (cl @p_197) :rule refl) (step t49.t13 (cl (= @p_198 (! (= veriT_vr25 veriT_vr27) :named @p_199))) :rule cong :premises (t49.t11 t49.t12)) (step t49.t14 (cl (= @p_200 (! (and @p_195 @p_199) :named @p_201))) :rule cong :premises (t49.t10 t49.t13)) (step t49.t15 (cl (= @p_202 (! (= @p_191 @p_201) :named @p_203))) :rule cong :premises (t49.t7 t49.t14)) (step t49 (cl (! (= @p_185 (! (forall ((veriT_vr24 Real) (veriT_vr25 Real) (veriT_vr26 Real) (veriT_vr27 Real)) @p_203) :named @p_205)) :named @p_204)) :rule bind) (step t50 (cl (not @p_204) (not @p_185) @p_205) :rule equiv_pos2) (step t51 (cl @p_205) :rule th_resolution :premises (axiom10 t49 t50)) (anchor :step t52 :args ((:= (veriT_vr24 Real) veriT_vr28) (:= (veriT_vr25 Real) veriT_vr29) (:= (veriT_vr26 Real) veriT_vr30) (:= (veriT_vr27 Real) veriT_vr31))) (step t52.t1 (cl (! (= veriT_vr24 veriT_vr28) :named @p_209)) :rule refl) (step t52.t2 (cl (! (= veriT_vr25 veriT_vr29) :named @p_212)) :rule refl) (step t52.t3 (cl (= @p_187 (! (pair$ veriT_vr28 veriT_vr29) :named @p_206))) :rule cong :premises (t52.t1 t52.t2)) (step t52.t4 (cl (! (= veriT_vr26 veriT_vr30) :named @p_210)) :rule refl) (step t52.t5 (cl (! (= veriT_vr27 veriT_vr31) :named @p_213)) :rule refl) (step t52.t6 (cl (= @p_189 (! (pair$ veriT_vr30 veriT_vr31) :named @p_207))) :rule cong :premises (t52.t4 t52.t5)) (step t52.t7 (cl (= @p_191 (! (= @p_206 @p_207) :named @p_208))) :rule cong :premises (t52.t3 t52.t6)) (step t52.t8 (cl @p_209) :rule refl) (step t52.t9 (cl @p_210) :rule refl) (step t52.t10 (cl (= @p_195 (! (= veriT_vr28 veriT_vr30) :named @p_211))) :rule cong :premises (t52.t8 t52.t9)) (step t52.t11 (cl @p_212) :rule refl) (step t52.t12 (cl @p_213) :rule refl) (step t52.t13 (cl (= @p_199 (! (= veriT_vr29 veriT_vr31) :named @p_214))) :rule cong :premises (t52.t11 t52.t12)) (step t52.t14 (cl (= @p_201 (! (and @p_211 @p_214) :named @p_215))) :rule cong :premises (t52.t10 t52.t13)) (step t52.t15 (cl (= @p_203 (! (= @p_208 @p_215) :named @p_216))) :rule cong :premises (t52.t7 t52.t14)) (step t52 (cl (! (= @p_205 (! (forall ((veriT_vr28 Real) (veriT_vr29 Real) (veriT_vr30 Real) (veriT_vr31 Real)) @p_216) :named @p_218)) :named @p_217)) :rule bind) (step t53 (cl (not @p_217) (not @p_205) @p_218) :rule equiv_pos2) (step t54 (cl @p_218) :rule th_resolution :premises (t51 t52 t53)) (step t55 (cl (! (= @p_219 (! (and @p_220 (! (not false) :named @p_228)) :named @p_222)) :named @p_221)) :rule bool_simplify) (step t56 (cl (! (not @p_221) :named @p_225) (! (not @p_219) :named @p_223) @p_222) :rule equiv_pos2) (step t57 (cl (not @p_223) @p_224) :rule not_not) (step t58 (cl @p_225 @p_224 @p_222) :rule th_resolution :premises (t57 t56)) (step t59 (cl @p_222) :rule th_resolution :premises (axiom11 t55 t58)) (step t60 (cl (! (= @p_222 (! (and @p_226 @p_227 @p_228) :named @p_230)) :named @p_229)) :rule ac_simp) (step t61 (cl (not @p_229) (not @p_222) @p_230) :rule equiv_pos2) (step t62 (cl @p_230) :rule th_resolution :premises (t59 t60 t61)) (step t63 (cl (! (= @p_228 true) :named @p_379)) :rule not_simplify) (step t64 (cl (= @p_230 (! (and @p_226 @p_227 true) :named @p_231))) :rule cong :premises (t63)) (step t65 (cl (= @p_231 @p_220)) :rule and_simplify) (step t66 (cl (! (= @p_230 @p_220) :named @p_232)) :rule trans :premises (t64 t65)) (step t67 (cl (not @p_232) (not @p_230) @p_220) :rule equiv_pos2) (step t68 (cl @p_220) :rule th_resolution :premises (t62 t66 t67)) (step t69 (cl @p_227) :rule and :premises (t68)) (step t70 (cl (or (! (not @p_218) :named @p_244) (! (forall ((veriT_vr28 Real) (veriT_vr29 Real) (veriT_vr30 Real) (veriT_vr31 Real)) (or (not @p_208) @p_211)) :named @p_533))) :rule qnt_cnf) (step t71 (cl (or (! (not @p_161) :named @p_241) (! (= (! (divide$ (! (- 0.0) :named @p_236) d$) :named @p_237) (! (- (! (divide$ 0.0 d$) :named @p_234)) :named @p_235)) :named @p_233))) :rule forall_inst :args ((:= veriT_vr18 0.0) (:= veriT_vr19 d$))) (anchor :step t72) (assume t72.h1 @p_233) (step t72.t2 (cl (! (= 0.0 @p_236) :named @p_249)) :rule minus_simplify) (step t72.t3 (cl (= @p_234 @p_237)) :rule cong :premises (t72.t2)) (step t72.t4 (cl (! (= @p_233 (! (= @p_234 @p_235) :named @p_238)) :named @p_239)) :rule cong :premises (t72.t3)) (step t72.t5 (cl (not @p_239) (! (not @p_233) :named @p_240) @p_238) :rule equiv_pos2) (step t72.t6 (cl @p_238) :rule th_resolution :premises (t72.h1 t72.t4 t72.t5)) (step t72 (cl @p_240 @p_238) :rule subproof :discharge (h1)) (step t73 (cl @p_241 @p_233) :rule or :premises (t71)) (step t74 (cl (! (or @p_241 @p_238) :named @p_243) (! (not @p_241) :named @p_242)) :rule or_neg) (step t75 (cl (not @p_242) @p_161) :rule not_not) (step t76 (cl @p_243 @p_161) :rule th_resolution :premises (t75 t74)) (step t77 (cl @p_243 (! (not @p_238) :named @p_527)) :rule or_neg) (step t78 (cl @p_243) :rule th_resolution :premises (t73 t72 t76 t77)) (step t79 (cl (not (! (not (! (not @p_139) :named @p_282)) :named @p_302)) @p_139) :rule not_not) (step t80 (cl @p_241 @p_238) :rule or :premises (t78)) (step t81 (cl @p_238) :rule resolution :premises (t80 t42)) (step t82 (cl (not (! (not @p_244) :named @p_538)) @p_218) :rule not_not) (step t83 (cl (not (! (not (! (not @p_184) :named @p_245)) :named @p_254)) @p_184) :rule not_not) (step t84 (cl (or @p_245 (! (= (! (times$ @p_236 d$) :named @p_250) (! (- (! (times$ 0.0 d$) :named @p_247)) :named @p_248)) :named @p_246))) :rule forall_inst :args ((:= veriT_vr22 0.0) (:= veriT_vr23 d$))) (anchor :step t85) (assume t85.h1 @p_246) (step t85.t2 (cl @p_249) :rule minus_simplify) (step t85.t3 (cl (= @p_250 @p_247)) :rule cong :premises (t85.t2)) (step t85.t4 (cl (! (= @p_246 (! (= @p_247 @p_248) :named @p_251)) :named @p_252)) :rule cong :premises (t85.t3)) (step t85.t5 (cl (not @p_252) (! (not @p_246) :named @p_253) @p_251) :rule equiv_pos2) (step t85.t6 (cl @p_251) :rule th_resolution :premises (t85.h1 t85.t4 t85.t5)) (step t85 (cl @p_253 @p_251) :rule subproof :discharge (h1)) (step t86 (cl @p_245 @p_246) :rule or :premises (t84)) (step t87 (cl (! (or @p_245 @p_251) :named @p_255) @p_254) :rule or_neg) (step t88 (cl @p_255 @p_184) :rule th_resolution :premises (t83 t87)) (step t89 (cl @p_255 (! (not @p_251) :named @p_531)) :rule or_neg) (step t90 (cl @p_255) :rule th_resolution :premises (t86 t85 t88 t89)) (step t91 (cl (or @p_241 (! (= (! (divide$ (! (- @p_2) :named @p_258) 2.0) :named @p_259) (! (- (! (divide$ @p_2 2.0) :named @p_366)) :named @p_257)) :named @p_256))) :rule forall_inst :args ((:= veriT_vr18 @p_2) (:= veriT_vr19 2.0))) (anchor :step t92) (assume t92.h1 @p_256) (step t92.t2 (cl (! (= @p_3 @p_258) :named @p_267)) :rule minus_simplify) (step t92.t3 (cl (= @p_259 (! (divide$ @p_3 2.0) :named @p_260))) :rule cong :premises (t92.t2)) (step t92.t4 (cl (! (= @p_256 (! (= @p_257 @p_260) :named @p_261)) :named @p_262)) :rule cong :premises (t92.t3)) (step t92.t5 (cl (not @p_262) (! (not @p_256) :named @p_263) @p_261) :rule equiv_pos2) (step t92.t6 (cl @p_261) :rule th_resolution :premises (t92.h1 t92.t4 t92.t5)) (step t92 (cl @p_263 @p_261) :rule subproof :discharge (h1)) (step t93 (cl @p_241 @p_256) :rule or :premises (t91)) (step t94 (cl (! (or @p_241 @p_261) :named @p_264) @p_242) :rule or_neg) (step t95 (cl @p_264 @p_161) :rule th_resolution :premises (t75 t94)) (step t96 (cl @p_264 (! (not @p_261) :named @p_563)) :rule or_neg) (step t97 (cl @p_264) :rule th_resolution :premises (t93 t92 t95 t96)) (step t98 (cl (or @p_241 (! (= (! (divide$ @p_258 0.0) :named @p_268) (! (- (! (divide$ @p_2 0.0) :named @p_305)) :named @p_266)) :named @p_265))) :rule forall_inst :args ((:= veriT_vr18 @p_2) (:= veriT_vr19 0.0))) (anchor :step t99) (assume t99.h1 @p_265) (step t99.t2 (cl @p_267) :rule minus_simplify) (step t99.t3 (cl (= @p_268 (! (divide$ @p_3 0.0) :named @p_269))) :rule cong :premises (t99.t2)) (step t99.t4 (cl (! (= @p_265 (! (= @p_266 @p_269) :named @p_270)) :named @p_271)) :rule cong :premises (t99.t3)) (step t99.t5 (cl (not @p_271) (! (not @p_265) :named @p_272) @p_270) :rule equiv_pos2) (step t99.t6 (cl @p_270) :rule th_resolution :premises (t99.h1 t99.t4 t99.t5)) (step t99 (cl @p_272 @p_270) :rule subproof :discharge (h1)) (step t100 (cl @p_241 @p_265) :rule or :premises (t98)) (step t101 (cl (! (or @p_241 @p_270) :named @p_273) @p_242) :rule or_neg) (step t102 (cl @p_273 @p_161) :rule th_resolution :premises (t75 t101)) (step t103 (cl @p_273 (! (not @p_270) :named @p_528)) :rule or_neg) (step t104 (cl @p_273) :rule th_resolution :premises (t100 t99 t102 t103)) (step t105 (cl (or @p_241 (! (= (! (divide$ @p_235 2.0) :named @p_562) (! (- (! (divide$ @p_234 2.0) :named @p_552)) :named @p_558)) :named @p_495))) :rule forall_inst :args ((:= veriT_vr18 @p_234) (:= veriT_vr19 2.0))) (step t106 (cl (or @p_241 (! (= (! (divide$ @p_236 2.0) :named @p_277) (! (- (! (divide$ 0.0 2.0) :named @p_275)) :named @p_276)) :named @p_274))) :rule forall_inst :args ((:= veriT_vr18 0.0) (:= veriT_vr19 2.0))) (anchor :step t107) (assume t107.h1 @p_274) (step t107.t2 (cl @p_249) :rule minus_simplify) (step t107.t3 (cl (= @p_277 @p_275)) :rule cong :premises (t107.t2)) (step t107.t4 (cl (! (= @p_274 (! (= @p_275 @p_276) :named @p_278)) :named @p_279)) :rule cong :premises (t107.t3)) (step t107.t5 (cl (not @p_279) (! (not @p_274) :named @p_280) @p_278) :rule equiv_pos2) (step t107.t6 (cl @p_278) :rule th_resolution :premises (t107.h1 t107.t4 t107.t5)) (step t107 (cl @p_280 @p_278) :rule subproof :discharge (h1)) (step t108 (cl @p_241 @p_274) :rule or :premises (t106)) (step t109 (cl (! (or @p_241 @p_278) :named @p_281) @p_242) :rule or_neg) (step t110 (cl @p_281 @p_161) :rule th_resolution :premises (t75 t109)) (step t111 (cl @p_281 (! (not @p_278) :named @p_510)) :rule or_neg) (step t112 (cl @p_281) :rule th_resolution :premises (t108 t107 t110 t111)) (step t113 (cl (or @p_282 (! (= (! (< (! (divide$ @p_36 @p_36) :named @p_283) @p_283) :named @p_289) (! (and (! (=> (! (< 0.0 @p_36) :named @p_286) (! (< @p_36 @p_36) :named @p_284)) :named @p_290) (! (=> (! (< @p_36 0.0) :named @p_287) @p_284) :named @p_293) (! (not (! (= @p_36 0.0) :named @p_499)) :named @p_288)) :named @p_296)) :named @p_285))) :rule forall_inst :args ((:= veriT_vr13 @p_36) (:= veriT_vr14 @p_36) (:= veriT_vr15 @p_36))) (anchor :step t114) (assume t114.h1 @p_285) (step t114.t2 (cl (= @p_289 false)) :rule comp_simplify) (step t114.t3 (cl (= @p_284 false)) :rule comp_simplify) (step t114.t4 (cl (= @p_290 (! (=> @p_286 false) :named @p_291))) :rule cong :premises (t114.t3)) (step t114.t5 (cl (= @p_291 (! (not @p_286) :named @p_292))) :rule implies_simplify) (step t114.t6 (cl (= @p_290 @p_292)) :rule trans :premises (t114.t4 t114.t5)) (step t114.t7 (cl (= @p_293 (! (=> @p_287 false) :named @p_294))) :rule cong :premises (t114.t3)) (step t114.t8 (cl (= @p_294 (! (not @p_287) :named @p_295))) :rule implies_simplify) (step t114.t9 (cl (= @p_293 @p_295)) :rule trans :premises (t114.t7 t114.t8)) (step t114.t10 (cl (= @p_296 (! (and @p_292 @p_295 @p_288) :named @p_297))) :rule cong :premises (t114.t6 t114.t9)) (step t114.t11 (cl (= @p_285 (! (= false @p_297) :named @p_298))) :rule cong :premises (t114.t2 t114.t10)) (step t114.t12 (cl (= @p_298 (! (not @p_297) :named @p_299))) :rule equiv_simplify) (step t114.t13 (cl (! (= @p_285 @p_299) :named @p_300)) :rule trans :premises (t114.t11 t114.t12)) (step t114.t14 (cl (not @p_300) (! (not @p_285) :named @p_301) @p_299) :rule equiv_pos2) (step t114.t15 (cl @p_299) :rule th_resolution :premises (t114.h1 t114.t13 t114.t14)) (step t114 (cl @p_301 @p_299) :rule subproof :discharge (h1)) (step t115 (cl @p_282 @p_285) :rule or :premises (t113)) (step t116 (cl (! (or @p_282 @p_299) :named @p_303) @p_302) :rule or_neg) (step t117 (cl @p_303 @p_139) :rule th_resolution :premises (t79 t116)) (step t118 (cl @p_303 (! (not @p_299) :named @p_304)) :rule or_neg) (step t119 (cl (not @p_304) @p_297) :rule not_not) (step t120 (cl @p_303 @p_297) :rule th_resolution :premises (t119 t118)) (step t121 (cl @p_303) :rule th_resolution :premises (t115 t114 t117 t120)) (step t122 (cl (or @p_282 (! (= (! (< (! (divide$ (! (diamond_y$ @p_3) :named @p_306) 0.0) :named @p_328) @p_305) :named @p_309) (! (and (! (=> (! (< 0.0 0.0) :named @p_307) (! (< @p_306 @p_2) :named @p_311)) :named @p_310) (! (=> @p_307 (! (< @p_2 @p_306) :named @p_314)) :named @p_313) (! (not (! (= 0.0 0.0) :named @p_316)) :named @p_317)) :named @p_319)) :named @p_308))) :rule forall_inst :args ((:= veriT_vr13 @p_306) (:= veriT_vr14 0.0) (:= veriT_vr15 @p_2))) (anchor :step t123) (assume t123.h1 @p_308) (step t123.t2 (cl (! (= @p_307 false) :named @p_331)) :rule comp_simplify) (step t123.t3 (cl (! (= @p_310 (! (=> false @p_311) :named @p_312)) :named @p_392)) :rule cong :premises (t123.t2)) (step t123.t4 (cl (! (= @p_312 true) :named @p_393)) :rule implies_simplify) (step t123.t5 (cl (! (= @p_310 true) :named @p_394)) :rule trans :premises (t123.t3 t123.t4)) (step t123.t6 (cl (! (= @p_313 (! (=> false @p_314) :named @p_315)) :named @p_389)) :rule cong :premises (t123.t2)) (step t123.t7 (cl (! (= @p_315 true) :named @p_390)) :rule implies_simplify) (step t123.t8 (cl (! (= @p_313 true) :named @p_391)) :rule trans :premises (t123.t6 t123.t7)) (step t123.t9 (cl (! (= @p_316 true) :named @p_338)) :rule eq_simplify) (step t123.t10 (cl (! (= @p_317 (! (not true) :named @p_318)) :named @p_339)) :rule cong :premises (t123.t9)) (step t123.t11 (cl (! (= @p_318 false) :named @p_340)) :rule not_simplify) (step t123.t12 (cl (! (= @p_317 false) :named @p_341)) :rule trans :premises (t123.t10 t123.t11)) (step t123.t13 (cl (= @p_319 (! (and true true false) :named @p_320))) :rule cong :premises (t123.t5 t123.t8 t123.t12)) (step t123.t14 (cl (! (= @p_320 (! (and false) :named @p_321)) :named @p_343)) :rule and_simplify) (step t123.t15 (cl (! (= @p_321 false) :named @p_344)) :rule and_simplify) (step t123.t16 (cl (= @p_319 false)) :rule trans :premises (t123.t13 t123.t14 t123.t15)) (step t123.t17 (cl (= @p_308 (! (= @p_309 false) :named @p_322))) :rule cong :premises (t123.t16)) (step t123.t18 (cl (= @p_322 (! (not @p_309) :named @p_323))) :rule equiv_simplify) (step t123.t19 (cl (! (= @p_308 @p_323) :named @p_324)) :rule trans :premises (t123.t17 t123.t18)) (step t123.t20 (cl (not @p_324) (! (not @p_308) :named @p_325) @p_323) :rule equiv_pos2) (step t123.t21 (cl @p_323) :rule th_resolution :premises (t123.h1 t123.t19 t123.t20)) (step t123 (cl @p_325 @p_323) :rule subproof :discharge (h1)) (step t124 (cl @p_282 @p_308) :rule or :premises (t122)) (step t125 (cl (! (or @p_282 @p_323) :named @p_326) @p_302) :rule or_neg) (step t126 (cl @p_326 @p_139) :rule th_resolution :premises (t79 t125)) (step t127 (cl @p_326 (! (not @p_323) :named @p_327)) :rule or_neg) (step t128 (cl (not @p_327) @p_309) :rule not_not) (step t129 (cl @p_326 @p_309) :rule th_resolution :premises (t128 t127)) (step t130 (cl @p_326) :rule th_resolution :premises (t124 t123 t126 t129)) (step t131 (cl (or @p_282 (! (= (! (< @p_328 @p_269) :named @p_330) (! (and (! (=> @p_307 (! (< @p_306 @p_3) :named @p_333)) :named @p_332) (! (=> @p_307 (! (< @p_3 @p_306) :named @p_336)) :named @p_335) @p_317) :named @p_342)) :named @p_329))) :rule forall_inst :args ((:= veriT_vr13 @p_306) (:= veriT_vr14 0.0) (:= veriT_vr15 @p_3))) (anchor :step t132) (assume t132.h1 @p_329) (step t132.t2 (cl @p_331) :rule comp_simplify) (step t132.t3 (cl (! (= @p_332 (! (=> false @p_333) :named @p_334)) :named @p_427)) :rule cong :premises (t132.t2)) (step t132.t4 (cl (! (= @p_334 true) :named @p_428)) :rule implies_simplify) (step t132.t5 (cl (! (= @p_332 true) :named @p_429)) :rule trans :premises (t132.t3 t132.t4)) (step t132.t6 (cl (! (= @p_335 (! (=> false @p_336) :named @p_337)) :named @p_424)) :rule cong :premises (t132.t2)) (step t132.t7 (cl (! (= @p_337 true) :named @p_425)) :rule implies_simplify) (step t132.t8 (cl (! (= @p_335 true) :named @p_426)) :rule trans :premises (t132.t6 t132.t7)) (step t132.t9 (cl @p_338) :rule eq_simplify) (step t132.t10 (cl @p_339) :rule cong :premises (t132.t9)) (step t132.t11 (cl @p_340) :rule not_simplify) (step t132.t12 (cl @p_341) :rule trans :premises (t132.t10 t132.t11)) (step t132.t13 (cl (= @p_342 @p_320)) :rule cong :premises (t132.t5 t132.t8 t132.t12)) (step t132.t14 (cl @p_343) :rule and_simplify) (step t132.t15 (cl @p_344) :rule and_simplify) (step t132.t16 (cl (= @p_342 false)) :rule trans :premises (t132.t13 t132.t14 t132.t15)) (step t132.t17 (cl (= @p_329 (! (= @p_330 false) :named @p_345))) :rule cong :premises (t132.t16)) (step t132.t18 (cl (= @p_345 (! (not @p_330) :named @p_346))) :rule equiv_simplify) (step t132.t19 (cl (! (= @p_329 @p_346) :named @p_347)) :rule trans :premises (t132.t17 t132.t18)) (step t132.t20 (cl (not @p_347) (! (not @p_329) :named @p_348) @p_346) :rule equiv_pos2) (step t132.t21 (cl @p_346) :rule th_resolution :premises (t132.h1 t132.t19 t132.t20)) (step t132 (cl @p_348 @p_346) :rule subproof :discharge (h1)) (step t133 (cl @p_282 @p_329) :rule or :premises (t131)) (step t134 (cl (! (or @p_282 @p_346) :named @p_349) @p_302) :rule or_neg) (step t135 (cl @p_349 @p_139) :rule th_resolution :premises (t79 t134)) (step t136 (cl @p_349 (! (not @p_346) :named @p_350)) :rule or_neg) (step t137 (cl (not @p_350) @p_330) :rule not_not) (step t138 (cl @p_349 @p_330) :rule th_resolution :premises (t137 t136)) (step t139 (cl @p_349) :rule th_resolution :premises (t133 t132 t135 t138)) (step t140 (cl (or @p_282 (! (= (! (< @p_328 (! (divide$ 2.0 0.0) :named @p_437)) :named @p_352) (! (and (! (=> @p_307 (! (< @p_306 2.0) :named @p_354)) :named @p_353) (! (=> @p_307 (! (< 2.0 @p_306) :named @p_357)) :named @p_356) @p_317) :named @p_359)) :named @p_351))) :rule forall_inst :args ((:= veriT_vr13 @p_306) (:= veriT_vr14 0.0) (:= veriT_vr15 2.0))) (anchor :step t141) (assume t141.h1 @p_351) (step t141.t2 (cl @p_331) :rule comp_simplify) (step t141.t3 (cl (! (= @p_353 (! (=> false @p_354) :named @p_355)) :named @p_443)) :rule cong :premises (t141.t2)) (step t141.t4 (cl (! (= @p_355 true) :named @p_444)) :rule implies_simplify) (step t141.t5 (cl (! (= @p_353 true) :named @p_445)) :rule trans :premises (t141.t3 t141.t4)) (step t141.t6 (cl (! (= @p_356 (! (=> false @p_357) :named @p_358)) :named @p_440)) :rule cong :premises (t141.t2)) (step t141.t7 (cl (! (= @p_358 true) :named @p_441)) :rule implies_simplify) (step t141.t8 (cl (! (= @p_356 true) :named @p_442)) :rule trans :premises (t141.t6 t141.t7)) (step t141.t9 (cl @p_338) :rule eq_simplify) (step t141.t10 (cl @p_339) :rule cong :premises (t141.t9)) (step t141.t11 (cl @p_340) :rule not_simplify) (step t141.t12 (cl @p_341) :rule trans :premises (t141.t10 t141.t11)) (step t141.t13 (cl (= @p_359 @p_320)) :rule cong :premises (t141.t5 t141.t8 t141.t12)) (step t141.t14 (cl @p_343) :rule and_simplify) (step t141.t15 (cl @p_344) :rule and_simplify) (step t141.t16 (cl (= @p_359 false)) :rule trans :premises (t141.t13 t141.t14 t141.t15)) (step t141.t17 (cl (= @p_351 (! (= @p_352 false) :named @p_360))) :rule cong :premises (t141.t16)) (step t141.t18 (cl (= @p_360 (! (not @p_352) :named @p_361))) :rule equiv_simplify) (step t141.t19 (cl (! (= @p_351 @p_361) :named @p_362)) :rule trans :premises (t141.t17 t141.t18)) (step t141.t20 (cl (not @p_362) (! (not @p_351) :named @p_363) @p_361) :rule equiv_pos2) (step t141.t21 (cl @p_361) :rule th_resolution :premises (t141.h1 t141.t19 t141.t20)) (step t141 (cl @p_363 @p_361) :rule subproof :discharge (h1)) (step t142 (cl @p_282 @p_351) :rule or :premises (t140)) (step t143 (cl (! (or @p_282 @p_361) :named @p_364) @p_302) :rule or_neg) (step t144 (cl @p_364 @p_139) :rule th_resolution :premises (t79 t143)) (step t145 (cl @p_364 (! (not @p_361) :named @p_365)) :rule or_neg) (step t146 (cl (not @p_365) @p_352) :rule not_not) (step t147 (cl @p_364 @p_352) :rule th_resolution :premises (t146 t145)) (step t148 (cl @p_364) :rule th_resolution :premises (t142 t141 t144 t147)) (step t149 (cl (or @p_282 (! (= (! (< @p_366 @p_275) :named @p_368) (! (and (! (=> (! (< 0.0 2.0) :named @p_370) (! (< @p_2 0.0) :named @p_369)) :named @p_371) (! (=> (! (< 2.0 0.0) :named @p_373) (! (< 0.0 @p_2) :named @p_375)) :named @p_374) (! (not (! (= 2.0 0.0) :named @p_377)) :named @p_378)) :named @p_380)) :named @p_367))) :rule forall_inst :args ((:= veriT_vr13 @p_2) (:= veriT_vr14 2.0) (:= veriT_vr15 0.0))) (anchor :step t150) (assume t150.h1 @p_367) (step t150.t2 (cl (! (= @p_370 true) :named @p_405)) :rule comp_simplify) (step t150.t3 (cl (= @p_371 (! (=> true @p_369) :named @p_372))) :rule cong :premises (t150.t2)) (step t150.t4 (cl (= @p_372 @p_369)) :rule implies_simplify) (step t150.t5 (cl (= @p_371 @p_369)) :rule trans :premises (t150.t3 t150.t4)) (step t150.t6 (cl (! (= @p_373 false) :named @p_408)) :rule comp_simplify) (step t150.t7 (cl (= @p_374 (! (=> false @p_375) :named @p_376))) :rule cong :premises (t150.t6)) (step t150.t8 (cl (= @p_376 true)) :rule implies_simplify) (step t150.t9 (cl (= @p_374 true)) :rule trans :premises (t150.t7 t150.t8)) (step t150.t10 (cl (! (= @p_377 false) :named @p_412)) :rule eq_simplify) (step t150.t11 (cl (! (= @p_378 @p_228) :named @p_413)) :rule cong :premises (t150.t10)) (step t150.t12 (cl @p_379) :rule not_simplify) (step t150.t13 (cl (! (= @p_378 true) :named @p_414)) :rule trans :premises (t150.t11 t150.t12)) (step t150.t14 (cl (= @p_380 (! (and @p_369 true true) :named @p_381))) :rule cong :premises (t150.t5 t150.t9 t150.t13)) (step t150.t15 (cl (= @p_381 (! (and @p_369) :named @p_382))) :rule and_simplify) (step t150.t16 (cl (= @p_382 @p_369)) :rule and_simplify) (step t150.t17 (cl (= @p_380 @p_369)) :rule trans :premises (t150.t14 t150.t15 t150.t16)) (step t150.t18 (cl (! (= @p_367 (! (= @p_368 @p_369) :named @p_383)) :named @p_384)) :rule cong :premises (t150.t17)) (step t150.t19 (cl (not @p_384) (! (not @p_367) :named @p_385) @p_383) :rule equiv_pos2) (step t150.t20 (cl @p_383) :rule th_resolution :premises (t150.h1 t150.t18 t150.t19)) (step t150 (cl @p_385 @p_383) :rule subproof :discharge (h1)) (step t151 (cl @p_282 @p_367) :rule or :premises (t149)) (step t152 (cl (! (or @p_282 @p_383) :named @p_386) @p_302) :rule or_neg) (step t153 (cl @p_386 @p_139) :rule th_resolution :premises (t79 t152)) (step t154 (cl @p_386 (! (not @p_383) :named @p_500)) :rule or_neg) (step t155 (cl @p_386) :rule th_resolution :premises (t151 t150 t153 t154)) (step t156 (cl (or @p_282 (! (= (! (< @p_305 @p_328) :named @p_388) (! (and @p_313 @p_310 @p_317) :named @p_395)) :named @p_387))) :rule forall_inst :args ((:= veriT_vr13 @p_2) (:= veriT_vr14 0.0) (:= veriT_vr15 @p_306))) (anchor :step t157) (assume t157.h1 @p_387) (step t157.t2 (cl @p_331) :rule comp_simplify) (step t157.t3 (cl @p_389) :rule cong :premises (t157.t2)) (step t157.t4 (cl @p_390) :rule implies_simplify) (step t157.t5 (cl @p_391) :rule trans :premises (t157.t3 t157.t4)) (step t157.t6 (cl @p_392) :rule cong :premises (t157.t2)) (step t157.t7 (cl @p_393) :rule implies_simplify) (step t157.t8 (cl @p_394) :rule trans :premises (t157.t6 t157.t7)) (step t157.t9 (cl @p_338) :rule eq_simplify) (step t157.t10 (cl @p_339) :rule cong :premises (t157.t9)) (step t157.t11 (cl @p_340) :rule not_simplify) (step t157.t12 (cl @p_341) :rule trans :premises (t157.t10 t157.t11)) (step t157.t13 (cl (= @p_395 @p_320)) :rule cong :premises (t157.t5 t157.t8 t157.t12)) (step t157.t14 (cl @p_343) :rule and_simplify) (step t157.t15 (cl @p_344) :rule and_simplify) (step t157.t16 (cl (= @p_395 false)) :rule trans :premises (t157.t13 t157.t14 t157.t15)) (step t157.t17 (cl (= @p_387 (! (= @p_388 false) :named @p_396))) :rule cong :premises (t157.t16)) (step t157.t18 (cl (= @p_396 (! (not @p_388) :named @p_397))) :rule equiv_simplify) (step t157.t19 (cl (! (= @p_387 @p_397) :named @p_398)) :rule trans :premises (t157.t17 t157.t18)) (step t157.t20 (cl (not @p_398) (! (not @p_387) :named @p_399) @p_397) :rule equiv_pos2) (step t157.t21 (cl @p_397) :rule th_resolution :premises (t157.h1 t157.t19 t157.t20)) (step t157 (cl @p_399 @p_397) :rule subproof :discharge (h1)) (step t158 (cl @p_282 @p_387) :rule or :premises (t156)) (step t159 (cl (! (or @p_282 @p_397) :named @p_400) @p_302) :rule or_neg) (step t160 (cl @p_400 @p_139) :rule th_resolution :premises (t79 t159)) (step t161 (cl @p_400 (! (not @p_397) :named @p_401)) :rule or_neg) (step t162 (cl (not @p_401) @p_388) :rule not_not) (step t163 (cl @p_400 @p_388) :rule th_resolution :premises (t162 t161)) (step t164 (cl @p_400) :rule th_resolution :premises (t158 t157 t160 t163)) (step t165 (cl (or @p_282 (! (= (! (< @p_260 (! (divide$ @p_36 2.0) :named @p_516)) :named @p_403) (! (and (! (=> @p_370 (! (< @p_3 @p_36) :named @p_404)) :named @p_406) (! (=> @p_373 (! (< @p_36 @p_3) :named @p_410)) :named @p_409) @p_378) :named @p_415)) :named @p_402))) :rule forall_inst :args ((:= veriT_vr13 @p_3) (:= veriT_vr14 2.0) (:= veriT_vr15 @p_36))) (anchor :step t166) (assume t166.h1 @p_402) (step t166.t2 (cl @p_405) :rule comp_simplify) (step t166.t3 (cl (= @p_406 (! (=> true @p_404) :named @p_407))) :rule cong :premises (t166.t2)) (step t166.t4 (cl (= @p_407 @p_404)) :rule implies_simplify) (step t166.t5 (cl (= @p_406 @p_404)) :rule trans :premises (t166.t3 t166.t4)) (step t166.t6 (cl @p_408) :rule comp_simplify) (step t166.t7 (cl (= @p_409 (! (=> false @p_410) :named @p_411))) :rule cong :premises (t166.t6)) (step t166.t8 (cl (= @p_411 true)) :rule implies_simplify) (step t166.t9 (cl (= @p_409 true)) :rule trans :premises (t166.t7 t166.t8)) (step t166.t10 (cl @p_412) :rule eq_simplify) (step t166.t11 (cl @p_413) :rule cong :premises (t166.t10)) (step t166.t12 (cl @p_379) :rule not_simplify) (step t166.t13 (cl @p_414) :rule trans :premises (t166.t11 t166.t12)) (step t166.t14 (cl (= @p_415 (! (and @p_404 true true) :named @p_416))) :rule cong :premises (t166.t5 t166.t9 t166.t13)) (step t166.t15 (cl (= @p_416 (! (and @p_404) :named @p_417))) :rule and_simplify) (step t166.t16 (cl (= @p_417 @p_404)) :rule and_simplify) (step t166.t17 (cl (= @p_415 @p_404)) :rule trans :premises (t166.t14 t166.t15 t166.t16)) (step t166.t18 (cl (! (= @p_402 (! (= @p_403 @p_404) :named @p_418)) :named @p_419)) :rule cong :premises (t166.t17)) (step t166.t19 (cl (not @p_419) (! (not @p_402) :named @p_420) @p_418) :rule equiv_pos2) (step t166.t20 (cl @p_418) :rule th_resolution :premises (t166.h1 t166.t18 t166.t19)) (step t166 (cl @p_420 @p_418) :rule subproof :discharge (h1)) (step t167 (cl @p_282 @p_402) :rule or :premises (t165)) (step t168 (cl (! (or @p_282 @p_418) :named @p_421) @p_302) :rule or_neg) (step t169 (cl @p_421 @p_139) :rule th_resolution :premises (t79 t168)) (step t170 (cl @p_421 (! (not @p_418) :named @p_501)) :rule or_neg) (step t171 (cl @p_421) :rule th_resolution :premises (t167 t166 t169 t170)) (step t172 (cl (or @p_282 (! (= (! (< @p_269 @p_328) :named @p_423) (! (and @p_335 @p_332 @p_317) :named @p_430)) :named @p_422))) :rule forall_inst :args ((:= veriT_vr13 @p_3) (:= veriT_vr14 0.0) (:= veriT_vr15 @p_306))) (anchor :step t173) (assume t173.h1 @p_422) (step t173.t2 (cl @p_331) :rule comp_simplify) (step t173.t3 (cl @p_424) :rule cong :premises (t173.t2)) (step t173.t4 (cl @p_425) :rule implies_simplify) (step t173.t5 (cl @p_426) :rule trans :premises (t173.t3 t173.t4)) (step t173.t6 (cl @p_427) :rule cong :premises (t173.t2)) (step t173.t7 (cl @p_428) :rule implies_simplify) (step t173.t8 (cl @p_429) :rule trans :premises (t173.t6 t173.t7)) (step t173.t9 (cl @p_338) :rule eq_simplify) (step t173.t10 (cl @p_339) :rule cong :premises (t173.t9)) (step t173.t11 (cl @p_340) :rule not_simplify) (step t173.t12 (cl @p_341) :rule trans :premises (t173.t10 t173.t11)) (step t173.t13 (cl (= @p_430 @p_320)) :rule cong :premises (t173.t5 t173.t8 t173.t12)) (step t173.t14 (cl @p_343) :rule and_simplify) (step t173.t15 (cl @p_344) :rule and_simplify) (step t173.t16 (cl (= @p_430 false)) :rule trans :premises (t173.t13 t173.t14 t173.t15)) (step t173.t17 (cl (= @p_422 (! (= @p_423 false) :named @p_431))) :rule cong :premises (t173.t16)) (step t173.t18 (cl (= @p_431 (! (not @p_423) :named @p_432))) :rule equiv_simplify) (step t173.t19 (cl (! (= @p_422 @p_432) :named @p_433)) :rule trans :premises (t173.t17 t173.t18)) (step t173.t20 (cl (not @p_433) (! (not @p_422) :named @p_434) @p_432) :rule equiv_pos2) (step t173.t21 (cl @p_432) :rule th_resolution :premises (t173.h1 t173.t19 t173.t20)) (step t173 (cl @p_434 @p_432) :rule subproof :discharge (h1)) (step t174 (cl @p_282 @p_422) :rule or :premises (t172)) (step t175 (cl (! (or @p_282 @p_432) :named @p_435) @p_302) :rule or_neg) (step t176 (cl @p_435 @p_139) :rule th_resolution :premises (t79 t175)) (step t177 (cl @p_435 (! (not @p_432) :named @p_436)) :rule or_neg) (step t178 (cl (not @p_436) @p_423) :rule not_not) (step t179 (cl @p_435 @p_423) :rule th_resolution :premises (t178 t177)) (step t180 (cl @p_435) :rule th_resolution :premises (t174 t173 t176 t179)) (step t181 (cl (or @p_282 (! (= (! (< @p_437 @p_328) :named @p_439) (! (and @p_356 @p_353 @p_317) :named @p_446)) :named @p_438))) :rule forall_inst :args ((:= veriT_vr13 2.0) (:= veriT_vr14 0.0) (:= veriT_vr15 @p_306))) (anchor :step t182) (assume t182.h1 @p_438) (step t182.t2 (cl @p_331) :rule comp_simplify) (step t182.t3 (cl @p_440) :rule cong :premises (t182.t2)) (step t182.t4 (cl @p_441) :rule implies_simplify) (step t182.t5 (cl @p_442) :rule trans :premises (t182.t3 t182.t4)) (step t182.t6 (cl @p_443) :rule cong :premises (t182.t2)) (step t182.t7 (cl @p_444) :rule implies_simplify) (step t182.t8 (cl @p_445) :rule trans :premises (t182.t6 t182.t7)) (step t182.t9 (cl @p_338) :rule eq_simplify) (step t182.t10 (cl @p_339) :rule cong :premises (t182.t9)) (step t182.t11 (cl @p_340) :rule not_simplify) (step t182.t12 (cl @p_341) :rule trans :premises (t182.t10 t182.t11)) (step t182.t13 (cl (= @p_446 @p_320)) :rule cong :premises (t182.t5 t182.t8 t182.t12)) (step t182.t14 (cl @p_343) :rule and_simplify) (step t182.t15 (cl @p_344) :rule and_simplify) (step t182.t16 (cl (= @p_446 false)) :rule trans :premises (t182.t13 t182.t14 t182.t15)) (step t182.t17 (cl (= @p_438 (! (= @p_439 false) :named @p_447))) :rule cong :premises (t182.t16)) (step t182.t18 (cl (= @p_447 (! (not @p_439) :named @p_448))) :rule equiv_simplify) (step t182.t19 (cl (! (= @p_438 @p_448) :named @p_449)) :rule trans :premises (t182.t17 t182.t18)) (step t182.t20 (cl (not @p_449) (! (not @p_438) :named @p_450) @p_448) :rule equiv_pos2) (step t182.t21 (cl @p_448) :rule th_resolution :premises (t182.h1 t182.t19 t182.t20)) (step t182 (cl @p_450 @p_448) :rule subproof :discharge (h1)) (step t183 (cl @p_282 @p_438) :rule or :premises (t181)) (step t184 (cl (! (or @p_282 @p_448) :named @p_451) @p_302) :rule or_neg) (step t185 (cl @p_451 @p_139) :rule th_resolution :premises (t79 t184)) (step t186 (cl @p_451 (! (not @p_448) :named @p_452)) :rule or_neg) (step t187 (cl (not @p_452) @p_439) :rule not_not) (step t188 (cl @p_451 @p_439) :rule th_resolution :premises (t187 t186)) (step t189 (cl @p_451) :rule th_resolution :premises (t183 t182 t185 t188)) (step t190 (cl (or @p_282 (! (= (! (< @p_275 @p_3) :named @p_455) (! (and (! (=> @p_370 @p_453) :named @p_456) (! (=> @p_373 (! (< d$ 0.0) :named @p_459)) :named @p_458) @p_378) :named @p_461)) :named @p_454))) :rule forall_inst :args ((:= veriT_vr13 0.0) (:= veriT_vr14 2.0) (:= veriT_vr15 d$))) (anchor :step t191) (assume t191.h1 @p_454) (step t191.t2 (cl @p_405) :rule comp_simplify) (step t191.t3 (cl (= @p_456 (! (=> true @p_453) :named @p_457))) :rule cong :premises (t191.t2)) (step t191.t4 (cl (= @p_457 @p_453)) :rule implies_simplify) (step t191.t5 (cl (= @p_456 @p_453)) :rule trans :premises (t191.t3 t191.t4)) (step t191.t6 (cl @p_408) :rule comp_simplify) (step t191.t7 (cl (= @p_458 (! (=> false @p_459) :named @p_460))) :rule cong :premises (t191.t6)) (step t191.t8 (cl (= @p_460 true)) :rule implies_simplify) (step t191.t9 (cl (= @p_458 true)) :rule trans :premises (t191.t7 t191.t8)) (step t191.t10 (cl @p_412) :rule eq_simplify) (step t191.t11 (cl @p_413) :rule cong :premises (t191.t10)) (step t191.t12 (cl @p_379) :rule not_simplify) (step t191.t13 (cl @p_414) :rule trans :premises (t191.t11 t191.t12)) (step t191.t14 (cl (= @p_461 (! (and @p_453 true true) :named @p_462))) :rule cong :premises (t191.t5 t191.t9 t191.t13)) (step t191.t15 (cl (= @p_462 (! (and @p_453) :named @p_463))) :rule and_simplify) (step t191.t16 (cl (= @p_463 @p_453)) :rule and_simplify) (step t191.t17 (cl (= @p_461 @p_453)) :rule trans :premises (t191.t14 t191.t15 t191.t16)) (step t191.t18 (cl (! (= @p_454 (! (= @p_455 @p_453) :named @p_464)) :named @p_465)) :rule cong :premises (t191.t17)) (step t191.t19 (cl (not @p_465) (! (not @p_454) :named @p_466) @p_464) :rule equiv_pos2) (step t191.t20 (cl @p_464) :rule th_resolution :premises (t191.h1 t191.t18 t191.t19)) (step t191 (cl @p_466 @p_464) :rule subproof :discharge (h1)) (step t192 (cl @p_282 @p_454) :rule or :premises (t190)) (step t193 (cl (! (or @p_282 @p_464) :named @p_467) @p_302) :rule or_neg) (step t194 (cl @p_467 @p_139) :rule th_resolution :premises (t79 t193)) (step t195 (cl @p_467 (! (not @p_464) :named @p_502)) :rule or_neg) (step t196 (cl @p_467) :rule th_resolution :premises (t192 t191 t194 t195)) (step t197 (cl (not (! (not (! (not @p_84) :named @p_468)) :named @p_479)) @p_84) :rule not_not) (step t198 (cl (or @p_468 (! (= @p_36 (! (- @p_3 (! (ite @p_369 @p_258 @p_2) :named @p_471)) :named @p_472)) :named @p_469))) :rule forall_inst :args ((:= veriT_vr9 @p_2))) (anchor :step t199) (assume t199.h1 @p_469) (step t199.t2 (cl @p_267) :rule minus_simplify) (step t199.t3 (cl (= @p_471 (! (ite @p_369 @p_3 @p_2) :named @p_470))) :rule cong :premises (t199.t2)) (step t199.t4 (cl (= @p_472 (! (- @p_3 @p_470) :named @p_473))) :rule cong :premises (t199.t3)) (step t199.t5 (cl (! (= @p_469 (! (= @p_36 @p_473) :named @p_476)) :named @p_474)) :rule cong :premises (t199.t4)) (step t199.t6 (cl (not @p_474) (! (not @p_469) :named @p_475) @p_476) :rule equiv_pos2) (step t199.t7 (cl @p_476) :rule th_resolution :premises (t199.h1 t199.t5 t199.t6)) (step t199.t8 (cl (! (= @p_476 (! (and (! (= @p_36 (- @p_3 @p_470)) :named @p_504) (! (ite @p_369 (! (= @p_3 @p_470) :named @p_507) (= @p_2 @p_470)) :named @p_505)) :named @p_477)) :named @p_478)) :rule ite_intro) (step t199.t9 (cl (not @p_478) (not @p_476) @p_477) :rule equiv_pos2) (step t199.t10 (cl @p_477) :rule th_resolution :premises (t199.t7 t199.t8 t199.t9)) (step t199 (cl @p_475 @p_477) :rule subproof :discharge (h1)) (step t200 (cl @p_468 @p_469) :rule or :premises (t198)) (step t201 (cl (! (or @p_468 @p_477) :named @p_480) @p_479) :rule or_neg) (step t202 (cl @p_480 @p_84) :rule th_resolution :premises (t197 t201)) (step t203 (cl @p_480 (! (not @p_477) :named @p_503)) :rule or_neg) (step t204 (cl @p_480) :rule th_resolution :premises (t200 t199 t202 t203)) (step t205 (cl (or (! (not @p_60) :named @p_508) (! (= (! (fun_app$ uub$ @p_7) :named @p_543) (! (pair$ @p_2 (! (times$ (- (* 2.0 @p_7) 1.0) @p_36) :named @p_535)) :named @p_534)) :named @p_509))) :rule forall_inst :args ((:= veriT_vr7 @p_7))) (step t206 (cl (or (! (not @p_34) :named @p_492) (! (= (! (fun_app$ uuc$ @p_7) :named @p_483) (! (pair$ (! (times$ (! (- @p_7 @p_7) :named @p_484) d$) :named @p_481) (! (diamond_y$ @p_481) :named @p_485)) :named @p_487)) :named @p_482))) :rule forall_inst :args ((:= veriT_vr1 @p_7))) (anchor :step t207) (assume t207.h1 @p_482) (step t207.t2 (cl (= 0.0 @p_484)) :rule minus_simplify) (step t207.t3 (cl (= @p_247 @p_481)) :rule cong :premises (t207.t2)) (step t207.t4 (cl (= @p_485 (! (diamond_y$ @p_247) :named @p_486))) :rule cong :premises (t207.t3)) (step t207.t5 (cl (= @p_487 (! (pair$ @p_247 @p_486) :named @p_488))) :rule cong :premises (t207.t3 t207.t4)) (step t207.t6 (cl (! (= @p_482 (! (= @p_483 @p_488) :named @p_489)) :named @p_490)) :rule cong :premises (t207.t5)) (step t207.t7 (cl (not @p_490) (! (not @p_482) :named @p_491) @p_489) :rule equiv_pos2) (step t207.t8 (cl @p_489) :rule th_resolution :premises (t207.h1 t207.t6 t207.t7)) (step t207 (cl @p_491 @p_489) :rule subproof :discharge (h1)) (step t208 (cl @p_492 @p_482) :rule or :premises (t206)) (step t209 (cl (! (or @p_492 @p_489) :named @p_494) (! (not @p_492) :named @p_493)) :rule or_neg) (step t210 (cl (not @p_493) @p_34) :rule not_not) (step t211 (cl @p_494 @p_34) :rule th_resolution :premises (t210 t209)) (step t212 (cl @p_494 (! (not @p_489) :named @p_544)) :rule or_neg) (step t213 (cl @p_494) :rule th_resolution :premises (t208 t207 t211 t212)) (step t214 (cl @p_245 @p_251) :rule or :premises (t90)) (step t215 (cl @p_251) :rule resolution :premises (t214 t48)) (step t216 (cl @p_241 @p_261) :rule or :premises (t97)) (step t217 (cl @p_261) :rule resolution :premises (t216 t42)) (step t218 (cl @p_241 @p_270) :rule or :premises (t104)) (step t219 (cl @p_270) :rule resolution :premises (t218 t42)) (step t220 (cl @p_241 @p_495) :rule or :premises (t105)) (step t221 (cl @p_495) :rule resolution :premises (t220 t42)) (step t222 (cl @p_241 @p_278) :rule or :premises (t112)) (step t223 (cl @p_278) :rule resolution :premises (t222 t42)) (step t224 (cl @p_297 (! (not @p_292) :named @p_496) (! (not @p_295) :named @p_497) (! (not @p_288) :named @p_498)) :rule and_neg) (step t225 (cl (not @p_496) @p_286) :rule not_not) (step t226 (cl @p_297 @p_286 @p_497 @p_498) :rule th_resolution :premises (t225 t224)) (step t227 (cl (not @p_497) @p_287) :rule not_not) (step t228 (cl @p_297 @p_286 @p_287 @p_498) :rule th_resolution :premises (t227 t226)) (step t229 (cl (not @p_498) @p_499) :rule not_not) (step t230 (cl @p_297 @p_286 @p_287 @p_499) :rule th_resolution :premises (t229 t228)) (step t231 (cl @p_282 @p_299) :rule or :premises (t121)) (step t232 (cl @p_299) :rule resolution :premises (t231 t36)) (step t233 (cl @p_282 @p_323) :rule or :premises (t130)) (step t234 (cl @p_323) :rule resolution :premises (t233 t36)) (step t235 (cl @p_282 @p_346) :rule or :premises (t139)) (step t236 (cl @p_346) :rule resolution :premises (t235 t36)) (step t237 (cl @p_282 @p_361) :rule or :premises (t148)) (step t238 (cl @p_361) :rule resolution :premises (t237 t36)) (step t239 (cl @p_500 @p_368 (! (not @p_369) :named @p_506)) :rule equiv_pos1) (step t240 (cl @p_282 @p_383) :rule or :premises (t155)) (step t241 (cl @p_383) :rule resolution :premises (t240 t36)) (step t242 (cl @p_282 @p_397) :rule or :premises (t164)) (step t243 (cl @p_397) :rule resolution :premises (t242 t36)) (step t244 (cl @p_501 (! (not @p_403) :named @p_519) @p_404) :rule equiv_pos2) (step t245 (cl @p_282 @p_418) :rule or :premises (t171)) (step t246 (cl @p_418) :rule resolution :premises (t245 t36)) (step t247 (cl @p_282 @p_432) :rule or :premises (t180)) (step t248 (cl @p_432) :rule resolution :premises (t247 t36)) (step t249 (cl @p_282 @p_448) :rule or :premises (t189)) (step t250 (cl @p_448) :rule resolution :premises (t249 t36)) (step t251 (cl @p_502 @p_455 (not @p_453)) :rule equiv_pos1) (step t252 (cl @p_282 @p_464) :rule or :premises (t196)) (step t253 (cl @p_502 @p_455) :rule resolution :premises (t251 axiom4)) (step t254 (cl @p_464) :rule resolution :premises (t252 t36)) (step t255 (cl @p_455) :rule resolution :premises (t253 t254)) (step t256 (cl @p_503 @p_504) :rule and_pos) (step t257 (cl (not @p_505) @p_506 @p_507) :rule ite_pos2) (step t258 (cl @p_503 @p_505) :rule and_pos) (step t259 (cl @p_468 @p_477) :rule or :premises (t204)) (step t260 (cl @p_477) :rule resolution :premises (t259 t27)) (step t261 (cl @p_504) :rule resolution :premises (t256 t260)) (step t262 (cl @p_505) :rule resolution :premises (t258 t260)) (step t263 (cl @p_508 @p_509) :rule or :premises (t205)) (step t264 (cl @p_509) :rule resolution :premises (t263 t21)) (step t265 (cl @p_492 @p_489) :rule or :premises (t213)) (step t266 (cl @p_489) :rule resolution :premises (t265 t15)) (step t267 (cl (! (= 2.0 2.0) :named @p_515)) :rule eq_reflexive) (step t268 (cl (not @p_455) @p_369 @p_510) :rule la_generic :args (2.0 2.0 1)) (step t269 (cl @p_369) :rule resolution :premises (t268 t223 t255)) (step t270 (cl @p_368) :rule resolution :premises (t239 t269 t241)) (step t271 (cl @p_507) :rule resolution :premises (t257 t269 t262)) (step t272 (cl (! (= @p_2 @p_2) :named @p_513)) :rule eq_reflexive) (step t273 (cl @p_295 (! (not @p_504) :named @p_511) (! (not @p_507) :named @p_512)) :rule la_generic :args (1.0 1 1)) (step t274 (cl @p_295) :rule resolution :premises (t273 t261 t271)) (step t275 (cl @p_292 @p_511 @p_512) :rule la_generic :args (1.0 (- 1) (- 1))) (step t276 (cl @p_292) :rule resolution :premises (t275 t261 t271)) (step t277 (cl @p_499) :rule resolution :premises (t230 t276 t232 t274)) (step t278 (cl (not @p_513) @p_288 (! (< @p_2 @p_36) :named @p_514) @p_506) :rule eq_congruent_pred) (step t279 (cl @p_288 @p_514 @p_506) :rule th_resolution :premises (t278 t272)) (step t280 (cl @p_514) :rule resolution :premises (t279 t277 t269)) (step t281 (cl @p_288 (! (not @p_515) :named @p_551) (! (= @p_516 @p_275) :named @p_517)) :rule eq_congruent) (step t282 (cl @p_288 @p_517) :rule th_resolution :premises (t281 t267)) (step t283 (cl (! (= @p_7 @p_7) :named @p_546)) :rule eq_reflexive) (step t284 (cl (! (not @p_404) :named @p_518) (not @p_514) @p_511 @p_512) :rule la_generic :args (1.0 1.0 (- 2) (- 2))) (step t285 (cl @p_518) :rule resolution :premises (t284 t280 t261 t271)) (step t286 (cl @p_519) :rule resolution :premises (t244 t285 t246)) (step t287 (cl (or (! (= @p_234 @p_437) :named @p_520) (! (not (! (<= @p_234 @p_437) :named @p_529)) :named @p_521) (! (not (! (<= @p_437 @p_234) :named @p_526)) :named @p_522))) :rule la_disequality) (step t288 (cl @p_520 @p_521 @p_522) :rule or :premises (t287)) (step t289 (cl (or (! (= @p_247 @p_437) :named @p_523) (! (not (! (<= @p_247 @p_437) :named @p_532)) :named @p_524) (! (not (! (<= @p_437 @p_247) :named @p_530)) :named @p_525))) :rule la_disequality) (step t290 (cl @p_523 @p_524 @p_525) :rule or :premises (t289)) (step t291 (cl @p_388 @p_423 @p_352 @p_526 @p_527 @p_528) :rule la_generic :args (1.0 1.0 2.0 2.0 1 1)) (step t292 (cl @p_526) :rule resolution :premises (t291 t81 t219 t238 t243 t248)) (step t293 (cl @p_309 @p_330 @p_439 @p_529 @p_527 @p_528) :rule la_generic :args (1.0 1.0 2.0 2.0 (- 1) (- 1))) (step t294 (cl @p_529) :rule resolution :premises (t293 t81 t219 t234 t236 t250)) (step t295 (cl @p_520) :rule resolution :premises (t288 t294 t292)) (step t296 (cl @p_388 @p_423 @p_352 @p_530 @p_531 @p_528) :rule la_generic :args (1.0 1.0 2.0 2.0 1 1)) (step t297 (cl @p_530) :rule resolution :premises (t296 t215 t219 t238 t243 t248)) (step t298 (cl @p_309 @p_330 @p_439 @p_532 @p_531 @p_528) :rule la_generic :args (1.0 1.0 2.0 2.0 (- 1) (- 1))) (step t299 (cl @p_532) :rule resolution :premises (t298 t215 t219 t234 t236 t250)) (step t300 (cl @p_523) :rule resolution :premises (t290 t299 t297)) (step t301 (cl @p_244 @p_533) :rule or :premises (t70)) (step t302 (cl (or (! (not @p_533) :named @p_536) (! (or (! (not (! (= @p_534 @p_488) :named @p_545)) :named @p_541) (! (= @p_2 @p_247) :named @p_542)) :named @p_537))) :rule forall_inst :args ((:= veriT_vr28 @p_247) (:= veriT_vr29 @p_486) (:= veriT_vr30 @p_2) (:= veriT_vr31 @p_535))) (step t303 (cl @p_536 @p_537) :rule or :premises (t302)) (step t304 (cl (! (or @p_244 @p_537) :named @p_539) @p_538) :rule or_neg) (step t305 (cl @p_539 @p_218) :rule th_resolution :premises (t82 t304)) (step t306 (cl @p_539 (! (not @p_537) :named @p_540)) :rule or_neg) (step t307 (cl @p_539) :rule th_resolution :premises (t301 t303 t305 t306)) (step t308 (cl @p_540 @p_541 @p_542) :rule or_pos) (step t309 (cl @p_244 @p_537) :rule or :premises (t307)) (step t310 (cl @p_537) :rule resolution :premises (t309 t54)) (step t311 (cl (! (not @p_509) :named @p_549) (not (! (= @p_543 @p_483) :named @p_547)) @p_544 @p_545) :rule eq_transitive) (step t312 (cl (! (not @p_227) :named @p_548) (not @p_546) @p_547) :rule eq_congruent) (step t313 (cl @p_548 @p_547) :rule th_resolution :premises (t312 t283)) (step t314 (cl @p_549 @p_544 @p_545 @p_548) :rule th_resolution :premises (t311 t313)) (step t315 (cl @p_545) :rule resolution :premises (t314 t69 t264 t266)) (step t316 (cl @p_542) :rule resolution :premises (t308 t315 t310)) (step t317 (cl (! (not @p_542) :named @p_555) (! (not @p_523) :named @p_556) (! (not @p_520) :named @p_557) (! (= @p_2 @p_234) :named @p_550)) :rule eq_transitive) (step t318 (cl (! (not @p_550) :named @p_553) @p_551 (! (= @p_366 @p_552) :named @p_554)) :rule eq_congruent) (step t319 (cl @p_553 @p_554) :rule th_resolution :premises (t318 t267)) (step t320 (cl @p_554 @p_555 @p_556 @p_557) :rule th_resolution :premises (t319 t317)) (step t321 (cl @p_555 @p_556 @p_557 @p_527 (! (= @p_2 @p_235) :named @p_565)) :rule eq_transitive) (step t322 (cl (not @p_554) (! (= @p_257 @p_558) :named @p_559)) :rule eq_congruent) (step t323 (cl @p_559 @p_555 @p_556 @p_557) :rule th_resolution :premises (t322 t320)) (step t324 (cl (! (not (! (= @p_366 @p_260) :named @p_564)) :named @p_560) (not @p_517) (! (not @p_368) :named @p_561) @p_403) :rule eq_congruent_pred) (step t325 (cl @p_560 @p_561 @p_403 @p_288) :rule th_resolution :premises (t324 t282)) (step t326 (cl (not (! (= @p_366 @p_562) :named @p_566)) (! (not @p_495) :named @p_568) (! (not @p_559) :named @p_569) @p_563 @p_564) :rule eq_transitive) (step t327 (cl (! (not @p_565) :named @p_567) @p_551 @p_566) :rule eq_congruent) (step t328 (cl @p_567 @p_566) :rule th_resolution :premises (t327 t267)) (step t329 (cl @p_566 @p_555 @p_556 @p_557 @p_527) :rule th_resolution :premises (t328 t321)) (step t330 (cl @p_568 @p_569 @p_563 @p_564 @p_555 @p_556 @p_557 @p_527) :rule th_resolution :premises (t326 t329)) (step t331 (cl @p_568 @p_563 @p_564 @p_555 @p_556 @p_557 @p_527 @p_555 @p_556 @p_557) :rule th_resolution :premises (t330 t323)) (step t332 (cl @p_568 @p_563 @p_564 @p_555 @p_556 @p_557 @p_527) :rule contraction :premises (t331)) (step t333 (cl @p_561 @p_403 @p_288 @p_568 @p_563 @p_555 @p_556 @p_557 @p_527) :rule th_resolution :premises (t325 t332)) (step t334 (cl) :rule resolution :premises (t333 t81 t217 t221 t277 t270 t286 t295 t300 t316)) 6120eaf40c2621e298051bc401bc258d4c6ef4d6 323 0 unsat (assume axiom1 (! (not (! (=> (! (forall ((?v0 Real_a_fun$) (?v1 B_list$)) (! (= (! (=> (! (and (! (= (! (rec_join$ ?v1) :named @p_3) ?v0) :named @p_68) (! (and (! (=> (! (and (! (= ?v1 nil$) :named @p_4) (! (= uu$ ?v0) :named @p_72)) :named @p_74) false) :named @p_76) (! (and (! (forall ((?v2 B$)) (! (=> (! (and (! (= ?v1 (! (cons$ ?v2 nil$) :named @p_8)) :named @p_5) (! (= ?v0 (! (coeff_cube_to_path$ ?v2) :named @p_1)) :named @p_82)) :named @p_84) false) :named @p_86)) :named @p_78) (! (forall ((?v2 B$) (?v3 B$) (?v4 B_list$)) (! (=> (! (and (! (= ?v1 (! (cons$ ?v2 (! (cons$ ?v3 ?v4) :named @p_2)) :named @p_9)) :named @p_6) (! (= ?v0 (! (joinpaths$ @p_1 (! (rec_join$ @p_2) :named @p_95)) :named @p_7)) :named @p_97)) :named @p_99) false) :named @p_101)) :named @p_88)) :named @p_103)) :named @p_105)) :named @p_107) false) :named @p_109) (! (=> (! (and (! (= @p_3 @p_3) :named @p_112) (! (and (! (=> (! (and @p_4 (! (= uu$ @p_3) :named @p_115)) :named @p_117) false) :named @p_119) (! (and (! (forall ((?v2 B$)) (! (=> (! (and @p_5 (! (= @p_3 @p_1) :named @p_125)) :named @p_127) false) :named @p_129)) :named @p_121) (! (forall ((?v2 B$) (?v3 B$) (?v4 B_list$)) (! (=> (! (and @p_6 (! (= @p_3 @p_7) :named @p_137)) :named @p_139) false) :named @p_141)) :named @p_131)) :named @p_143)) :named @p_145)) :named @p_147) false) :named @p_149)) :named @p_151)) :named @p_53) (! (= (! (forall ((?v0 B_list$) (?v1 Real_a_fun$)) (! (=> (! (and (! (= (! (rec_join$ ?v0) :named @p_10) ?v1) :named @p_19) (! (and (! (=> (! (and (! (= nil$ ?v0) :named @p_11) (! (= uu$ ?v1) :named @p_20)) :named @p_22) false) :named @p_24) (! (and (! (forall ((?v2 B$)) (! (=> (! (and (! (= @p_8 ?v0) :named @p_17) (! (= @p_1 ?v1) :named @p_27)) :named @p_29) false) :named @p_31)) :named @p_25) (! (forall ((?v2 B$) (?v3 B$) (?v4 B_list$)) (! (=> (! (and (! (= @p_9 ?v0) :named @p_18) (! (= @p_7 ?v1) :named @p_35)) :named @p_37) false) :named @p_39)) :named @p_33)) :named @p_41)) :named @p_43)) :named @p_45) false) :named @p_47)) :named @p_14) (! (forall ((?v0 B_list$)) (! (=> (! (and (! (= @p_10 @p_10) :named @p_15) (! (and (! (=> (! (and @p_11 (! (= uu$ @p_10) :named @p_21)) :named @p_23) false) :named @p_16) (! (and (! (forall ((?v1 B$)) (! (=> (! (and (! (= ?v0 (! (cons$ ?v1 nil$) :named @p_162)) :named @p_163) (! (= @p_10 (! (coeff_cube_to_path$ ?v1) :named @p_12)) :named @p_165)) :named @p_166) false) :named @p_167)) :named @p_161) (! (forall ((?v1 B$) (?v2 B$) (?v3 B_list$)) (! (=> (! (and (! (= ?v0 (! (cons$ ?v1 (! (cons$ ?v2 ?v3) :named @p_13)) :named @p_169)) :named @p_170) (! (= @p_10 (! (joinpaths$ @p_12 (! (rec_join$ @p_13) :named @p_175)) :named @p_176)) :named @p_177)) :named @p_178) false) :named @p_179)) :named @p_168)) :named @p_180)) :named @p_181)) :named @p_182) false) :named @p_183)) :named @p_51)) :named @p_49)) :named @p_52)) :named @p_55)) (anchor :step t2 :args ((?v0 B_list$) (:= (?v1 Real_a_fun$) @p_10))) (step t2.t1 (cl @p_19) :rule refl) (step t2.t2 (cl (= @p_19 @p_15)) :rule cong :premises (t2.t1)) (step t2.t3 (cl @p_19) :rule refl) (step t2.t4 (cl (= @p_20 @p_21)) :rule cong :premises (t2.t3)) (step t2.t5 (cl (= @p_22 @p_23)) :rule cong :premises (t2.t4)) (step t2.t6 (cl (= @p_24 @p_16)) :rule cong :premises (t2.t5)) (anchor :step t2.t7 :args ((?v2 B$))) (step t2.t7.t1 (cl @p_19) :rule refl) (step t2.t7.t2 (cl (= @p_27 (! (= @p_1 @p_10) :named @p_28))) :rule cong :premises (t2.t7.t1)) (step t2.t7.t3 (cl (= @p_29 (! (and @p_17 @p_28) :named @p_30))) :rule cong :premises (t2.t7.t2)) (step t2.t7.t4 (cl (= @p_31 (! (=> @p_30 false) :named @p_32))) :rule cong :premises (t2.t7.t3)) (step t2.t7 (cl (= @p_25 (! (forall ((?v2 B$)) @p_32) :named @p_26))) :rule bind) (anchor :step t2.t8 :args ((?v2 B$) (?v3 B$) (?v4 B_list$))) (step t2.t8.t1 (cl @p_19) :rule refl) (step t2.t8.t2 (cl (= @p_35 (! (= @p_7 @p_10) :named @p_36))) :rule cong :premises (t2.t8.t1)) (step t2.t8.t3 (cl (= @p_37 (! (and @p_18 @p_36) :named @p_38))) :rule cong :premises (t2.t8.t2)) (step t2.t8.t4 (cl (= @p_39 (! (=> @p_38 false) :named @p_40))) :rule cong :premises (t2.t8.t3)) (step t2.t8 (cl (= @p_33 (! (forall ((?v2 B$) (?v3 B$) (?v4 B_list$)) @p_40) :named @p_34))) :rule bind) (step t2.t9 (cl (= @p_41 (! (and @p_26 @p_34) :named @p_42))) :rule cong :premises (t2.t7 t2.t8)) (step t2.t10 (cl (= @p_43 (! (and @p_16 @p_42) :named @p_44))) :rule cong :premises (t2.t6 t2.t9)) (step t2.t11 (cl (= @p_45 (! (and @p_15 @p_44) :named @p_46))) :rule cong :premises (t2.t2 t2.t10)) (step t2.t12 (cl (= @p_47 (! (=> @p_46 false) :named @p_48))) :rule cong :premises (t2.t11)) (step t2 (cl (= @p_14 (! (forall ((?v0 B_list$)) @p_48) :named @p_50))) :rule onepoint) (step t3 (cl (= @p_49 (! (= @p_50 @p_51) :named @p_54))) :rule cong :premises (t2)) (step t4 (cl (= @p_52 (! (=> @p_53 @p_54) :named @p_56))) :rule cong :premises (t3)) (step t5 (cl (! (= @p_55 (! (not @p_56) :named @p_58)) :named @p_57)) :rule cong :premises (t4)) (step t6 (cl (! (not @p_57) :named @p_60) (! (not @p_55) :named @p_59) @p_58) :rule equiv_pos2) (step t7 (cl (not @p_59) @p_52) :rule not_not) (step t8 (cl @p_60 @p_52 @p_58) :rule th_resolution :premises (t7 t6)) (step t9 (cl @p_58) :rule th_resolution :premises (axiom1 t5 t8)) (anchor :step t10 :args ((:= (?v0 Real_a_fun$) veriT_vr2) (:= (?v1 B_list$) veriT_vr3))) (step t10.t1 (cl (! (= ?v1 veriT_vr3) :named @p_70)) :rule refl) (step t10.t2 (cl (! (= @p_3 (! (rec_join$ veriT_vr3) :named @p_63)) :named @p_111)) :rule cong :premises (t10.t1)) (step t10.t3 (cl (! (= ?v0 veriT_vr2) :named @p_71)) :rule refl) (step t10.t4 (cl (= @p_68 (! (= veriT_vr2 @p_63) :named @p_69))) :rule cong :premises (t10.t2 t10.t3)) (step t10.t5 (cl @p_70) :rule refl) (step t10.t6 (cl (! (= @p_4 (! (= nil$ veriT_vr3) :named @p_64)) :named @p_114)) :rule cong :premises (t10.t5)) (step t10.t7 (cl @p_71) :rule refl) (step t10.t8 (cl (= @p_72 (! (= uu$ veriT_vr2) :named @p_73))) :rule cong :premises (t10.t7)) (step t10.t9 (cl (= @p_74 (! (and @p_64 @p_73) :named @p_75))) :rule cong :premises (t10.t6 t10.t8)) (step t10.t10 (cl (= @p_76 (! (=> @p_75 false) :named @p_77))) :rule cong :premises (t10.t9)) (anchor :step t10.t11 :args ((:= (?v2 B$) veriT_vr4))) (step t10.t11.t1 (cl @p_70) :rule refl) (step t10.t11.t2 (cl (! (= ?v2 veriT_vr4) :named @p_81)) :rule refl) (step t10.t11.t3 (cl (! (= @p_8 (! (cons$ veriT_vr4 nil$) :named @p_80)) :named @p_123)) :rule cong :premises (t10.t11.t2)) (step t10.t11.t4 (cl (! (= @p_5 (! (= veriT_vr3 @p_80) :named @p_65)) :named @p_124)) :rule cong :premises (t10.t11.t1 t10.t11.t3)) (step t10.t11.t5 (cl @p_71) :rule refl) (step t10.t11.t6 (cl @p_81) :rule refl) (step t10.t11.t7 (cl (! (= @p_1 (! (coeff_cube_to_path$ veriT_vr4) :named @p_61)) :named @p_91)) :rule cong :premises (t10.t11.t6)) (step t10.t11.t8 (cl (= @p_82 (! (= veriT_vr2 @p_61) :named @p_83))) :rule cong :premises (t10.t11.t5 t10.t11.t7)) (step t10.t11.t9 (cl (= @p_84 (! (and @p_65 @p_83) :named @p_85))) :rule cong :premises (t10.t11.t4 t10.t11.t8)) (step t10.t11.t10 (cl (= @p_86 (! (=> @p_85 false) :named @p_87))) :rule cong :premises (t10.t11.t9)) (step t10.t11 (cl (= @p_78 (! (forall ((veriT_vr4 B$)) @p_87) :named @p_79))) :rule bind) (anchor :step t10.t12 :args ((:= (?v2 B$) veriT_vr4) (:= (?v3 B$) veriT_vr5) (:= (?v4 B_list$) veriT_vr6))) (step t10.t12.t1 (cl @p_70) :rule refl) (step t10.t12.t2 (cl @p_81) :rule refl) (step t10.t12.t3 (cl (! (= ?v3 veriT_vr5) :named @p_92)) :rule refl) (step t10.t12.t4 (cl (! (= ?v4 veriT_vr6) :named @p_93)) :rule refl) (step t10.t12.t5 (cl (! (= @p_2 (! (cons$ veriT_vr5 veriT_vr6) :named @p_62)) :named @p_94)) :rule cong :premises (t10.t12.t3 t10.t12.t4)) (step t10.t12.t6 (cl (! (= @p_9 (! (cons$ veriT_vr4 @p_62) :named @p_90)) :named @p_133)) :rule cong :premises (t10.t12.t2 t10.t12.t5)) (step t10.t12.t7 (cl (! (= @p_6 (! (= veriT_vr3 @p_90) :named @p_66)) :named @p_134)) :rule cong :premises (t10.t12.t1 t10.t12.t6)) (step t10.t12.t8 (cl @p_71) :rule refl) (step t10.t12.t9 (cl @p_81) :rule refl) (step t10.t12.t10 (cl @p_91) :rule cong :premises (t10.t12.t9)) (step t10.t12.t11 (cl @p_92) :rule refl) (step t10.t12.t12 (cl @p_93) :rule refl) (step t10.t12.t13 (cl @p_94) :rule cong :premises (t10.t12.t11 t10.t12.t12)) (step t10.t12.t14 (cl (! (= @p_95 (! (rec_join$ @p_62) :named @p_96)) :named @p_135)) :rule cong :premises (t10.t12.t13)) (step t10.t12.t15 (cl (! (= @p_7 (! (joinpaths$ @p_61 @p_96) :named @p_67)) :named @p_136)) :rule cong :premises (t10.t12.t10 t10.t12.t14)) (step t10.t12.t16 (cl (= @p_97 (! (= veriT_vr2 @p_67) :named @p_98))) :rule cong :premises (t10.t12.t8 t10.t12.t15)) (step t10.t12.t17 (cl (= @p_99 (! (and @p_66 @p_98) :named @p_100))) :rule cong :premises (t10.t12.t7 t10.t12.t16)) (step t10.t12.t18 (cl (= @p_101 (! (=> @p_100 false) :named @p_102))) :rule cong :premises (t10.t12.t17)) (step t10.t12 (cl (= @p_88 (! (forall ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$)) @p_102) :named @p_89))) :rule bind) (step t10.t13 (cl (= @p_103 (! (and @p_79 @p_89) :named @p_104))) :rule cong :premises (t10.t11 t10.t12)) (step t10.t14 (cl (= @p_105 (! (and @p_77 @p_104) :named @p_106))) :rule cong :premises (t10.t10 t10.t13)) (step t10.t15 (cl (= @p_107 (! (and @p_69 @p_106) :named @p_108))) :rule cong :premises (t10.t4 t10.t14)) (step t10.t16 (cl (= @p_109 (! (=> @p_108 false) :named @p_110))) :rule cong :premises (t10.t15)) (step t10.t17 (cl @p_70) :rule refl) (step t10.t18 (cl @p_111) :rule cong :premises (t10.t17)) (step t10.t19 (cl @p_70) :rule refl) (step t10.t20 (cl @p_111) :rule cong :premises (t10.t19)) (step t10.t21 (cl (= @p_112 (! (= @p_63 @p_63) :named @p_113))) :rule cong :premises (t10.t18 t10.t20)) (step t10.t22 (cl @p_70) :rule refl) (step t10.t23 (cl @p_114) :rule cong :premises (t10.t22)) (step t10.t24 (cl @p_70) :rule refl) (step t10.t25 (cl @p_111) :rule cong :premises (t10.t24)) (step t10.t26 (cl (= @p_115 (! (= uu$ @p_63) :named @p_116))) :rule cong :premises (t10.t25)) (step t10.t27 (cl (= @p_117 (! (and @p_64 @p_116) :named @p_118))) :rule cong :premises (t10.t23 t10.t26)) (step t10.t28 (cl (= @p_119 (! (=> @p_118 false) :named @p_120))) :rule cong :premises (t10.t27)) (anchor :step t10.t29 :args ((:= (?v2 B$) veriT_vr4))) (step t10.t29.t1 (cl @p_70) :rule refl) (step t10.t29.t2 (cl @p_81) :rule refl) (step t10.t29.t3 (cl @p_123) :rule cong :premises (t10.t29.t2)) (step t10.t29.t4 (cl @p_124) :rule cong :premises (t10.t29.t1 t10.t29.t3)) (step t10.t29.t5 (cl @p_70) :rule refl) (step t10.t29.t6 (cl @p_111) :rule cong :premises (t10.t29.t5)) (step t10.t29.t7 (cl @p_81) :rule refl) (step t10.t29.t8 (cl @p_91) :rule cong :premises (t10.t29.t7)) (step t10.t29.t9 (cl (= @p_125 (! (= @p_63 @p_61) :named @p_126))) :rule cong :premises (t10.t29.t6 t10.t29.t8)) (step t10.t29.t10 (cl (= @p_127 (! (and @p_65 @p_126) :named @p_128))) :rule cong :premises (t10.t29.t4 t10.t29.t9)) (step t10.t29.t11 (cl (= @p_129 (! (=> @p_128 false) :named @p_130))) :rule cong :premises (t10.t29.t10)) (step t10.t29 (cl (= @p_121 (! (forall ((veriT_vr4 B$)) @p_130) :named @p_122))) :rule bind) (anchor :step t10.t30 :args ((:= (?v2 B$) veriT_vr4) (:= (?v3 B$) veriT_vr5) (:= (?v4 B_list$) veriT_vr6))) (step t10.t30.t1 (cl @p_70) :rule refl) (step t10.t30.t2 (cl @p_81) :rule refl) (step t10.t30.t3 (cl @p_92) :rule refl) (step t10.t30.t4 (cl @p_93) :rule refl) (step t10.t30.t5 (cl @p_94) :rule cong :premises (t10.t30.t3 t10.t30.t4)) (step t10.t30.t6 (cl @p_133) :rule cong :premises (t10.t30.t2 t10.t30.t5)) (step t10.t30.t7 (cl @p_134) :rule cong :premises (t10.t30.t1 t10.t30.t6)) (step t10.t30.t8 (cl @p_70) :rule refl) (step t10.t30.t9 (cl @p_111) :rule cong :premises (t10.t30.t8)) (step t10.t30.t10 (cl @p_81) :rule refl) (step t10.t30.t11 (cl @p_91) :rule cong :premises (t10.t30.t10)) (step t10.t30.t12 (cl @p_92) :rule refl) (step t10.t30.t13 (cl @p_93) :rule refl) (step t10.t30.t14 (cl @p_94) :rule cong :premises (t10.t30.t12 t10.t30.t13)) (step t10.t30.t15 (cl @p_135) :rule cong :premises (t10.t30.t14)) (step t10.t30.t16 (cl @p_136) :rule cong :premises (t10.t30.t11 t10.t30.t15)) (step t10.t30.t17 (cl (= @p_137 (! (= @p_63 @p_67) :named @p_138))) :rule cong :premises (t10.t30.t9 t10.t30.t16)) (step t10.t30.t18 (cl (= @p_139 (! (and @p_66 @p_138) :named @p_140))) :rule cong :premises (t10.t30.t7 t10.t30.t17)) (step t10.t30.t19 (cl (= @p_141 (! (=> @p_140 false) :named @p_142))) :rule cong :premises (t10.t30.t18)) (step t10.t30 (cl (= @p_131 (! (forall ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$)) @p_142) :named @p_132))) :rule bind) (step t10.t31 (cl (= @p_143 (! (and @p_122 @p_132) :named @p_144))) :rule cong :premises (t10.t29 t10.t30)) (step t10.t32 (cl (= @p_145 (! (and @p_120 @p_144) :named @p_146))) :rule cong :premises (t10.t28 t10.t31)) (step t10.t33 (cl (= @p_147 (! (and @p_113 @p_146) :named @p_148))) :rule cong :premises (t10.t21 t10.t32)) (step t10.t34 (cl (= @p_149 (! (=> @p_148 false) :named @p_150))) :rule cong :premises (t10.t33)) (step t10.t35 (cl (= @p_151 (! (= @p_110 @p_150) :named @p_152))) :rule cong :premises (t10.t16 t10.t34)) (step t10 (cl (= @p_53 (! (forall ((veriT_vr2 Real_a_fun$) (veriT_vr3 B_list$)) @p_152) :named @p_184))) :rule bind) (anchor :step t11 :args ((:= (?v0 B_list$) veriT_vr3))) (step t11.t1 (cl (! (= ?v0 veriT_vr3) :named @p_153)) :rule refl) (step t11.t2 (cl (! (= @p_10 @p_63) :named @p_154)) :rule cong :premises (t11.t1)) (step t11.t3 (cl @p_153) :rule refl) (step t11.t4 (cl @p_154) :rule cong :premises (t11.t3)) (step t11.t5 (cl (! (= @p_15 @p_113) :named @p_156)) :rule cong :premises (t11.t2 t11.t4)) (step t11.t6 (cl @p_153) :rule refl) (step t11.t7 (cl (! (= @p_11 @p_64) :named @p_157)) :rule cong :premises (t11.t6)) (step t11.t8 (cl @p_153) :rule refl) (step t11.t9 (cl @p_154) :rule cong :premises (t11.t8)) (step t11.t10 (cl (! (= @p_21 @p_116) :named @p_158)) :rule cong :premises (t11.t9)) (step t11.t11 (cl (! (= @p_23 @p_118) :named @p_159)) :rule cong :premises (t11.t7 t11.t10)) (step t11.t12 (cl (! (= @p_16 @p_120) :named @p_160)) :rule cong :premises (t11.t11)) (anchor :step t11.t13 :args ((:= (?v2 B$) veriT_vr4))) (step t11.t13.t1 (cl @p_81) :rule refl) (step t11.t13.t2 (cl @p_123) :rule cong :premises (t11.t13.t1)) (step t11.t13.t3 (cl @p_153) :rule refl) (step t11.t13.t4 (cl (= @p_17 @p_65)) :rule cong :premises (t11.t13.t2 t11.t13.t3)) (step t11.t13.t5 (cl @p_81) :rule refl) (step t11.t13.t6 (cl @p_91) :rule cong :premises (t11.t13.t5)) (step t11.t13.t7 (cl @p_153) :rule refl) (step t11.t13.t8 (cl @p_154) :rule cong :premises (t11.t13.t7)) (step t11.t13.t9 (cl (= @p_28 @p_126)) :rule cong :premises (t11.t13.t6 t11.t13.t8)) (step t11.t13.t10 (cl (= @p_30 @p_128)) :rule cong :premises (t11.t13.t4 t11.t13.t9)) (step t11.t13.t11 (cl (= @p_32 @p_130)) :rule cong :premises (t11.t13.t10)) (step t11.t13 (cl (= @p_26 @p_122)) :rule bind) (anchor :step t11.t14 :args ((:= (?v2 B$) veriT_vr4) (:= (?v3 B$) veriT_vr5) (:= (?v4 B_list$) veriT_vr6))) (step t11.t14.t1 (cl @p_81) :rule refl) (step t11.t14.t2 (cl @p_92) :rule refl) (step t11.t14.t3 (cl @p_93) :rule refl) (step t11.t14.t4 (cl @p_94) :rule cong :premises (t11.t14.t2 t11.t14.t3)) (step t11.t14.t5 (cl @p_133) :rule cong :premises (t11.t14.t1 t11.t14.t4)) (step t11.t14.t6 (cl @p_153) :rule refl) (step t11.t14.t7 (cl (= @p_18 @p_66)) :rule cong :premises (t11.t14.t5 t11.t14.t6)) (step t11.t14.t8 (cl @p_81) :rule refl) (step t11.t14.t9 (cl @p_91) :rule cong :premises (t11.t14.t8)) (step t11.t14.t10 (cl @p_92) :rule refl) (step t11.t14.t11 (cl @p_93) :rule refl) (step t11.t14.t12 (cl @p_94) :rule cong :premises (t11.t14.t10 t11.t14.t11)) (step t11.t14.t13 (cl @p_135) :rule cong :premises (t11.t14.t12)) (step t11.t14.t14 (cl @p_136) :rule cong :premises (t11.t14.t9 t11.t14.t13)) (step t11.t14.t15 (cl @p_153) :rule refl) (step t11.t14.t16 (cl @p_154) :rule cong :premises (t11.t14.t15)) (step t11.t14.t17 (cl (= @p_36 @p_138)) :rule cong :premises (t11.t14.t14 t11.t14.t16)) (step t11.t14.t18 (cl (= @p_38 @p_140)) :rule cong :premises (t11.t14.t7 t11.t14.t17)) (step t11.t14.t19 (cl (= @p_40 @p_142)) :rule cong :premises (t11.t14.t18)) (step t11.t14 (cl (= @p_34 @p_132)) :rule bind) (step t11.t15 (cl (= @p_42 @p_144)) :rule cong :premises (t11.t13 t11.t14)) (step t11.t16 (cl (= @p_44 @p_146)) :rule cong :premises (t11.t12 t11.t15)) (step t11.t17 (cl (= @p_46 @p_148)) :rule cong :premises (t11.t5 t11.t16)) (step t11.t18 (cl (= @p_48 @p_150)) :rule cong :premises (t11.t17)) (step t11 (cl (= @p_50 (! (forall ((veriT_vr3 B_list$)) @p_150) :named @p_155))) :rule bind) (anchor :step t12 :args ((:= (?v0 B_list$) veriT_vr3))) (step t12.t1 (cl @p_153) :rule refl) (step t12.t2 (cl @p_154) :rule cong :premises (t12.t1)) (step t12.t3 (cl @p_153) :rule refl) (step t12.t4 (cl @p_154) :rule cong :premises (t12.t3)) (step t12.t5 (cl @p_156) :rule cong :premises (t12.t2 t12.t4)) (step t12.t6 (cl @p_153) :rule refl) (step t12.t7 (cl @p_157) :rule cong :premises (t12.t6)) (step t12.t8 (cl @p_153) :rule refl) (step t12.t9 (cl @p_154) :rule cong :premises (t12.t8)) (step t12.t10 (cl @p_158) :rule cong :premises (t12.t9)) (step t12.t11 (cl @p_159) :rule cong :premises (t12.t7 t12.t10)) (step t12.t12 (cl @p_160) :rule cong :premises (t12.t11)) (anchor :step t12.t13 :args ((:= (?v1 B$) veriT_vr4))) (step t12.t13.t1 (cl @p_153) :rule refl) (step t12.t13.t2 (cl (! (= ?v1 veriT_vr4) :named @p_164)) :rule refl) (step t12.t13.t3 (cl (= @p_162 @p_80)) :rule cong :premises (t12.t13.t2)) (step t12.t13.t4 (cl (= @p_163 @p_65)) :rule cong :premises (t12.t13.t1 t12.t13.t3)) (step t12.t13.t5 (cl @p_153) :rule refl) (step t12.t13.t6 (cl @p_154) :rule cong :premises (t12.t13.t5)) (step t12.t13.t7 (cl @p_164) :rule refl) (step t12.t13.t8 (cl (! (= @p_12 @p_61) :named @p_171)) :rule cong :premises (t12.t13.t7)) (step t12.t13.t9 (cl (= @p_165 @p_126)) :rule cong :premises (t12.t13.t6 t12.t13.t8)) (step t12.t13.t10 (cl (= @p_166 @p_128)) :rule cong :premises (t12.t13.t4 t12.t13.t9)) (step t12.t13.t11 (cl (= @p_167 @p_130)) :rule cong :premises (t12.t13.t10)) (step t12.t13 (cl (= @p_161 @p_122)) :rule bind) (anchor :step t12.t14 :args ((:= (?v1 B$) veriT_vr4) (:= (?v2 B$) veriT_vr5) (:= (?v3 B_list$) veriT_vr6))) (step t12.t14.t1 (cl @p_153) :rule refl) (step t12.t14.t2 (cl @p_164) :rule refl) (step t12.t14.t3 (cl (! (= ?v2 veriT_vr5) :named @p_172)) :rule refl) (step t12.t14.t4 (cl (! (= ?v3 veriT_vr6) :named @p_173)) :rule refl) (step t12.t14.t5 (cl (! (= @p_13 @p_62) :named @p_174)) :rule cong :premises (t12.t14.t3 t12.t14.t4)) (step t12.t14.t6 (cl (= @p_169 @p_90)) :rule cong :premises (t12.t14.t2 t12.t14.t5)) (step t12.t14.t7 (cl (= @p_170 @p_66)) :rule cong :premises (t12.t14.t1 t12.t14.t6)) (step t12.t14.t8 (cl @p_153) :rule refl) (step t12.t14.t9 (cl @p_154) :rule cong :premises (t12.t14.t8)) (step t12.t14.t10 (cl @p_164) :rule refl) (step t12.t14.t11 (cl @p_171) :rule cong :premises (t12.t14.t10)) (step t12.t14.t12 (cl @p_172) :rule refl) (step t12.t14.t13 (cl @p_173) :rule refl) (step t12.t14.t14 (cl @p_174) :rule cong :premises (t12.t14.t12 t12.t14.t13)) (step t12.t14.t15 (cl (= @p_175 @p_96)) :rule cong :premises (t12.t14.t14)) (step t12.t14.t16 (cl (= @p_176 @p_67)) :rule cong :premises (t12.t14.t11 t12.t14.t15)) (step t12.t14.t17 (cl (= @p_177 @p_138)) :rule cong :premises (t12.t14.t9 t12.t14.t16)) (step t12.t14.t18 (cl (= @p_178 @p_140)) :rule cong :premises (t12.t14.t7 t12.t14.t17)) (step t12.t14.t19 (cl (= @p_179 @p_142)) :rule cong :premises (t12.t14.t18)) (step t12.t14 (cl (= @p_168 @p_132)) :rule bind) (step t12.t15 (cl (= @p_180 @p_144)) :rule cong :premises (t12.t13 t12.t14)) (step t12.t16 (cl (= @p_181 @p_146)) :rule cong :premises (t12.t12 t12.t15)) (step t12.t17 (cl (= @p_182 @p_148)) :rule cong :premises (t12.t5 t12.t16)) (step t12.t18 (cl (= @p_183 @p_150)) :rule cong :premises (t12.t17)) (step t12 (cl (= @p_51 @p_155)) :rule bind) (step t13 (cl (= @p_54 (! (= @p_155 @p_155) :named @p_185))) :rule cong :premises (t11 t12)) (step t14 (cl (= @p_56 (! (=> @p_184 @p_185) :named @p_186))) :rule cong :premises (t10 t13)) (step t15 (cl (! (= @p_58 (! (not @p_186) :named @p_188)) :named @p_187)) :rule cong :premises (t14)) (step t16 (cl (! (not @p_187) :named @p_190) (! (not @p_58) :named @p_189) @p_188) :rule equiv_pos2) (step t17 (cl (not @p_189) @p_56) :rule not_not) (step t18 (cl @p_190 @p_56 @p_188) :rule th_resolution :premises (t17 t16)) (step t19 (cl @p_188) :rule th_resolution :premises (t9 t15 t18)) (step t20 (cl (! (= @p_188 (! (and @p_184 (! (not @p_185) :named @p_203)) :named @p_192)) :named @p_191)) :rule bool_simplify) (step t21 (cl (! (not @p_191) :named @p_194) (! (not @p_188) :named @p_193) @p_192) :rule equiv_pos2) (step t22 (cl (not @p_193) @p_186) :rule not_not) (step t23 (cl @p_194 @p_186 @p_192) :rule th_resolution :premises (t22 t21)) (step t24 (cl @p_192) :rule th_resolution :premises (t19 t20 t23)) (anchor :step t25 :args ((veriT_vr2 Real_a_fun$) (veriT_vr3 B_list$))) (step t25.t1 (cl (= @p_108 (! (and @p_69 @p_77 @p_79 @p_89) :named @p_195))) :rule ac_simp) (step t25.t2 (cl (= @p_110 (! (=> @p_195 false) :named @p_196))) :rule cong :premises (t25.t1)) (step t25.t3 (cl (! (= @p_148 (! (and @p_113 @p_120 @p_122 @p_132) :named @p_197)) :named @p_200)) :rule ac_simp) (step t25.t4 (cl (! (= @p_150 (! (=> @p_197 false) :named @p_198)) :named @p_201)) :rule cong :premises (t25.t3)) (step t25.t5 (cl (= @p_152 (! (= @p_196 @p_198) :named @p_199))) :rule cong :premises (t25.t2 t25.t4)) (step t25 (cl (= @p_184 (! (forall ((veriT_vr2 Real_a_fun$) (veriT_vr3 B_list$)) @p_199) :named @p_205))) :rule bind) (anchor :step t26 :args ((veriT_vr3 B_list$))) (step t26.t1 (cl @p_200) :rule ac_simp) (step t26.t2 (cl @p_201) :rule cong :premises (t26.t1)) (step t26 (cl (= @p_155 (! (forall ((veriT_vr3 B_list$)) @p_198) :named @p_202))) :rule bind) (step t27 (cl (= @p_185 (! (= @p_202 @p_202) :named @p_204))) :rule cong :premises (t26 t26)) (step t28 (cl (= @p_203 (! (not @p_204) :named @p_206))) :rule cong :premises (t27)) (step t29 (cl (! (= @p_192 (! (and @p_205 @p_206) :named @p_208)) :named @p_207)) :rule ac_simp :premises (t25 t28)) (step t30 (cl (not @p_207) (not @p_192) @p_208) :rule equiv_pos2) (step t31 (cl @p_208) :rule th_resolution :premises (t24 t29 t30)) (anchor :step t32 :args ((veriT_vr2 Real_a_fun$) (veriT_vr3 B_list$))) (step t32.t1 (cl (= @p_77 (! (not @p_75) :named @p_209))) :rule implies_simplify) (anchor :step t32.t2 :args ((veriT_vr4 B$))) (step t32.t2.t1 (cl (= @p_87 (! (not @p_85) :named @p_211))) :rule implies_simplify) (step t32.t2 (cl (= @p_79 (! (forall ((veriT_vr4 B$)) @p_211) :named @p_210))) :rule bind) (anchor :step t32.t3 :args ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$))) (step t32.t3.t1 (cl (= @p_102 (! (not @p_100) :named @p_213))) :rule implies_simplify) (step t32.t3 (cl (= @p_89 (! (forall ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$)) @p_213) :named @p_212))) :rule bind) (step t32.t4 (cl (= @p_195 (! (and @p_69 @p_209 @p_210 @p_212) :named @p_214))) :rule cong :premises (t32.t1 t32.t2 t32.t3)) (step t32.t5 (cl (= @p_196 (! (=> @p_214 false) :named @p_215))) :rule cong :premises (t32.t4)) (step t32.t6 (cl (= @p_215 (! (not @p_214) :named @p_216))) :rule implies_simplify) (step t32.t7 (cl (= @p_196 @p_216)) :rule trans :premises (t32.t5 t32.t6)) (step t32.t8 (cl (! (= @p_113 true) :named @p_228)) :rule eq_simplify) (step t32.t9 (cl (! (= @p_120 (! (not @p_118) :named @p_217)) :named @p_229)) :rule implies_simplify) (anchor :step t32.t10 :args ((veriT_vr4 B$))) (step t32.t10.t1 (cl (! (= @p_130 (! (not @p_128) :named @p_219)) :named @p_231)) :rule implies_simplify) (step t32.t10 (cl (! (= @p_122 (! (forall ((veriT_vr4 B$)) @p_219) :named @p_218)) :named @p_230)) :rule bind) (anchor :step t32.t11 :args ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$))) (step t32.t11.t1 (cl (! (= @p_142 (! (not @p_140) :named @p_221)) :named @p_233)) :rule implies_simplify) (step t32.t11 (cl (! (= @p_132 (! (forall ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$)) @p_221) :named @p_220)) :named @p_232)) :rule bind) (step t32.t12 (cl (! (= @p_197 (! (and true @p_217 @p_218 @p_220) :named @p_222)) :named @p_234)) :rule cong :premises (t32.t8 t32.t9 t32.t10 t32.t11)) (step t32.t13 (cl (! (= @p_222 (! (and @p_217 @p_218 @p_220) :named @p_223)) :named @p_235)) :rule and_simplify) (step t32.t14 (cl (! (= @p_197 @p_223) :named @p_236)) :rule trans :premises (t32.t12 t32.t13)) (step t32.t15 (cl (! (= @p_198 (! (=> @p_223 false) :named @p_224)) :named @p_237)) :rule cong :premises (t32.t14)) (step t32.t16 (cl (! (= @p_224 (! (not @p_223) :named @p_225)) :named @p_238)) :rule implies_simplify) (step t32.t17 (cl (! (= @p_198 @p_225) :named @p_239)) :rule trans :premises (t32.t15 t32.t16)) (step t32.t18 (cl (= @p_199 (! (= @p_216 @p_225) :named @p_226))) :rule cong :premises (t32.t7 t32.t17)) (step t32.t19 (cl (= @p_226 (! (= @p_214 @p_223) :named @p_227))) :rule equiv_simplify) (step t32.t20 (cl (= @p_199 @p_227)) :rule trans :premises (t32.t18 t32.t19)) (step t32 (cl (= @p_205 (! (forall ((veriT_vr2 Real_a_fun$) (veriT_vr3 B_list$)) @p_227) :named @p_243))) :rule bind) (anchor :step t33 :args ((veriT_vr3 B_list$))) (step t33.t1 (cl @p_228) :rule eq_simplify) (step t33.t2 (cl @p_229) :rule implies_simplify) (anchor :step t33.t3 :args ((veriT_vr4 B$))) (step t33.t3.t1 (cl @p_231) :rule implies_simplify) (step t33.t3 (cl @p_230) :rule bind) (anchor :step t33.t4 :args ((veriT_vr4 B$) (veriT_vr5 B$) (veriT_vr6 B_list$))) (step t33.t4.t1 (cl @p_233) :rule implies_simplify) (step t33.t4 (cl @p_232) :rule bind) (step t33.t5 (cl @p_234) :rule cong :premises (t33.t1 t33.t2 t33.t3 t33.t4)) (step t33.t6 (cl @p_235) :rule and_simplify) (step t33.t7 (cl @p_236) :rule trans :premises (t33.t5 t33.t6)) (step t33.t8 (cl @p_237) :rule cong :premises (t33.t7)) (step t33.t9 (cl @p_238) :rule implies_simplify) (step t33.t10 (cl @p_239) :rule trans :premises (t33.t8 t33.t9)) (step t33 (cl (= @p_202 (! (forall ((veriT_vr3 B_list$)) @p_225) :named @p_240))) :rule bind) (step t34 (cl (= @p_204 (! (= @p_240 @p_240) :named @p_241))) :rule cong :premises (t33 t33)) (step t35 (cl (= @p_241 true)) :rule equiv_simplify) (step t36 (cl (= @p_204 true)) :rule trans :premises (t34 t35)) (step t37 (cl (= @p_206 (! (not true) :named @p_242))) :rule cong :premises (t36)) (step t38 (cl (= @p_242 false)) :rule not_simplify) (step t39 (cl (= @p_206 false)) :rule trans :premises (t37 t38)) (step t40 (cl (= @p_208 (! (and @p_243 false) :named @p_244))) :rule cong :premises (t32 t39)) (step t41 (cl (= @p_244 false)) :rule and_simplify) (step t42 (cl (! (= @p_208 false) :named @p_245)) :rule trans :premises (t40 t41)) (step t43 (cl (not @p_245) (not @p_208) false) :rule equiv_pos2) (step t44 (cl false) :rule th_resolution :premises (t31 t42 t43)) (step t45 (cl (not false)) :rule false) (step t46 (cl) :rule resolution :premises (t44 t45)) f76f257aace4d1c925379d3446177d02a43e00b9 36 0 unsat (assume axiom4 (! (forall ((?v0 Int)) (! (and (! (pred$e ?v0) :named @p_4) (! (or (! (pred$d (! (cons$d ?v0 nil$d) :named @p_7)) :named @p_1) (! (not @p_1) :named @p_11)) :named @p_13)) :named @p_15)) :named @p_2)) (assume axiom5 (not (! (pred$e 1) :named @p_26))) (anchor :step t3 :args ((:= (?v0 Int) veriT_vr8))) (step t3.t1 (cl (! (= ?v0 veriT_vr8) :named @p_6)) :rule refl) (step t3.t2 (cl (= @p_4 (! (pred$e veriT_vr8) :named @p_5))) :rule cong :premises (t3.t1)) (step t3.t3 (cl @p_6) :rule refl) (step t3.t4 (cl (! (= @p_7 (! (cons$d veriT_vr8 nil$d) :named @p_8)) :named @p_9)) :rule cong :premises (t3.t3)) (step t3.t5 (cl (! (= @p_1 (! (pred$d @p_8) :named @p_3)) :named @p_10)) :rule cong :premises (t3.t4)) (step t3.t6 (cl @p_6) :rule refl) (step t3.t7 (cl @p_9) :rule cong :premises (t3.t6)) (step t3.t8 (cl @p_10) :rule cong :premises (t3.t7)) (step t3.t9 (cl (= @p_11 (! (not @p_3) :named @p_12))) :rule cong :premises (t3.t8)) (step t3.t10 (cl (= @p_13 (! (or @p_3 @p_12) :named @p_14))) :rule cong :premises (t3.t5 t3.t9)) (step t3.t11 (cl (= @p_15 (! (and @p_5 @p_14) :named @p_16))) :rule cong :premises (t3.t2 t3.t10)) (step t3 (cl (! (= @p_2 (! (forall ((veriT_vr8 Int)) @p_16) :named @p_18)) :named @p_17)) :rule bind) (step t4 (cl (not @p_17) (not @p_2) @p_18) :rule equiv_pos2) (step t5 (cl @p_18) :rule th_resolution :premises (axiom4 t3 t4)) (anchor :step t6 :args ((veriT_vr8 Int))) (step t6.t1 (cl (= @p_14 true)) :rule or_simplify) (step t6.t2 (cl (= @p_16 (! (and @p_5 true) :named @p_19))) :rule cong :premises (t6.t1)) (step t6.t3 (cl (= @p_19 (! (and @p_5) :named @p_20))) :rule and_simplify) (step t6.t4 (cl (= @p_20 @p_5)) :rule and_simplify) (step t6.t5 (cl (= @p_16 @p_5)) :rule trans :premises (t6.t2 t6.t3 t6.t4)) (step t6 (cl (! (= @p_18 (! (forall ((veriT_vr8 Int)) @p_5) :named @p_22)) :named @p_21)) :rule bind) (step t7 (cl (not @p_21) (not @p_18) @p_22) :rule equiv_pos2) (step t8 (cl @p_22) :rule th_resolution :premises (t5 t6 t7)) (anchor :step t9 :args ((:= (veriT_vr8 Int) veriT_vr9))) (step t9.t1 (cl (= veriT_vr8 veriT_vr9)) :rule refl) (step t9.t2 (cl (= @p_5 (! (pred$e veriT_vr9) :named @p_23))) :rule cong :premises (t9.t1)) (step t9 (cl (! (= @p_22 (! (forall ((veriT_vr9 Int)) @p_23) :named @p_25)) :named @p_24)) :rule bind) (step t10 (cl (not @p_24) (not @p_22) @p_25) :rule equiv_pos2) (step t11 (cl @p_25) :rule th_resolution :premises (t8 t9 t10)) (step t12 (cl (or (! (not @p_25) :named @p_27) @p_26)) :rule forall_inst :args ((:= veriT_vr9 1))) (step t13 (cl @p_27 @p_26) :rule or :premises (t12)) (step t14 (cl) :rule resolution :premises (t13 t11 axiom5)) 9aecf5ce2acf5c9a6a97abe747fcd03e35b92209 158 0 unsat (assume axiom0 (! (forall ((?v0 Int)) (! (= (! (g$ (! (some$ ?v0) :named @p_5)) :named @p_7) (! (g$a (! (cons$ ?v0 nil$) :named @p_10)) :named @p_12)) :named @p_14)) :named @p_4)) (assume axiom1 (forall ((?v0 Bool)) (= (g$b (some$a ?v0)) (g$c (cons$a ?v0 nil$a))))) (assume axiom6 (! (forall ((?v0 Bool_list$)) (! (= (! (g$c ?v0) :named @p_27) (! (size$ ?v0) :named @p_30)) :named @p_32)) :named @p_26)) (assume axiom7 (! (forall ((?v0 Int_list$)) (! (= (! (g$a ?v0) :named @p_43) (! (size$a ?v0) :named @p_46)) :named @p_48)) :named @p_42)) (assume axiom12 (! (= zero$ (! (size$ nil$a) :named @p_118)) :named @p_143)) (assume axiom13 (! (= zero$ (! (size$a nil$) :named @p_133)) :named @p_144)) (assume axiom14 (forall ((?v0 Bool) (?v1 Bool_list$)) (= (size$ (cons$a ?v0 ?v1)) (! (plus$ (! (size$ ?v1) :named @p_60) (! (suc$ zero$) :named @p_1)) :named @p_3)))) (assume axiom15 (! (forall ((?v0 Int) (?v1 Int_list$)) (! (= (! (size$a (! (cons$ ?v0 ?v1) :named @p_96)) :named @p_98) (! (plus$ (! (size$a ?v1) :named @p_101) @p_1) :named @p_103)) :named @p_105)) :named @p_95)) (assume axiom16 (not (! (= (! (g$ (some$ 3)) :named @p_124) (! (g$b (some$a true)) :named @p_2)) :named @p_141))) (step t10 (cl (and (= (g$b (some$a false)) (g$c (! (cons$a false nil$a) :named @p_119))) (! (= @p_2 (! (g$c (! (cons$a true nil$a) :named @p_121)) :named @p_122)) :named @p_117))) :rule bfun_elim :premises (axiom1)) (step t11 (cl (! (forall ((?v1 Bool_list$)) (! (and (! (= @p_3 (! (size$ (! (cons$a false ?v1) :named @p_63)) :named @p_65)) :named @p_67) (! (= @p_3 (! (size$ (! (cons$a true ?v1) :named @p_71)) :named @p_73)) :named @p_75)) :named @p_77)) :named @p_58)) :rule bfun_elim :premises (axiom14)) (anchor :step t12 :args ((:= (?v0 Int) veriT_vr0))) (step t12.t1 (cl (! (= ?v0 veriT_vr0) :named @p_9)) :rule refl) (step t12.t2 (cl (= @p_5 (! (some$ veriT_vr0) :named @p_6))) :rule cong :premises (t12.t1)) (step t12.t3 (cl (= @p_7 (! (g$ @p_6) :named @p_8))) :rule cong :premises (t12.t2)) (step t12.t4 (cl @p_9) :rule refl) (step t12.t5 (cl (= @p_10 (! (cons$ veriT_vr0 nil$) :named @p_11))) :rule cong :premises (t12.t4)) (step t12.t6 (cl (= @p_12 (! (g$a @p_11) :named @p_13))) :rule cong :premises (t12.t5)) (step t12.t7 (cl (= @p_14 (! (= @p_8 @p_13) :named @p_15))) :rule cong :premises (t12.t3 t12.t6)) (step t12 (cl (! (= @p_4 (! (forall ((veriT_vr0 Int)) @p_15) :named @p_17)) :named @p_16)) :rule bind) (step t13 (cl (not @p_16) (not @p_4) @p_17) :rule equiv_pos2) (step t14 (cl @p_17) :rule th_resolution :premises (axiom0 t12 t13)) (anchor :step t15 :args ((:= (veriT_vr0 Int) veriT_vr1))) (step t15.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_20)) :rule refl) (step t15.t2 (cl (= @p_6 (! (some$ veriT_vr1) :named @p_18))) :rule cong :premises (t15.t1)) (step t15.t3 (cl (= @p_8 (! (g$ @p_18) :named @p_19))) :rule cong :premises (t15.t2)) (step t15.t4 (cl @p_20) :rule refl) (step t15.t5 (cl (= @p_11 (! (cons$ veriT_vr1 nil$) :named @p_21))) :rule cong :premises (t15.t4)) (step t15.t6 (cl (= @p_13 (! (g$a @p_21) :named @p_22))) :rule cong :premises (t15.t5)) (step t15.t7 (cl (= @p_15 (! (= @p_19 @p_22) :named @p_23))) :rule cong :premises (t15.t3 t15.t6)) (step t15 (cl (! (= @p_17 (! (forall ((veriT_vr1 Int)) @p_23) :named @p_25)) :named @p_24)) :rule bind) (step t16 (cl (not @p_24) (not @p_17) @p_25) :rule equiv_pos2) (step t17 (cl @p_25) :rule th_resolution :premises (t14 t15 t16)) (anchor :step t18 :args ((:= (?v0 Bool_list$) veriT_vr2))) (step t18.t1 (cl (! (= ?v0 veriT_vr2) :named @p_29)) :rule refl) (step t18.t2 (cl (= @p_27 (! (g$c veriT_vr2) :named @p_28))) :rule cong :premises (t18.t1)) (step t18.t3 (cl @p_29) :rule refl) (step t18.t4 (cl (= @p_30 (! (size$ veriT_vr2) :named @p_31))) :rule cong :premises (t18.t3)) (step t18.t5 (cl (= @p_32 (! (= @p_28 @p_31) :named @p_33))) :rule cong :premises (t18.t2 t18.t4)) (step t18 (cl (! (= @p_26 (! (forall ((veriT_vr2 Bool_list$)) @p_33) :named @p_35)) :named @p_34)) :rule bind) (step t19 (cl (not @p_34) (not @p_26) @p_35) :rule equiv_pos2) (step t20 (cl @p_35) :rule th_resolution :premises (axiom6 t18 t19)) (anchor :step t21 :args ((:= (veriT_vr2 Bool_list$) veriT_vr3))) (step t21.t1 (cl (! (= veriT_vr2 veriT_vr3) :named @p_37)) :rule refl) (step t21.t2 (cl (= @p_28 (! (g$c veriT_vr3) :named @p_36))) :rule cong :premises (t21.t1)) (step t21.t3 (cl @p_37) :rule refl) (step t21.t4 (cl (= @p_31 (! (size$ veriT_vr3) :named @p_38))) :rule cong :premises (t21.t3)) (step t21.t5 (cl (= @p_33 (! (= @p_36 @p_38) :named @p_39))) :rule cong :premises (t21.t2 t21.t4)) (step t21 (cl (! (= @p_35 (! (forall ((veriT_vr3 Bool_list$)) @p_39) :named @p_41)) :named @p_40)) :rule bind) (step t22 (cl (not @p_40) (not @p_35) @p_41) :rule equiv_pos2) (step t23 (cl @p_41) :rule th_resolution :premises (t20 t21 t22)) (anchor :step t24 :args ((:= (?v0 Int_list$) veriT_vr4))) (step t24.t1 (cl (! (= ?v0 veriT_vr4) :named @p_45)) :rule refl) (step t24.t2 (cl (= @p_43 (! (g$a veriT_vr4) :named @p_44))) :rule cong :premises (t24.t1)) (step t24.t3 (cl @p_45) :rule refl) (step t24.t4 (cl (= @p_46 (! (size$a veriT_vr4) :named @p_47))) :rule cong :premises (t24.t3)) (step t24.t5 (cl (= @p_48 (! (= @p_44 @p_47) :named @p_49))) :rule cong :premises (t24.t2 t24.t4)) (step t24 (cl (! (= @p_42 (! (forall ((veriT_vr4 Int_list$)) @p_49) :named @p_51)) :named @p_50)) :rule bind) (step t25 (cl (not @p_50) (not @p_42) @p_51) :rule equiv_pos2) (step t26 (cl @p_51) :rule th_resolution :premises (axiom7 t24 t25)) (anchor :step t27 :args ((:= (veriT_vr4 Int_list$) veriT_vr5))) (step t27.t1 (cl (! (= veriT_vr4 veriT_vr5) :named @p_53)) :rule refl) (step t27.t2 (cl (= @p_44 (! (g$a veriT_vr5) :named @p_52))) :rule cong :premises (t27.t1)) (step t27.t3 (cl @p_53) :rule refl) (step t27.t4 (cl (= @p_47 (! (size$a veriT_vr5) :named @p_54))) :rule cong :premises (t27.t3)) (step t27.t5 (cl (= @p_49 (! (= @p_52 @p_54) :named @p_55))) :rule cong :premises (t27.t2 t27.t4)) (step t27 (cl (! (= @p_51 (! (forall ((veriT_vr5 Int_list$)) @p_55) :named @p_57)) :named @p_56)) :rule bind) (step t28 (cl (not @p_56) (not @p_51) @p_57) :rule equiv_pos2) (step t29 (cl @p_57) :rule th_resolution :premises (t26 t27 t28)) (anchor :step t30 :args ((:= (?v1 Bool_list$) veriT_vr20))) (step t30.t1 (cl (! (= ?v1 veriT_vr20) :named @p_62)) :rule refl) (step t30.t2 (cl (! (= @p_60 (! (size$ veriT_vr20) :named @p_61)) :named @p_69)) :rule cong :premises (t30.t1)) (step t30.t3 (cl (! (= @p_3 (! (plus$ @p_61 @p_1) :named @p_59)) :named @p_70)) :rule cong :premises (t30.t2)) (step t30.t4 (cl @p_62) :rule refl) (step t30.t5 (cl (= @p_63 (! (cons$a false veriT_vr20) :named @p_64))) :rule cong :premises (t30.t4)) (step t30.t6 (cl (= @p_65 (! (size$ @p_64) :named @p_66))) :rule cong :premises (t30.t5)) (step t30.t7 (cl (= @p_67 (! (= @p_59 @p_66) :named @p_68))) :rule cong :premises (t30.t3 t30.t6)) (step t30.t8 (cl @p_62) :rule refl) (step t30.t9 (cl @p_69) :rule cong :premises (t30.t8)) (step t30.t10 (cl @p_70) :rule cong :premises (t30.t9)) (step t30.t11 (cl @p_62) :rule refl) (step t30.t12 (cl (= @p_71 (! (cons$a true veriT_vr20) :named @p_72))) :rule cong :premises (t30.t11)) (step t30.t13 (cl (= @p_73 (! (size$ @p_72) :named @p_74))) :rule cong :premises (t30.t12)) (step t30.t14 (cl (= @p_75 (! (= @p_59 @p_74) :named @p_76))) :rule cong :premises (t30.t10 t30.t13)) (step t30.t15 (cl (= @p_77 (! (and @p_68 @p_76) :named @p_78))) :rule cong :premises (t30.t7 t30.t14)) (step t30 (cl (! (= @p_58 (! (forall ((veriT_vr20 Bool_list$)) @p_78) :named @p_80)) :named @p_79)) :rule bind) (step t31 (cl (not @p_79) (not @p_58) @p_80) :rule equiv_pos2) (step t32 (cl @p_80) :rule th_resolution :premises (t11 t30 t31)) (anchor :step t33 :args ((:= (veriT_vr20 Bool_list$) veriT_vr21))) (step t33.t1 (cl (! (= veriT_vr20 veriT_vr21) :named @p_83)) :rule refl) (step t33.t2 (cl (! (= @p_61 (! (size$ veriT_vr21) :named @p_82)) :named @p_87)) :rule cong :premises (t33.t1)) (step t33.t3 (cl (! (= @p_59 (! (plus$ @p_82 @p_1) :named @p_81)) :named @p_88)) :rule cong :premises (t33.t2)) (step t33.t4 (cl @p_83) :rule refl) (step t33.t5 (cl (= @p_64 (! (cons$a false veriT_vr21) :named @p_84))) :rule cong :premises (t33.t4)) (step t33.t6 (cl (= @p_66 (! (size$ @p_84) :named @p_85))) :rule cong :premises (t33.t5)) (step t33.t7 (cl (= @p_68 (! (= @p_81 @p_85) :named @p_86))) :rule cong :premises (t33.t3 t33.t6)) (step t33.t8 (cl @p_83) :rule refl) (step t33.t9 (cl @p_87) :rule cong :premises (t33.t8)) (step t33.t10 (cl @p_88) :rule cong :premises (t33.t9)) (step t33.t11 (cl @p_83) :rule refl) (step t33.t12 (cl (= @p_72 (! (cons$a true veriT_vr21) :named @p_89))) :rule cong :premises (t33.t11)) (step t33.t13 (cl (= @p_74 (! (size$ @p_89) :named @p_90))) :rule cong :premises (t33.t12)) (step t33.t14 (cl (= @p_76 (! (= @p_81 @p_90) :named @p_91))) :rule cong :premises (t33.t10 t33.t13)) (step t33.t15 (cl (= @p_78 (! (and @p_86 @p_91) :named @p_92))) :rule cong :premises (t33.t7 t33.t14)) (step t33 (cl (! (= @p_80 (! (forall ((veriT_vr21 Bool_list$)) @p_92) :named @p_94)) :named @p_93)) :rule bind) (step t34 (cl (not @p_93) (not @p_80) @p_94) :rule equiv_pos2) (step t35 (cl @p_94) :rule th_resolution :premises (t32 t33 t34)) (anchor :step t36 :args ((:= (?v0 Int) veriT_vr22) (:= (?v1 Int_list$) veriT_vr23))) (step t36.t1 (cl (= ?v0 veriT_vr22)) :rule refl) (step t36.t2 (cl (! (= ?v1 veriT_vr23) :named @p_100)) :rule refl) (step t36.t3 (cl (= @p_96 (! (cons$ veriT_vr22 veriT_vr23) :named @p_97))) :rule cong :premises (t36.t1 t36.t2)) (step t36.t4 (cl (= @p_98 (! (size$a @p_97) :named @p_99))) :rule cong :premises (t36.t3)) (step t36.t5 (cl @p_100) :rule refl) (step t36.t6 (cl (= @p_101 (! (size$a veriT_vr23) :named @p_102))) :rule cong :premises (t36.t5)) (step t36.t7 (cl (= @p_103 (! (plus$ @p_102 @p_1) :named @p_104))) :rule cong :premises (t36.t6)) (step t36.t8 (cl (= @p_105 (! (= @p_99 @p_104) :named @p_106))) :rule cong :premises (t36.t4 t36.t7)) (step t36 (cl (! (= @p_95 (! (forall ((veriT_vr22 Int) (veriT_vr23 Int_list$)) @p_106) :named @p_108)) :named @p_107)) :rule bind) (step t37 (cl (not @p_107) (not @p_95) @p_108) :rule equiv_pos2) (step t38 (cl @p_108) :rule th_resolution :premises (axiom15 t36 t37)) (anchor :step t39 :args ((:= (veriT_vr22 Int) veriT_vr24) (:= (veriT_vr23 Int_list$) veriT_vr25))) (step t39.t1 (cl (= veriT_vr22 veriT_vr24)) :rule refl) (step t39.t2 (cl (! (= veriT_vr23 veriT_vr25) :named @p_111)) :rule refl) (step t39.t3 (cl (= @p_97 (! (cons$ veriT_vr24 veriT_vr25) :named @p_109))) :rule cong :premises (t39.t1 t39.t2)) (step t39.t4 (cl (= @p_99 (! (size$a @p_109) :named @p_110))) :rule cong :premises (t39.t3)) (step t39.t5 (cl @p_111) :rule refl) (step t39.t6 (cl (= @p_102 (! (size$a veriT_vr25) :named @p_112))) :rule cong :premises (t39.t5)) (step t39.t7 (cl (= @p_104 (! (plus$ @p_112 @p_1) :named @p_113))) :rule cong :premises (t39.t6)) (step t39.t8 (cl (= @p_106 (! (= @p_110 @p_113) :named @p_114))) :rule cong :premises (t39.t4 t39.t7)) (step t39 (cl (! (= @p_108 (! (forall ((veriT_vr24 Int) (veriT_vr25 Int_list$)) @p_114) :named @p_116)) :named @p_115)) :rule bind) (step t40 (cl (not @p_115) (not @p_108) @p_116) :rule equiv_pos2) (step t41 (cl @p_116) :rule th_resolution :premises (t38 t39 t40)) (step t42 (cl @p_117) :rule and :premises (t10)) (step t43 (cl (or (! (not @p_94) :named @p_127) (! (and (= (! (plus$ @p_118 @p_1) :named @p_120) (size$ @p_119)) (! (= @p_120 (! (size$ @p_121) :named @p_123)) :named @p_126)) :named @p_125))) :rule forall_inst :args ((:= veriT_vr21 nil$a))) (step t44 (cl (or (! (not @p_41) :named @p_128) (! (= @p_122 @p_123) :named @p_129))) :rule forall_inst :args ((:= veriT_vr3 @p_121))) (step t45 (cl (or (! (not @p_25) :named @p_130) (! (= @p_124 (! (g$a (! (cons$ 3 nil$) :named @p_132)) :named @p_134)) :named @p_131))) :rule forall_inst :args ((:= veriT_vr1 3))) (step t46 (cl (not @p_125) @p_126) :rule and_pos) (step t47 (cl @p_127 @p_125) :rule or :premises (t43)) (step t48 (cl @p_125) :rule resolution :premises (t47 t35)) (step t49 (cl @p_126) :rule resolution :premises (t46 t48)) (step t50 (cl @p_128 @p_129) :rule or :premises (t44)) (step t51 (cl @p_129) :rule resolution :premises (t50 t23)) (step t52 (cl @p_130 @p_131) :rule or :premises (t45)) (step t53 (cl @p_131) :rule resolution :premises (t52 t17)) (step t54 (cl (or (! (not @p_116) :named @p_136) (! (= (! (size$a @p_132) :named @p_135) (! (plus$ @p_133 @p_1) :named @p_140)) :named @p_137))) :rule forall_inst :args ((:= veriT_vr24 3) (:= veriT_vr25 nil$))) (step t55 (cl (or (! (not @p_57) :named @p_138) (! (= @p_134 @p_135) :named @p_139))) :rule forall_inst :args ((:= veriT_vr5 @p_132))) (step t56 (cl @p_136 @p_137) :rule or :premises (t54)) (step t57 (cl @p_137) :rule resolution :premises (t56 t41)) (step t58 (cl @p_138 @p_139) :rule or :premises (t55)) (step t59 (cl @p_139) :rule resolution :premises (t58 t29)) (step t60 (cl (! (not @p_131) :named @p_150) (! (not @p_139) :named @p_151) (! (not @p_137) :named @p_152) (not (! (= @p_120 @p_140) :named @p_142)) (! (not @p_126) :named @p_153) (! (not @p_129) :named @p_154) (! (not @p_117) :named @p_155) @p_141) :rule eq_transitive) (step t61 (cl (not (! (= @p_118 @p_133) :named @p_145)) (! (not (! (= @p_1 @p_1) :named @p_149)) :named @p_146) @p_142) :rule eq_congruent) (step t62 (cl (! (not @p_143) :named @p_147) (! (not @p_144) :named @p_148) @p_145) :rule eq_transitive) (step t63 (cl @p_146 @p_142 @p_147 @p_148) :rule th_resolution :premises (t61 t62)) (step t64 (cl @p_149) :rule eq_reflexive) (step t65 (cl @p_142 @p_147 @p_148) :rule th_resolution :premises (t63 t64)) (step t66 (cl @p_150 @p_151 @p_152 @p_153 @p_154 @p_155 @p_141 @p_147 @p_148) :rule th_resolution :premises (t60 t65)) (step t67 (cl) :rule resolution :premises (t66 t42 axiom12 axiom13 axiom16 t49 t51 t53 t57 t59)) 1a2d4d1ee4565edc7b401dfc82d8d10a78382c1c 910 0 unsat (define-fun veriT_sk0 () V$ (! (choice ((veriT_vr65 V$)) (not (! (not (! (= x2$ (! (rraise$ veriT_vr65) :named @p_401)) :named @p_402)) :named @p_400))) :named @p_414)) (define-fun veriT_sk1 () Abort$ (! (choice ((veriT_vr66 Abort$)) (not (! (not (! (= x2$ (! (rabort$ veriT_vr66) :named @p_404)) :named @p_405)) :named @p_403))) :named @p_418)) (define-fun veriT_sk3 () V_list_v_result$ (! (choice ((veriT_vr73 V_list_v_result$)) (! (= (! (fun_evaluate$ st$a env$ (cons$ e$ nil$)) :named @p_3) (! (pair$ (! (fst$ @p_3) :named @p_378) veriT_vr73) :named @p_461)) :named @p_460)) :named @p_465)) (define-fun veriT_sk11 () V_list_v_result$ (! (choice ((veriT_vr108 V_list_v_result$)) (! (= (! (fix_clock$ st$a @p_3) :named @p_470) (! (pair$ st$ veriT_vr108) :named @p_503)) :named @p_502)) :named @p_515)) (assume axiom0 (! (forall ((?v0 V$)) (! (= (! (fun_app$ uua$ ?v0) :named @p_9) (! (fun_app$ (! (fun_evaluate_match$ st$ env$ ?v0 pes$) :named @p_12) ?v0) :named @p_14)) :named @p_16)) :named @p_8)) (assume axiom1 (! (forall ((?v0 Abort$)) (! (= (! (fun_app$a uub$ ?v0) :named @p_28) (! (pair$ st$ (! (rerr$ (! (rabort$ ?v0) :named @p_31)) :named @p_33)) :named @p_35)) :named @p_37)) :named @p_27)) (assume axiom2 (! (forall ((?v0 Astate$) (?v1 Astate$) (?v2 Nat$)) (! (= (! (fun_app$b (! (uu$ ?v0 ?v1) :named @p_5) ?v2) :named @p_53) (! (ite (! (less_eq$ (! (clock$ ?v1) :named @p_1) (! (clock$ ?v0) :named @p_2)) :named @p_57) @p_1 @p_2) :named @p_61)) :named @p_63)) :named @p_49)) (assume axiom3 (! (= @p_470 (! (pair$ st$ r$) :named @p_609)) :named @p_628)) (assume axiom4 (! (less_eq$ (! (clock$ @p_378) :named @p_371) (! (clock$ st$a) :named @p_7)) :named @p_369)) (assume axiom5 (! (forall ((?v0 Nat$) (?v1 Nat$) (?v2 Nat$)) (! (=> (! (and (! (less_eq$ ?v0 ?v1) :named @p_81) (! (less_eq$ ?v2 ?v0) :named @p_84)) :named @p_86) (! (less_eq$ ?v2 ?v1) :named @p_90)) :named @p_92)) :named @p_80)) (assume axiom6 (! (forall ((?v0 Astate$) (?v1 Astate_v_list_v_result_prod$)) (! (= (! (= ?v0 (! (fst$ ?v1) :named @p_107)) :named @p_109) (! (exists ((?v2 V_list_v_result$)) (! (= ?v1 (! (pair$ ?v0 ?v2) :named @p_115)) :named @p_117)) :named @p_111)) :named @p_119)) :named @p_106)) (assume axiom7 (! (forall ((?v0 V_error_result$)) (! (=> (! (and (! (forall ((?v1 V$)) (! (=> (! (= ?v0 (! (rraise$ ?v1) :named @p_174)) :named @p_6) false) :named @p_177)) :named @p_172) (! (forall ((?v1 Abort$)) (! (=> (! (= ?v0 (! (rabort$ ?v1) :named @p_182)) :named @p_184) false) :named @p_186)) :named @p_179)) :named @p_188) false) :named @p_190)) :named @p_171)) (assume axiom8 (! (forall ((?v0 V_astate_v_list_v_result_prod_fun$) (?v1 Abort_astate_v_list_v_result_prod_fun$) (?v2 V$)) (! (= (! (case_error_result$ ?v0 ?v1 (! (rraise$ ?v2) :named @p_217)) :named @p_219) (! (fun_app$ ?v0 ?v2) :named @p_223)) :named @p_225)) :named @p_216)) (assume axiom9 (! (forall ((?v0 V_astate_v_list_v_result_prod_fun$) (?v1 Abort_astate_v_list_v_result_prod_fun$) (?v2 Abort$)) (! (= (! (case_error_result$ ?v0 ?v1 (! (rabort$ ?v2) :named @p_238)) :named @p_240) (! (fun_app$a ?v1 ?v2) :named @p_244)) :named @p_246)) :named @p_237)) (assume axiom10 (! (forall ((?v0 Astate$) (?v1 Astate$) (?v2 V_list_v_result$) (?v3 Astate$)) (! (=> (! (= (! (fix_clock$ ?v0 (! (pair$ ?v1 ?v2) :named @p_259)) :named @p_4) (! (pair$ ?v3 ?v2) :named @p_263)) :named @p_265) (! (less_eq$ (! (clock$ ?v3) :named @p_268) @p_1) :named @p_272)) :named @p_274)) :named @p_258)) (assume axiom11 (! (forall ((?v0 Astate$) (?v1 Astate$) (?v2 V_list_v_result$)) (! (= @p_4 (! (pair$ (! (update_clock$ @p_5 ?v1) :named @p_297) ?v2) :named @p_300)) :named @p_302)) :named @p_291)) (assume axiom12 (! (forall ((?v0 V_error_result$) (?v1 V$)) (! (=> (! (and (! (= r$ (! (rerr$ ?v0) :named @p_319)) :named @p_321) @p_6) :named @p_326) (! (less_eq$ (! (clock$ (! (fst$ (! (fun_app$ (! (fun_evaluate_match$ st$ env$ ?v1 pes$) :named @p_329) ?v1) :named @p_331)) :named @p_333)) :named @p_335) (! (clock$ st$) :named @p_318)) :named @p_337)) :named @p_339)) :named @p_317)) (assume axiom13 (! (not (! (=> (! (= r$ (! (rerr$ x2$) :named @p_615)) :named @p_359) (! (less_eq$ (! (clock$ (! (fst$ (! (case_error_result$ uua$ uub$ x2$) :named @p_602)) :named @p_530)) :named @p_370) @p_7) :named @p_360)) :named @p_364)) :named @p_358)) (anchor :step t15 :args ((:= (?v0 V$) veriT_vr0))) (step t15.t1 (cl (! (= ?v0 veriT_vr0) :named @p_11)) :rule refl) (step t15.t2 (cl (= @p_9 (! (fun_app$ uua$ veriT_vr0) :named @p_10))) :rule cong :premises (t15.t1)) (step t15.t3 (cl @p_11) :rule refl) (step t15.t4 (cl (= @p_12 (! (fun_evaluate_match$ st$ env$ veriT_vr0 pes$) :named @p_13))) :rule cong :premises (t15.t3)) (step t15.t5 (cl @p_11) :rule refl) (step t15.t6 (cl (= @p_14 (! (fun_app$ @p_13 veriT_vr0) :named @p_15))) :rule cong :premises (t15.t4 t15.t5)) (step t15.t7 (cl (= @p_16 (! (= @p_10 @p_15) :named @p_17))) :rule cong :premises (t15.t2 t15.t6)) (step t15 (cl (! (= @p_8 (! (forall ((veriT_vr0 V$)) @p_17) :named @p_19)) :named @p_18)) :rule bind) (step t16 (cl (not @p_18) (not @p_8) @p_19) :rule equiv_pos2) (step t17 (cl @p_19) :rule th_resolution :premises (axiom0 t15 t16)) (anchor :step t18 :args ((:= (veriT_vr0 V$) veriT_vr1))) (step t18.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_21)) :rule refl) (step t18.t2 (cl (= @p_10 (! (fun_app$ uua$ veriT_vr1) :named @p_20))) :rule cong :premises (t18.t1)) (step t18.t3 (cl @p_21) :rule refl) (step t18.t4 (cl (= @p_13 (! (fun_evaluate_match$ st$ env$ veriT_vr1 pes$) :named @p_22))) :rule cong :premises (t18.t3)) (step t18.t5 (cl @p_21) :rule refl) (step t18.t6 (cl (= @p_15 (! (fun_app$ @p_22 veriT_vr1) :named @p_23))) :rule cong :premises (t18.t4 t18.t5)) (step t18.t7 (cl (= @p_17 (! (= @p_20 @p_23) :named @p_24))) :rule cong :premises (t18.t2 t18.t6)) (step t18 (cl (! (= @p_19 (! (forall ((veriT_vr1 V$)) @p_24) :named @p_26)) :named @p_25)) :rule bind) (step t19 (cl (not @p_25) (not @p_19) @p_26) :rule equiv_pos2) (step t20 (cl @p_26) :rule th_resolution :premises (t17 t18 t19)) (anchor :step t21 :args ((:= (?v0 Abort$) veriT_vr2))) (step t21.t1 (cl (! (= ?v0 veriT_vr2) :named @p_30)) :rule refl) (step t21.t2 (cl (= @p_28 (! (fun_app$a uub$ veriT_vr2) :named @p_29))) :rule cong :premises (t21.t1)) (step t21.t3 (cl @p_30) :rule refl) (step t21.t4 (cl (= @p_31 (! (rabort$ veriT_vr2) :named @p_32))) :rule cong :premises (t21.t3)) (step t21.t5 (cl (= @p_33 (! (rerr$ @p_32) :named @p_34))) :rule cong :premises (t21.t4)) (step t21.t6 (cl (= @p_35 (! (pair$ st$ @p_34) :named @p_36))) :rule cong :premises (t21.t5)) (step t21.t7 (cl (= @p_37 (! (= @p_29 @p_36) :named @p_38))) :rule cong :premises (t21.t2 t21.t6)) (step t21 (cl (! (= @p_27 (! (forall ((veriT_vr2 Abort$)) @p_38) :named @p_40)) :named @p_39)) :rule bind) (step t22 (cl (not @p_39) (not @p_27) @p_40) :rule equiv_pos2) (step t23 (cl @p_40) :rule th_resolution :premises (axiom1 t21 t22)) (anchor :step t24 :args ((:= (veriT_vr2 Abort$) veriT_vr3))) (step t24.t1 (cl (! (= veriT_vr2 veriT_vr3) :named @p_42)) :rule refl) (step t24.t2 (cl (= @p_29 (! (fun_app$a uub$ veriT_vr3) :named @p_41))) :rule cong :premises (t24.t1)) (step t24.t3 (cl @p_42) :rule refl) (step t24.t4 (cl (= @p_32 (! (rabort$ veriT_vr3) :named @p_43))) :rule cong :premises (t24.t3)) (step t24.t5 (cl (= @p_34 (! (rerr$ @p_43) :named @p_44))) :rule cong :premises (t24.t4)) (step t24.t6 (cl (= @p_36 (! (pair$ st$ @p_44) :named @p_45))) :rule cong :premises (t24.t5)) (step t24.t7 (cl (= @p_38 (! (= @p_41 @p_45) :named @p_46))) :rule cong :premises (t24.t2 t24.t6)) (step t24 (cl (! (= @p_40 (! (forall ((veriT_vr3 Abort$)) @p_46) :named @p_48)) :named @p_47)) :rule bind) (step t25 (cl (not @p_47) (not @p_40) @p_48) :rule equiv_pos2) (step t26 (cl @p_48) :rule th_resolution :premises (t23 t24 t25)) (anchor :step t27 :args ((:= (?v0 Astate$) veriT_vr4) (:= (?v1 Astate$) veriT_vr5) (:= (?v2 Nat$) veriT_vr6))) (step t27.t1 (cl (! (= ?v0 veriT_vr4) :named @p_56)) :rule refl) (step t27.t2 (cl (! (= ?v1 veriT_vr5) :named @p_55)) :rule refl) (step t27.t3 (cl (= @p_5 (! (uu$ veriT_vr4 veriT_vr5) :named @p_52))) :rule cong :premises (t27.t1 t27.t2)) (step t27.t4 (cl (= ?v2 veriT_vr6)) :rule refl) (step t27.t5 (cl (= @p_53 (! (fun_app$b @p_52 veriT_vr6) :named @p_54))) :rule cong :premises (t27.t3 t27.t4)) (step t27.t6 (cl @p_55) :rule refl) (step t27.t7 (cl (! (= @p_1 (! (clock$ veriT_vr5) :named @p_50)) :named @p_59)) :rule cong :premises (t27.t6)) (step t27.t8 (cl @p_56) :rule refl) (step t27.t9 (cl (! (= @p_2 (! (clock$ veriT_vr4) :named @p_51)) :named @p_60)) :rule cong :premises (t27.t8)) (step t27.t10 (cl (= @p_57 (! (less_eq$ @p_50 @p_51) :named @p_58))) :rule cong :premises (t27.t7 t27.t9)) (step t27.t11 (cl @p_55) :rule refl) (step t27.t12 (cl @p_59) :rule cong :premises (t27.t11)) (step t27.t13 (cl @p_56) :rule refl) (step t27.t14 (cl @p_60) :rule cong :premises (t27.t13)) (step t27.t15 (cl (= @p_61 (! (ite @p_58 @p_50 @p_51) :named @p_62))) :rule cong :premises (t27.t10 t27.t12 t27.t14)) (step t27.t16 (cl (= @p_63 (! (= @p_54 @p_62) :named @p_64))) :rule cong :premises (t27.t5 t27.t15)) (step t27 (cl (! (= @p_49 (! (forall ((veriT_vr4 Astate$) (veriT_vr5 Astate$) (veriT_vr6 Nat$)) @p_64) :named @p_66)) :named @p_65)) :rule bind) (step t28 (cl (not @p_65) (not @p_49) @p_66) :rule equiv_pos2) (step t29 (cl @p_66) :rule th_resolution :premises (axiom2 t27 t28)) (anchor :step t30 :args ((:= (veriT_vr4 Astate$) veriT_vr7) (:= (veriT_vr5 Astate$) veriT_vr8) (:= (veriT_vr6 Nat$) veriT_vr9))) (step t30.t1 (cl (! (= veriT_vr4 veriT_vr7) :named @p_72)) :rule refl) (step t30.t2 (cl (! (= veriT_vr5 veriT_vr8) :named @p_71)) :rule refl) (step t30.t3 (cl (= @p_52 (! (uu$ veriT_vr7 veriT_vr8) :named @p_69))) :rule cong :premises (t30.t1 t30.t2)) (step t30.t4 (cl (= veriT_vr6 veriT_vr9)) :rule refl) (step t30.t5 (cl (= @p_54 (! (fun_app$b @p_69 veriT_vr9) :named @p_70))) :rule cong :premises (t30.t3 t30.t4)) (step t30.t6 (cl @p_71) :rule refl) (step t30.t7 (cl (! (= @p_50 (! (clock$ veriT_vr8) :named @p_67)) :named @p_74)) :rule cong :premises (t30.t6)) (step t30.t8 (cl @p_72) :rule refl) (step t30.t9 (cl (! (= @p_51 (! (clock$ veriT_vr7) :named @p_68)) :named @p_75)) :rule cong :premises (t30.t8)) (step t30.t10 (cl (= @p_58 (! (less_eq$ @p_67 @p_68) :named @p_73))) :rule cong :premises (t30.t7 t30.t9)) (step t30.t11 (cl @p_71) :rule refl) (step t30.t12 (cl @p_74) :rule cong :premises (t30.t11)) (step t30.t13 (cl @p_72) :rule refl) (step t30.t14 (cl @p_75) :rule cong :premises (t30.t13)) (step t30.t15 (cl (= @p_62 (! (ite @p_73 @p_67 @p_68) :named @p_76))) :rule cong :premises (t30.t10 t30.t12 t30.t14)) (step t30.t16 (cl (= @p_64 (! (= @p_70 @p_76) :named @p_77))) :rule cong :premises (t30.t5 t30.t15)) (step t30 (cl (! (= @p_66 (! (forall ((veriT_vr7 Astate$) (veriT_vr8 Astate$) (veriT_vr9 Nat$)) @p_77) :named @p_79)) :named @p_78)) :rule bind) (step t31 (cl (not @p_78) (not @p_66) @p_79) :rule equiv_pos2) (step t32 (cl @p_79) :rule th_resolution :premises (t29 t30 t31)) (anchor :step t33 :args ((:= (?v0 Nat$) veriT_vr10) (:= (?v1 Nat$) veriT_vr11) (:= (?v2 Nat$) veriT_vr12))) (step t33.t1 (cl (! (= ?v0 veriT_vr10) :named @p_83)) :rule refl) (step t33.t2 (cl (! (= ?v1 veriT_vr11) :named @p_89)) :rule refl) (step t33.t3 (cl (= @p_81 (! (less_eq$ veriT_vr10 veriT_vr11) :named @p_82))) :rule cong :premises (t33.t1 t33.t2)) (step t33.t4 (cl (! (= ?v2 veriT_vr12) :named @p_88)) :rule refl) (step t33.t5 (cl @p_83) :rule refl) (step t33.t6 (cl (= @p_84 (! (less_eq$ veriT_vr12 veriT_vr10) :named @p_85))) :rule cong :premises (t33.t4 t33.t5)) (step t33.t7 (cl (= @p_86 (! (and @p_82 @p_85) :named @p_87))) :rule cong :premises (t33.t3 t33.t6)) (step t33.t8 (cl @p_88) :rule refl) (step t33.t9 (cl @p_89) :rule refl) (step t33.t10 (cl (= @p_90 (! (less_eq$ veriT_vr12 veriT_vr11) :named @p_91))) :rule cong :premises (t33.t8 t33.t9)) (step t33.t11 (cl (= @p_92 (! (=> @p_87 @p_91) :named @p_93))) :rule cong :premises (t33.t7 t33.t10)) (step t33 (cl (! (= @p_80 (! (forall ((veriT_vr10 Nat$) (veriT_vr11 Nat$) (veriT_vr12 Nat$)) @p_93) :named @p_95)) :named @p_94)) :rule bind) (step t34 (cl (not @p_94) (not @p_80) @p_95) :rule equiv_pos2) (step t35 (cl @p_95) :rule th_resolution :premises (axiom5 t33 t34)) (anchor :step t36 :args ((:= (veriT_vr10 Nat$) veriT_vr13) (:= (veriT_vr11 Nat$) veriT_vr14) (:= (veriT_vr12 Nat$) veriT_vr15))) (step t36.t1 (cl (! (= veriT_vr10 veriT_vr13) :named @p_97)) :rule refl) (step t36.t2 (cl (! (= veriT_vr11 veriT_vr14) :named @p_101)) :rule refl) (step t36.t3 (cl (= @p_82 (! (less_eq$ veriT_vr13 veriT_vr14) :named @p_96))) :rule cong :premises (t36.t1 t36.t2)) (step t36.t4 (cl (! (= veriT_vr12 veriT_vr15) :named @p_100)) :rule refl) (step t36.t5 (cl @p_97) :rule refl) (step t36.t6 (cl (= @p_85 (! (less_eq$ veriT_vr15 veriT_vr13) :named @p_98))) :rule cong :premises (t36.t4 t36.t5)) (step t36.t7 (cl (= @p_87 (! (and @p_96 @p_98) :named @p_99))) :rule cong :premises (t36.t3 t36.t6)) (step t36.t8 (cl @p_100) :rule refl) (step t36.t9 (cl @p_101) :rule refl) (step t36.t10 (cl (= @p_91 (! (less_eq$ veriT_vr15 veriT_vr14) :named @p_102))) :rule cong :premises (t36.t8 t36.t9)) (step t36.t11 (cl (= @p_93 (! (=> @p_99 @p_102) :named @p_103))) :rule cong :premises (t36.t7 t36.t10)) (step t36 (cl (! (= @p_95 (! (forall ((veriT_vr13 Nat$) (veriT_vr14 Nat$) (veriT_vr15 Nat$)) @p_103) :named @p_105)) :named @p_104)) :rule bind) (step t37 (cl (not @p_104) (not @p_95) @p_105) :rule equiv_pos2) (step t38 (cl @p_105) :rule th_resolution :premises (t35 t36 t37)) (anchor :step t39 :args ((:= (?v0 Astate$) veriT_vr16) (:= (?v1 Astate_v_list_v_result_prod$) veriT_vr17))) (step t39.t1 (cl (! (= ?v0 veriT_vr16) :named @p_114)) :rule refl) (step t39.t2 (cl (! (= ?v1 veriT_vr17) :named @p_113)) :rule refl) (step t39.t3 (cl (= @p_107 (! (fst$ veriT_vr17) :named @p_108))) :rule cong :premises (t39.t2)) (step t39.t4 (cl (= @p_109 (! (= veriT_vr16 @p_108) :named @p_110))) :rule cong :premises (t39.t1 t39.t3)) (anchor :step t39.t5 :args ((:= (?v2 V_list_v_result$) veriT_vr18))) (step t39.t5.t1 (cl @p_113) :rule refl) (step t39.t5.t2 (cl @p_114) :rule refl) (step t39.t5.t3 (cl (= ?v2 veriT_vr18)) :rule refl) (step t39.t5.t4 (cl (= @p_115 (! (pair$ veriT_vr16 veriT_vr18) :named @p_116))) :rule cong :premises (t39.t5.t2 t39.t5.t3)) (step t39.t5.t5 (cl (= @p_117 (! (= veriT_vr17 @p_116) :named @p_118))) :rule cong :premises (t39.t5.t1 t39.t5.t4)) (step t39.t5 (cl (= @p_111 (! (exists ((veriT_vr18 V_list_v_result$)) @p_118) :named @p_112))) :rule bind) (step t39.t6 (cl (= @p_119 (! (= @p_110 @p_112) :named @p_120))) :rule cong :premises (t39.t4 t39.t5)) (step t39 (cl (! (= @p_106 (! (forall ((veriT_vr16 Astate$) (veriT_vr17 Astate_v_list_v_result_prod$)) @p_120) :named @p_122)) :named @p_121)) :rule bind) (step t40 (cl (not @p_121) (not @p_106) @p_122) :rule equiv_pos2) (step t41 (cl @p_122) :rule th_resolution :premises (axiom6 t39 t40)) (anchor :step t42 :args ((veriT_vr16 Astate$) (veriT_vr17 Astate_v_list_v_result_prod$))) (step t42.t1 (cl (= @p_120 (! (and (! (=> @p_110 @p_112) :named @p_133) (! (=> @p_112 @p_110) :named @p_140)) :named @p_123))) :rule connective_def) (step t42 (cl (! (= @p_122 (! (forall ((veriT_vr16 Astate$) (veriT_vr17 Astate_v_list_v_result_prod$)) @p_123) :named @p_125)) :named @p_124)) :rule bind) (step t43 (cl (not @p_124) (not @p_122) @p_125) :rule equiv_pos2) (step t44 (cl @p_125) :rule th_resolution :premises (t41 t42 t43)) (anchor :step t45 :args ((:= (veriT_vr16 Astate$) veriT_vr19) (:= (veriT_vr17 Astate_v_list_v_result_prod$) veriT_vr20))) (step t45.t1 (cl (! (= veriT_vr16 veriT_vr19) :named @p_130)) :rule refl) (step t45.t2 (cl (! (= veriT_vr17 veriT_vr20) :named @p_129)) :rule refl) (step t45.t3 (cl (! (= @p_108 (! (fst$ veriT_vr20) :named @p_127)) :named @p_138)) :rule cong :premises (t45.t2)) (step t45.t4 (cl (! (= @p_110 (! (= veriT_vr19 @p_127) :named @p_126)) :named @p_139)) :rule cong :premises (t45.t1 t45.t3)) (anchor :step t45.t5 :args ((:= (veriT_vr18 V_list_v_result$) veriT_vr21))) (step t45.t5.t1 (cl @p_129) :rule refl) (step t45.t5.t2 (cl @p_130) :rule refl) (step t45.t5.t3 (cl (= veriT_vr18 veriT_vr21)) :rule refl) (step t45.t5.t4 (cl (= @p_116 (! (pair$ veriT_vr19 veriT_vr21) :named @p_131))) :rule cong :premises (t45.t5.t2 t45.t5.t3)) (step t45.t5.t5 (cl (= @p_118 (! (= veriT_vr20 @p_131) :named @p_132))) :rule cong :premises (t45.t5.t1 t45.t5.t4)) (step t45.t5 (cl (= @p_112 (! (exists ((veriT_vr21 V_list_v_result$)) @p_132) :named @p_128))) :rule bind) (step t45.t6 (cl (= @p_133 (! (=> @p_126 @p_128) :named @p_134))) :rule cong :premises (t45.t4 t45.t5)) (anchor :step t45.t7 :args ((:= (veriT_vr18 V_list_v_result$) veriT_vr22))) (step t45.t7.t1 (cl @p_129) :rule refl) (step t45.t7.t2 (cl @p_130) :rule refl) (step t45.t7.t3 (cl (= veriT_vr18 veriT_vr22)) :rule refl) (step t45.t7.t4 (cl (= @p_116 (! (pair$ veriT_vr19 veriT_vr22) :named @p_136))) :rule cong :premises (t45.t7.t2 t45.t7.t3)) (step t45.t7.t5 (cl (= @p_118 (! (= veriT_vr20 @p_136) :named @p_137))) :rule cong :premises (t45.t7.t1 t45.t7.t4)) (step t45.t7 (cl (= @p_112 (! (exists ((veriT_vr22 V_list_v_result$)) @p_137) :named @p_135))) :rule bind) (step t45.t8 (cl @p_130) :rule refl) (step t45.t9 (cl @p_129) :rule refl) (step t45.t10 (cl @p_138) :rule cong :premises (t45.t9)) (step t45.t11 (cl @p_139) :rule cong :premises (t45.t8 t45.t10)) (step t45.t12 (cl (= @p_140 (! (=> @p_135 @p_126) :named @p_141))) :rule cong :premises (t45.t7 t45.t11)) (step t45.t13 (cl (= @p_123 (! (and @p_134 @p_141) :named @p_142))) :rule cong :premises (t45.t6 t45.t12)) (step t45 (cl (! (= @p_125 (! (forall ((veriT_vr19 Astate$) (veriT_vr20 Astate_v_list_v_result_prod$)) @p_142) :named @p_144)) :named @p_143)) :rule bind) (step t46 (cl (not @p_143) (not @p_125) @p_144) :rule equiv_pos2) (step t47 (cl @p_144) :rule th_resolution :premises (t44 t45 t46)) (anchor :step t48 :args ((:= (veriT_vr19 Astate$) veriT_vr23) (:= (veriT_vr20 Astate_v_list_v_result_prod$) veriT_vr24))) (step t48.t1 (cl (! (= veriT_vr19 veriT_vr23) :named @p_149)) :rule refl) (step t48.t2 (cl (! (= veriT_vr20 veriT_vr24) :named @p_148)) :rule refl) (step t48.t3 (cl (! (= @p_127 (! (fst$ veriT_vr24) :named @p_147)) :named @p_153)) :rule cong :premises (t48.t2)) (step t48.t4 (cl (! (= @p_126 (! (= veriT_vr23 @p_147) :named @p_146)) :named @p_154)) :rule cong :premises (t48.t1 t48.t3)) (anchor :step t48.t5 :args ((:= (veriT_vr21 V_list_v_result$) veriT_vr25))) (step t48.t5.t1 (cl @p_148) :rule refl) (step t48.t5.t2 (cl @p_149) :rule refl) (step t48.t5.t3 (cl (= veriT_vr21 veriT_vr25)) :rule refl) (step t48.t5.t4 (cl (= @p_131 (! (pair$ veriT_vr23 veriT_vr25) :named @p_150))) :rule cong :premises (t48.t5.t2 t48.t5.t3)) (step t48.t5.t5 (cl (= @p_132 (! (= veriT_vr24 @p_150) :named @p_151))) :rule cong :premises (t48.t5.t1 t48.t5.t4)) (step t48.t5 (cl (= @p_128 (! (exists ((veriT_vr25 V_list_v_result$)) @p_151) :named @p_145))) :rule bind) (step t48.t6 (cl (= @p_134 (! (=> @p_146 @p_145) :named @p_152))) :rule cong :premises (t48.t4 t48.t5)) (anchor :step t48.t7 :args ((:= (veriT_vr22 V_list_v_result$) veriT_vr25))) (step t48.t7.t1 (cl @p_148) :rule refl) (step t48.t7.t2 (cl @p_149) :rule refl) (step t48.t7.t3 (cl (= veriT_vr22 veriT_vr25)) :rule refl) (step t48.t7.t4 (cl (= @p_136 @p_150)) :rule cong :premises (t48.t7.t2 t48.t7.t3)) (step t48.t7.t5 (cl (= @p_137 @p_151)) :rule cong :premises (t48.t7.t1 t48.t7.t4)) (step t48.t7 (cl (= @p_135 @p_145)) :rule bind) (step t48.t8 (cl @p_149) :rule refl) (step t48.t9 (cl @p_148) :rule refl) (step t48.t10 (cl @p_153) :rule cong :premises (t48.t9)) (step t48.t11 (cl @p_154) :rule cong :premises (t48.t8 t48.t10)) (step t48.t12 (cl (= @p_141 (! (=> @p_145 @p_146) :named @p_155))) :rule cong :premises (t48.t7 t48.t11)) (step t48.t13 (cl (= @p_142 (! (and @p_152 @p_155) :named @p_156))) :rule cong :premises (t48.t6 t48.t12)) (step t48 (cl (! (= @p_144 (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$)) @p_156) :named @p_158)) :named @p_157)) :rule bind) (step t49 (cl (not @p_157) (not @p_144) @p_158) :rule equiv_pos2) (step t50 (cl @p_158) :rule th_resolution :premises (t47 t48 t49)) (anchor :step t51 :args ((:= (veriT_vr23 Astate$) veriT_vr23) (:= (veriT_vr24 Astate_v_list_v_result_prod$) veriT_vr24))) (anchor :step t51.t1 :args ((:= (veriT_vr25 V_list_v_result$) veriT_vr26))) (step t51.t1.t1 (cl (= veriT_vr25 veriT_vr26)) :rule refl) (step t51.t1.t2 (cl (= @p_150 (! (pair$ veriT_vr23 veriT_vr26) :named @p_160))) :rule cong :premises (t51.t1.t1)) (step t51.t1.t3 (cl (= @p_151 (! (= veriT_vr24 @p_160) :named @p_161))) :rule cong :premises (t51.t1.t2)) (step t51.t1 (cl (= @p_145 (! (exists ((veriT_vr26 V_list_v_result$)) @p_161) :named @p_159))) :rule bind) (step t51.t2 (cl (= @p_155 (! (=> @p_159 @p_146) :named @p_162))) :rule cong :premises (t51.t1)) (step t51.t3 (cl (= @p_156 (! (and @p_152 @p_162) :named @p_163))) :rule cong :premises (t51.t2)) (step t51 (cl (! (= @p_158 (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$)) @p_163) :named @p_165)) :named @p_164)) :rule bind) (step t52 (cl (not @p_164) (not @p_158) @p_165) :rule equiv_pos2) (step t53 (cl @p_165) :rule th_resolution :premises (t50 t51 t52)) (anchor :step t54 :args ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$))) (step t54.t1 (cl (= @p_159 (! (not (forall ((veriT_vr26 V_list_v_result$)) (! (not @p_161) :named @p_367))) :named @p_166))) :rule connective_def) (step t54.t2 (cl (= @p_162 (! (=> @p_166 @p_146) :named @p_167))) :rule cong :premises (t54.t1)) (step t54.t3 (cl (= @p_163 (! (and @p_152 @p_167) :named @p_168))) :rule cong :premises (t54.t2)) (step t54 (cl (! (= @p_165 (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$)) @p_168) :named @p_170)) :named @p_169)) :rule bind) (step t55 (cl (not @p_169) (not @p_165) @p_170) :rule equiv_pos2) (step t56 (cl @p_170) :rule th_resolution :premises (t53 t54 t55)) (anchor :step t57 :args ((:= (?v0 V_error_result$) veriT_vr27))) (anchor :step t57.t1 :args ((:= (?v1 V$) veriT_vr28))) (step t57.t1.t1 (cl (! (= ?v0 veriT_vr27) :named @p_181)) :rule refl) (step t57.t1.t2 (cl (= ?v1 veriT_vr28)) :rule refl) (step t57.t1.t3 (cl (= @p_174 (! (rraise$ veriT_vr28) :named @p_175))) :rule cong :premises (t57.t1.t2)) (step t57.t1.t4 (cl (= @p_6 (! (= veriT_vr27 @p_175) :named @p_176))) :rule cong :premises (t57.t1.t1 t57.t1.t3)) (step t57.t1.t5 (cl (= @p_177 (! (=> @p_176 false) :named @p_178))) :rule cong :premises (t57.t1.t4)) (step t57.t1 (cl (= @p_172 (! (forall ((veriT_vr28 V$)) @p_178) :named @p_173))) :rule bind) (anchor :step t57.t2 :args ((:= (?v1 Abort$) veriT_vr29))) (step t57.t2.t1 (cl @p_181) :rule refl) (step t57.t2.t2 (cl (= ?v1 veriT_vr29)) :rule refl) (step t57.t2.t3 (cl (= @p_182 (! (rabort$ veriT_vr29) :named @p_183))) :rule cong :premises (t57.t2.t2)) (step t57.t2.t4 (cl (= @p_184 (! (= veriT_vr27 @p_183) :named @p_185))) :rule cong :premises (t57.t2.t1 t57.t2.t3)) (step t57.t2.t5 (cl (= @p_186 (! (=> @p_185 false) :named @p_187))) :rule cong :premises (t57.t2.t4)) (step t57.t2 (cl (= @p_179 (! (forall ((veriT_vr29 Abort$)) @p_187) :named @p_180))) :rule bind) (step t57.t3 (cl (= @p_188 (! (and @p_173 @p_180) :named @p_189))) :rule cong :premises (t57.t1 t57.t2)) (step t57.t4 (cl (= @p_190 (! (=> @p_189 false) :named @p_191))) :rule cong :premises (t57.t3)) (step t57 (cl (! (= @p_171 (! (forall ((veriT_vr27 V_error_result$)) @p_191) :named @p_193)) :named @p_192)) :rule bind) (step t58 (cl (not @p_192) (not @p_171) @p_193) :rule equiv_pos2) (step t59 (cl @p_193) :rule th_resolution :premises (axiom7 t57 t58)) (anchor :step t60 :args ((veriT_vr27 V_error_result$))) (anchor :step t60.t1 :args ((veriT_vr28 V$))) (step t60.t1.t1 (cl (= @p_178 (! (not @p_176) :named @p_195))) :rule implies_simplify) (step t60.t1 (cl (= @p_173 (! (forall ((veriT_vr28 V$)) @p_195) :named @p_194))) :rule bind) (anchor :step t60.t2 :args ((veriT_vr29 Abort$))) (step t60.t2.t1 (cl (= @p_187 (! (not @p_185) :named @p_197))) :rule implies_simplify) (step t60.t2 (cl (= @p_180 (! (forall ((veriT_vr29 Abort$)) @p_197) :named @p_196))) :rule bind) (step t60.t3 (cl (= @p_189 (! (and @p_194 @p_196) :named @p_198))) :rule cong :premises (t60.t1 t60.t2)) (step t60.t4 (cl (= @p_191 (! (=> @p_198 false) :named @p_199))) :rule cong :premises (t60.t3)) (step t60.t5 (cl (= @p_199 (! (not @p_198) :named @p_200))) :rule implies_simplify) (step t60.t6 (cl (= @p_191 @p_200)) :rule trans :premises (t60.t4 t60.t5)) (step t60 (cl (! (= @p_193 (! (forall ((veriT_vr27 V_error_result$)) @p_200) :named @p_202)) :named @p_201)) :rule bind) (step t61 (cl (not @p_201) (not @p_193) @p_202) :rule equiv_pos2) (step t62 (cl @p_202) :rule th_resolution :premises (t59 t60 t61)) (anchor :step t63 :args ((:= (veriT_vr27 V_error_result$) veriT_vr30))) (anchor :step t63.t1 :args ((:= (veriT_vr28 V$) veriT_vr31))) (step t63.t1.t1 (cl (! (= veriT_vr27 veriT_vr30) :named @p_208)) :rule refl) (step t63.t1.t2 (cl (= veriT_vr28 veriT_vr31)) :rule refl) (step t63.t1.t3 (cl (= @p_175 (! (rraise$ veriT_vr31) :named @p_204))) :rule cong :premises (t63.t1.t2)) (step t63.t1.t4 (cl (= @p_176 (! (= veriT_vr30 @p_204) :named @p_205))) :rule cong :premises (t63.t1.t1 t63.t1.t3)) (step t63.t1.t5 (cl (= @p_195 (! (not @p_205) :named @p_206))) :rule cong :premises (t63.t1.t4)) (step t63.t1 (cl (= @p_194 (! (forall ((veriT_vr31 V$)) @p_206) :named @p_203))) :rule bind) (anchor :step t63.t2 :args ((:= (veriT_vr29 Abort$) veriT_vr32))) (step t63.t2.t1 (cl @p_208) :rule refl) (step t63.t2.t2 (cl (= veriT_vr29 veriT_vr32)) :rule refl) (step t63.t2.t3 (cl (= @p_183 (! (rabort$ veriT_vr32) :named @p_209))) :rule cong :premises (t63.t2.t2)) (step t63.t2.t4 (cl (= @p_185 (! (= veriT_vr30 @p_209) :named @p_210))) :rule cong :premises (t63.t2.t1 t63.t2.t3)) (step t63.t2.t5 (cl (= @p_197 (! (not @p_210) :named @p_211))) :rule cong :premises (t63.t2.t4)) (step t63.t2 (cl (= @p_196 (! (forall ((veriT_vr32 Abort$)) @p_211) :named @p_207))) :rule bind) (step t63.t3 (cl (= @p_198 (! (and @p_203 @p_207) :named @p_212))) :rule cong :premises (t63.t1 t63.t2)) (step t63.t4 (cl (= @p_200 (! (not @p_212) :named @p_213))) :rule cong :premises (t63.t3)) (step t63 (cl (! (= @p_202 (! (forall ((veriT_vr30 V_error_result$)) @p_213) :named @p_215)) :named @p_214)) :rule bind) (step t64 (cl (not @p_214) (not @p_202) @p_215) :rule equiv_pos2) (step t65 (cl @p_215) :rule th_resolution :premises (t62 t63 t64)) (anchor :step t66 :args ((:= (?v0 V_astate_v_list_v_result_prod_fun$) veriT_vr33) (:= (?v1 Abort_astate_v_list_v_result_prod_fun$) veriT_vr34) (:= (?v2 V$) veriT_vr35))) (step t66.t1 (cl (! (= ?v0 veriT_vr33) :named @p_221)) :rule refl) (step t66.t2 (cl (= ?v1 veriT_vr34)) :rule refl) (step t66.t3 (cl (! (= ?v2 veriT_vr35) :named @p_222)) :rule refl) (step t66.t4 (cl (= @p_217 (! (rraise$ veriT_vr35) :named @p_218))) :rule cong :premises (t66.t3)) (step t66.t5 (cl (= @p_219 (! (case_error_result$ veriT_vr33 veriT_vr34 @p_218) :named @p_220))) :rule cong :premises (t66.t1 t66.t2 t66.t4)) (step t66.t6 (cl @p_221) :rule refl) (step t66.t7 (cl @p_222) :rule refl) (step t66.t8 (cl (= @p_223 (! (fun_app$ veriT_vr33 veriT_vr35) :named @p_224))) :rule cong :premises (t66.t6 t66.t7)) (step t66.t9 (cl (= @p_225 (! (= @p_220 @p_224) :named @p_226))) :rule cong :premises (t66.t5 t66.t8)) (step t66 (cl (! (= @p_216 (! (forall ((veriT_vr33 V_astate_v_list_v_result_prod_fun$) (veriT_vr34 Abort_astate_v_list_v_result_prod_fun$) (veriT_vr35 V$)) @p_226) :named @p_228)) :named @p_227)) :rule bind) (step t67 (cl (not @p_227) (not @p_216) @p_228) :rule equiv_pos2) (step t68 (cl @p_228) :rule th_resolution :premises (axiom8 t66 t67)) (anchor :step t69 :args ((:= (veriT_vr33 V_astate_v_list_v_result_prod_fun$) veriT_vr36) (:= (veriT_vr34 Abort_astate_v_list_v_result_prod_fun$) veriT_vr37) (:= (veriT_vr35 V$) veriT_vr38))) (step t69.t1 (cl (! (= veriT_vr33 veriT_vr36) :named @p_231)) :rule refl) (step t69.t2 (cl (= veriT_vr34 veriT_vr37)) :rule refl) (step t69.t3 (cl (! (= veriT_vr35 veriT_vr38) :named @p_232)) :rule refl) (step t69.t4 (cl (= @p_218 (! (rraise$ veriT_vr38) :named @p_229))) :rule cong :premises (t69.t3)) (step t69.t5 (cl (= @p_220 (! (case_error_result$ veriT_vr36 veriT_vr37 @p_229) :named @p_230))) :rule cong :premises (t69.t1 t69.t2 t69.t4)) (step t69.t6 (cl @p_231) :rule refl) (step t69.t7 (cl @p_232) :rule refl) (step t69.t8 (cl (= @p_224 (! (fun_app$ veriT_vr36 veriT_vr38) :named @p_233))) :rule cong :premises (t69.t6 t69.t7)) (step t69.t9 (cl (= @p_226 (! (= @p_230 @p_233) :named @p_234))) :rule cong :premises (t69.t5 t69.t8)) (step t69 (cl (! (= @p_228 (! (forall ((veriT_vr36 V_astate_v_list_v_result_prod_fun$) (veriT_vr37 Abort_astate_v_list_v_result_prod_fun$) (veriT_vr38 V$)) @p_234) :named @p_236)) :named @p_235)) :rule bind) (step t70 (cl (not @p_235) (not @p_228) @p_236) :rule equiv_pos2) (step t71 (cl @p_236) :rule th_resolution :premises (t68 t69 t70)) (anchor :step t72 :args ((:= (?v0 V_astate_v_list_v_result_prod_fun$) veriT_vr39) (:= (?v1 Abort_astate_v_list_v_result_prod_fun$) veriT_vr40) (:= (?v2 Abort$) veriT_vr41))) (step t72.t1 (cl (= ?v0 veriT_vr39)) :rule refl) (step t72.t2 (cl (! (= ?v1 veriT_vr40) :named @p_242)) :rule refl) (step t72.t3 (cl (! (= ?v2 veriT_vr41) :named @p_243)) :rule refl) (step t72.t4 (cl (= @p_238 (! (rabort$ veriT_vr41) :named @p_239))) :rule cong :premises (t72.t3)) (step t72.t5 (cl (= @p_240 (! (case_error_result$ veriT_vr39 veriT_vr40 @p_239) :named @p_241))) :rule cong :premises (t72.t1 t72.t2 t72.t4)) (step t72.t6 (cl @p_242) :rule refl) (step t72.t7 (cl @p_243) :rule refl) (step t72.t8 (cl (= @p_244 (! (fun_app$a veriT_vr40 veriT_vr41) :named @p_245))) :rule cong :premises (t72.t6 t72.t7)) (step t72.t9 (cl (= @p_246 (! (= @p_241 @p_245) :named @p_247))) :rule cong :premises (t72.t5 t72.t8)) (step t72 (cl (! (= @p_237 (! (forall ((veriT_vr39 V_astate_v_list_v_result_prod_fun$) (veriT_vr40 Abort_astate_v_list_v_result_prod_fun$) (veriT_vr41 Abort$)) @p_247) :named @p_249)) :named @p_248)) :rule bind) (step t73 (cl (not @p_248) (not @p_237) @p_249) :rule equiv_pos2) (step t74 (cl @p_249) :rule th_resolution :premises (axiom9 t72 t73)) (anchor :step t75 :args ((:= (veriT_vr39 V_astate_v_list_v_result_prod_fun$) veriT_vr42) (:= (veriT_vr40 Abort_astate_v_list_v_result_prod_fun$) veriT_vr43) (:= (veriT_vr41 Abort$) veriT_vr44))) (step t75.t1 (cl (= veriT_vr39 veriT_vr42)) :rule refl) (step t75.t2 (cl (! (= veriT_vr40 veriT_vr43) :named @p_252)) :rule refl) (step t75.t3 (cl (! (= veriT_vr41 veriT_vr44) :named @p_253)) :rule refl) (step t75.t4 (cl (= @p_239 (! (rabort$ veriT_vr44) :named @p_250))) :rule cong :premises (t75.t3)) (step t75.t5 (cl (= @p_241 (! (case_error_result$ veriT_vr42 veriT_vr43 @p_250) :named @p_251))) :rule cong :premises (t75.t1 t75.t2 t75.t4)) (step t75.t6 (cl @p_252) :rule refl) (step t75.t7 (cl @p_253) :rule refl) (step t75.t8 (cl (= @p_245 (! (fun_app$a veriT_vr43 veriT_vr44) :named @p_254))) :rule cong :premises (t75.t6 t75.t7)) (step t75.t9 (cl (= @p_247 (! (= @p_251 @p_254) :named @p_255))) :rule cong :premises (t75.t5 t75.t8)) (step t75 (cl (! (= @p_249 (! (forall ((veriT_vr42 V_astate_v_list_v_result_prod_fun$) (veriT_vr43 Abort_astate_v_list_v_result_prod_fun$) (veriT_vr44 Abort$)) @p_255) :named @p_257)) :named @p_256)) :rule bind) (step t76 (cl (not @p_256) (not @p_249) @p_257) :rule equiv_pos2) (step t77 (cl @p_257) :rule th_resolution :premises (t74 t75 t76)) (anchor :step t78 :args ((:= (?v0 Astate$) veriT_vr45) (:= (?v1 Astate$) veriT_vr46) (:= (?v2 V_list_v_result$) veriT_vr47) (:= (?v3 Astate$) veriT_vr48))) (step t78.t1 (cl (= ?v0 veriT_vr45)) :rule refl) (step t78.t2 (cl (! (= ?v1 veriT_vr46) :named @p_270)) :rule refl) (step t78.t3 (cl (! (= ?v2 veriT_vr47) :named @p_262)) :rule refl) (step t78.t4 (cl (= @p_259 (! (pair$ veriT_vr46 veriT_vr47) :named @p_260))) :rule cong :premises (t78.t2 t78.t3)) (step t78.t5 (cl (= @p_4 (! (fix_clock$ veriT_vr45 @p_260) :named @p_261))) :rule cong :premises (t78.t1 t78.t4)) (step t78.t6 (cl (! (= ?v3 veriT_vr48) :named @p_267)) :rule refl) (step t78.t7 (cl @p_262) :rule refl) (step t78.t8 (cl (= @p_263 (! (pair$ veriT_vr48 veriT_vr47) :named @p_264))) :rule cong :premises (t78.t6 t78.t7)) (step t78.t9 (cl (= @p_265 (! (= @p_261 @p_264) :named @p_266))) :rule cong :premises (t78.t5 t78.t8)) (step t78.t10 (cl @p_267) :rule refl) (step t78.t11 (cl (= @p_268 (! (clock$ veriT_vr48) :named @p_269))) :rule cong :premises (t78.t10)) (step t78.t12 (cl @p_270) :rule refl) (step t78.t13 (cl (= @p_1 (! (clock$ veriT_vr46) :named @p_271))) :rule cong :premises (t78.t12)) (step t78.t14 (cl (= @p_272 (! (less_eq$ @p_269 @p_271) :named @p_273))) :rule cong :premises (t78.t11 t78.t13)) (step t78.t15 (cl (= @p_274 (! (=> @p_266 @p_273) :named @p_275))) :rule cong :premises (t78.t9 t78.t14)) (step t78 (cl (! (= @p_258 (! (forall ((veriT_vr45 Astate$) (veriT_vr46 Astate$) (veriT_vr47 V_list_v_result$) (veriT_vr48 Astate$)) @p_275) :named @p_277)) :named @p_276)) :rule bind) (step t79 (cl (not @p_276) (not @p_258) @p_277) :rule equiv_pos2) (step t80 (cl @p_277) :rule th_resolution :premises (axiom10 t78 t79)) (anchor :step t81 :args ((:= (veriT_vr45 Astate$) veriT_vr49) (:= (veriT_vr46 Astate$) veriT_vr50) (:= (veriT_vr47 V_list_v_result$) veriT_vr51) (:= (veriT_vr48 Astate$) veriT_vr52))) (step t81.t1 (cl (= veriT_vr45 veriT_vr49)) :rule refl) (step t81.t2 (cl (! (= veriT_vr46 veriT_vr50) :named @p_285)) :rule refl) (step t81.t3 (cl (! (= veriT_vr47 veriT_vr51) :named @p_280)) :rule refl) (step t81.t4 (cl (= @p_260 (! (pair$ veriT_vr50 veriT_vr51) :named @p_278))) :rule cong :premises (t81.t2 t81.t3)) (step t81.t5 (cl (= @p_261 (! (fix_clock$ veriT_vr49 @p_278) :named @p_279))) :rule cong :premises (t81.t1 t81.t4)) (step t81.t6 (cl (! (= veriT_vr48 veriT_vr52) :named @p_283)) :rule refl) (step t81.t7 (cl @p_280) :rule refl) (step t81.t8 (cl (= @p_264 (! (pair$ veriT_vr52 veriT_vr51) :named @p_281))) :rule cong :premises (t81.t6 t81.t7)) (step t81.t9 (cl (= @p_266 (! (= @p_279 @p_281) :named @p_282))) :rule cong :premises (t81.t5 t81.t8)) (step t81.t10 (cl @p_283) :rule refl) (step t81.t11 (cl (= @p_269 (! (clock$ veriT_vr52) :named @p_284))) :rule cong :premises (t81.t10)) (step t81.t12 (cl @p_285) :rule refl) (step t81.t13 (cl (= @p_271 (! (clock$ veriT_vr50) :named @p_286))) :rule cong :premises (t81.t12)) (step t81.t14 (cl (= @p_273 (! (less_eq$ @p_284 @p_286) :named @p_287))) :rule cong :premises (t81.t11 t81.t13)) (step t81.t15 (cl (= @p_275 (! (=> @p_282 @p_287) :named @p_288))) :rule cong :premises (t81.t9 t81.t14)) (step t81 (cl (! (= @p_277 (! (forall ((veriT_vr49 Astate$) (veriT_vr50 Astate$) (veriT_vr51 V_list_v_result$) (veriT_vr52 Astate$)) @p_288) :named @p_290)) :named @p_289)) :rule bind) (step t82 (cl (not @p_289) (not @p_277) @p_290) :rule equiv_pos2) (step t83 (cl @p_290) :rule th_resolution :premises (t80 t81 t82)) (anchor :step t84 :args ((:= (?v0 Astate$) veriT_vr53) (:= (?v1 Astate$) veriT_vr54) (:= (?v2 V_list_v_result$) veriT_vr55))) (step t84.t1 (cl (! (= ?v0 veriT_vr53) :named @p_294)) :rule refl) (step t84.t2 (cl (! (= ?v1 veriT_vr54) :named @p_295)) :rule refl) (step t84.t3 (cl (! (= ?v2 veriT_vr55) :named @p_299)) :rule refl) (step t84.t4 (cl (= @p_259 (! (pair$ veriT_vr54 veriT_vr55) :named @p_292))) :rule cong :premises (t84.t2 t84.t3)) (step t84.t5 (cl (= @p_4 (! (fix_clock$ veriT_vr53 @p_292) :named @p_293))) :rule cong :premises (t84.t1 t84.t4)) (step t84.t6 (cl @p_294) :rule refl) (step t84.t7 (cl @p_295) :rule refl) (step t84.t8 (cl (= @p_5 (! (uu$ veriT_vr53 veriT_vr54) :named @p_296))) :rule cong :premises (t84.t6 t84.t7)) (step t84.t9 (cl @p_295) :rule refl) (step t84.t10 (cl (= @p_297 (! (update_clock$ @p_296 veriT_vr54) :named @p_298))) :rule cong :premises (t84.t8 t84.t9)) (step t84.t11 (cl @p_299) :rule refl) (step t84.t12 (cl (= @p_300 (! (pair$ @p_298 veriT_vr55) :named @p_301))) :rule cong :premises (t84.t10 t84.t11)) (step t84.t13 (cl (= @p_302 (! (= @p_293 @p_301) :named @p_303))) :rule cong :premises (t84.t5 t84.t12)) (step t84 (cl (! (= @p_291 (! (forall ((veriT_vr53 Astate$) (veriT_vr54 Astate$) (veriT_vr55 V_list_v_result$)) @p_303) :named @p_305)) :named @p_304)) :rule bind) (step t85 (cl (not @p_304) (not @p_291) @p_305) :rule equiv_pos2) (step t86 (cl @p_305) :rule th_resolution :premises (axiom11 t84 t85)) (anchor :step t87 :args ((:= (veriT_vr53 Astate$) veriT_vr56) (:= (veriT_vr54 Astate$) veriT_vr57) (:= (veriT_vr55 V_list_v_result$) veriT_vr58))) (step t87.t1 (cl (! (= veriT_vr53 veriT_vr56) :named @p_308)) :rule refl) (step t87.t2 (cl (! (= veriT_vr54 veriT_vr57) :named @p_309)) :rule refl) (step t87.t3 (cl (! (= veriT_vr55 veriT_vr58) :named @p_312)) :rule refl) (step t87.t4 (cl (= @p_292 (! (pair$ veriT_vr57 veriT_vr58) :named @p_306))) :rule cong :premises (t87.t2 t87.t3)) (step t87.t5 (cl (= @p_293 (! (fix_clock$ veriT_vr56 @p_306) :named @p_307))) :rule cong :premises (t87.t1 t87.t4)) (step t87.t6 (cl @p_308) :rule refl) (step t87.t7 (cl @p_309) :rule refl) (step t87.t8 (cl (= @p_296 (! (uu$ veriT_vr56 veriT_vr57) :named @p_310))) :rule cong :premises (t87.t6 t87.t7)) (step t87.t9 (cl @p_309) :rule refl) (step t87.t10 (cl (= @p_298 (! (update_clock$ @p_310 veriT_vr57) :named @p_311))) :rule cong :premises (t87.t8 t87.t9)) (step t87.t11 (cl @p_312) :rule refl) (step t87.t12 (cl (= @p_301 (! (pair$ @p_311 veriT_vr58) :named @p_313))) :rule cong :premises (t87.t10 t87.t11)) (step t87.t13 (cl (= @p_303 (! (= @p_307 @p_313) :named @p_314))) :rule cong :premises (t87.t5 t87.t12)) (step t87 (cl (! (= @p_305 (! (forall ((veriT_vr56 Astate$) (veriT_vr57 Astate$) (veriT_vr58 V_list_v_result$)) @p_314) :named @p_316)) :named @p_315)) :rule bind) (step t88 (cl (not @p_315) (not @p_305) @p_316) :rule equiv_pos2) (step t89 (cl @p_316) :rule th_resolution :premises (t86 t87 t88)) (anchor :step t90 :args ((:= (?v0 V_error_result$) veriT_vr59) (:= (?v1 V$) veriT_vr60))) (step t90.t1 (cl (! (= ?v0 veriT_vr59) :named @p_323)) :rule refl) (step t90.t2 (cl (= @p_319 (! (rerr$ veriT_vr59) :named @p_320))) :rule cong :premises (t90.t1)) (step t90.t3 (cl (= @p_321 (! (= r$ @p_320) :named @p_322))) :rule cong :premises (t90.t2)) (step t90.t4 (cl @p_323) :rule refl) (step t90.t5 (cl (! (= ?v1 veriT_vr60) :named @p_328)) :rule refl) (step t90.t6 (cl (= @p_174 (! (rraise$ veriT_vr60) :named @p_324))) :rule cong :premises (t90.t5)) (step t90.t7 (cl (= @p_6 (! (= veriT_vr59 @p_324) :named @p_325))) :rule cong :premises (t90.t4 t90.t6)) (step t90.t8 (cl (= @p_326 (! (and @p_322 @p_325) :named @p_327))) :rule cong :premises (t90.t3 t90.t7)) (step t90.t9 (cl @p_328) :rule refl) (step t90.t10 (cl (= @p_329 (! (fun_evaluate_match$ st$ env$ veriT_vr60 pes$) :named @p_330))) :rule cong :premises (t90.t9)) (step t90.t11 (cl @p_328) :rule refl) (step t90.t12 (cl (= @p_331 (! (fun_app$ @p_330 veriT_vr60) :named @p_332))) :rule cong :premises (t90.t10 t90.t11)) (step t90.t13 (cl (= @p_333 (! (fst$ @p_332) :named @p_334))) :rule cong :premises (t90.t12)) (step t90.t14 (cl (= @p_335 (! (clock$ @p_334) :named @p_336))) :rule cong :premises (t90.t13)) (step t90.t15 (cl (= @p_337 (! (less_eq$ @p_336 @p_318) :named @p_338))) :rule cong :premises (t90.t14)) (step t90.t16 (cl (= @p_339 (! (=> @p_327 @p_338) :named @p_340))) :rule cong :premises (t90.t8 t90.t15)) (step t90 (cl (! (= @p_317 (! (forall ((veriT_vr59 V_error_result$) (veriT_vr60 V$)) @p_340) :named @p_342)) :named @p_341)) :rule bind) (step t91 (cl (not @p_341) (not @p_317) @p_342) :rule equiv_pos2) (step t92 (cl @p_342) :rule th_resolution :premises (axiom12 t90 t91)) (anchor :step t93 :args ((:= (veriT_vr59 V_error_result$) veriT_vr61) (:= (veriT_vr60 V$) veriT_vr62))) (step t93.t1 (cl (! (= veriT_vr59 veriT_vr61) :named @p_345)) :rule refl) (step t93.t2 (cl (= @p_320 (! (rerr$ veriT_vr61) :named @p_343))) :rule cong :premises (t93.t1)) (step t93.t3 (cl (= @p_322 (! (= r$ @p_343) :named @p_344))) :rule cong :premises (t93.t2)) (step t93.t4 (cl @p_345) :rule refl) (step t93.t5 (cl (! (= veriT_vr60 veriT_vr62) :named @p_349)) :rule refl) (step t93.t6 (cl (= @p_324 (! (rraise$ veriT_vr62) :named @p_346))) :rule cong :premises (t93.t5)) (step t93.t7 (cl (= @p_325 (! (= veriT_vr61 @p_346) :named @p_347))) :rule cong :premises (t93.t4 t93.t6)) (step t93.t8 (cl (= @p_327 (! (and @p_344 @p_347) :named @p_348))) :rule cong :premises (t93.t3 t93.t7)) (step t93.t9 (cl @p_349) :rule refl) (step t93.t10 (cl (= @p_330 (! (fun_evaluate_match$ st$ env$ veriT_vr62 pes$) :named @p_350))) :rule cong :premises (t93.t9)) (step t93.t11 (cl @p_349) :rule refl) (step t93.t12 (cl (= @p_332 (! (fun_app$ @p_350 veriT_vr62) :named @p_351))) :rule cong :premises (t93.t10 t93.t11)) (step t93.t13 (cl (= @p_334 (! (fst$ @p_351) :named @p_352))) :rule cong :premises (t93.t12)) (step t93.t14 (cl (= @p_336 (! (clock$ @p_352) :named @p_353))) :rule cong :premises (t93.t13)) (step t93.t15 (cl (= @p_338 (! (less_eq$ @p_353 @p_318) :named @p_354))) :rule cong :premises (t93.t14)) (step t93.t16 (cl (= @p_340 (! (=> @p_348 @p_354) :named @p_355))) :rule cong :premises (t93.t8 t93.t15)) (step t93 (cl (! (= @p_342 (! (forall ((veriT_vr61 V_error_result$) (veriT_vr62 V$)) @p_355) :named @p_357)) :named @p_356)) :rule bind) (step t94 (cl (not @p_356) (not @p_342) @p_357) :rule equiv_pos2) (step t95 (cl @p_357) :rule th_resolution :premises (t92 t93 t94)) (step t96 (cl (! (= @p_358 (! (and @p_359 (! (not @p_360) :named @p_366)) :named @p_362)) :named @p_361)) :rule bool_simplify) (step t97 (cl (! (not @p_361) :named @p_365) (! (not @p_358) :named @p_363) @p_362) :rule equiv_pos2) (step t98 (cl (not @p_363) @p_364) :rule not_not) (step t99 (cl @p_365 @p_364 @p_362) :rule th_resolution :premises (t98 t97)) (step t100 (cl @p_362) :rule th_resolution :premises (axiom13 t96 t99)) (step t101 (cl @p_359) :rule and :premises (t100)) (step t102 (cl @p_366) :rule and :premises (t100)) (step t103 (cl (or (! (not @p_105) :named @p_368) (! (forall ((veriT_vr13 Nat$) (veriT_vr14 Nat$) (veriT_vr15 Nat$)) (or (not @p_96) (not @p_98) @p_102)) :named @p_573))) :rule qnt_cnf) (step t104 (cl (or (! (not @p_170) :named @p_431) (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$) (veriT_vr26 V_list_v_result$)) (or @p_367 @p_146)) :named @p_629))) :rule qnt_cnf) (step t105 (cl (or @p_368 (! (=> (! (and @p_369 (! (less_eq$ @p_370 @p_371) :named @p_373)) :named @p_372) @p_360) :named @p_374))) :rule forall_inst :args ((:= veriT_vr13 @p_371) (:= veriT_vr14 @p_7) (:= veriT_vr15 @p_370))) (step t106 (cl @p_372 (! (not @p_369) :named @p_574) (! (not @p_373) :named @p_375)) :rule and_neg) (step t107 (cl (! (not @p_374) :named @p_376) (! (not @p_372) :named @p_377) @p_360) :rule implies_pos) (step t108 (cl @p_368 @p_374) :rule or :premises (t105)) (step t109 (cl @p_372 @p_375) :rule resolution :premises (t106 axiom4)) (step t110 (cl @p_376 @p_377) :rule resolution :premises (t107 t102)) (step t111 (cl @p_374) :rule resolution :premises (t108 t38)) (step t112 (cl @p_377) :rule resolution :premises (t110 t111)) (step t113 (cl @p_375) :rule resolution :premises (t109 t112)) (step t114 (cl (not (! (not @p_368) :named @p_578)) @p_105) :rule not_not) (step t115 (cl (or (! (not @p_316) :named @p_547) (! (= (fix_clock$ st$a (pair$ @p_378 r$)) (pair$ (! (update_clock$ (uu$ st$a @p_378) @p_378) :named @p_561) r$)) :named @p_548))) :rule forall_inst :args ((:= veriT_vr56 st$a) (:= veriT_vr57 @p_378) (:= veriT_vr58 r$))) (step t116 (cl (or (! (not @p_215) :named @p_427) (! (not (! (and (! (forall ((veriT_vr31 V$)) (! (not (! (= x2$ @p_204) :named @p_382)) :named @p_384)) :named @p_380) (! (forall ((veriT_vr32 Abort$)) (! (not (! (= x2$ @p_209) :named @p_388)) :named @p_390)) :named @p_386)) :named @p_392)) :named @p_379))) :rule forall_inst :args ((:= veriT_vr30 x2$))) (anchor :step t117) (assume t117.h1 @p_379) (anchor :step t117.t2 :args ((:= (veriT_vr31 V$) veriT_vr63))) (step t117.t2.t1 (cl (= veriT_vr31 veriT_vr63)) :rule refl) (step t117.t2.t2 (cl (= @p_204 (! (rraise$ veriT_vr63) :named @p_381))) :rule cong :premises (t117.t2.t1)) (step t117.t2.t3 (cl (= @p_382 (! (= x2$ @p_381) :named @p_383))) :rule cong :premises (t117.t2.t2)) (step t117.t2.t4 (cl (= @p_384 (! (not @p_383) :named @p_385))) :rule cong :premises (t117.t2.t3)) (step t117.t2 (cl (= @p_380 (! (forall ((veriT_vr63 V$)) @p_385) :named @p_393))) :rule bind) (anchor :step t117.t3 :args ((:= (veriT_vr32 Abort$) veriT_vr64))) (step t117.t3.t1 (cl (= veriT_vr32 veriT_vr64)) :rule refl) (step t117.t3.t2 (cl (= @p_209 (! (rabort$ veriT_vr64) :named @p_387))) :rule cong :premises (t117.t3.t1)) (step t117.t3.t3 (cl (= @p_388 (! (= x2$ @p_387) :named @p_389))) :rule cong :premises (t117.t3.t2)) (step t117.t3.t4 (cl (= @p_390 (! (not @p_389) :named @p_391))) :rule cong :premises (t117.t3.t3)) (step t117.t3 (cl (= @p_386 (! (forall ((veriT_vr64 Abort$)) @p_391) :named @p_394))) :rule bind) (step t117.t4 (cl (= @p_392 (! (and @p_393 @p_394) :named @p_395))) :rule cong :premises (t117.t2 t117.t3)) (step t117.t5 (cl (! (= @p_379 (! (not @p_395) :named @p_398)) :named @p_396)) :rule cong :premises (t117.t4)) (step t117.t6 (cl (! (not @p_396) :named @p_399) (! (not @p_379) :named @p_397) @p_398) :rule equiv_pos2) (step t117.t7 (cl (! (not @p_397) :named @p_426) @p_392) :rule not_not) (step t117.t8 (cl @p_399 @p_392 @p_398) :rule th_resolution :premises (t117.t7 t117.t6)) (step t117.t9 (cl @p_398) :rule th_resolution :premises (t117.h1 t117.t5 t117.t8)) (anchor :step t117.t10 :args ((:= (veriT_vr63 V$) veriT_vr65))) (step t117.t10.t1 (cl (= veriT_vr63 veriT_vr65)) :rule refl) (step t117.t10.t2 (cl (= @p_381 @p_401)) :rule cong :premises (t117.t10.t1)) (step t117.t10.t3 (cl (= @p_383 @p_402)) :rule cong :premises (t117.t10.t2)) (step t117.t10.t4 (cl (= @p_385 @p_400)) :rule cong :premises (t117.t10.t3)) (step t117.t10 (cl (= @p_393 (! (forall ((veriT_vr65 V$)) @p_400) :named @p_406))) :rule bind) (anchor :step t117.t11 :args ((:= (veriT_vr64 Abort$) veriT_vr66))) (step t117.t11.t1 (cl (= veriT_vr64 veriT_vr66)) :rule refl) (step t117.t11.t2 (cl (= @p_387 @p_404)) :rule cong :premises (t117.t11.t1)) (step t117.t11.t3 (cl (= @p_389 @p_405)) :rule cong :premises (t117.t11.t2)) (step t117.t11.t4 (cl (= @p_391 @p_403)) :rule cong :premises (t117.t11.t3)) (step t117.t11 (cl (= @p_394 (! (forall ((veriT_vr66 Abort$)) @p_403) :named @p_407))) :rule bind) (step t117.t12 (cl (= @p_395 (! (and @p_406 @p_407) :named @p_408))) :rule cong :premises (t117.t10 t117.t11)) (step t117.t13 (cl (! (= @p_398 (! (not @p_408) :named @p_410)) :named @p_409)) :rule cong :premises (t117.t12)) (step t117.t14 (cl (! (not @p_409) :named @p_412) (! (not @p_398) :named @p_411) @p_410) :rule equiv_pos2) (step t117.t15 (cl (not @p_411) @p_395) :rule not_not) (step t117.t16 (cl @p_412 @p_395 @p_410) :rule th_resolution :premises (t117.t15 t117.t14)) (step t117.t17 (cl @p_410) :rule th_resolution :premises (t117.t9 t117.t13 t117.t16)) (anchor :step t117.t18 :args ((:= (veriT_vr65 V$) veriT_sk0))) (step t117.t18.t1 (cl (= veriT_vr65 veriT_sk0)) :rule refl) (step t117.t18.t2 (cl (= @p_401 (! (rraise$ veriT_sk0) :named @p_415))) :rule cong :premises (t117.t18.t1)) (step t117.t18.t3 (cl (= @p_402 (! (= x2$ @p_415) :named @p_416))) :rule cong :premises (t117.t18.t2)) (step t117.t18.t4 (cl (= @p_400 (! (not @p_416) :named @p_413))) :rule cong :premises (t117.t18.t3)) (step t117.t18 (cl (= @p_406 @p_413)) :rule sko_forall) (anchor :step t117.t19 :args ((:= (veriT_vr66 Abort$) veriT_sk1))) (step t117.t19.t1 (cl (= veriT_vr66 veriT_sk1)) :rule refl) (step t117.t19.t2 (cl (= @p_404 (! (rabort$ veriT_sk1) :named @p_419))) :rule cong :premises (t117.t19.t1)) (step t117.t19.t3 (cl (= @p_405 (! (= x2$ @p_419) :named @p_420))) :rule cong :premises (t117.t19.t2)) (step t117.t19.t4 (cl (= @p_403 (! (not @p_420) :named @p_417))) :rule cong :premises (t117.t19.t3)) (step t117.t19 (cl (= @p_407 @p_417)) :rule sko_forall) (step t117.t20 (cl (= @p_408 (! (and @p_413 @p_417) :named @p_421))) :rule cong :premises (t117.t18 t117.t19)) (step t117.t21 (cl (! (= @p_410 (! (not @p_421) :named @p_422)) :named @p_423)) :rule cong :premises (t117.t20)) (step t117.t22 (cl (! (not @p_423) :named @p_425) (! (not @p_410) :named @p_424) @p_422) :rule equiv_pos2) (step t117.t23 (cl (not @p_424) @p_408) :rule not_not) (step t117.t24 (cl @p_425 @p_408 @p_422) :rule th_resolution :premises (t117.t23 t117.t22)) (step t117.t25 (cl @p_422) :rule th_resolution :premises (t117.t17 t117.t21 t117.t24)) (step t117 (cl @p_397 @p_422) :rule subproof :discharge (h1)) (step t118 (cl @p_426 @p_392) :rule not_not) (step t119 (cl @p_392 @p_422) :rule th_resolution :premises (t118 t117)) (step t120 (cl @p_427 @p_379) :rule or :premises (t116)) (step t121 (cl (! (or @p_427 @p_422) :named @p_429) (! (not @p_427) :named @p_428)) :rule or_neg) (step t122 (cl (not @p_428) @p_215) :rule not_not) (step t123 (cl @p_429 @p_215) :rule th_resolution :premises (t122 t121)) (step t124 (cl @p_429 (! (not @p_422) :named @p_430)) :rule or_neg) (step t125 (cl (not @p_430) @p_421) :rule not_not) (step t126 (cl @p_429 @p_421) :rule th_resolution :premises (t125 t124)) (step t127 (cl @p_429) :rule th_resolution :premises (t120 t119 t123 t126)) (step t128 (cl (not (! (not @p_431) :named @p_468)) @p_170) :rule not_not) (step t129 (cl (or @p_431 (! (and (! (=> (! (= @p_378 @p_378) :named @p_432) (! (exists ((veriT_vr25 V_list_v_result$)) (! (= @p_3 (! (pair$ @p_378 veriT_vr25) :named @p_435)) :named @p_437)) :named @p_434)) :named @p_439) (! (=> (! (not (! (forall ((veriT_vr26 V_list_v_result$)) (! (not (! (= @p_3 (! (pair$ @p_378 veriT_vr26) :named @p_442)) :named @p_443)) :named @p_444)) :named @p_441)) :named @p_446) @p_432) :named @p_448)) :named @p_433))) :rule forall_inst :args ((:= veriT_vr23 @p_378) (:= veriT_vr24 @p_3))) (anchor :step t130) (assume t130.h1 @p_433) (anchor :step t130.t2 :args ((:= (veriT_vr25 V_list_v_result$) veriT_vr72))) (step t130.t2.t1 (cl (= veriT_vr25 veriT_vr72)) :rule refl) (step t130.t2.t2 (cl (= @p_435 (! (pair$ @p_378 veriT_vr72) :named @p_436))) :rule cong :premises (t130.t2.t1)) (step t130.t2.t3 (cl (= @p_437 (! (= @p_3 @p_436) :named @p_438))) :rule cong :premises (t130.t2.t2)) (step t130.t2 (cl (= @p_434 (! (exists ((veriT_vr72 V_list_v_result$)) @p_438) :named @p_440))) :rule bind) (step t130.t3 (cl (= @p_439 (! (=> @p_432 @p_440) :named @p_450))) :rule cong :premises (t130.t2)) (anchor :step t130.t4 :args ((:= (veriT_vr26 V_list_v_result$) veriT_vr72))) (step t130.t4.t1 (cl (= veriT_vr26 veriT_vr72)) :rule refl) (step t130.t4.t2 (cl (= @p_442 @p_436)) :rule cong :premises (t130.t4.t1)) (step t130.t4.t3 (cl (= @p_443 @p_438)) :rule cong :premises (t130.t4.t2)) (step t130.t4.t4 (cl (= @p_444 (! (not @p_438) :named @p_445))) :rule cong :premises (t130.t4.t3)) (step t130.t4 (cl (= @p_441 (! (forall ((veriT_vr72 V_list_v_result$)) @p_445) :named @p_447))) :rule bind) (step t130.t5 (cl (= @p_446 (! (not @p_447) :named @p_449))) :rule cong :premises (t130.t4)) (step t130.t6 (cl (= @p_448 (! (=> @p_449 @p_432) :named @p_451))) :rule cong :premises (t130.t5)) (step t130.t7 (cl (! (= @p_433 (! (and @p_450 @p_451) :named @p_454)) :named @p_452)) :rule cong :premises (t130.t3 t130.t6)) (step t130.t8 (cl (not @p_452) (! (not @p_433) :named @p_453) @p_454) :rule equiv_pos2) (step t130.t9 (cl @p_454) :rule th_resolution :premises (t130.h1 t130.t7 t130.t8)) (step t130.t10 (cl (= @p_432 true)) :rule eq_simplify) (step t130.t11 (cl (= @p_450 (! (=> true @p_440) :named @p_455))) :rule cong :premises (t130.t10)) (step t130.t12 (cl (= @p_455 @p_440)) :rule implies_simplify) (step t130.t13 (cl (= @p_450 @p_440)) :rule trans :premises (t130.t11 t130.t12)) (step t130.t14 (cl (= @p_451 (! (=> @p_449 true) :named @p_456))) :rule cong :premises (t130.t10)) (step t130.t15 (cl (= @p_456 true)) :rule implies_simplify) (step t130.t16 (cl (= @p_451 true)) :rule trans :premises (t130.t14 t130.t15)) (step t130.t17 (cl (= @p_454 (! (and @p_440 true) :named @p_457))) :rule cong :premises (t130.t13 t130.t16)) (step t130.t18 (cl (= @p_457 (! (and @p_440) :named @p_458))) :rule and_simplify) (step t130.t19 (cl (= @p_458 @p_440)) :rule and_simplify) (step t130.t20 (cl (! (= @p_454 @p_440) :named @p_459)) :rule trans :premises (t130.t17 t130.t18 t130.t19)) (step t130.t21 (cl (not @p_459) (not @p_454) @p_440) :rule equiv_pos2) (step t130.t22 (cl @p_440) :rule th_resolution :premises (t130.t9 t130.t20 t130.t21)) (anchor :step t130.t23 :args ((:= (veriT_vr72 V_list_v_result$) veriT_vr73))) (step t130.t23.t1 (cl (= veriT_vr72 veriT_vr73)) :rule refl) (step t130.t23.t2 (cl (= @p_436 @p_461)) :rule cong :premises (t130.t23.t1)) (step t130.t23.t3 (cl (= @p_438 @p_460)) :rule cong :premises (t130.t23.t2)) (step t130.t23 (cl (! (= @p_440 (! (exists ((veriT_vr73 V_list_v_result$)) @p_460) :named @p_463)) :named @p_462)) :rule bind) (step t130.t24 (cl (not @p_462) (not @p_440) @p_463) :rule equiv_pos2) (step t130.t25 (cl @p_463) :rule th_resolution :premises (t130.t22 t130.t23 t130.t24)) (anchor :step t130.t26 :args ((:= (veriT_vr73 V_list_v_result$) veriT_sk3))) (step t130.t26.t1 (cl (= veriT_vr73 veriT_sk3)) :rule refl) (step t130.t26.t2 (cl (= @p_461 (! (pair$ @p_378 veriT_sk3) :named @p_466))) :rule cong :premises (t130.t26.t1)) (step t130.t26.t3 (cl (= @p_460 (! (= @p_3 @p_466) :named @p_464))) :rule cong :premises (t130.t26.t2)) (step t130.t26 (cl (! (= @p_463 @p_464) :named @p_467)) :rule sko_ex) (step t130.t27 (cl (not @p_467) (not @p_463) @p_464) :rule equiv_pos2) (step t130.t28 (cl @p_464) :rule th_resolution :premises (t130.t25 t130.t26 t130.t27)) (step t130 (cl @p_453 @p_464) :rule subproof :discharge (h1)) (step t131 (cl @p_431 @p_433) :rule or :premises (t129)) (step t132 (cl (! (or @p_431 @p_464) :named @p_469) @p_468) :rule or_neg) (step t133 (cl @p_469 @p_170) :rule th_resolution :premises (t128 t132)) (step t134 (cl @p_469 (! (not @p_464) :named @p_595)) :rule or_neg) (step t135 (cl @p_469) :rule th_resolution :premises (t131 t130 t133 t134)) (step t136 (cl (or @p_431 (! (and (! (=> (! (= st$ (! (fst$ @p_470) :named @p_650)) :named @p_471) (! (exists ((veriT_vr25 V_list_v_result$)) (! (= @p_470 (! (pair$ st$ veriT_vr25) :named @p_474)) :named @p_476)) :named @p_473)) :named @p_478) (! (=> (! (not (! (forall ((veriT_vr26 V_list_v_result$)) (! (not (! (= @p_470 (! (pair$ st$ veriT_vr26) :named @p_481)) :named @p_482)) :named @p_483)) :named @p_480)) :named @p_485) @p_471) :named @p_487)) :named @p_472))) :rule forall_inst :args ((:= veriT_vr23 st$) (:= veriT_vr24 @p_470))) (anchor :step t137) (assume t137.h1 @p_472) (anchor :step t137.t2 :args ((:= (veriT_vr25 V_list_v_result$) veriT_vr106))) (step t137.t2.t1 (cl (= veriT_vr25 veriT_vr106)) :rule refl) (step t137.t2.t2 (cl (= @p_474 (! (pair$ st$ veriT_vr106) :named @p_475))) :rule cong :premises (t137.t2.t1)) (step t137.t2.t3 (cl (= @p_476 (! (= @p_470 @p_475) :named @p_477))) :rule cong :premises (t137.t2.t2)) (step t137.t2 (cl (= @p_473 (! (exists ((veriT_vr106 V_list_v_result$)) @p_477) :named @p_479))) :rule bind) (step t137.t3 (cl (= @p_478 (! (=> @p_471 @p_479) :named @p_489))) :rule cong :premises (t137.t2)) (anchor :step t137.t4 :args ((:= (veriT_vr26 V_list_v_result$) veriT_vr106))) (step t137.t4.t1 (cl (= veriT_vr26 veriT_vr106)) :rule refl) (step t137.t4.t2 (cl (= @p_481 @p_475)) :rule cong :premises (t137.t4.t1)) (step t137.t4.t3 (cl (= @p_482 @p_477)) :rule cong :premises (t137.t4.t2)) (step t137.t4.t4 (cl (= @p_483 (! (not @p_477) :named @p_484))) :rule cong :premises (t137.t4.t3)) (step t137.t4 (cl (= @p_480 (! (forall ((veriT_vr106 V_list_v_result$)) @p_484) :named @p_486))) :rule bind) (step t137.t5 (cl (= @p_485 (! (not @p_486) :named @p_488))) :rule cong :premises (t137.t4)) (step t137.t6 (cl (= @p_487 (! (=> @p_488 @p_471) :named @p_490))) :rule cong :premises (t137.t5)) (step t137.t7 (cl (! (= @p_472 (! (and @p_489 @p_490) :named @p_493)) :named @p_491)) :rule cong :premises (t137.t3 t137.t6)) (step t137.t8 (cl (not @p_491) (! (not @p_472) :named @p_492) @p_493) :rule equiv_pos2) (step t137.t9 (cl @p_493) :rule th_resolution :premises (t137.h1 t137.t7 t137.t8)) (anchor :step t137.t10 :args ((:= (veriT_vr106 V_list_v_result$) veriT_vr107))) (step t137.t10.t1 (cl (= veriT_vr106 veriT_vr107)) :rule refl) (step t137.t10.t2 (cl (= @p_475 (! (pair$ st$ veriT_vr107) :named @p_494))) :rule cong :premises (t137.t10.t1)) (step t137.t10.t3 (cl (= @p_477 (! (= @p_470 @p_494) :named @p_495))) :rule cong :premises (t137.t10.t2)) (step t137.t10.t4 (cl (= @p_484 (! (not @p_495) :named @p_496))) :rule cong :premises (t137.t10.t3)) (step t137.t10 (cl (= @p_486 (! (forall ((veriT_vr107 V_list_v_result$)) @p_496) :named @p_497))) :rule bind) (step t137.t11 (cl (= @p_488 (! (not @p_497) :named @p_498))) :rule cong :premises (t137.t10)) (step t137.t12 (cl (= @p_490 (! (=> @p_498 @p_471) :named @p_499))) :rule cong :premises (t137.t11)) (step t137.t13 (cl (! (= @p_493 (! (and @p_489 @p_499) :named @p_501)) :named @p_500)) :rule cong :premises (t137.t12)) (step t137.t14 (cl (not @p_500) (not @p_493) @p_501) :rule equiv_pos2) (step t137.t15 (cl @p_501) :rule th_resolution :premises (t137.t9 t137.t13 t137.t14)) (anchor :step t137.t16 :args ((:= (veriT_vr106 V_list_v_result$) veriT_vr108))) (step t137.t16.t1 (cl (= veriT_vr106 veriT_vr108)) :rule refl) (step t137.t16.t2 (cl (= @p_475 @p_503)) :rule cong :premises (t137.t16.t1)) (step t137.t16.t3 (cl (= @p_477 @p_502)) :rule cong :premises (t137.t16.t2)) (step t137.t16 (cl (= @p_479 (! (exists ((veriT_vr108 V_list_v_result$)) @p_502) :named @p_504))) :rule bind) (step t137.t17 (cl (= @p_489 (! (=> @p_471 @p_504) :named @p_510))) :rule cong :premises (t137.t16)) (anchor :step t137.t18 :args ((:= (veriT_vr107 V_list_v_result$) veriT_vr109))) (step t137.t18.t1 (cl (= veriT_vr107 veriT_vr109)) :rule refl) (step t137.t18.t2 (cl (= @p_494 (! (pair$ st$ veriT_vr109) :named @p_505))) :rule cong :premises (t137.t18.t1)) (step t137.t18.t3 (cl (= @p_495 (! (= @p_470 @p_505) :named @p_506))) :rule cong :premises (t137.t18.t2)) (step t137.t18.t4 (cl (= @p_496 (! (not @p_506) :named @p_507))) :rule cong :premises (t137.t18.t3)) (step t137.t18 (cl (= @p_497 (! (forall ((veriT_vr109 V_list_v_result$)) @p_507) :named @p_508))) :rule bind) (step t137.t19 (cl (= @p_498 (! (not @p_508) :named @p_509))) :rule cong :premises (t137.t18)) (step t137.t20 (cl (= @p_499 (! (=> @p_509 @p_471) :named @p_511))) :rule cong :premises (t137.t19)) (step t137.t21 (cl (! (= @p_501 (! (and @p_510 @p_511) :named @p_513)) :named @p_512)) :rule cong :premises (t137.t17 t137.t20)) (step t137.t22 (cl (not @p_512) (not @p_501) @p_513) :rule equiv_pos2) (step t137.t23 (cl @p_513) :rule th_resolution :premises (t137.t15 t137.t21 t137.t22)) (anchor :step t137.t24 :args ((:= (veriT_vr108 V_list_v_result$) veriT_sk11))) (step t137.t24.t1 (cl (= veriT_vr108 veriT_sk11)) :rule refl) (step t137.t24.t2 (cl (= @p_503 (! (pair$ st$ veriT_sk11) :named @p_516))) :rule cong :premises (t137.t24.t1)) (step t137.t24.t3 (cl (= @p_502 (! (= @p_470 @p_516) :named @p_514))) :rule cong :premises (t137.t24.t2)) (step t137.t24 (cl (= @p_504 @p_514)) :rule sko_ex) (step t137.t25 (cl (= @p_510 (! (=> @p_471 @p_514) :named @p_517))) :rule cong :premises (t137.t24)) (step t137.t26 (cl (! (= @p_513 (! (and @p_517 @p_511) :named @p_519)) :named @p_518)) :rule cong :premises (t137.t25)) (step t137.t27 (cl (not @p_518) (not @p_513) @p_519) :rule equiv_pos2) (step t137.t28 (cl @p_519) :rule th_resolution :premises (t137.t23 t137.t26 t137.t27)) (anchor :step t137.t29 :args ((:= (veriT_vr109 V_list_v_result$) veriT_vr110))) (step t137.t29.t1 (cl (= veriT_vr109 veriT_vr110)) :rule refl) (step t137.t29.t2 (cl (= @p_505 (! (pair$ st$ veriT_vr110) :named @p_521))) :rule cong :premises (t137.t29.t1)) (step t137.t29.t3 (cl (= @p_506 (! (= @p_470 @p_521) :named @p_522))) :rule cong :premises (t137.t29.t2)) (step t137.t29.t4 (cl (= @p_507 (! (not @p_522) :named @p_523))) :rule cong :premises (t137.t29.t3)) (step t137.t29 (cl (= @p_508 (! (forall ((veriT_vr110 V_list_v_result$)) @p_523) :named @p_520))) :rule bind) (step t137.t30 (cl (= @p_509 (! (not @p_520) :named @p_524))) :rule cong :premises (t137.t29)) (step t137.t31 (cl (= @p_511 (! (=> @p_524 @p_471) :named @p_525))) :rule cong :premises (t137.t30)) (step t137.t32 (cl (! (= @p_519 (! (and @p_517 @p_525) :named @p_526)) :named @p_527)) :rule cong :premises (t137.t31)) (step t137.t33 (cl (not @p_527) (not @p_519) @p_526) :rule equiv_pos2) (step t137.t34 (cl @p_526) :rule th_resolution :premises (t137.t28 t137.t32 t137.t33)) (step t137 (cl @p_492 @p_526) :rule subproof :discharge (h1)) (step t138 (cl @p_431 @p_472) :rule or :premises (t136)) (step t139 (cl (! (or @p_431 @p_526) :named @p_528) @p_468) :rule or_neg) (step t140 (cl @p_528 @p_170) :rule th_resolution :premises (t128 t139)) (step t141 (cl @p_528 (! (not @p_526) :named @p_553)) :rule or_neg) (step t142 (cl @p_528) :rule th_resolution :premises (t138 t137 t140 t141)) (step t143 (cl (not (! (not (! (not @p_79) :named @p_529)) :named @p_537)) @p_79) :rule not_not) (step t144 (cl (or @p_529 (! (= (! (fun_app$b (! (uu$ @p_378 @p_530) :named @p_655) @p_371) :named @p_532) (! (ite @p_373 @p_370 @p_371) :named @p_533)) :named @p_531))) :rule forall_inst :args ((:= veriT_vr7 @p_378) (:= veriT_vr8 @p_530) (:= veriT_vr9 @p_371))) (anchor :step t145) (assume t145.h1 @p_531) (step t145.t2 (cl (! (= @p_531 (! (and (! (= @p_532 @p_533) :named @p_555) (! (ite @p_373 (= @p_370 @p_533) (! (= @p_371 @p_533) :named @p_557)) :named @p_556)) :named @p_534)) :named @p_535)) :rule ite_intro) (step t145.t3 (cl (not @p_535) (! (not @p_531) :named @p_536) @p_534) :rule equiv_pos2) (step t145.t4 (cl @p_534) :rule th_resolution :premises (t145.h1 t145.t2 t145.t3)) (step t145 (cl @p_536 @p_534) :rule subproof :discharge (h1)) (step t146 (cl @p_529 @p_531) :rule or :premises (t144)) (step t147 (cl (! (or @p_529 @p_534) :named @p_538) @p_537) :rule or_neg) (step t148 (cl @p_538 @p_79) :rule th_resolution :premises (t143 t147)) (step t149 (cl @p_538 (! (not @p_534) :named @p_554)) :rule or_neg) (step t150 (cl @p_538) :rule th_resolution :premises (t146 t145 t148 t149)) (step t151 (cl (or @p_529 (! (= (! (fun_app$b (! (uu$ @p_378 st$) :named @p_656) @p_371) :named @p_540) (! (ite (! (less_eq$ @p_318 @p_371) :named @p_542) @p_318 @p_371) :named @p_541)) :named @p_539))) :rule forall_inst :args ((:= veriT_vr7 @p_378) (:= veriT_vr8 st$) (:= veriT_vr9 @p_371))) (anchor :step t152) (assume t152.h1 @p_539) (step t152.t2 (cl (! (= @p_539 (! (and (! (= @p_540 @p_541) :named @p_560) (ite @p_542 (! (= @p_318 @p_541) :named @p_662) (= @p_371 @p_541))) :named @p_543)) :named @p_544)) :rule ite_intro) (step t152.t3 (cl (not @p_544) (! (not @p_539) :named @p_545) @p_543) :rule equiv_pos2) (step t152.t4 (cl @p_543) :rule th_resolution :premises (t152.h1 t152.t2 t152.t3)) (step t152 (cl @p_545 @p_543) :rule subproof :discharge (h1)) (step t153 (cl @p_529 @p_539) :rule or :premises (t151)) (step t154 (cl (! (or @p_529 @p_543) :named @p_546) @p_537) :rule or_neg) (step t155 (cl @p_546 @p_79) :rule th_resolution :premises (t143 t154)) (step t156 (cl @p_546 (! (not @p_543) :named @p_559)) :rule or_neg) (step t157 (cl @p_546) :rule th_resolution :premises (t153 t152 t155 t156)) (step t158 (cl @p_547 @p_548) :rule or :premises (t115)) (step t159 (cl @p_548) :rule resolution :premises (t158 t89)) (step t160 (cl @p_421 (! (not @p_413) :named @p_549) (! (not @p_417) :named @p_550)) :rule and_neg) (step t161 (cl (not @p_549) @p_416) :rule not_not) (step t162 (cl @p_421 @p_416 @p_550) :rule th_resolution :premises (t161 t160)) (step t163 (cl (not @p_550) @p_420) :rule not_not) (step t164 (cl @p_421 @p_416 @p_420) :rule th_resolution :premises (t163 t162)) (step t165 (cl @p_427 @p_422) :rule or :premises (t127)) (step t166 (cl @p_422) :rule resolution :premises (t165 t65)) (step t167 (cl @p_431 @p_464) :rule or :premises (t135)) (step t168 (cl @p_464) :rule resolution :premises (t167 t56)) (step t169 (cl (! (not @p_525) :named @p_552) (! (not @p_524) :named @p_551) @p_471) :rule implies_pos) (step t170 (cl (not @p_551) @p_520) :rule not_not) (step t171 (cl @p_552 @p_520 @p_471) :rule th_resolution :premises (t170 t169)) (step t172 (cl @p_553 @p_525) :rule and_pos) (step t173 (cl @p_431 @p_526) :rule or :premises (t142)) (step t174 (cl @p_526) :rule resolution :premises (t173 t56)) (step t175 (cl @p_525) :rule resolution :premises (t172 t174)) (step t176 (cl @p_554 @p_555) :rule and_pos) (step t177 (cl (! (not @p_556) :named @p_558) @p_373 @p_557) :rule ite_pos1) (step t178 (cl @p_554 @p_556) :rule and_pos) (step t179 (cl @p_529 @p_534) :rule or :premises (t150)) (step t180 (cl @p_558 @p_557) :rule resolution :premises (t177 t113)) (step t181 (cl @p_534) :rule resolution :premises (t179 t32)) (step t182 (cl @p_555) :rule resolution :premises (t176 t181)) (step t183 (cl @p_556) :rule resolution :premises (t178 t181)) (step t184 (cl @p_557) :rule resolution :premises (t180 t183)) (step t185 (cl @p_559 @p_560) :rule and_pos) (step t186 (cl @p_529 @p_543) :rule or :premises (t157)) (step t187 (cl @p_543) :rule resolution :premises (t186 t32)) (step t188 (cl @p_560) :rule resolution :premises (t185 t187)) (step t189 (cl (! (not (! (= st$ @p_530) :named @p_651)) :named @p_654) (! (= @p_318 @p_370) :named @p_663)) :rule eq_congruent) (step t190 (cl (or (! (not @p_357) :named @p_565) (! (=> (! (and @p_359 @p_416) :named @p_562) (! (less_eq$ (! (clock$ (! (fst$ (! (fun_app$ (fun_evaluate_match$ st$ env$ veriT_sk0 pes$) veriT_sk0) :named @p_583)) :named @p_618)) :named @p_619) @p_318) :named @p_564)) :named @p_563))) :rule forall_inst :args ((:= veriT_vr61 x2$) (:= veriT_vr62 veriT_sk0))) (step t191 (cl (or @p_547 (! (= (! (fix_clock$ st$a @p_466) :named @p_596) (! (pair$ @p_561 veriT_sk3) :named @p_676)) :named @p_566))) :rule forall_inst :args ((:= veriT_vr56 st$a) (:= veriT_vr57 @p_378) (:= veriT_vr58 veriT_sk3))) (step t192 (cl (or (! (not @p_290) :named @p_569) (! (=> @p_548 (! (less_eq$ (! (clock$ @p_561) :named @p_575) @p_371) :named @p_568)) :named @p_567))) :rule forall_inst :args ((:= veriT_vr49 st$a) (:= veriT_vr50 @p_378) (:= veriT_vr51 r$) (:= veriT_vr52 @p_561))) (step t193 (cl (or (! (not @p_236) :named @p_571) (! (= (! (case_error_result$ uua$ uub$ @p_415) :named @p_621) (! (fun_app$ uua$ veriT_sk0) :named @p_584)) :named @p_572))) :rule forall_inst :args ((:= veriT_vr36 uua$) (:= veriT_vr37 uub$) (:= veriT_vr38 veriT_sk0))) (step t194 (cl @p_562 (! (not @p_359) :named @p_614) @p_413) :rule and_neg) (step t195 (cl (not @p_563) (not @p_562) @p_564) :rule implies_pos) (step t196 (cl @p_565 @p_563) :rule or :premises (t190)) (step t197 (cl @p_562 @p_413) :rule resolution :premises (t194 t101)) (step t198 (cl @p_563) :rule resolution :premises (t196 t95)) (step t199 (cl @p_547 @p_566) :rule or :premises (t191)) (step t200 (cl @p_566) :rule resolution :premises (t199 t89)) (step t201 (cl (! (not @p_567) :named @p_570) (not @p_548) @p_568) :rule implies_pos) (step t202 (cl @p_569 @p_567) :rule or :premises (t192)) (step t203 (cl @p_570 @p_568) :rule resolution :premises (t201 t159)) (step t204 (cl @p_567) :rule resolution :premises (t202 t83)) (step t205 (cl @p_568) :rule resolution :premises (t203 t204)) (step t206 (cl @p_571 @p_572) :rule or :premises (t193)) (step t207 (cl @p_572) :rule resolution :premises (t206 t71)) (step t208 (cl @p_368 @p_573) :rule or :premises (t103)) (step t209 (cl (or (! (not @p_573) :named @p_576) (! (or @p_574 (! (not @p_568) :named @p_581) (! (less_eq$ @p_575 @p_7) :named @p_582)) :named @p_577))) :rule forall_inst :args ((:= veriT_vr13 @p_371) (:= veriT_vr14 @p_7) (:= veriT_vr15 @p_575))) (step t210 (cl @p_576 @p_577) :rule or :premises (t209)) (step t211 (cl (! (or @p_368 @p_577) :named @p_579) @p_578) :rule or_neg) (step t212 (cl @p_579 @p_105) :rule th_resolution :premises (t114 t211)) (step t213 (cl @p_579 (! (not @p_577) :named @p_580)) :rule or_neg) (step t214 (cl @p_579) :rule th_resolution :premises (t208 t210 t212 t213)) (step t215 (cl @p_580 @p_574 @p_581 @p_582) :rule or_pos) (step t216 (cl @p_368 @p_577) :rule or :premises (t214)) (step t217 (cl @p_580 @p_582) :rule resolution :premises (t215 axiom4 t205)) (step t218 (cl @p_577) :rule resolution :premises (t216 t38)) (step t219 (cl @p_582) :rule resolution :premises (t217 t218)) (step t220 (cl (or (! (not @p_257) :named @p_585) (! (= (! (case_error_result$ uua$ uub$ @p_419) :named @p_603) (! (fun_app$a uub$ veriT_sk1) :named @p_599)) :named @p_586))) :rule forall_inst :args ((:= veriT_vr42 uua$) (:= veriT_vr43 uub$) (:= veriT_vr44 veriT_sk1))) (step t221 (cl (or @p_368 (! (=> (! (and @p_582 (! (less_eq$ @p_370 @p_575) :named @p_588)) :named @p_587) @p_360) :named @p_589))) :rule forall_inst :args ((:= veriT_vr13 @p_575) (:= veriT_vr14 @p_7) (:= veriT_vr15 @p_370))) (step t222 (cl (or (! (not @p_26) :named @p_593) (! (= @p_583 @p_584) :named @p_594))) :rule forall_inst :args ((:= veriT_vr1 veriT_sk0))) (step t223 (cl @p_585 @p_586) :rule or :premises (t220)) (step t224 (cl @p_586) :rule resolution :premises (t223 t77)) (step t225 (cl @p_587 (not @p_582) (! (not @p_588) :named @p_590)) :rule and_neg) (step t226 (cl (! (not @p_589) :named @p_591) (! (not @p_587) :named @p_592) @p_360) :rule implies_pos) (step t227 (cl @p_368 @p_589) :rule or :premises (t221)) (step t228 (cl @p_587 @p_590) :rule resolution :premises (t225 t219)) (step t229 (cl @p_591 @p_592) :rule resolution :premises (t226 t102)) (step t230 (cl @p_589) :rule resolution :premises (t227 t38)) (step t231 (cl @p_592) :rule resolution :premises (t229 t230)) (step t232 (cl @p_590) :rule resolution :premises (t228 t231)) (step t233 (cl @p_593 @p_594) :rule or :premises (t222)) (step t234 (cl @p_594) :rule resolution :premises (t233 t20)) (step t235 (cl (not (! (= st$a st$a) :named @p_597)) @p_595 (! (= @p_470 @p_596) :named @p_598)) :rule eq_congruent) (step t236 (cl @p_597) :rule eq_reflexive) (step t237 (cl @p_595 @p_598) :rule th_resolution :premises (t235 t236)) (step t238 (cl (or (! (not @p_48) :named @p_600) (! (= @p_599 (! (pair$ st$ (! (rerr$ @p_419) :named @p_608)) :named @p_610)) :named @p_601))) :rule forall_inst :args ((:= veriT_vr3 veriT_sk1))) (step t239 (cl @p_600 @p_601) :rule or :premises (t238)) (step t240 (cl @p_601) :rule resolution :premises (t239 t26)) (step t241 (cl (! (not (! (= uua$ uua$) :named @p_604)) :named @p_623) (! (not (! (= uub$ uub$) :named @p_607)) :named @p_605) @p_417 (! (= @p_602 @p_603) :named @p_606)) :rule eq_congruent) (step t242 (cl @p_604) :rule eq_reflexive) (step t243 (cl @p_605 @p_417 @p_606) :rule th_resolution :premises (t241 t242)) (step t244 (cl @p_607) :rule eq_reflexive) (step t245 (cl @p_417 @p_606) :rule th_resolution :premises (t243 t244)) (step t246 (cl (not (! (= st$ st$) :named @p_611)) (! (not (! (= r$ @p_608) :named @p_616)) :named @p_612) (! (= @p_609 @p_610) :named @p_613)) :rule eq_congruent) (step t247 (cl @p_611) :rule eq_reflexive) (step t248 (cl @p_612 @p_613) :rule th_resolution :premises (t246 t247)) (step t249 (cl @p_614 (not (! (= @p_615 @p_608) :named @p_617)) @p_616) :rule eq_transitive) (step t250 (cl @p_417 @p_617) :rule eq_congruent) (step t251 (cl @p_614 @p_616 @p_417) :rule th_resolution :premises (t249 t250)) (step t252 (cl @p_613 @p_614 @p_417) :rule th_resolution :premises (t248 t251)) (step t253 (cl (not (! (= @p_530 @p_618) :named @p_620)) (! (= @p_370 @p_619) :named @p_627)) :rule eq_congruent) (step t254 (cl (not (! (= @p_602 @p_583) :named @p_622)) @p_620) :rule eq_congruent) (step t255 (cl (not (! (= @p_602 @p_621) :named @p_624)) (! (not @p_572) :named @p_625) (! (not @p_594) :named @p_626) @p_622) :rule eq_transitive) (step t256 (cl @p_623 @p_605 @p_413 @p_624) :rule eq_congruent) (step t257 (cl @p_605 @p_413 @p_624) :rule th_resolution :premises (t256 t242)) (step t258 (cl @p_413 @p_624) :rule th_resolution :premises (t257 t244)) (step t259 (cl @p_625 @p_626 @p_622 @p_413) :rule th_resolution :premises (t255 t258)) (step t260 (cl @p_620 @p_625 @p_626 @p_413) :rule th_resolution :premises (t254 t259)) (step t261 (cl @p_627 @p_625 @p_626 @p_413) :rule th_resolution :premises (t253 t260)) (step t262 (cl (or @p_524 (! (not @p_628) :named @p_630))) :rule forall_inst :args ((:= veriT_vr110 r$))) (step t263 (cl @p_431 @p_629) :rule or :premises (t104)) (step t264 (cl @p_524 @p_630) :rule or :premises (t262)) (step t265 (cl @p_524) :rule resolution :premises (t264 axiom3)) (step t266 (cl @p_471) :rule resolution :premises (t171 t265 t175)) (step t267 (cl (or @p_529 (! (= (! (fun_app$b (! (uu$ @p_378 @p_561) :named @p_631) (! (clock$ (update_clock$ @p_631 @p_561)) :named @p_632)) :named @p_634) (! (ite @p_568 @p_575 @p_371) :named @p_635)) :named @p_633))) :rule forall_inst :args ((:= veriT_vr7 @p_378) (:= veriT_vr8 @p_561) (:= veriT_vr9 @p_632))) (anchor :step t268) (assume t268.h1 @p_633) (step t268.t2 (cl (! (= @p_633 (! (and (= @p_634 @p_635) (! (ite @p_568 (! (= @p_575 @p_635) :named @p_647) (= @p_371 @p_635)) :named @p_646)) :named @p_636)) :named @p_637)) :rule ite_intro) (step t268.t3 (cl (not @p_637) (! (not @p_633) :named @p_638) @p_636) :rule equiv_pos2) (step t268.t4 (cl @p_636) :rule th_resolution :premises (t268.h1 t268.t2 t268.t3)) (step t268 (cl @p_638 @p_636) :rule subproof :discharge (h1)) (step t269 (cl @p_529 @p_633) :rule or :premises (t267)) (step t270 (cl (! (or @p_529 @p_636) :named @p_639) @p_537) :rule or_neg) (step t271 (cl @p_639 @p_79) :rule th_resolution :premises (t143 t270)) (step t272 (cl @p_639 (! (not @p_636) :named @p_648)) :rule or_neg) (step t273 (cl @p_639) :rule th_resolution :premises (t269 t268 t271 t272)) (step t274 (cl (or @p_529 (! (= @p_635 (! (fun_app$b @p_631 @p_371) :named @p_641)) :named @p_640))) :rule forall_inst :args ((:= veriT_vr7 @p_378) (:= veriT_vr8 @p_561) (:= veriT_vr9 @p_371))) (anchor :step t275) (assume t275.h1 @p_640) (step t275.t2 (cl (! (= @p_640 (! (= @p_635 @p_641) :named @p_642)) :named @p_643)) :rule ite_intro) (step t275.t3 (cl (not @p_643) (! (not @p_640) :named @p_644) @p_642) :rule equiv_pos2) (step t275.t4 (cl @p_642) :rule th_resolution :premises (t275.h1 t275.t2 t275.t3)) (step t275 (cl @p_644 @p_642) :rule subproof :discharge (h1)) (step t276 (cl @p_529 @p_640) :rule or :premises (t274)) (step t277 (cl (! (or @p_529 @p_642) :named @p_645) @p_537) :rule or_neg) (step t278 (cl @p_645 @p_79) :rule th_resolution :premises (t143 t277)) (step t279 (cl @p_645 (! (not @p_642) :named @p_696)) :rule or_neg) (step t280 (cl @p_645) :rule th_resolution :premises (t276 t275 t278 t279)) (step t281 (cl (! (not @p_646) :named @p_649) @p_581 @p_647) :rule ite_pos2) (step t282 (cl @p_648 @p_646) :rule and_pos) (step t283 (cl @p_529 @p_636) :rule or :premises (t273)) (step t284 (cl @p_649 @p_647) :rule resolution :premises (t281 t205)) (step t285 (cl @p_636) :rule resolution :premises (t283 t32)) (step t286 (cl @p_646) :rule resolution :premises (t282 t285)) (step t287 (cl @p_647) :rule resolution :premises (t284 t286)) (step t288 (cl @p_529 @p_642) :rule or :premises (t280)) (step t289 (cl @p_642) :rule resolution :premises (t288 t32)) (step t290 (cl (! (= @p_371 @p_371) :named @p_671)) :rule eq_reflexive) (step t291 (cl (! (not @p_471) :named @p_661) (not (! (= @p_530 @p_650) :named @p_652)) @p_651) :rule eq_transitive) (step t292 (cl (not (! (= @p_470 @p_602) :named @p_653)) @p_652) :rule eq_congruent) (step t293 (cl @p_630 (not @p_613) (! (not @p_601) :named @p_657) (! (not @p_586) :named @p_658) (! (not @p_606) :named @p_659) @p_653) :rule eq_transitive) (step t294 (cl (! (not @p_432) :named @p_674) @p_654 (! (= @p_655 @p_656) :named @p_670)) :rule eq_congruent) (step t295 (cl @p_432) :rule eq_reflexive) (step t296 (cl @p_630 @p_657 @p_658 @p_659 @p_653 @p_614 @p_417) :rule th_resolution :premises (t293 t252)) (step t297 (cl @p_630 @p_657 @p_658 @p_653 @p_614 @p_417 @p_417) :rule th_resolution :premises (t296 t245)) (step t298 (cl @p_630 @p_657 @p_658 @p_653 @p_614 @p_417) :rule contraction :premises (t297)) (step t299 (cl @p_652 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t292 t298)) (step t300 (cl (not (! (= @p_371 @p_370) :named @p_664)) (! (not (! (= @p_7 @p_7) :named @p_660)) :named @p_675) @p_574 @p_360) :rule eq_congruent_pred) (step t301 (cl @p_660) :rule eq_reflexive) (step t302 (cl @p_661 @p_651 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t291 t299)) (step t303 (cl (! (not @p_557) :named @p_665) (! (not @p_555) :named @p_666) (! (not (! (= @p_532 @p_540) :named @p_672)) :named @p_667) (! (not @p_560) :named @p_668) (! (not @p_662) :named @p_669) (not @p_663) @p_664) :rule eq_transitive) (step t304 (cl @p_663 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t189 t302)) (step t305 (cl @p_665 @p_666 @p_667 @p_668 @p_669 @p_664 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t303 t304)) (step t306 (cl (! (not @p_670) :named @p_673) (! (not @p_671) :named @p_699) @p_672) :rule eq_congruent) (step t307 (cl @p_673 @p_672) :rule th_resolution :premises (t306 t290)) (step t308 (cl @p_674 @p_670 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t294 t302)) (step t309 (cl @p_670 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t308 t295)) (step t310 (cl @p_672 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t307 t309)) (step t311 (cl @p_665 @p_666 @p_668 @p_669 @p_664 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t305 t310)) (step t312 (cl @p_665 @p_666 @p_668 @p_669 @p_664 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule contraction :premises (t311)) (step t313 (cl @p_675 @p_574 @p_360 @p_665 @p_666 @p_668 @p_669 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t300 t312)) (step t314 (cl @p_574 @p_360 @p_665 @p_666 @p_668 @p_669 @p_661 @p_630 @p_657 @p_658 @p_614 @p_417) :rule th_resolution :premises (t313 t301)) (step t315 (cl @p_669 @p_417) :rule resolution :premises (t314 axiom3 axiom4 t101 t102 t266 t182 t184 t188 t224 t240)) (step t316 (cl (or (! (not @p_629) :named @p_677) (! (or (! (not (! (= @p_676 @p_676) :named @p_682)) :named @p_683) (! (= @p_561 (! (fst$ @p_676) :named @p_691)) :named @p_681)) :named @p_678))) :rule forall_inst :args ((:= veriT_vr23 @p_561) (:= veriT_vr24 @p_676) (:= veriT_vr26 veriT_sk3))) (step t317 (cl @p_677 @p_678) :rule or :premises (t316)) (step t318 (cl (! (or @p_431 @p_678) :named @p_679) @p_468) :rule or_neg) (step t319 (cl @p_679 @p_170) :rule th_resolution :premises (t128 t318)) (step t320 (cl @p_679 (! (not @p_678) :named @p_680)) :rule or_neg) (step t321 (cl @p_679) :rule th_resolution :premises (t263 t317 t319 t320)) (anchor :step t322) (assume t322.h1 @p_678) (step t322.t2 (cl (= @p_682 true)) :rule eq_simplify) (step t322.t3 (cl (= @p_683 (! (not true) :named @p_684))) :rule cong :premises (t322.t2)) (step t322.t4 (cl (= @p_684 false)) :rule not_simplify) (step t322.t5 (cl (= @p_683 false)) :rule trans :premises (t322.t3 t322.t4)) (step t322.t6 (cl (= @p_678 (! (or false @p_681) :named @p_685))) :rule cong :premises (t322.t5)) (step t322.t7 (cl (= @p_685 (! (or @p_681) :named @p_686))) :rule or_simplify) (step t322.t8 (cl (= @p_686 @p_681)) :rule or_simplify) (step t322.t9 (cl (! (= @p_678 @p_681) :named @p_687)) :rule trans :premises (t322.t6 t322.t7 t322.t8)) (step t322.t10 (cl (not @p_687) @p_680 @p_681) :rule equiv_pos2) (step t322.t11 (cl @p_681) :rule th_resolution :premises (t322.h1 t322.t9 t322.t10)) (step t322 (cl @p_680 @p_681) :rule subproof :discharge (h1)) (step t323 (cl @p_431 @p_678) :rule or :premises (t321)) (step t324 (cl (! (or @p_431 @p_681) :named @p_688) @p_468) :rule or_neg) (step t325 (cl @p_688 @p_170) :rule th_resolution :premises (t128 t324)) (step t326 (cl @p_688 (! (not @p_681) :named @p_693)) :rule or_neg) (step t327 (cl @p_688) :rule th_resolution :premises (t323 t322 t325 t326)) (step t328 (cl @p_431 @p_681) :rule or :premises (t327)) (step t329 (cl @p_681) :rule resolution :premises (t328 t56)) (step t330 (cl (not @p_598) (! (not @p_566) :named @p_689) (! (= @p_470 @p_676) :named @p_690)) :rule eq_transitive) (step t331 (cl @p_689 @p_690 @p_595) :rule th_resolution :premises (t330 t237)) (step t332 (cl (not @p_690) (! (= @p_650 @p_691) :named @p_692)) :rule eq_congruent) (step t333 (cl @p_692 @p_689 @p_595) :rule th_resolution :premises (t332 t331)) (step t334 (cl @p_661 (not @p_692) @p_693 (! (= st$ @p_561) :named @p_694)) :rule eq_transitive) (step t335 (cl @p_661 @p_693 @p_694 @p_689 @p_595) :rule th_resolution :premises (t334 t333)) (step t336 (cl (! (not @p_694) :named @p_702) (! (= @p_318 @p_575) :named @p_695)) :rule eq_congruent) (step t337 (cl @p_695 @p_661 @p_693 @p_689 @p_595) :rule th_resolution :premises (t336 t335)) (step t338 (cl (! (not @p_695) :named @p_704) (! (not @p_647) :named @p_697) @p_696 (! (not (! (= @p_540 @p_641) :named @p_700)) :named @p_698) @p_668 @p_662) :rule eq_transitive) (step t339 (cl @p_697 @p_696 @p_698 @p_668 @p_662 @p_661 @p_693 @p_689 @p_595) :rule th_resolution :premises (t338 t337)) (step t340 (cl (! (not (! (= @p_656 @p_631) :named @p_703)) :named @p_701) @p_699 @p_700) :rule eq_congruent) (step t341 (cl @p_701 @p_700) :rule th_resolution :premises (t340 t290)) (step t342 (cl @p_674 @p_702 @p_703) :rule eq_congruent) (step t343 (cl @p_702 @p_703) :rule th_resolution :premises (t342 t295)) (step t344 (cl @p_703 @p_661 @p_693 @p_689 @p_595) :rule th_resolution :premises (t343 t335)) (step t345 (cl @p_700 @p_661 @p_693 @p_689 @p_595) :rule th_resolution :premises (t341 t344)) (step t346 (cl @p_697 @p_696 @p_668 @p_662 @p_661 @p_693 @p_689 @p_595 @p_661 @p_693 @p_689 @p_595) :rule th_resolution :premises (t339 t345)) (step t347 (cl @p_697 @p_696 @p_668 @p_662 @p_661 @p_693 @p_689 @p_595) :rule contraction :premises (t346)) (step t348 (cl @p_662) :rule resolution :premises (t347 t266 t188 t200 t287 t289 t168 t329)) (step t349 (cl @p_417) :rule resolution :premises (t315 t348)) (step t350 (cl @p_416) :rule resolution :premises (t164 t349 t166)) (step t351 (cl @p_562) :rule resolution :premises (t197 t350)) (step t352 (cl @p_564) :rule resolution :premises (t195 t351 t198)) (step t353 (cl (not @p_627) @p_704 (! (not @p_564) :named @p_705) @p_588) :rule eq_congruent_pred) (step t354 (cl @p_704 @p_705 @p_588 @p_625 @p_626 @p_413) :rule th_resolution :premises (t353 t261)) (step t355 (cl @p_705 @p_588 @p_625 @p_626 @p_413 @p_661 @p_693 @p_689 @p_595) :rule th_resolution :premises (t354 t337)) (step t356 (cl) :rule resolution :premises (t355 t350 t168 t266 t352 t200 t207 t232 t234 t329)) +eae55ce4deb2476399eb5222073e987ca2cc4536 3015 0 +unsat +(define-fun veriT_sk0 () A_b_c_M_state_fun$ (! (choice ((veriT_vr57 A_b_c_M_state_fun$)) (not (forall ((veriT_vr58 A_b_c_M_state_fun$)) (! (=> (! (forall ((veriT_vr59 A$) (veriT_vr60 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_vr57 veriT_vr59) :named @p_552) veriT_vr60) :named @p_544)) :named @p_542) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ veriT_vr58 veriT_vr59) :named @p_554) veriT_vr60) :named @p_543)) :named @p_556) @p_542) :named @p_561) (! (forall ((veriT_vr61 B$) (veriT_vr62 C$)) (! (= (! (is_res$ @p_543 (! (pair$ veriT_vr61 veriT_vr62) :named @p_545)) :named @p_566) (! (is_res$ @p_544 @p_545) :named @p_570)) :named @p_571)) :named @p_562)) :named @p_572)) :named @p_573)) :named @p_551) (! (forall ((veriT_vr63 D$)) (! (or (! (is_fail$a (! (run$a (! (b$ veriT_vr57) :named @p_575) veriT_vr63) :named @p_546)) :named @p_548) (! (exists ((veriT_vr64 E$) (veriT_vr65 D$)) (! (and (! (is_res$a @p_546 (! (pair$a veriT_vr64 veriT_vr65) :named @p_580)) :named @p_581) (! (is_fail$b (! (run$b (! (c$ veriT_vr64 veriT_vr57) :named @p_583) veriT_vr65) :named @p_585)) :named @p_586)) :named @p_587)) :named @p_576) (! (and (! (and (! (=> (! (or (! (is_fail$a (! (run$a (! (b$ veriT_vr58) :named @p_588) veriT_vr63) :named @p_547)) :named @p_549) (! (exists ((veriT_vr66 E$) (veriT_vr67 D$)) (! (and (! (is_res$a @p_547 (! (pair$a veriT_vr66 veriT_vr67) :named @p_592)) :named @p_593) (! (is_fail$b (! (run$b (! (c$ veriT_vr66 veriT_vr58) :named @p_595) veriT_vr67) :named @p_597)) :named @p_598)) :named @p_599)) :named @p_589)) :named @p_600) (! (or @p_548 (! (exists ((veriT_vr68 E$) (veriT_vr69 D$)) (! (and (! (is_res$a @p_546 (! (pair$a veriT_vr68 veriT_vr69) :named @p_603)) :named @p_604) (! (is_fail$b (! (run$b (! (c$ veriT_vr68 veriT_vr57) :named @p_606) veriT_vr69) :named @p_608)) :named @p_609)) :named @p_610)) :named @p_602)) :named @p_611)) :named @p_613) (! (=> (! (or @p_548 (! (exists ((veriT_vr70 E$) (veriT_vr71 D$)) (! (and (! (is_res$a @p_546 (! (pair$a veriT_vr70 veriT_vr71) :named @p_615)) :named @p_616) (! (is_fail$b (! (run$b (! (c$ veriT_vr70 veriT_vr57) :named @p_618) veriT_vr71) :named @p_620)) :named @p_621)) :named @p_622)) :named @p_614)) :named @p_623) (! (or @p_549 (! (exists ((veriT_vr72 E$) (veriT_vr73 D$)) (! (and (! (is_res$a @p_547 (! (pair$a veriT_vr72 veriT_vr73) :named @p_626)) :named @p_627) (! (is_fail$b (! (run$b (! (c$ veriT_vr72 veriT_vr58) :named @p_629) veriT_vr73) :named @p_631)) :named @p_632)) :named @p_633)) :named @p_625)) :named @p_634)) :named @p_636)) :named @p_637) (! (forall ((veriT_vr74 F$) (veriT_vr75 D$)) (! (and (! (=> (! (or @p_549 (! (exists ((veriT_vr76 E$) (veriT_vr77 D$)) (! (and (! (is_res$a @p_547 (! (pair$a veriT_vr76 veriT_vr77) :named @p_640)) :named @p_641) (! (is_res$b (! (run$b (! (c$ veriT_vr76 veriT_vr58) :named @p_643) veriT_vr77) :named @p_645) (! (pair$b veriT_vr74 veriT_vr75) :named @p_550)) :named @p_646)) :named @p_647)) :named @p_639)) :named @p_648) (! (or @p_548 (! (exists ((veriT_vr78 E$) (veriT_vr79 D$)) (! (and (! (is_res$a @p_546 (! (pair$a veriT_vr78 veriT_vr79) :named @p_650)) :named @p_651) (! (is_res$b (! (run$b (! (c$ veriT_vr78 veriT_vr57) :named @p_653) veriT_vr79) :named @p_655) @p_550) :named @p_659)) :named @p_660)) :named @p_649)) :named @p_661)) :named @p_663) (! (=> (! (or @p_548 (! (exists ((veriT_vr80 E$) (veriT_vr81 D$)) (! (and (! (is_res$a @p_546 (! (pair$a veriT_vr80 veriT_vr81) :named @p_665)) :named @p_666) (! (is_res$b (! (run$b (! (c$ veriT_vr80 veriT_vr57) :named @p_668) veriT_vr81) :named @p_670) @p_550) :named @p_671)) :named @p_672)) :named @p_664)) :named @p_673) (! (or @p_549 (! (exists ((veriT_vr82 E$) (veriT_vr83 D$)) (! (and (! (is_res$a @p_547 (! (pair$a veriT_vr82 veriT_vr83) :named @p_675)) :named @p_676) (! (is_res$b (! (run$b (! (c$ veriT_vr82 veriT_vr58) :named @p_678) veriT_vr83) :named @p_680) @p_550) :named @p_681)) :named @p_682)) :named @p_674)) :named @p_683)) :named @p_685)) :named @p_686)) :named @p_638)) :named @p_687)) :named @p_688)) :named @p_574)) :named @p_689)))) :named @p_696)) +(define-fun veriT_sk1 () A_b_c_M_state_fun$ (! (choice ((veriT_vr58 A_b_c_M_state_fun$)) (not (=> (forall ((veriT_vr59 A$) (veriT_vr60 C$)) (or (! (is_fail$ (! (run$ (fun_app$ @p_696 veriT_vr59) veriT_vr60) :named @p_698)) :named @p_697) (and (= @p_556 @p_697) (forall ((veriT_vr61 B$) (veriT_vr62 C$)) (= @p_566 (is_res$ @p_698 @p_545)))))) (forall ((veriT_vr63 D$)) (or (! (is_fail$a (! (run$a (! (b$ @p_696) :named @p_721) veriT_vr63) :named @p_699)) :named @p_700) (! (exists ((veriT_vr64 E$) (veriT_vr65 D$)) (and (is_res$a @p_699 @p_580) (is_fail$b (run$b (c$ veriT_vr64 @p_696) veriT_vr65)))) :named @p_704) (and (and (=> @p_600 (! (or @p_700 (exists ((veriT_vr68 E$) (veriT_vr69 D$)) (and (is_res$a @p_699 @p_603) (is_fail$b (run$b (c$ veriT_vr68 @p_696) veriT_vr69))))) :named @p_707)) (=> (! (or @p_700 (exists ((veriT_vr70 E$) (veriT_vr71 D$)) (and (is_res$a @p_699 @p_615) (! (is_fail$b (run$b (c$ veriT_vr70 @p_696) veriT_vr71)) :named @p_722)))) :named @p_708) @p_634)) (forall ((veriT_vr74 F$) (veriT_vr75 D$)) (and (=> @p_648 (! (or @p_700 (exists ((veriT_vr78 E$) (veriT_vr79 D$)) (and (is_res$a @p_699 @p_650) (! (is_res$b (! (run$b (c$ veriT_vr78 @p_696) veriT_vr79) :named @p_737) @p_550) :named @p_730)))) :named @p_710)) (=> (! (or @p_700 (exists ((veriT_vr80 E$) (veriT_vr81 D$)) (and (is_res$a @p_699 @p_665) (! (is_res$b (! (run$b (c$ veriT_vr80 @p_696) veriT_vr81) :named @p_740) @p_550) :named @p_732)))) :named @p_711) @p_683))))))))) :named @p_705)) +(define-fun veriT_sk2 () D$ (! (choice ((veriT_vr63 D$)) (not (or @p_700 @p_704 (and (and (=> (or (! (is_fail$a (! (run$a (! (b$ @p_705) :named @p_715) veriT_vr63) :named @p_706)) :named @p_709) (exists ((veriT_vr66 E$) (veriT_vr67 D$)) (and (is_res$a @p_706 @p_592) (! (is_fail$b (run$b (c$ veriT_vr66 @p_705) veriT_vr67)) :named @p_716)))) @p_707) (=> @p_708 (or @p_709 (exists ((veriT_vr72 E$) (veriT_vr73 D$)) (and (is_res$a @p_706 @p_626) (is_fail$b (run$b (c$ veriT_vr72 @p_705) veriT_vr73))))))) (forall ((veriT_vr74 F$) (veriT_vr75 D$)) (and (=> (or @p_709 (exists ((veriT_vr76 E$) (veriT_vr77 D$)) (and (is_res$a @p_706 @p_640) (! (is_res$b (! (run$b (c$ veriT_vr76 @p_705) veriT_vr77) :named @p_727) @p_550) :named @p_729)))) @p_710) (=> @p_711 (or @p_709 (exists ((veriT_vr82 E$) (veriT_vr83 D$)) (and (is_res$a @p_706 @p_675) (! (is_res$b (! (run$b (c$ veriT_vr82 @p_705) veriT_vr83) :named @p_742) @p_550) :named @p_734))))))))))) :named @p_713)) +(define-fun veriT_sk3 () E$ (! (choice ((veriT_vr66 E$)) (exists ((veriT_vr67 D$)) (and (is_res$a (! (run$a @p_715 @p_713) :named @p_717) @p_592) @p_716))) :named @p_718)) +(define-fun veriT_sk4 () D$ (! (choice ((veriT_vr67 D$)) (and (is_res$a @p_717 (pair$a @p_718 veriT_vr67)) (is_fail$b (run$b (c$ @p_718 @p_705) veriT_vr67)))) :named @p_719)) +(define-fun veriT_sk5 () E$ (! (choice ((veriT_vr70 E$)) (exists ((veriT_vr71 D$)) (and (is_res$a (! (run$a @p_721 @p_713) :named @p_723) @p_615) @p_722))) :named @p_724)) +(define-fun veriT_sk6 () D$ (! (choice ((veriT_vr71 D$)) (and (is_res$a @p_723 (pair$a @p_724 veriT_vr71)) (is_fail$b (run$b (c$ @p_724 @p_696) veriT_vr71)))) :named @p_725)) +(define-fun veriT_sk7 () F$ (! (choice ((veriT_vr74 F$)) (not (forall ((veriT_vr75 D$)) (and (=> (or (! (is_fail$a @p_717) :named @p_733) (exists ((veriT_vr76 E$) (veriT_vr77 D$)) (and (! (is_res$a @p_717 @p_640) :named @p_728) @p_729))) (or (! (is_fail$a @p_723) :named @p_731) (exists ((veriT_vr78 E$) (veriT_vr79 D$)) (and (! (is_res$a @p_723 @p_650) :named @p_736) @p_730)))) (=> (or @p_731 (exists ((veriT_vr80 E$) (veriT_vr81 D$)) (and (! (is_res$a @p_723 @p_665) :named @p_739) @p_732))) (or @p_733 (exists ((veriT_vr82 E$) (veriT_vr83 D$)) (and (! (is_res$a @p_717 @p_675) :named @p_741) @p_734)))))))) :named @p_735)) +(define-fun veriT_sk8 () D$ (! (choice ((veriT_vr75 D$)) (not (and (=> (or @p_733 (exists ((veriT_vr76 E$) (veriT_vr77 D$)) (and @p_728 (is_res$b @p_727 (! (pair$b @p_735 veriT_vr75) :named @p_738))))) (or @p_731 (exists ((veriT_vr78 E$) (veriT_vr79 D$)) (and @p_736 (is_res$b @p_737 @p_738))))) (=> (or @p_731 (exists ((veriT_vr80 E$) (veriT_vr81 D$)) (and @p_739 (is_res$b @p_740 @p_738)))) (or @p_733 (exists ((veriT_vr82 E$) (veriT_vr83 D$)) (and @p_741 (is_res$b @p_742 @p_738)))))))) :named @p_746)) +(define-fun veriT_sk9 () E$ (! (choice ((veriT_vr76 E$)) (exists ((veriT_vr77 D$)) (and @p_728 (is_res$b @p_727 (! (pair$b @p_735 @p_746) :named @p_744))))) :named @p_743)) +(define-fun veriT_sk10 () D$ (! (choice ((veriT_vr77 D$)) (and (is_res$a @p_717 (pair$a @p_743 veriT_vr77)) (is_res$b (run$b (c$ @p_743 @p_705) veriT_vr77) @p_744))) :named @p_745)) +(define-fun veriT_sk11 () E$ (! (choice ((veriT_vr80 E$)) (exists ((veriT_vr81 D$)) (and @p_739 (is_res$b @p_740 @p_744)))) :named @p_748)) +(define-fun veriT_sk12 () D$ (! (choice ((veriT_vr81 D$)) (and (is_res$a @p_723 (pair$a @p_748 veriT_vr81)) (is_res$b (run$b (c$ @p_748 @p_696) veriT_vr81) @p_744))) :named @p_749)) +(define-fun veriT_sk13 () A$ (! (choice ((veriT_vr122 A$)) (not (forall ((veriT_vr123 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 veriT_vr122) :named @p_1255) veriT_vr123) :named @p_1191)) :named @p_1189) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 veriT_vr122) :named @p_1257) veriT_vr123) :named @p_1190)) :named @p_1259) @p_1189) :named @p_1263) (! (forall ((veriT_vr124 B$) (veriT_vr125 C$)) (! (= (! (is_res$ @p_1190 (! (pair$ veriT_vr124 veriT_vr125) :named @p_1192)) :named @p_1267) (! (is_res$ @p_1191 @p_1192) :named @p_1271)) :named @p_1272)) :named @p_1264)) :named @p_1273)) :named @p_1254)))) :named @p_1193)) +(define-fun veriT_sk14 () C$ (! (choice ((veriT_vr123 C$)) (not (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 @p_1193) :named @p_1201) veriT_vr123) :named @p_1196)) :named @p_1194) (and (= (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 @p_1193) :named @p_1200) veriT_vr123) :named @p_1195)) @p_1194) (forall ((veriT_vr124 B$) (veriT_vr125 C$)) (= (is_res$ @p_1195 @p_1192) (is_res$ @p_1196 @p_1192))))))) :named @p_1197)) +(define-fun veriT_sk15 () B$ (! (choice ((veriT_vr124 B$)) (not (forall ((veriT_vr125 C$)) (= (is_res$ (! (run$ @p_1200 @p_1197) :named @p_1202) @p_1192) (is_res$ (! (run$ @p_1201 @p_1197) :named @p_1204) @p_1192))))) :named @p_1203)) +(define-fun veriT_sk16 () C$ (! (choice ((veriT_vr125 C$)) (not (= (is_res$ @p_1202 (! (pair$ @p_1203 veriT_vr125) :named @p_1205)) (is_res$ @p_1204 @p_1205)))) :named @p_1300)) +(define-fun veriT_sk25 () A$ (! (choice ((veriT_vr171 A$)) (not (forall ((veriT_vr172 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 veriT_vr171) :named @p_1468) veriT_vr172) :named @p_1393)) :named @p_1391) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 veriT_vr171) :named @p_1470) veriT_vr172) :named @p_1392)) :named @p_1472) @p_1391) :named @p_1476) (! (forall ((veriT_vr173 B$) (veriT_vr174 C$)) (! (= (! (is_res$ @p_1392 (! (pair$ veriT_vr173 veriT_vr174) :named @p_1394)) :named @p_1480) (! (is_res$ @p_1393 @p_1394) :named @p_1484)) :named @p_1485)) :named @p_1477)) :named @p_1486)) :named @p_1467)))) :named @p_1395)) +(define-fun veriT_sk26 () C$ (! (choice ((veriT_vr172 C$)) (not (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 @p_1395) :named @p_1403) veriT_vr172) :named @p_1398)) :named @p_1396) (and (= (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 @p_1395) :named @p_1402) veriT_vr172) :named @p_1397)) @p_1396) (forall ((veriT_vr173 B$) (veriT_vr174 C$)) (= (is_res$ @p_1397 @p_1394) (is_res$ @p_1398 @p_1394))))))) :named @p_1399)) +(define-fun veriT_sk27 () B$ (! (choice ((veriT_vr173 B$)) (not (forall ((veriT_vr174 C$)) (= (is_res$ (! (run$ @p_1402 @p_1399) :named @p_1404) @p_1394) (is_res$ (! (run$ @p_1403 @p_1399) :named @p_1406) @p_1394))))) :named @p_1405)) +(define-fun veriT_sk28 () C$ (! (choice ((veriT_vr174 C$)) (not (= (is_res$ @p_1404 (! (pair$ @p_1405 veriT_vr174) :named @p_1407)) (is_res$ @p_1406 @p_1407)))) :named @p_1513)) +(define-fun veriT_sk29 () A$ (! (choice ((veriT_vr185 A$)) (not (forall ((veriT_vr186 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 veriT_vr185) :named @p_1602) veriT_vr186) :named @p_1538)) :named @p_1536) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 veriT_vr185) :named @p_1604) veriT_vr186) :named @p_1537)) :named @p_1606) @p_1536) :named @p_1610) (! (forall ((veriT_vr187 B$) (veriT_vr188 C$)) (! (= (! (is_res$ @p_1537 (! (pair$ veriT_vr187 veriT_vr188) :named @p_1539)) :named @p_1614) (! (is_res$ @p_1538 @p_1539) :named @p_1618)) :named @p_1619)) :named @p_1611)) :named @p_1620)) :named @p_1601)))) :named @p_1540)) +(define-fun veriT_sk30 () C$ (! (choice ((veriT_vr186 C$)) (not (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 @p_1540) :named @p_1548) veriT_vr186) :named @p_1543)) :named @p_1541) (and (= (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 @p_1540) :named @p_1547) veriT_vr186) :named @p_1542)) @p_1541) (forall ((veriT_vr187 B$) (veriT_vr188 C$)) (= (is_res$ @p_1542 @p_1539) (is_res$ @p_1543 @p_1539))))))) :named @p_1544)) +(define-fun veriT_sk31 () B$ (! (choice ((veriT_vr187 B$)) (not (forall ((veriT_vr188 C$)) (= (is_res$ (! (run$ @p_1547 @p_1544) :named @p_1549) @p_1539) (is_res$ (! (run$ @p_1548 @p_1544) :named @p_1551) @p_1539))))) :named @p_1550)) +(define-fun veriT_sk32 () C$ (! (choice ((veriT_vr188 C$)) (not (= (is_res$ @p_1549 (! (pair$ @p_1550 veriT_vr188) :named @p_1552)) (is_res$ @p_1551 @p_1552)))) :named @p_1647)) +(define-fun veriT_sk41 () A$ (! (choice ((veriT_vr243 A$)) (not (forall ((veriT_vr244 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 veriT_vr243) :named @p_1825) veriT_vr244) :named @p_1761)) :named @p_1759) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 veriT_vr243) :named @p_1827) veriT_vr244) :named @p_1760)) :named @p_1829) @p_1759) :named @p_1833) (! (forall ((veriT_vr245 B$) (veriT_vr246 C$)) (! (= (! (is_res$ @p_1760 (! (pair$ veriT_vr245 veriT_vr246) :named @p_1762)) :named @p_1837) (! (is_res$ @p_1761 @p_1762) :named @p_1841)) :named @p_1842)) :named @p_1834)) :named @p_1843)) :named @p_1824)))) :named @p_1763)) +(define-fun veriT_sk42 () C$ (! (choice ((veriT_vr244 C$)) (not (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 @p_1763) :named @p_1771) veriT_vr244) :named @p_1766)) :named @p_1764) (and (= (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 @p_1763) :named @p_1770) veriT_vr244) :named @p_1765)) @p_1764) (forall ((veriT_vr245 B$) (veriT_vr246 C$)) (= (is_res$ @p_1765 @p_1762) (is_res$ @p_1766 @p_1762))))))) :named @p_1767)) +(define-fun veriT_sk43 () B$ (! (choice ((veriT_vr245 B$)) (not (forall ((veriT_vr246 C$)) (= (is_res$ (! (run$ @p_1770 @p_1767) :named @p_1772) @p_1762) (is_res$ (! (run$ @p_1771 @p_1767) :named @p_1774) @p_1762))))) :named @p_1773)) +(define-fun veriT_sk44 () C$ (! (choice ((veriT_vr246 C$)) (not (= (is_res$ @p_1772 (! (pair$ @p_1773 veriT_vr246) :named @p_1775)) (is_res$ @p_1774 @p_1775)))) :named @p_1870)) +(assume axiom0 (! (not (! (=> (! (and (! (forall ((?v0 A_b_c_M_state_fun$) (?v1 A_b_c_M_state_fun$)) (! (=> (! (forall ((?v2 A$) (?v3 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ ?v0 ?v2) :named @p_34) ?v3) :named @p_3)) :named @p_1) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ ?v1 ?v2) :named @p_37) ?v3) :named @p_2)) :named @p_40) @p_1) :named @p_46) (! (forall ((?v4 B$) (?v5 C$)) (! (= (! (is_res$ @p_2 (! (pair$ ?v4 ?v5) :named @p_4)) :named @p_53) (! (is_res$ @p_3 @p_4) :named @p_58)) :named @p_60)) :named @p_48)) :named @p_62)) :named @p_64)) :named @p_17) (! (forall ((?v2 D$)) (! (or (! (is_fail$a (! (run$a (! (b$ ?v0) :named @p_68) ?v2) :named @p_7)) :named @p_5) (! (and (! (= (! (is_fail$a (! (run$a (! (b$ ?v1) :named @p_70) ?v2) :named @p_6)) :named @p_19) @p_5) :named @p_77) (! (forall ((?v3 E$) (?v4 D$)) (! (= (! (is_res$a @p_6 (! (pair$a ?v3 ?v4) :named @p_8)) :named @p_20) (! (is_res$a @p_7 @p_8) :named @p_18)) :named @p_88)) :named @p_79)) :named @p_90)) :named @p_92)) :named @p_66)) :named @p_94)) :named @p_24) (! (forall ((?v0 E$) (?v1 A_b_c_M_state_fun$) (?v2 A_b_c_M_state_fun$)) (! (=> (! (forall ((?v3 A$) (?v4 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ ?v1 ?v3) :named @p_102) ?v4) :named @p_11)) :named @p_9) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ ?v2 ?v3) :named @p_104) ?v4) :named @p_10)) :named @p_106) @p_9) :named @p_111) (! (forall ((?v5 B$) (?v6 C$)) (! (= (! (is_res$ @p_10 (! (pair$ ?v5 ?v6) :named @p_12)) :named @p_116) (! (is_res$ @p_11 @p_12) :named @p_120)) :named @p_121)) :named @p_112)) :named @p_122)) :named @p_123)) :named @p_101) (! (forall ((?v3 D$)) (! (or (! (is_fail$b (! (run$b (! (c$ ?v0 ?v1) :named @p_126) ?v3) :named @p_15)) :named @p_13) (! (and (! (= (! (is_fail$b (! (run$b (! (c$ ?v0 ?v2) :named @p_129) ?v3) :named @p_14)) :named @p_132) @p_13) :named @p_137) (! (forall ((?v4 F$) (?v5 D$)) (! (= (! (is_res$b @p_14 (! (pair$b ?v4 ?v5) :named @p_16)) :named @p_143) (! (is_res$b @p_15 @p_16) :named @p_148)) :named @p_150)) :named @p_139)) :named @p_152)) :named @p_154)) :named @p_124)) :named @p_156)) :named @p_96)) :named @p_158) (! (forall ((?v0 A_b_c_M_state_fun$) (?v1 A_b_c_M_state_fun$)) (! (=> @p_17 (! (forall ((?v2 D$)) (! (or @p_5 (! (or (! (exists ((?v3 E$) (?v4 D$)) (! (and @p_18 (! (is_fail$b (! (run$b (! (c$ ?v3 ?v0) :named @p_176) ?v4) :named @p_177)) :named @p_179)) :named @p_181)) :named @p_21) (! (and (! (= (! (or @p_19 (! (exists ((?v3 E$) (?v4 D$)) (! (and @p_20 (! (is_fail$b (! (run$b (! (c$ ?v3 ?v1) :named @p_187) ?v4) :named @p_188)) :named @p_190)) :named @p_192)) :named @p_184)) :named @p_194) (! (or @p_5 @p_21) :named @p_201)) :named @p_203) (! (forall ((?v3 F$) (?v4 D$)) (! (= (! (or @p_19 (! (exists ((?v5 E$) (?v6 D$)) (! (and (! (is_res$a @p_6 (! (pair$a ?v5 ?v6) :named @p_22)) :named @p_209) (! (is_res$b (! (run$b (! (c$ ?v5 ?v1) :named @p_212) ?v6) :named @p_214) (! (pair$b ?v3 ?v4) :named @p_23)) :named @p_216)) :named @p_218)) :named @p_207)) :named @p_220) (! (or @p_5 (! (exists ((?v5 E$) (?v6 D$)) (! (and (! (is_res$a @p_7 @p_22) :named @p_225) (! (is_res$b (! (run$b (! (c$ ?v5 ?v0) :named @p_227) ?v6) :named @p_228) @p_23) :named @p_232)) :named @p_234)) :named @p_222)) :named @p_236)) :named @p_238)) :named @p_205)) :named @p_240)) :named @p_242)) :named @p_244)) :named @p_173)) :named @p_246)) :named @p_161)) :named @p_248)) :named @p_251)) +(anchor :step t2 :args ((:= (?v0 A_b_c_M_state_fun$) veriT_vr0) (:= (?v1 A_b_c_M_state_fun$) veriT_vr1))) +(anchor :step t2.t1 :args ((:= (?v2 A$) veriT_vr2) (:= (?v3 C$) veriT_vr3))) +(step t2.t1.t1 (cl (! (= ?v0 veriT_vr0) :named @p_42)) :rule refl) +(step t2.t1.t2 (cl (! (= ?v2 veriT_vr2) :named @p_36)) :rule refl) +(step t2.t1.t3 (cl (! (= @p_34 (! (fun_app$ veriT_vr0 veriT_vr2) :named @p_35)) :named @p_43)) :rule cong :premises (t2.t1.t1 t2.t1.t2)) +(step t2.t1.t4 (cl (! (= ?v3 veriT_vr3) :named @p_39)) :rule refl) +(step t2.t1.t5 (cl (! (= @p_3 (! (run$ @p_35 veriT_vr3) :named @p_27)) :named @p_44)) :rule cong :premises (t2.t1.t3 t2.t1.t4)) +(step t2.t1.t6 (cl (! (= @p_1 (! (is_fail$ @p_27) :named @p_25)) :named @p_45)) :rule cong :premises (t2.t1.t5)) +(step t2.t1.t7 (cl (! (= ?v1 veriT_vr1) :named @p_50)) :rule refl) +(step t2.t1.t8 (cl @p_36) :rule refl) +(step t2.t1.t9 (cl (! (= @p_37 (! (fun_app$ veriT_vr1 veriT_vr2) :named @p_38)) :named @p_51)) :rule cong :premises (t2.t1.t7 t2.t1.t8)) +(step t2.t1.t10 (cl @p_39) :rule refl) +(step t2.t1.t11 (cl (! (= @p_2 (! (run$ @p_38 veriT_vr3) :named @p_26)) :named @p_52)) :rule cong :premises (t2.t1.t9 t2.t1.t10)) +(step t2.t1.t12 (cl (! (= @p_40 (! (is_fail$ @p_26) :named @p_41)) :named @p_165)) :rule cong :premises (t2.t1.t11)) +(step t2.t1.t13 (cl @p_42) :rule refl) +(step t2.t1.t14 (cl @p_36) :rule refl) +(step t2.t1.t15 (cl @p_43) :rule cong :premises (t2.t1.t13 t2.t1.t14)) +(step t2.t1.t16 (cl @p_39) :rule refl) +(step t2.t1.t17 (cl @p_44) :rule cong :premises (t2.t1.t15 t2.t1.t16)) +(step t2.t1.t18 (cl @p_45) :rule cong :premises (t2.t1.t17)) +(step t2.t1.t19 (cl (! (= @p_46 (! (= @p_41 @p_25) :named @p_47)) :named @p_166)) :rule cong :premises (t2.t1.t12 t2.t1.t18)) +(anchor :step t2.t1.t20 :args ((:= (?v4 B$) veriT_vr4) (:= (?v5 C$) veriT_vr5))) +(step t2.t1.t20.t1 (cl @p_50) :rule refl) +(step t2.t1.t20.t2 (cl @p_36) :rule refl) +(step t2.t1.t20.t3 (cl @p_51) :rule cong :premises (t2.t1.t20.t1 t2.t1.t20.t2)) +(step t2.t1.t20.t4 (cl @p_39) :rule refl) +(step t2.t1.t20.t5 (cl @p_52) :rule cong :premises (t2.t1.t20.t3 t2.t1.t20.t4)) +(step t2.t1.t20.t6 (cl (! (= ?v4 veriT_vr4) :named @p_55)) :rule refl) +(step t2.t1.t20.t7 (cl (! (= ?v5 veriT_vr5) :named @p_56)) :rule refl) +(step t2.t1.t20.t8 (cl (! (= @p_4 (! (pair$ veriT_vr4 veriT_vr5) :named @p_28)) :named @p_57)) :rule cong :premises (t2.t1.t20.t6 t2.t1.t20.t7)) +(step t2.t1.t20.t9 (cl (! (= @p_53 (! (is_res$ @p_26 @p_28) :named @p_54)) :named @p_168)) :rule cong :premises (t2.t1.t20.t5 t2.t1.t20.t8)) +(step t2.t1.t20.t10 (cl @p_42) :rule refl) +(step t2.t1.t20.t11 (cl @p_36) :rule refl) +(step t2.t1.t20.t12 (cl @p_43) :rule cong :premises (t2.t1.t20.t10 t2.t1.t20.t11)) +(step t2.t1.t20.t13 (cl @p_39) :rule refl) +(step t2.t1.t20.t14 (cl @p_44) :rule cong :premises (t2.t1.t20.t12 t2.t1.t20.t13)) +(step t2.t1.t20.t15 (cl @p_55) :rule refl) +(step t2.t1.t20.t16 (cl @p_56) :rule refl) +(step t2.t1.t20.t17 (cl @p_57) :rule cong :premises (t2.t1.t20.t15 t2.t1.t20.t16)) +(step t2.t1.t20.t18 (cl (! (= @p_58 (! (is_res$ @p_27 @p_28) :named @p_59)) :named @p_169)) :rule cong :premises (t2.t1.t20.t14 t2.t1.t20.t17)) +(step t2.t1.t20.t19 (cl (! (= @p_60 (! (= @p_54 @p_59) :named @p_61)) :named @p_170)) :rule cong :premises (t2.t1.t20.t9 t2.t1.t20.t18)) +(step t2.t1.t20 (cl (! (= @p_48 (! (forall ((veriT_vr4 B$) (veriT_vr5 C$)) @p_61) :named @p_49)) :named @p_167)) :rule bind) +(step t2.t1.t21 (cl (! (= @p_62 (! (and @p_47 @p_49) :named @p_63)) :named @p_171)) :rule cong :premises (t2.t1.t19 t2.t1.t20)) +(step t2.t1.t22 (cl (! (= @p_64 (! (or @p_25 @p_63) :named @p_65)) :named @p_172)) :rule cong :premises (t2.t1.t6 t2.t1.t21)) +(step t2.t1 (cl (! (= @p_17 (! (forall ((veriT_vr2 A$) (veriT_vr3 C$)) @p_65) :named @p_33)) :named @p_164)) :rule bind) +(anchor :step t2.t2 :args ((:= (?v2 D$) veriT_vr6))) +(step t2.t2.t1 (cl @p_42) :rule refl) +(step t2.t2.t2 (cl (! (= @p_68 (! (b$ veriT_vr0) :named @p_69)) :named @p_74)) :rule cong :premises (t2.t2.t1)) +(step t2.t2.t3 (cl (! (= ?v2 veriT_vr6) :named @p_72)) :rule refl) +(step t2.t2.t4 (cl (! (= @p_7 (! (run$a @p_69 veriT_vr6) :named @p_31)) :named @p_75)) :rule cong :premises (t2.t2.t2 t2.t2.t3)) +(step t2.t2.t5 (cl (! (= @p_5 (! (is_fail$a @p_31) :named @p_29)) :named @p_76)) :rule cong :premises (t2.t2.t4)) +(step t2.t2.t6 (cl @p_50) :rule refl) +(step t2.t2.t7 (cl (! (= @p_70 (! (b$ veriT_vr1) :named @p_71)) :named @p_81)) :rule cong :premises (t2.t2.t6)) +(step t2.t2.t8 (cl @p_72) :rule refl) +(step t2.t2.t9 (cl (! (= @p_6 (! (run$a @p_71 veriT_vr6) :named @p_30)) :named @p_82)) :rule cong :premises (t2.t2.t7 t2.t2.t8)) +(step t2.t2.t10 (cl (! (= @p_19 (! (is_fail$a @p_30) :named @p_73)) :named @p_183)) :rule cong :premises (t2.t2.t9)) +(step t2.t2.t11 (cl @p_42) :rule refl) +(step t2.t2.t12 (cl @p_74) :rule cong :premises (t2.t2.t11)) +(step t2.t2.t13 (cl @p_72) :rule refl) +(step t2.t2.t14 (cl @p_75) :rule cong :premises (t2.t2.t12 t2.t2.t13)) +(step t2.t2.t15 (cl @p_76) :rule cong :premises (t2.t2.t14)) +(step t2.t2.t16 (cl (= @p_77 (! (= @p_73 @p_29) :named @p_78))) :rule cong :premises (t2.t2.t10 t2.t2.t15)) +(anchor :step t2.t2.t17 :args ((:= (?v3 E$) veriT_vr7) (:= (?v4 D$) veriT_vr8))) +(step t2.t2.t17.t1 (cl @p_50) :rule refl) +(step t2.t2.t17.t2 (cl @p_81) :rule cong :premises (t2.t2.t17.t1)) +(step t2.t2.t17.t3 (cl @p_72) :rule refl) +(step t2.t2.t17.t4 (cl @p_82) :rule cong :premises (t2.t2.t17.t2 t2.t2.t17.t3)) +(step t2.t2.t17.t5 (cl (! (= ?v3 veriT_vr7) :named @p_84)) :rule refl) +(step t2.t2.t17.t6 (cl (! (= ?v4 veriT_vr8) :named @p_85)) :rule refl) +(step t2.t2.t17.t7 (cl (! (= @p_8 (! (pair$a veriT_vr7 veriT_vr8) :named @p_32)) :named @p_86)) :rule cong :premises (t2.t2.t17.t5 t2.t2.t17.t6)) +(step t2.t2.t17.t8 (cl (! (= @p_20 (! (is_res$a @p_30 @p_32) :named @p_83)) :named @p_186)) :rule cong :premises (t2.t2.t17.t4 t2.t2.t17.t7)) +(step t2.t2.t17.t9 (cl @p_42) :rule refl) +(step t2.t2.t17.t10 (cl @p_74) :rule cong :premises (t2.t2.t17.t9)) +(step t2.t2.t17.t11 (cl @p_72) :rule refl) +(step t2.t2.t17.t12 (cl @p_75) :rule cong :premises (t2.t2.t17.t10 t2.t2.t17.t11)) +(step t2.t2.t17.t13 (cl @p_84) :rule refl) +(step t2.t2.t17.t14 (cl @p_85) :rule refl) +(step t2.t2.t17.t15 (cl @p_86) :rule cong :premises (t2.t2.t17.t13 t2.t2.t17.t14)) +(step t2.t2.t17.t16 (cl (! (= @p_18 (! (is_res$a @p_31 @p_32) :named @p_87)) :named @p_175)) :rule cong :premises (t2.t2.t17.t12 t2.t2.t17.t15)) +(step t2.t2.t17.t17 (cl (= @p_88 (! (= @p_83 @p_87) :named @p_89))) :rule cong :premises (t2.t2.t17.t8 t2.t2.t17.t16)) +(step t2.t2.t17 (cl (= @p_79 (! (forall ((veriT_vr7 E$) (veriT_vr8 D$)) @p_89) :named @p_80))) :rule bind) +(step t2.t2.t18 (cl (= @p_90 (! (and @p_78 @p_80) :named @p_91))) :rule cong :premises (t2.t2.t16 t2.t2.t17)) +(step t2.t2.t19 (cl (= @p_92 (! (or @p_29 @p_91) :named @p_93))) :rule cong :premises (t2.t2.t5 t2.t2.t18)) +(step t2.t2 (cl (= @p_66 (! (forall ((veriT_vr6 D$)) @p_93) :named @p_67))) :rule bind) +(step t2.t3 (cl (= @p_94 (! (=> @p_33 @p_67) :named @p_95))) :rule cong :premises (t2.t1 t2.t2)) +(step t2 (cl (= @p_24 (! (forall ((veriT_vr0 A_b_c_M_state_fun$) (veriT_vr1 A_b_c_M_state_fun$)) @p_95) :named @p_159))) :rule bind) +(anchor :step t3 :args ((:= (?v0 E$) veriT_vr7) (:= (?v1 A_b_c_M_state_fun$) veriT_vr0) (:= (?v2 A_b_c_M_state_fun$) veriT_vr1))) +(anchor :step t3.t1 :args ((:= (?v3 A$) veriT_vr2) (:= (?v4 C$) veriT_vr3))) +(step t3.t1.t1 (cl (! (= ?v1 veriT_vr0) :named @p_107)) :rule refl) +(step t3.t1.t2 (cl (! (= ?v3 veriT_vr2) :named @p_103)) :rule refl) +(step t3.t1.t3 (cl (! (= @p_102 @p_35) :named @p_108)) :rule cong :premises (t3.t1.t1 t3.t1.t2)) +(step t3.t1.t4 (cl (! (= ?v4 veriT_vr3) :named @p_105)) :rule refl) +(step t3.t1.t5 (cl (! (= @p_11 @p_27) :named @p_109)) :rule cong :premises (t3.t1.t3 t3.t1.t4)) +(step t3.t1.t6 (cl (! (= @p_9 @p_25) :named @p_110)) :rule cong :premises (t3.t1.t5)) +(step t3.t1.t7 (cl (! (= ?v2 veriT_vr1) :named @p_113)) :rule refl) +(step t3.t1.t8 (cl @p_103) :rule refl) +(step t3.t1.t9 (cl (! (= @p_104 @p_38) :named @p_114)) :rule cong :premises (t3.t1.t7 t3.t1.t8)) +(step t3.t1.t10 (cl @p_105) :rule refl) +(step t3.t1.t11 (cl (! (= @p_10 @p_26) :named @p_115)) :rule cong :premises (t3.t1.t9 t3.t1.t10)) +(step t3.t1.t12 (cl (= @p_106 @p_41)) :rule cong :premises (t3.t1.t11)) +(step t3.t1.t13 (cl @p_107) :rule refl) +(step t3.t1.t14 (cl @p_103) :rule refl) +(step t3.t1.t15 (cl @p_108) :rule cong :premises (t3.t1.t13 t3.t1.t14)) +(step t3.t1.t16 (cl @p_105) :rule refl) +(step t3.t1.t17 (cl @p_109) :rule cong :premises (t3.t1.t15 t3.t1.t16)) +(step t3.t1.t18 (cl @p_110) :rule cong :premises (t3.t1.t17)) +(step t3.t1.t19 (cl (= @p_111 @p_47)) :rule cong :premises (t3.t1.t12 t3.t1.t18)) +(anchor :step t3.t1.t20 :args ((:= (?v5 B$) veriT_vr4) (:= (?v6 C$) veriT_vr5))) +(step t3.t1.t20.t1 (cl @p_113) :rule refl) +(step t3.t1.t20.t2 (cl @p_103) :rule refl) +(step t3.t1.t20.t3 (cl @p_114) :rule cong :premises (t3.t1.t20.t1 t3.t1.t20.t2)) +(step t3.t1.t20.t4 (cl @p_105) :rule refl) +(step t3.t1.t20.t5 (cl @p_115) :rule cong :premises (t3.t1.t20.t3 t3.t1.t20.t4)) +(step t3.t1.t20.t6 (cl (! (= ?v5 veriT_vr4) :named @p_117)) :rule refl) +(step t3.t1.t20.t7 (cl (! (= ?v6 veriT_vr5) :named @p_118)) :rule refl) +(step t3.t1.t20.t8 (cl (! (= @p_12 @p_28) :named @p_119)) :rule cong :premises (t3.t1.t20.t6 t3.t1.t20.t7)) +(step t3.t1.t20.t9 (cl (= @p_116 @p_54)) :rule cong :premises (t3.t1.t20.t5 t3.t1.t20.t8)) +(step t3.t1.t20.t10 (cl @p_107) :rule refl) +(step t3.t1.t20.t11 (cl @p_103) :rule refl) +(step t3.t1.t20.t12 (cl @p_108) :rule cong :premises (t3.t1.t20.t10 t3.t1.t20.t11)) +(step t3.t1.t20.t13 (cl @p_105) :rule refl) +(step t3.t1.t20.t14 (cl @p_109) :rule cong :premises (t3.t1.t20.t12 t3.t1.t20.t13)) +(step t3.t1.t20.t15 (cl @p_117) :rule refl) +(step t3.t1.t20.t16 (cl @p_118) :rule refl) +(step t3.t1.t20.t17 (cl @p_119) :rule cong :premises (t3.t1.t20.t15 t3.t1.t20.t16)) +(step t3.t1.t20.t18 (cl (= @p_120 @p_59)) :rule cong :premises (t3.t1.t20.t14 t3.t1.t20.t17)) +(step t3.t1.t20.t19 (cl (= @p_121 @p_61)) :rule cong :premises (t3.t1.t20.t9 t3.t1.t20.t18)) +(step t3.t1.t20 (cl (= @p_112 @p_49)) :rule bind) +(step t3.t1.t21 (cl (= @p_122 @p_63)) :rule cong :premises (t3.t1.t19 t3.t1.t20)) +(step t3.t1.t22 (cl (= @p_123 @p_65)) :rule cong :premises (t3.t1.t6 t3.t1.t21)) +(step t3.t1 (cl (= @p_101 @p_33)) :rule bind) +(anchor :step t3.t2 :args ((:= (?v3 D$) veriT_vr6))) +(step t3.t2.t1 (cl (! (= ?v0 veriT_vr7) :named @p_128)) :rule refl) +(step t3.t2.t2 (cl @p_107) :rule refl) +(step t3.t2.t3 (cl (! (= @p_126 (! (c$ veriT_vr7 veriT_vr0) :named @p_127)) :named @p_134)) :rule cong :premises (t3.t2.t1 t3.t2.t2)) +(step t3.t2.t4 (cl (! (= ?v3 veriT_vr6) :named @p_131)) :rule refl) +(step t3.t2.t5 (cl (! (= @p_15 (! (run$b @p_127 veriT_vr6) :named @p_99)) :named @p_135)) :rule cong :premises (t3.t2.t3 t3.t2.t4)) +(step t3.t2.t6 (cl (! (= @p_13 (! (is_fail$b @p_99) :named @p_97)) :named @p_136)) :rule cong :premises (t3.t2.t5)) +(step t3.t2.t7 (cl @p_128) :rule refl) +(step t3.t2.t8 (cl @p_113) :rule refl) +(step t3.t2.t9 (cl (! (= @p_129 (! (c$ veriT_vr7 veriT_vr1) :named @p_130)) :named @p_141)) :rule cong :premises (t3.t2.t7 t3.t2.t8)) +(step t3.t2.t10 (cl @p_131) :rule refl) +(step t3.t2.t11 (cl (! (= @p_14 (! (run$b @p_130 veriT_vr6) :named @p_98)) :named @p_142)) :rule cong :premises (t3.t2.t9 t3.t2.t10)) +(step t3.t2.t12 (cl (= @p_132 (! (is_fail$b @p_98) :named @p_133))) :rule cong :premises (t3.t2.t11)) +(step t3.t2.t13 (cl @p_128) :rule refl) +(step t3.t2.t14 (cl @p_107) :rule refl) +(step t3.t2.t15 (cl @p_134) :rule cong :premises (t3.t2.t13 t3.t2.t14)) +(step t3.t2.t16 (cl @p_131) :rule refl) +(step t3.t2.t17 (cl @p_135) :rule cong :premises (t3.t2.t15 t3.t2.t16)) +(step t3.t2.t18 (cl @p_136) :rule cong :premises (t3.t2.t17)) +(step t3.t2.t19 (cl (= @p_137 (! (= @p_133 @p_97) :named @p_138))) :rule cong :premises (t3.t2.t12 t3.t2.t18)) +(anchor :step t3.t2.t20 :args ((:= (?v4 F$) veriT_vr9) (:= (?v5 D$) veriT_vr8))) +(step t3.t2.t20.t1 (cl @p_128) :rule refl) +(step t3.t2.t20.t2 (cl @p_113) :rule refl) +(step t3.t2.t20.t3 (cl @p_141) :rule cong :premises (t3.t2.t20.t1 t3.t2.t20.t2)) +(step t3.t2.t20.t4 (cl @p_131) :rule refl) +(step t3.t2.t20.t5 (cl @p_142) :rule cong :premises (t3.t2.t20.t3 t3.t2.t20.t4)) +(step t3.t2.t20.t6 (cl (! (= ?v4 veriT_vr9) :named @p_145)) :rule refl) +(step t3.t2.t20.t7 (cl (! (= ?v5 veriT_vr8) :named @p_146)) :rule refl) +(step t3.t2.t20.t8 (cl (! (= @p_16 (! (pair$b veriT_vr9 veriT_vr8) :named @p_100)) :named @p_147)) :rule cong :premises (t3.t2.t20.t6 t3.t2.t20.t7)) +(step t3.t2.t20.t9 (cl (= @p_143 (! (is_res$b @p_98 @p_100) :named @p_144))) :rule cong :premises (t3.t2.t20.t5 t3.t2.t20.t8)) +(step t3.t2.t20.t10 (cl @p_128) :rule refl) +(step t3.t2.t20.t11 (cl @p_107) :rule refl) +(step t3.t2.t20.t12 (cl @p_134) :rule cong :premises (t3.t2.t20.t10 t3.t2.t20.t11)) +(step t3.t2.t20.t13 (cl @p_131) :rule refl) +(step t3.t2.t20.t14 (cl @p_135) :rule cong :premises (t3.t2.t20.t12 t3.t2.t20.t13)) +(step t3.t2.t20.t15 (cl @p_145) :rule refl) +(step t3.t2.t20.t16 (cl @p_146) :rule refl) +(step t3.t2.t20.t17 (cl @p_147) :rule cong :premises (t3.t2.t20.t15 t3.t2.t20.t16)) +(step t3.t2.t20.t18 (cl (= @p_148 (! (is_res$b @p_99 @p_100) :named @p_149))) :rule cong :premises (t3.t2.t20.t14 t3.t2.t20.t17)) +(step t3.t2.t20.t19 (cl (= @p_150 (! (= @p_144 @p_149) :named @p_151))) :rule cong :premises (t3.t2.t20.t9 t3.t2.t20.t18)) +(step t3.t2.t20 (cl (= @p_139 (! (forall ((veriT_vr9 F$) (veriT_vr8 D$)) @p_151) :named @p_140))) :rule bind) +(step t3.t2.t21 (cl (= @p_152 (! (and @p_138 @p_140) :named @p_153))) :rule cong :premises (t3.t2.t19 t3.t2.t20)) +(step t3.t2.t22 (cl (= @p_154 (! (or @p_97 @p_153) :named @p_155))) :rule cong :premises (t3.t2.t6 t3.t2.t21)) +(step t3.t2 (cl (= @p_124 (! (forall ((veriT_vr6 D$)) @p_155) :named @p_125))) :rule bind) +(step t3.t3 (cl (= @p_156 (! (=> @p_33 @p_125) :named @p_157))) :rule cong :premises (t3.t1 t3.t2)) +(step t3 (cl (= @p_96 (! (forall ((veriT_vr7 E$) (veriT_vr0 A_b_c_M_state_fun$) (veriT_vr1 A_b_c_M_state_fun$)) @p_157) :named @p_160))) :rule bind) +(step t4 (cl (= @p_158 (! (and @p_159 @p_160) :named @p_249))) :rule cong :premises (t2 t3)) +(anchor :step t5 :args ((:= (?v0 A_b_c_M_state_fun$) veriT_vr0) (:= (?v1 A_b_c_M_state_fun$) veriT_vr1))) +(anchor :step t5.t1 :args ((:= (?v2 A$) veriT_vr2) (:= (?v3 C$) veriT_vr3))) +(step t5.t1.t1 (cl @p_42) :rule refl) +(step t5.t1.t2 (cl @p_36) :rule refl) +(step t5.t1.t3 (cl @p_43) :rule cong :premises (t5.t1.t1 t5.t1.t2)) +(step t5.t1.t4 (cl @p_39) :rule refl) +(step t5.t1.t5 (cl @p_44) :rule cong :premises (t5.t1.t3 t5.t1.t4)) +(step t5.t1.t6 (cl @p_45) :rule cong :premises (t5.t1.t5)) +(step t5.t1.t7 (cl @p_50) :rule refl) +(step t5.t1.t8 (cl @p_36) :rule refl) +(step t5.t1.t9 (cl @p_51) :rule cong :premises (t5.t1.t7 t5.t1.t8)) +(step t5.t1.t10 (cl @p_39) :rule refl) +(step t5.t1.t11 (cl @p_52) :rule cong :premises (t5.t1.t9 t5.t1.t10)) +(step t5.t1.t12 (cl @p_165) :rule cong :premises (t5.t1.t11)) +(step t5.t1.t13 (cl @p_42) :rule refl) +(step t5.t1.t14 (cl @p_36) :rule refl) +(step t5.t1.t15 (cl @p_43) :rule cong :premises (t5.t1.t13 t5.t1.t14)) +(step t5.t1.t16 (cl @p_39) :rule refl) +(step t5.t1.t17 (cl @p_44) :rule cong :premises (t5.t1.t15 t5.t1.t16)) +(step t5.t1.t18 (cl @p_45) :rule cong :premises (t5.t1.t17)) +(step t5.t1.t19 (cl @p_166) :rule cong :premises (t5.t1.t12 t5.t1.t18)) +(anchor :step t5.t1.t20 :args ((:= (?v4 B$) veriT_vr4) (:= (?v5 C$) veriT_vr5))) +(step t5.t1.t20.t1 (cl @p_50) :rule refl) +(step t5.t1.t20.t2 (cl @p_36) :rule refl) +(step t5.t1.t20.t3 (cl @p_51) :rule cong :premises (t5.t1.t20.t1 t5.t1.t20.t2)) +(step t5.t1.t20.t4 (cl @p_39) :rule refl) +(step t5.t1.t20.t5 (cl @p_52) :rule cong :premises (t5.t1.t20.t3 t5.t1.t20.t4)) +(step t5.t1.t20.t6 (cl @p_55) :rule refl) +(step t5.t1.t20.t7 (cl @p_56) :rule refl) +(step t5.t1.t20.t8 (cl @p_57) :rule cong :premises (t5.t1.t20.t6 t5.t1.t20.t7)) +(step t5.t1.t20.t9 (cl @p_168) :rule cong :premises (t5.t1.t20.t5 t5.t1.t20.t8)) +(step t5.t1.t20.t10 (cl @p_42) :rule refl) +(step t5.t1.t20.t11 (cl @p_36) :rule refl) +(step t5.t1.t20.t12 (cl @p_43) :rule cong :premises (t5.t1.t20.t10 t5.t1.t20.t11)) +(step t5.t1.t20.t13 (cl @p_39) :rule refl) +(step t5.t1.t20.t14 (cl @p_44) :rule cong :premises (t5.t1.t20.t12 t5.t1.t20.t13)) +(step t5.t1.t20.t15 (cl @p_55) :rule refl) +(step t5.t1.t20.t16 (cl @p_56) :rule refl) +(step t5.t1.t20.t17 (cl @p_57) :rule cong :premises (t5.t1.t20.t15 t5.t1.t20.t16)) +(step t5.t1.t20.t18 (cl @p_169) :rule cong :premises (t5.t1.t20.t14 t5.t1.t20.t17)) +(step t5.t1.t20.t19 (cl @p_170) :rule cong :premises (t5.t1.t20.t9 t5.t1.t20.t18)) +(step t5.t1.t20 (cl @p_167) :rule bind) +(step t5.t1.t21 (cl @p_171) :rule cong :premises (t5.t1.t19 t5.t1.t20)) +(step t5.t1.t22 (cl @p_172) :rule cong :premises (t5.t1.t6 t5.t1.t21)) +(step t5.t1 (cl @p_164) :rule bind) +(anchor :step t5.t2 :args ((:= (?v2 D$) veriT_vr6))) +(step t5.t2.t1 (cl @p_42) :rule refl) +(step t5.t2.t2 (cl @p_74) :rule cong :premises (t5.t2.t1)) +(step t5.t2.t3 (cl @p_72) :rule refl) +(step t5.t2.t4 (cl @p_75) :rule cong :premises (t5.t2.t2 t5.t2.t3)) +(step t5.t2.t5 (cl @p_76) :rule cong :premises (t5.t2.t4)) +(anchor :step t5.t2.t6 :args ((:= (?v3 E$) veriT_vr7) (:= (?v4 D$) veriT_vr8))) +(step t5.t2.t6.t1 (cl @p_42) :rule refl) +(step t5.t2.t6.t2 (cl @p_74) :rule cong :premises (t5.t2.t6.t1)) +(step t5.t2.t6.t3 (cl @p_72) :rule refl) +(step t5.t2.t6.t4 (cl @p_75) :rule cong :premises (t5.t2.t6.t2 t5.t2.t6.t3)) +(step t5.t2.t6.t5 (cl @p_84) :rule refl) +(step t5.t2.t6.t6 (cl @p_85) :rule refl) +(step t5.t2.t6.t7 (cl @p_86) :rule cong :premises (t5.t2.t6.t5 t5.t2.t6.t6)) +(step t5.t2.t6.t8 (cl @p_175) :rule cong :premises (t5.t2.t6.t4 t5.t2.t6.t7)) +(step t5.t2.t6.t9 (cl @p_84) :rule refl) +(step t5.t2.t6.t10 (cl @p_42) :rule refl) +(step t5.t2.t6.t11 (cl (! (= @p_176 @p_127) :named @p_197)) :rule cong :premises (t5.t2.t6.t9 t5.t2.t6.t10)) +(step t5.t2.t6.t12 (cl @p_85) :rule refl) +(step t5.t2.t6.t13 (cl (! (= @p_177 (! (run$b @p_127 veriT_vr8) :named @p_178)) :named @p_198)) :rule cong :premises (t5.t2.t6.t11 t5.t2.t6.t12)) +(step t5.t2.t6.t14 (cl (! (= @p_179 (! (is_fail$b @p_178) :named @p_180)) :named @p_199)) :rule cong :premises (t5.t2.t6.t13)) +(step t5.t2.t6.t15 (cl (! (= @p_181 (! (and @p_87 @p_180) :named @p_182)) :named @p_200)) :rule cong :premises (t5.t2.t6.t8 t5.t2.t6.t14)) +(step t5.t2.t6 (cl (! (= @p_21 (! (exists ((veriT_vr7 E$) (veriT_vr8 D$)) @p_182) :named @p_162)) :named @p_196)) :rule bind) +(step t5.t2.t7 (cl @p_50) :rule refl) +(step t5.t2.t8 (cl @p_81) :rule cong :premises (t5.t2.t7)) +(step t5.t2.t9 (cl @p_72) :rule refl) +(step t5.t2.t10 (cl @p_82) :rule cong :premises (t5.t2.t8 t5.t2.t9)) +(step t5.t2.t11 (cl @p_183) :rule cong :premises (t5.t2.t10)) +(anchor :step t5.t2.t12 :args ((:= (?v3 E$) veriT_vr7) (:= (?v4 D$) veriT_vr8))) +(step t5.t2.t12.t1 (cl @p_50) :rule refl) +(step t5.t2.t12.t2 (cl @p_81) :rule cong :premises (t5.t2.t12.t1)) +(step t5.t2.t12.t3 (cl @p_72) :rule refl) +(step t5.t2.t12.t4 (cl @p_82) :rule cong :premises (t5.t2.t12.t2 t5.t2.t12.t3)) +(step t5.t2.t12.t5 (cl @p_84) :rule refl) +(step t5.t2.t12.t6 (cl @p_85) :rule refl) +(step t5.t2.t12.t7 (cl @p_86) :rule cong :premises (t5.t2.t12.t5 t5.t2.t12.t6)) +(step t5.t2.t12.t8 (cl @p_186) :rule cong :premises (t5.t2.t12.t4 t5.t2.t12.t7)) +(step t5.t2.t12.t9 (cl @p_84) :rule refl) +(step t5.t2.t12.t10 (cl @p_50) :rule refl) +(step t5.t2.t12.t11 (cl (= @p_187 @p_130)) :rule cong :premises (t5.t2.t12.t9 t5.t2.t12.t10)) +(step t5.t2.t12.t12 (cl @p_85) :rule refl) +(step t5.t2.t12.t13 (cl (= @p_188 (! (run$b @p_130 veriT_vr8) :named @p_189))) :rule cong :premises (t5.t2.t12.t11 t5.t2.t12.t12)) +(step t5.t2.t12.t14 (cl (= @p_190 (! (is_fail$b @p_189) :named @p_191))) :rule cong :premises (t5.t2.t12.t13)) +(step t5.t2.t12.t15 (cl (= @p_192 (! (and @p_83 @p_191) :named @p_193))) :rule cong :premises (t5.t2.t12.t8 t5.t2.t12.t14)) +(step t5.t2.t12 (cl (= @p_184 (! (exists ((veriT_vr7 E$) (veriT_vr8 D$)) @p_193) :named @p_185))) :rule bind) +(step t5.t2.t13 (cl (= @p_194 (! (or @p_73 @p_185) :named @p_195))) :rule cong :premises (t5.t2.t11 t5.t2.t12)) +(step t5.t2.t14 (cl @p_42) :rule refl) +(step t5.t2.t15 (cl @p_74) :rule cong :premises (t5.t2.t14)) +(step t5.t2.t16 (cl @p_72) :rule refl) +(step t5.t2.t17 (cl @p_75) :rule cong :premises (t5.t2.t15 t5.t2.t16)) +(step t5.t2.t18 (cl @p_76) :rule cong :premises (t5.t2.t17)) +(anchor :step t5.t2.t19 :args ((:= (?v3 E$) veriT_vr7) (:= (?v4 D$) veriT_vr8))) +(step t5.t2.t19.t1 (cl @p_42) :rule refl) +(step t5.t2.t19.t2 (cl @p_74) :rule cong :premises (t5.t2.t19.t1)) +(step t5.t2.t19.t3 (cl @p_72) :rule refl) +(step t5.t2.t19.t4 (cl @p_75) :rule cong :premises (t5.t2.t19.t2 t5.t2.t19.t3)) +(step t5.t2.t19.t5 (cl @p_84) :rule refl) +(step t5.t2.t19.t6 (cl @p_85) :rule refl) +(step t5.t2.t19.t7 (cl @p_86) :rule cong :premises (t5.t2.t19.t5 t5.t2.t19.t6)) +(step t5.t2.t19.t8 (cl @p_175) :rule cong :premises (t5.t2.t19.t4 t5.t2.t19.t7)) +(step t5.t2.t19.t9 (cl @p_84) :rule refl) +(step t5.t2.t19.t10 (cl @p_42) :rule refl) +(step t5.t2.t19.t11 (cl @p_197) :rule cong :premises (t5.t2.t19.t9 t5.t2.t19.t10)) +(step t5.t2.t19.t12 (cl @p_85) :rule refl) +(step t5.t2.t19.t13 (cl @p_198) :rule cong :premises (t5.t2.t19.t11 t5.t2.t19.t12)) +(step t5.t2.t19.t14 (cl @p_199) :rule cong :premises (t5.t2.t19.t13)) +(step t5.t2.t19.t15 (cl @p_200) :rule cong :premises (t5.t2.t19.t8 t5.t2.t19.t14)) +(step t5.t2.t19 (cl @p_196) :rule bind) +(step t5.t2.t20 (cl (= @p_201 (! (or @p_29 @p_162) :named @p_202))) :rule cong :premises (t5.t2.t18 t5.t2.t19)) +(step t5.t2.t21 (cl (= @p_203 (! (= @p_195 @p_202) :named @p_204))) :rule cong :premises (t5.t2.t13 t5.t2.t20)) +(anchor :step t5.t2.t22 :args ((:= (?v3 F$) veriT_vr9) (:= (?v4 D$) veriT_vr8))) +(step t5.t2.t22.t1 (cl @p_50) :rule refl) +(step t5.t2.t22.t2 (cl @p_81) :rule cong :premises (t5.t2.t22.t1)) +(step t5.t2.t22.t3 (cl @p_72) :rule refl) +(step t5.t2.t22.t4 (cl @p_82) :rule cong :premises (t5.t2.t22.t2 t5.t2.t22.t3)) +(step t5.t2.t22.t5 (cl @p_183) :rule cong :premises (t5.t2.t22.t4)) +(anchor :step t5.t2.t22.t6 :args ((:= (?v5 E$) veriT_vr7) (:= (?v6 D$) veriT_vr10))) +(step t5.t2.t22.t6.t1 (cl @p_50) :rule refl) +(step t5.t2.t22.t6.t2 (cl @p_81) :rule cong :premises (t5.t2.t22.t6.t1)) +(step t5.t2.t22.t6.t3 (cl @p_72) :rule refl) +(step t5.t2.t22.t6.t4 (cl @p_82) :rule cong :premises (t5.t2.t22.t6.t2 t5.t2.t22.t6.t3)) +(step t5.t2.t22.t6.t5 (cl (! (= ?v5 veriT_vr7) :named @p_211)) :rule refl) +(step t5.t2.t22.t6.t6 (cl (! (= ?v6 veriT_vr10) :named @p_213)) :rule refl) +(step t5.t2.t22.t6.t7 (cl (! (= @p_22 (! (pair$a veriT_vr7 veriT_vr10) :named @p_163)) :named @p_224)) :rule cong :premises (t5.t2.t22.t6.t5 t5.t2.t22.t6.t6)) +(step t5.t2.t22.t6.t8 (cl (= @p_209 (! (is_res$a @p_30 @p_163) :named @p_210))) :rule cong :premises (t5.t2.t22.t6.t4 t5.t2.t22.t6.t7)) +(step t5.t2.t22.t6.t9 (cl @p_211) :rule refl) +(step t5.t2.t22.t6.t10 (cl @p_50) :rule refl) +(step t5.t2.t22.t6.t11 (cl (= @p_212 @p_130)) :rule cong :premises (t5.t2.t22.t6.t9 t5.t2.t22.t6.t10)) +(step t5.t2.t22.t6.t12 (cl @p_213) :rule refl) +(step t5.t2.t22.t6.t13 (cl (= @p_214 (! (run$b @p_130 veriT_vr10) :named @p_215))) :rule cong :premises (t5.t2.t22.t6.t11 t5.t2.t22.t6.t12)) +(step t5.t2.t22.t6.t14 (cl (! (= ?v3 veriT_vr9) :named @p_230)) :rule refl) +(step t5.t2.t22.t6.t15 (cl @p_85) :rule refl) +(step t5.t2.t22.t6.t16 (cl (! (= @p_23 @p_100) :named @p_231)) :rule cong :premises (t5.t2.t22.t6.t14 t5.t2.t22.t6.t15)) +(step t5.t2.t22.t6.t17 (cl (= @p_216 (! (is_res$b @p_215 @p_100) :named @p_217))) :rule cong :premises (t5.t2.t22.t6.t13 t5.t2.t22.t6.t16)) +(step t5.t2.t22.t6.t18 (cl (= @p_218 (! (and @p_210 @p_217) :named @p_219))) :rule cong :premises (t5.t2.t22.t6.t8 t5.t2.t22.t6.t17)) +(step t5.t2.t22.t6 (cl (= @p_207 (! (exists ((veriT_vr7 E$) (veriT_vr10 D$)) @p_219) :named @p_208))) :rule bind) +(step t5.t2.t22.t7 (cl (= @p_220 (! (or @p_73 @p_208) :named @p_221))) :rule cong :premises (t5.t2.t22.t5 t5.t2.t22.t6)) +(step t5.t2.t22.t8 (cl @p_42) :rule refl) +(step t5.t2.t22.t9 (cl @p_74) :rule cong :premises (t5.t2.t22.t8)) +(step t5.t2.t22.t10 (cl @p_72) :rule refl) +(step t5.t2.t22.t11 (cl @p_75) :rule cong :premises (t5.t2.t22.t9 t5.t2.t22.t10)) +(step t5.t2.t22.t12 (cl @p_76) :rule cong :premises (t5.t2.t22.t11)) +(anchor :step t5.t2.t22.t13 :args ((:= (?v5 E$) veriT_vr7) (:= (?v6 D$) veriT_vr10))) +(step t5.t2.t22.t13.t1 (cl @p_42) :rule refl) +(step t5.t2.t22.t13.t2 (cl @p_74) :rule cong :premises (t5.t2.t22.t13.t1)) +(step t5.t2.t22.t13.t3 (cl @p_72) :rule refl) +(step t5.t2.t22.t13.t4 (cl @p_75) :rule cong :premises (t5.t2.t22.t13.t2 t5.t2.t22.t13.t3)) +(step t5.t2.t22.t13.t5 (cl @p_211) :rule refl) +(step t5.t2.t22.t13.t6 (cl @p_213) :rule refl) +(step t5.t2.t22.t13.t7 (cl @p_224) :rule cong :premises (t5.t2.t22.t13.t5 t5.t2.t22.t13.t6)) +(step t5.t2.t22.t13.t8 (cl (= @p_225 (! (is_res$a @p_31 @p_163) :named @p_226))) :rule cong :premises (t5.t2.t22.t13.t4 t5.t2.t22.t13.t7)) +(step t5.t2.t22.t13.t9 (cl @p_211) :rule refl) +(step t5.t2.t22.t13.t10 (cl @p_42) :rule refl) +(step t5.t2.t22.t13.t11 (cl (= @p_227 @p_127)) :rule cong :premises (t5.t2.t22.t13.t9 t5.t2.t22.t13.t10)) +(step t5.t2.t22.t13.t12 (cl @p_213) :rule refl) +(step t5.t2.t22.t13.t13 (cl (= @p_228 (! (run$b @p_127 veriT_vr10) :named @p_229))) :rule cong :premises (t5.t2.t22.t13.t11 t5.t2.t22.t13.t12)) +(step t5.t2.t22.t13.t14 (cl @p_230) :rule refl) +(step t5.t2.t22.t13.t15 (cl @p_85) :rule refl) +(step t5.t2.t22.t13.t16 (cl @p_231) :rule cong :premises (t5.t2.t22.t13.t14 t5.t2.t22.t13.t15)) +(step t5.t2.t22.t13.t17 (cl (= @p_232 (! (is_res$b @p_229 @p_100) :named @p_233))) :rule cong :premises (t5.t2.t22.t13.t13 t5.t2.t22.t13.t16)) +(step t5.t2.t22.t13.t18 (cl (= @p_234 (! (and @p_226 @p_233) :named @p_235))) :rule cong :premises (t5.t2.t22.t13.t8 t5.t2.t22.t13.t17)) +(step t5.t2.t22.t13 (cl (= @p_222 (! (exists ((veriT_vr7 E$) (veriT_vr10 D$)) @p_235) :named @p_223))) :rule bind) +(step t5.t2.t22.t14 (cl (= @p_236 (! (or @p_29 @p_223) :named @p_237))) :rule cong :premises (t5.t2.t22.t12 t5.t2.t22.t13)) +(step t5.t2.t22.t15 (cl (= @p_238 (! (= @p_221 @p_237) :named @p_239))) :rule cong :premises (t5.t2.t22.t7 t5.t2.t22.t14)) +(step t5.t2.t22 (cl (= @p_205 (! (forall ((veriT_vr9 F$) (veriT_vr8 D$)) @p_239) :named @p_206))) :rule bind) +(step t5.t2.t23 (cl (= @p_240 (! (and @p_204 @p_206) :named @p_241))) :rule cong :premises (t5.t2.t21 t5.t2.t22)) +(step t5.t2.t24 (cl (= @p_242 (! (or @p_162 @p_241) :named @p_243))) :rule cong :premises (t5.t2.t6 t5.t2.t23)) +(step t5.t2.t25 (cl (= @p_244 (! (or @p_29 @p_243) :named @p_245))) :rule cong :premises (t5.t2.t5 t5.t2.t24)) +(step t5.t2 (cl (= @p_173 (! (forall ((veriT_vr6 D$)) @p_245) :named @p_174))) :rule bind) +(step t5.t3 (cl (= @p_246 (! (=> @p_33 @p_174) :named @p_247))) :rule cong :premises (t5.t1 t5.t2)) +(step t5 (cl (= @p_161 (! (forall ((veriT_vr0 A_b_c_M_state_fun$) (veriT_vr1 A_b_c_M_state_fun$)) @p_247) :named @p_250))) :rule bind) +(step t6 (cl (= @p_248 (! (=> @p_249 @p_250) :named @p_252))) :rule cong :premises (t4 t5)) +(step t7 (cl (! (= @p_251 (! (not @p_252) :named @p_254)) :named @p_253)) :rule cong :premises (t6)) +(step t8 (cl (! (not @p_253) :named @p_256) (! (not @p_251) :named @p_255) @p_254) :rule equiv_pos2) +(step t9 (cl (not @p_255) @p_248) :rule not_not) +(step t10 (cl @p_256 @p_248 @p_254) :rule th_resolution :premises (t9 t8)) +(step t11 (cl @p_254) :rule th_resolution :premises (axiom0 t7 t10)) +(step t12 (cl (! (= @p_254 (! (and @p_249 (! (not @p_250) :named @p_264)) :named @p_258)) :named @p_257)) :rule bool_simplify) +(step t13 (cl (! (not @p_257) :named @p_260) (! (not @p_254) :named @p_259) @p_258) :rule equiv_pos2) +(step t14 (cl (not @p_259) @p_252) :rule not_not) +(step t15 (cl @p_260 @p_252 @p_258) :rule th_resolution :premises (t14 t13)) +(step t16 (cl @p_258) :rule th_resolution :premises (t11 t12 t15)) +(anchor :step t17 :args ((veriT_vr0 A_b_c_M_state_fun$) (veriT_vr1 A_b_c_M_state_fun$))) +(anchor :step t17.t1 :args ((veriT_vr6 D$))) +(step t17.t1.t1 (cl (= @p_245 (! (or @p_29 @p_162 @p_241) :named @p_262))) :rule ac_simp) +(step t17.t1 (cl (= @p_174 (! (forall ((veriT_vr6 D$)) @p_262) :named @p_261))) :rule bind) +(step t17.t2 (cl (= @p_247 (! (=> @p_33 @p_261) :named @p_263))) :rule cong :premises (t17.t1)) +(step t17 (cl (= @p_250 (! (forall ((veriT_vr0 A_b_c_M_state_fun$) (veriT_vr1 A_b_c_M_state_fun$)) @p_263) :named @p_265))) :rule bind) +(step t18 (cl (= @p_264 (! (not @p_265) :named @p_266))) :rule cong :premises (t17)) +(step t19 (cl (! (= @p_258 (! (and @p_159 @p_160 @p_266) :named @p_268)) :named @p_267)) :rule ac_simp :premises (t18)) +(step t20 (cl (not @p_267) (not @p_258) @p_268) :rule equiv_pos2) +(step t21 (cl @p_268) :rule th_resolution :premises (t16 t19 t20)) +(anchor :step t22 :args ((:= (veriT_vr7 E$) veriT_vr11) (:= (veriT_vr0 A_b_c_M_state_fun$) veriT_vr12) (:= (veriT_vr1 A_b_c_M_state_fun$) veriT_vr13))) +(anchor :step t22.t1 :args ((:= (veriT_vr2 A$) veriT_vr14) (:= (veriT_vr3 C$) veriT_vr15))) +(step t22.t1.t1 (cl (! (= veriT_vr0 veriT_vr12) :named @p_283)) :rule refl) +(step t22.t1.t2 (cl (! (= veriT_vr2 veriT_vr14) :named @p_279)) :rule refl) +(step t22.t1.t3 (cl (! (= @p_35 (! (fun_app$ veriT_vr12 veriT_vr14) :named @p_278)) :named @p_284)) :rule cong :premises (t22.t1.t1 t22.t1.t2)) +(step t22.t1.t4 (cl (! (= veriT_vr3 veriT_vr15) :named @p_281)) :rule refl) +(step t22.t1.t5 (cl (! (= @p_27 (! (run$ @p_278 veriT_vr15) :named @p_271)) :named @p_285)) :rule cong :premises (t22.t1.t3 t22.t1.t4)) +(step t22.t1.t6 (cl (! (= @p_25 (! (is_fail$ @p_271) :named @p_269)) :named @p_286)) :rule cong :premises (t22.t1.t5)) +(step t22.t1.t7 (cl (! (= veriT_vr1 veriT_vr13) :named @p_289)) :rule refl) +(step t22.t1.t8 (cl @p_279) :rule refl) +(step t22.t1.t9 (cl (! (= @p_38 (! (fun_app$ veriT_vr13 veriT_vr14) :named @p_280)) :named @p_290)) :rule cong :premises (t22.t1.t7 t22.t1.t8)) +(step t22.t1.t10 (cl @p_281) :rule refl) +(step t22.t1.t11 (cl (! (= @p_26 (! (run$ @p_280 veriT_vr15) :named @p_270)) :named @p_291)) :rule cong :premises (t22.t1.t9 t22.t1.t10)) +(step t22.t1.t12 (cl (= @p_41 (! (is_fail$ @p_270) :named @p_282))) :rule cong :premises (t22.t1.t11)) +(step t22.t1.t13 (cl @p_283) :rule refl) +(step t22.t1.t14 (cl @p_279) :rule refl) +(step t22.t1.t15 (cl @p_284) :rule cong :premises (t22.t1.t13 t22.t1.t14)) +(step t22.t1.t16 (cl @p_281) :rule refl) +(step t22.t1.t17 (cl @p_285) :rule cong :premises (t22.t1.t15 t22.t1.t16)) +(step t22.t1.t18 (cl @p_286) :rule cong :premises (t22.t1.t17)) +(step t22.t1.t19 (cl (= @p_47 (! (= @p_282 @p_269) :named @p_287))) :rule cong :premises (t22.t1.t12 t22.t1.t18)) +(anchor :step t22.t1.t20 :args ((:= (veriT_vr4 B$) veriT_vr16) (:= (veriT_vr5 C$) veriT_vr17))) +(step t22.t1.t20.t1 (cl @p_289) :rule refl) +(step t22.t1.t20.t2 (cl @p_279) :rule refl) +(step t22.t1.t20.t3 (cl @p_290) :rule cong :premises (t22.t1.t20.t1 t22.t1.t20.t2)) +(step t22.t1.t20.t4 (cl @p_281) :rule refl) +(step t22.t1.t20.t5 (cl @p_291) :rule cong :premises (t22.t1.t20.t3 t22.t1.t20.t4)) +(step t22.t1.t20.t6 (cl (! (= veriT_vr4 veriT_vr16) :named @p_293)) :rule refl) +(step t22.t1.t20.t7 (cl (! (= veriT_vr5 veriT_vr17) :named @p_294)) :rule refl) +(step t22.t1.t20.t8 (cl (! (= @p_28 (! (pair$ veriT_vr16 veriT_vr17) :named @p_272)) :named @p_295)) :rule cong :premises (t22.t1.t20.t6 t22.t1.t20.t7)) +(step t22.t1.t20.t9 (cl (= @p_54 (! (is_res$ @p_270 @p_272) :named @p_292))) :rule cong :premises (t22.t1.t20.t5 t22.t1.t20.t8)) +(step t22.t1.t20.t10 (cl @p_283) :rule refl) +(step t22.t1.t20.t11 (cl @p_279) :rule refl) +(step t22.t1.t20.t12 (cl @p_284) :rule cong :premises (t22.t1.t20.t10 t22.t1.t20.t11)) +(step t22.t1.t20.t13 (cl @p_281) :rule refl) +(step t22.t1.t20.t14 (cl @p_285) :rule cong :premises (t22.t1.t20.t12 t22.t1.t20.t13)) +(step t22.t1.t20.t15 (cl @p_293) :rule refl) +(step t22.t1.t20.t16 (cl @p_294) :rule refl) +(step t22.t1.t20.t17 (cl @p_295) :rule cong :premises (t22.t1.t20.t15 t22.t1.t20.t16)) +(step t22.t1.t20.t18 (cl (= @p_59 (! (is_res$ @p_271 @p_272) :named @p_296))) :rule cong :premises (t22.t1.t20.t14 t22.t1.t20.t17)) +(step t22.t1.t20.t19 (cl (= @p_61 (! (= @p_292 @p_296) :named @p_297))) :rule cong :premises (t22.t1.t20.t9 t22.t1.t20.t18)) +(step t22.t1.t20 (cl (= @p_49 (! (forall ((veriT_vr16 B$) (veriT_vr17 C$)) @p_297) :named @p_288))) :rule bind) +(step t22.t1.t21 (cl (= @p_63 (! (and @p_287 @p_288) :named @p_298))) :rule cong :premises (t22.t1.t19 t22.t1.t20)) +(step t22.t1.t22 (cl (= @p_65 (! (or @p_269 @p_298) :named @p_299))) :rule cong :premises (t22.t1.t6 t22.t1.t21)) +(step t22.t1 (cl (= @p_33 (! (forall ((veriT_vr14 A$) (veriT_vr15 C$)) @p_299) :named @p_277))) :rule bind) +(anchor :step t22.t2 :args ((:= (veriT_vr6 D$) veriT_vr10))) +(step t22.t2.t1 (cl (! (= veriT_vr7 veriT_vr11) :named @p_302)) :rule refl) +(step t22.t2.t2 (cl @p_283) :rule refl) +(step t22.t2.t3 (cl (! (= @p_127 (! (c$ veriT_vr11 veriT_vr12) :named @p_301)) :named @p_306)) :rule cong :premises (t22.t2.t1 t22.t2.t2)) +(step t22.t2.t4 (cl (! (= veriT_vr6 veriT_vr10) :named @p_304)) :rule refl) +(step t22.t2.t5 (cl (! (= @p_99 (! (run$b @p_301 veriT_vr10) :named @p_275)) :named @p_307)) :rule cong :premises (t22.t2.t3 t22.t2.t4)) +(step t22.t2.t6 (cl (! (= @p_97 (! (is_fail$b @p_275) :named @p_273)) :named @p_308)) :rule cong :premises (t22.t2.t5)) +(step t22.t2.t7 (cl @p_302) :rule refl) +(step t22.t2.t8 (cl @p_289) :rule refl) +(step t22.t2.t9 (cl (! (= @p_130 (! (c$ veriT_vr11 veriT_vr13) :named @p_303)) :named @p_311)) :rule cong :premises (t22.t2.t7 t22.t2.t8)) +(step t22.t2.t10 (cl @p_304) :rule refl) +(step t22.t2.t11 (cl (! (= @p_98 (! (run$b @p_303 veriT_vr10) :named @p_274)) :named @p_312)) :rule cong :premises (t22.t2.t9 t22.t2.t10)) +(step t22.t2.t12 (cl (= @p_133 (! (is_fail$b @p_274) :named @p_305))) :rule cong :premises (t22.t2.t11)) +(step t22.t2.t13 (cl @p_302) :rule refl) +(step t22.t2.t14 (cl @p_283) :rule refl) +(step t22.t2.t15 (cl @p_306) :rule cong :premises (t22.t2.t13 t22.t2.t14)) +(step t22.t2.t16 (cl @p_304) :rule refl) +(step t22.t2.t17 (cl @p_307) :rule cong :premises (t22.t2.t15 t22.t2.t16)) +(step t22.t2.t18 (cl @p_308) :rule cong :premises (t22.t2.t17)) +(step t22.t2.t19 (cl (= @p_138 (! (= @p_305 @p_273) :named @p_309))) :rule cong :premises (t22.t2.t12 t22.t2.t18)) +(anchor :step t22.t2.t20 :args ((:= (veriT_vr9 F$) veriT_vr9) (:= (veriT_vr8 D$) veriT_vr18))) +(step t22.t2.t20.t1 (cl @p_302) :rule refl) +(step t22.t2.t20.t2 (cl @p_289) :rule refl) +(step t22.t2.t20.t3 (cl @p_311) :rule cong :premises (t22.t2.t20.t1 t22.t2.t20.t2)) +(step t22.t2.t20.t4 (cl @p_304) :rule refl) +(step t22.t2.t20.t5 (cl @p_312) :rule cong :premises (t22.t2.t20.t3 t22.t2.t20.t4)) +(step t22.t2.t20.t6 (cl (! (= veriT_vr8 veriT_vr18) :named @p_314)) :rule refl) +(step t22.t2.t20.t7 (cl (! (= @p_100 (! (pair$b veriT_vr9 veriT_vr18) :named @p_276)) :named @p_315)) :rule cong :premises (t22.t2.t20.t6)) +(step t22.t2.t20.t8 (cl (= @p_144 (! (is_res$b @p_274 @p_276) :named @p_313))) :rule cong :premises (t22.t2.t20.t5 t22.t2.t20.t7)) +(step t22.t2.t20.t9 (cl @p_302) :rule refl) +(step t22.t2.t20.t10 (cl @p_283) :rule refl) +(step t22.t2.t20.t11 (cl @p_306) :rule cong :premises (t22.t2.t20.t9 t22.t2.t20.t10)) +(step t22.t2.t20.t12 (cl @p_304) :rule refl) +(step t22.t2.t20.t13 (cl @p_307) :rule cong :premises (t22.t2.t20.t11 t22.t2.t20.t12)) +(step t22.t2.t20.t14 (cl @p_314) :rule refl) +(step t22.t2.t20.t15 (cl @p_315) :rule cong :premises (t22.t2.t20.t14)) +(step t22.t2.t20.t16 (cl (= @p_149 (! (is_res$b @p_275 @p_276) :named @p_316))) :rule cong :premises (t22.t2.t20.t13 t22.t2.t20.t15)) +(step t22.t2.t20.t17 (cl (= @p_151 (! (= @p_313 @p_316) :named @p_317))) :rule cong :premises (t22.t2.t20.t8 t22.t2.t20.t16)) +(step t22.t2.t20 (cl (= @p_140 (! (forall ((veriT_vr9 F$) (veriT_vr18 D$)) @p_317) :named @p_310))) :rule bind) +(step t22.t2.t21 (cl (= @p_153 (! (and @p_309 @p_310) :named @p_318))) :rule cong :premises (t22.t2.t19 t22.t2.t20)) +(step t22.t2.t22 (cl (= @p_155 (! (or @p_273 @p_318) :named @p_319))) :rule cong :premises (t22.t2.t6 t22.t2.t21)) +(step t22.t2 (cl (= @p_125 (! (forall ((veriT_vr10 D$)) @p_319) :named @p_300))) :rule bind) +(step t22.t3 (cl (= @p_157 (! (=> @p_277 @p_300) :named @p_320))) :rule cong :premises (t22.t1 t22.t2)) +(step t22 (cl (= @p_160 (! (forall ((veriT_vr11 E$) (veriT_vr12 A_b_c_M_state_fun$) (veriT_vr13 A_b_c_M_state_fun$)) @p_320) :named @p_422))) :rule bind) +(anchor :step t23 :args ((:= (veriT_vr0 A_b_c_M_state_fun$) veriT_vr19) (:= (veriT_vr1 A_b_c_M_state_fun$) veriT_vr20))) +(anchor :step t23.t1 :args ((:= (veriT_vr2 A$) veriT_vr21) (:= (veriT_vr3 C$) veriT_vr22))) +(step t23.t1.t1 (cl (! (= veriT_vr0 veriT_vr19) :named @p_336)) :rule refl) +(step t23.t1.t2 (cl (! (= veriT_vr2 veriT_vr21) :named @p_332)) :rule refl) +(step t23.t1.t3 (cl (! (= @p_35 (! (fun_app$ veriT_vr19 veriT_vr21) :named @p_331)) :named @p_337)) :rule cong :premises (t23.t1.t1 t23.t1.t2)) +(step t23.t1.t4 (cl (! (= veriT_vr3 veriT_vr22) :named @p_334)) :rule refl) +(step t23.t1.t5 (cl (! (= @p_27 (! (run$ @p_331 veriT_vr22) :named @p_323)) :named @p_338)) :rule cong :premises (t23.t1.t3 t23.t1.t4)) +(step t23.t1.t6 (cl (! (= @p_25 (! (is_fail$ @p_323) :named @p_321)) :named @p_339)) :rule cong :premises (t23.t1.t5)) +(step t23.t1.t7 (cl (! (= veriT_vr1 veriT_vr20) :named @p_342)) :rule refl) +(step t23.t1.t8 (cl @p_332) :rule refl) +(step t23.t1.t9 (cl (! (= @p_38 (! (fun_app$ veriT_vr20 veriT_vr21) :named @p_333)) :named @p_343)) :rule cong :premises (t23.t1.t7 t23.t1.t8)) +(step t23.t1.t10 (cl @p_334) :rule refl) +(step t23.t1.t11 (cl (! (= @p_26 (! (run$ @p_333 veriT_vr22) :named @p_322)) :named @p_344)) :rule cong :premises (t23.t1.t9 t23.t1.t10)) +(step t23.t1.t12 (cl (= @p_41 (! (is_fail$ @p_322) :named @p_335))) :rule cong :premises (t23.t1.t11)) +(step t23.t1.t13 (cl @p_336) :rule refl) +(step t23.t1.t14 (cl @p_332) :rule refl) +(step t23.t1.t15 (cl @p_337) :rule cong :premises (t23.t1.t13 t23.t1.t14)) +(step t23.t1.t16 (cl @p_334) :rule refl) +(step t23.t1.t17 (cl @p_338) :rule cong :premises (t23.t1.t15 t23.t1.t16)) +(step t23.t1.t18 (cl @p_339) :rule cong :premises (t23.t1.t17)) +(step t23.t1.t19 (cl (= @p_47 (! (= @p_335 @p_321) :named @p_340))) :rule cong :premises (t23.t1.t12 t23.t1.t18)) +(anchor :step t23.t1.t20 :args ((:= (veriT_vr4 B$) veriT_vr23) (:= (veriT_vr5 C$) veriT_vr24))) +(step t23.t1.t20.t1 (cl @p_342) :rule refl) +(step t23.t1.t20.t2 (cl @p_332) :rule refl) +(step t23.t1.t20.t3 (cl @p_343) :rule cong :premises (t23.t1.t20.t1 t23.t1.t20.t2)) +(step t23.t1.t20.t4 (cl @p_334) :rule refl) +(step t23.t1.t20.t5 (cl @p_344) :rule cong :premises (t23.t1.t20.t3 t23.t1.t20.t4)) +(step t23.t1.t20.t6 (cl (! (= veriT_vr4 veriT_vr23) :named @p_346)) :rule refl) +(step t23.t1.t20.t7 (cl (! (= veriT_vr5 veriT_vr24) :named @p_347)) :rule refl) +(step t23.t1.t20.t8 (cl (! (= @p_28 (! (pair$ veriT_vr23 veriT_vr24) :named @p_324)) :named @p_348)) :rule cong :premises (t23.t1.t20.t6 t23.t1.t20.t7)) +(step t23.t1.t20.t9 (cl (= @p_54 (! (is_res$ @p_322 @p_324) :named @p_345))) :rule cong :premises (t23.t1.t20.t5 t23.t1.t20.t8)) +(step t23.t1.t20.t10 (cl @p_336) :rule refl) +(step t23.t1.t20.t11 (cl @p_332) :rule refl) +(step t23.t1.t20.t12 (cl @p_337) :rule cong :premises (t23.t1.t20.t10 t23.t1.t20.t11)) +(step t23.t1.t20.t13 (cl @p_334) :rule refl) +(step t23.t1.t20.t14 (cl @p_338) :rule cong :premises (t23.t1.t20.t12 t23.t1.t20.t13)) +(step t23.t1.t20.t15 (cl @p_346) :rule refl) +(step t23.t1.t20.t16 (cl @p_347) :rule refl) +(step t23.t1.t20.t17 (cl @p_348) :rule cong :premises (t23.t1.t20.t15 t23.t1.t20.t16)) +(step t23.t1.t20.t18 (cl (= @p_59 (! (is_res$ @p_323 @p_324) :named @p_349))) :rule cong :premises (t23.t1.t20.t14 t23.t1.t20.t17)) +(step t23.t1.t20.t19 (cl (= @p_61 (! (= @p_345 @p_349) :named @p_350))) :rule cong :premises (t23.t1.t20.t9 t23.t1.t20.t18)) +(step t23.t1.t20 (cl (= @p_49 (! (forall ((veriT_vr23 B$) (veriT_vr24 C$)) @p_350) :named @p_341))) :rule bind) +(step t23.t1.t21 (cl (= @p_63 (! (and @p_340 @p_341) :named @p_351))) :rule cong :premises (t23.t1.t19 t23.t1.t20)) +(step t23.t1.t22 (cl (= @p_65 (! (or @p_321 @p_351) :named @p_352))) :rule cong :premises (t23.t1.t6 t23.t1.t21)) +(step t23.t1 (cl (= @p_33 (! (forall ((veriT_vr21 A$) (veriT_vr22 C$)) @p_352) :named @p_330))) :rule bind) +(anchor :step t23.t2 :args ((:= (veriT_vr6 D$) veriT_vr25))) +(step t23.t2.t1 (cl @p_336) :rule refl) +(step t23.t2.t2 (cl (! (= @p_69 (! (b$ veriT_vr19) :named @p_354)) :named @p_356)) :rule cong :premises (t23.t2.t1)) +(step t23.t2.t3 (cl (! (= veriT_vr6 veriT_vr25) :named @p_357)) :rule refl) +(step t23.t2.t4 (cl (! (= @p_31 (! (run$a @p_354 veriT_vr25) :named @p_325)) :named @p_358)) :rule cong :premises (t23.t2.t2 t23.t2.t3)) +(step t23.t2.t5 (cl (! (= @p_29 (! (is_fail$a @p_325) :named @p_327)) :named @p_380)) :rule cong :premises (t23.t2.t4)) +(anchor :step t23.t2.t6 :args ((:= (veriT_vr7 E$) veriT_vr26) (:= (veriT_vr8 D$) veriT_vr27))) +(step t23.t2.t6.t1 (cl @p_336) :rule refl) +(step t23.t2.t6.t2 (cl @p_356) :rule cong :premises (t23.t2.t6.t1)) +(step t23.t2.t6.t3 (cl @p_357) :rule refl) +(step t23.t2.t6.t4 (cl @p_358) :rule cong :premises (t23.t2.t6.t2 t23.t2.t6.t3)) +(step t23.t2.t6.t5 (cl (! (= veriT_vr7 veriT_vr26) :named @p_361)) :rule refl) +(step t23.t2.t6.t6 (cl (! (= veriT_vr8 veriT_vr27) :named @p_363)) :rule refl) +(step t23.t2.t6.t7 (cl (= @p_32 (! (pair$a veriT_vr26 veriT_vr27) :named @p_359))) :rule cong :premises (t23.t2.t6.t5 t23.t2.t6.t6)) +(step t23.t2.t6.t8 (cl (= @p_87 (! (is_res$a @p_325 @p_359) :named @p_360))) :rule cong :premises (t23.t2.t6.t4 t23.t2.t6.t7)) +(step t23.t2.t6.t9 (cl @p_361) :rule refl) +(step t23.t2.t6.t10 (cl @p_336) :rule refl) +(step t23.t2.t6.t11 (cl (= @p_127 (! (c$ veriT_vr26 veriT_vr19) :named @p_362))) :rule cong :premises (t23.t2.t6.t9 t23.t2.t6.t10)) +(step t23.t2.t6.t12 (cl @p_363) :rule refl) +(step t23.t2.t6.t13 (cl (= @p_178 (! (run$b @p_362 veriT_vr27) :named @p_364))) :rule cong :premises (t23.t2.t6.t11 t23.t2.t6.t12)) +(step t23.t2.t6.t14 (cl (= @p_180 (! (is_fail$b @p_364) :named @p_365))) :rule cong :premises (t23.t2.t6.t13)) +(step t23.t2.t6.t15 (cl (= @p_182 (! (and @p_360 @p_365) :named @p_366))) :rule cong :premises (t23.t2.t6.t8 t23.t2.t6.t14)) +(step t23.t2.t6 (cl (= @p_162 (! (exists ((veriT_vr26 E$) (veriT_vr27 D$)) @p_366) :named @p_355))) :rule bind) +(step t23.t2.t7 (cl @p_342) :rule refl) +(step t23.t2.t8 (cl (! (= @p_71 (! (b$ veriT_vr20) :named @p_367)) :named @p_369)) :rule cong :premises (t23.t2.t7)) +(step t23.t2.t9 (cl @p_357) :rule refl) +(step t23.t2.t10 (cl (! (= @p_30 (! (run$a @p_367 veriT_vr25) :named @p_326)) :named @p_370)) :rule cong :premises (t23.t2.t8 t23.t2.t9)) +(step t23.t2.t11 (cl (! (= @p_73 (! (is_fail$a @p_326) :named @p_328)) :named @p_393)) :rule cong :premises (t23.t2.t10)) +(anchor :step t23.t2.t12 :args ((:= (veriT_vr7 E$) veriT_vr28) (:= (veriT_vr8 D$) veriT_vr29))) +(step t23.t2.t12.t1 (cl @p_342) :rule refl) +(step t23.t2.t12.t2 (cl @p_369) :rule cong :premises (t23.t2.t12.t1)) +(step t23.t2.t12.t3 (cl @p_357) :rule refl) +(step t23.t2.t12.t4 (cl @p_370) :rule cong :premises (t23.t2.t12.t2 t23.t2.t12.t3)) +(step t23.t2.t12.t5 (cl (! (= veriT_vr7 veriT_vr28) :named @p_373)) :rule refl) +(step t23.t2.t12.t6 (cl (! (= veriT_vr8 veriT_vr29) :named @p_375)) :rule refl) +(step t23.t2.t12.t7 (cl (= @p_32 (! (pair$a veriT_vr28 veriT_vr29) :named @p_371))) :rule cong :premises (t23.t2.t12.t5 t23.t2.t12.t6)) +(step t23.t2.t12.t8 (cl (= @p_83 (! (is_res$a @p_326 @p_371) :named @p_372))) :rule cong :premises (t23.t2.t12.t4 t23.t2.t12.t7)) +(step t23.t2.t12.t9 (cl @p_373) :rule refl) +(step t23.t2.t12.t10 (cl @p_342) :rule refl) +(step t23.t2.t12.t11 (cl (= @p_130 (! (c$ veriT_vr28 veriT_vr20) :named @p_374))) :rule cong :premises (t23.t2.t12.t9 t23.t2.t12.t10)) +(step t23.t2.t12.t12 (cl @p_375) :rule refl) +(step t23.t2.t12.t13 (cl (= @p_189 (! (run$b @p_374 veriT_vr29) :named @p_376))) :rule cong :premises (t23.t2.t12.t11 t23.t2.t12.t12)) +(step t23.t2.t12.t14 (cl (= @p_191 (! (is_fail$b @p_376) :named @p_377))) :rule cong :premises (t23.t2.t12.t13)) +(step t23.t2.t12.t15 (cl (= @p_193 (! (and @p_372 @p_377) :named @p_378))) :rule cong :premises (t23.t2.t12.t8 t23.t2.t12.t14)) +(step t23.t2.t12 (cl (= @p_185 (! (exists ((veriT_vr28 E$) (veriT_vr29 D$)) @p_378) :named @p_368))) :rule bind) +(step t23.t2.t13 (cl (= @p_195 (! (or @p_328 @p_368) :named @p_379))) :rule cong :premises (t23.t2.t11 t23.t2.t12)) +(step t23.t2.t14 (cl @p_336) :rule refl) +(step t23.t2.t15 (cl @p_356) :rule cong :premises (t23.t2.t14)) +(step t23.t2.t16 (cl @p_357) :rule refl) +(step t23.t2.t17 (cl @p_358) :rule cong :premises (t23.t2.t15 t23.t2.t16)) +(step t23.t2.t18 (cl @p_380) :rule cong :premises (t23.t2.t17)) +(anchor :step t23.t2.t19 :args ((:= (veriT_vr7 E$) veriT_vr30) (:= (veriT_vr8 D$) veriT_vr31))) +(step t23.t2.t19.t1 (cl @p_336) :rule refl) +(step t23.t2.t19.t2 (cl @p_356) :rule cong :premises (t23.t2.t19.t1)) +(step t23.t2.t19.t3 (cl @p_357) :rule refl) +(step t23.t2.t19.t4 (cl @p_358) :rule cong :premises (t23.t2.t19.t2 t23.t2.t19.t3)) +(step t23.t2.t19.t5 (cl (! (= veriT_vr7 veriT_vr30) :named @p_384)) :rule refl) +(step t23.t2.t19.t6 (cl (! (= veriT_vr8 veriT_vr31) :named @p_386)) :rule refl) +(step t23.t2.t19.t7 (cl (= @p_32 (! (pair$a veriT_vr30 veriT_vr31) :named @p_382))) :rule cong :premises (t23.t2.t19.t5 t23.t2.t19.t6)) +(step t23.t2.t19.t8 (cl (= @p_87 (! (is_res$a @p_325 @p_382) :named @p_383))) :rule cong :premises (t23.t2.t19.t4 t23.t2.t19.t7)) +(step t23.t2.t19.t9 (cl @p_384) :rule refl) +(step t23.t2.t19.t10 (cl @p_336) :rule refl) +(step t23.t2.t19.t11 (cl (= @p_127 (! (c$ veriT_vr30 veriT_vr19) :named @p_385))) :rule cong :premises (t23.t2.t19.t9 t23.t2.t19.t10)) +(step t23.t2.t19.t12 (cl @p_386) :rule refl) +(step t23.t2.t19.t13 (cl (= @p_178 (! (run$b @p_385 veriT_vr31) :named @p_387))) :rule cong :premises (t23.t2.t19.t11 t23.t2.t19.t12)) +(step t23.t2.t19.t14 (cl (= @p_180 (! (is_fail$b @p_387) :named @p_388))) :rule cong :premises (t23.t2.t19.t13)) +(step t23.t2.t19.t15 (cl (= @p_182 (! (and @p_383 @p_388) :named @p_389))) :rule cong :premises (t23.t2.t19.t8 t23.t2.t19.t14)) +(step t23.t2.t19 (cl (= @p_162 (! (exists ((veriT_vr30 E$) (veriT_vr31 D$)) @p_389) :named @p_381))) :rule bind) +(step t23.t2.t20 (cl (= @p_202 (! (or @p_327 @p_381) :named @p_390))) :rule cong :premises (t23.t2.t18 t23.t2.t19)) +(step t23.t2.t21 (cl (= @p_204 (! (= @p_379 @p_390) :named @p_391))) :rule cong :premises (t23.t2.t13 t23.t2.t20)) +(anchor :step t23.t2.t22 :args ((:= (veriT_vr9 F$) veriT_vr32) (:= (veriT_vr8 D$) veriT_vr33))) +(step t23.t2.t22.t1 (cl @p_342) :rule refl) +(step t23.t2.t22.t2 (cl @p_369) :rule cong :premises (t23.t2.t22.t1)) +(step t23.t2.t22.t3 (cl @p_357) :rule refl) +(step t23.t2.t22.t4 (cl @p_370) :rule cong :premises (t23.t2.t22.t2 t23.t2.t22.t3)) +(step t23.t2.t22.t5 (cl @p_393) :rule cong :premises (t23.t2.t22.t4)) +(anchor :step t23.t2.t22.t6 :args ((:= (veriT_vr7 E$) veriT_vr34) (:= (veriT_vr10 D$) veriT_vr35))) +(step t23.t2.t22.t6.t1 (cl @p_342) :rule refl) +(step t23.t2.t22.t6.t2 (cl @p_369) :rule cong :premises (t23.t2.t22.t6.t1)) +(step t23.t2.t22.t6.t3 (cl @p_357) :rule refl) +(step t23.t2.t22.t6.t4 (cl @p_370) :rule cong :premises (t23.t2.t22.t6.t2 t23.t2.t22.t6.t3)) +(step t23.t2.t22.t6.t5 (cl (! (= veriT_vr7 veriT_vr34) :named @p_397)) :rule refl) +(step t23.t2.t22.t6.t6 (cl (! (= veriT_vr10 veriT_vr35) :named @p_399)) :rule refl) +(step t23.t2.t22.t6.t7 (cl (= @p_163 (! (pair$a veriT_vr34 veriT_vr35) :named @p_395))) :rule cong :premises (t23.t2.t22.t6.t5 t23.t2.t22.t6.t6)) +(step t23.t2.t22.t6.t8 (cl (= @p_210 (! (is_res$a @p_326 @p_395) :named @p_396))) :rule cong :premises (t23.t2.t22.t6.t4 t23.t2.t22.t6.t7)) +(step t23.t2.t22.t6.t9 (cl @p_397) :rule refl) +(step t23.t2.t22.t6.t10 (cl @p_342) :rule refl) +(step t23.t2.t22.t6.t11 (cl (= @p_130 (! (c$ veriT_vr34 veriT_vr20) :named @p_398))) :rule cong :premises (t23.t2.t22.t6.t9 t23.t2.t22.t6.t10)) +(step t23.t2.t22.t6.t12 (cl @p_399) :rule refl) +(step t23.t2.t22.t6.t13 (cl (= @p_215 (! (run$b @p_398 veriT_vr35) :named @p_400))) :rule cong :premises (t23.t2.t22.t6.t11 t23.t2.t22.t6.t12)) +(step t23.t2.t22.t6.t14 (cl (! (= veriT_vr9 veriT_vr32) :named @p_411)) :rule refl) +(step t23.t2.t22.t6.t15 (cl (! (= veriT_vr8 veriT_vr33) :named @p_412)) :rule refl) +(step t23.t2.t22.t6.t16 (cl (! (= @p_100 (! (pair$b veriT_vr32 veriT_vr33) :named @p_329)) :named @p_413)) :rule cong :premises (t23.t2.t22.t6.t14 t23.t2.t22.t6.t15)) +(step t23.t2.t22.t6.t17 (cl (= @p_217 (! (is_res$b @p_400 @p_329) :named @p_401))) :rule cong :premises (t23.t2.t22.t6.t13 t23.t2.t22.t6.t16)) +(step t23.t2.t22.t6.t18 (cl (= @p_219 (! (and @p_396 @p_401) :named @p_402))) :rule cong :premises (t23.t2.t22.t6.t8 t23.t2.t22.t6.t17)) +(step t23.t2.t22.t6 (cl (= @p_208 (! (exists ((veriT_vr34 E$) (veriT_vr35 D$)) @p_402) :named @p_394))) :rule bind) +(step t23.t2.t22.t7 (cl (= @p_221 (! (or @p_328 @p_394) :named @p_403))) :rule cong :premises (t23.t2.t22.t5 t23.t2.t22.t6)) +(step t23.t2.t22.t8 (cl @p_336) :rule refl) +(step t23.t2.t22.t9 (cl @p_356) :rule cong :premises (t23.t2.t22.t8)) +(step t23.t2.t22.t10 (cl @p_357) :rule refl) +(step t23.t2.t22.t11 (cl @p_358) :rule cong :premises (t23.t2.t22.t9 t23.t2.t22.t10)) +(step t23.t2.t22.t12 (cl @p_380) :rule cong :premises (t23.t2.t22.t11)) +(anchor :step t23.t2.t22.t13 :args ((:= (veriT_vr7 E$) veriT_vr36) (:= (veriT_vr10 D$) veriT_vr37))) +(step t23.t2.t22.t13.t1 (cl @p_336) :rule refl) +(step t23.t2.t22.t13.t2 (cl @p_356) :rule cong :premises (t23.t2.t22.t13.t1)) +(step t23.t2.t22.t13.t3 (cl @p_357) :rule refl) +(step t23.t2.t22.t13.t4 (cl @p_358) :rule cong :premises (t23.t2.t22.t13.t2 t23.t2.t22.t13.t3)) +(step t23.t2.t22.t13.t5 (cl (! (= veriT_vr7 veriT_vr36) :named @p_407)) :rule refl) +(step t23.t2.t22.t13.t6 (cl (! (= veriT_vr10 veriT_vr37) :named @p_409)) :rule refl) +(step t23.t2.t22.t13.t7 (cl (= @p_163 (! (pair$a veriT_vr36 veriT_vr37) :named @p_405))) :rule cong :premises (t23.t2.t22.t13.t5 t23.t2.t22.t13.t6)) +(step t23.t2.t22.t13.t8 (cl (= @p_226 (! (is_res$a @p_325 @p_405) :named @p_406))) :rule cong :premises (t23.t2.t22.t13.t4 t23.t2.t22.t13.t7)) +(step t23.t2.t22.t13.t9 (cl @p_407) :rule refl) +(step t23.t2.t22.t13.t10 (cl @p_336) :rule refl) +(step t23.t2.t22.t13.t11 (cl (= @p_127 (! (c$ veriT_vr36 veriT_vr19) :named @p_408))) :rule cong :premises (t23.t2.t22.t13.t9 t23.t2.t22.t13.t10)) +(step t23.t2.t22.t13.t12 (cl @p_409) :rule refl) +(step t23.t2.t22.t13.t13 (cl (= @p_229 (! (run$b @p_408 veriT_vr37) :named @p_410))) :rule cong :premises (t23.t2.t22.t13.t11 t23.t2.t22.t13.t12)) +(step t23.t2.t22.t13.t14 (cl @p_411) :rule refl) +(step t23.t2.t22.t13.t15 (cl @p_412) :rule refl) +(step t23.t2.t22.t13.t16 (cl @p_413) :rule cong :premises (t23.t2.t22.t13.t14 t23.t2.t22.t13.t15)) +(step t23.t2.t22.t13.t17 (cl (= @p_233 (! (is_res$b @p_410 @p_329) :named @p_414))) :rule cong :premises (t23.t2.t22.t13.t13 t23.t2.t22.t13.t16)) +(step t23.t2.t22.t13.t18 (cl (= @p_235 (! (and @p_406 @p_414) :named @p_415))) :rule cong :premises (t23.t2.t22.t13.t8 t23.t2.t22.t13.t17)) +(step t23.t2.t22.t13 (cl (= @p_223 (! (exists ((veriT_vr36 E$) (veriT_vr37 D$)) @p_415) :named @p_404))) :rule bind) +(step t23.t2.t22.t14 (cl (= @p_237 (! (or @p_327 @p_404) :named @p_416))) :rule cong :premises (t23.t2.t22.t12 t23.t2.t22.t13)) +(step t23.t2.t22.t15 (cl (= @p_239 (! (= @p_403 @p_416) :named @p_417))) :rule cong :premises (t23.t2.t22.t7 t23.t2.t22.t14)) +(step t23.t2.t22 (cl (= @p_206 (! (forall ((veriT_vr32 F$) (veriT_vr33 D$)) @p_417) :named @p_392))) :rule bind) +(step t23.t2.t23 (cl (= @p_241 (! (and @p_391 @p_392) :named @p_418))) :rule cong :premises (t23.t2.t21 t23.t2.t22)) +(step t23.t2.t24 (cl (= @p_262 (! (or @p_327 @p_355 @p_418) :named @p_419))) :rule cong :premises (t23.t2.t5 t23.t2.t6 t23.t2.t23)) +(step t23.t2 (cl (= @p_261 (! (forall ((veriT_vr25 D$)) @p_419) :named @p_353))) :rule bind) +(step t23.t3 (cl (= @p_263 (! (=> @p_330 @p_353) :named @p_420))) :rule cong :premises (t23.t1 t23.t2)) +(step t23 (cl (= @p_265 (! (forall ((veriT_vr19 A_b_c_M_state_fun$) (veriT_vr20 A_b_c_M_state_fun$)) @p_420) :named @p_421))) :rule bind) +(step t24 (cl (= @p_266 (! (not @p_421) :named @p_423))) :rule cong :premises (t23)) +(step t25 (cl (! (= @p_268 (! (and @p_159 @p_422 @p_423) :named @p_425)) :named @p_424)) :rule cong :premises (t22 t24)) +(step t26 (cl (not @p_424) (not @p_268) @p_425) :rule equiv_pos2) +(step t27 (cl @p_425) :rule th_resolution :premises (t21 t25 t26)) +(anchor :step t28 :args ((veriT_vr19 A_b_c_M_state_fun$) (veriT_vr20 A_b_c_M_state_fun$))) +(anchor :step t28.t1 :args ((veriT_vr25 D$))) +(step t28.t1.t1 (cl (= @p_391 (! (and (! (=> @p_379 @p_390) :named @p_612) (! (=> @p_390 @p_379) :named @p_635)) :named @p_427))) :rule connective_def) +(anchor :step t28.t1.t2 :args ((veriT_vr32 F$) (veriT_vr33 D$))) +(step t28.t1.t2.t1 (cl (= @p_417 (! (and (! (=> @p_403 @p_416) :named @p_662) (! (=> @p_416 @p_403) :named @p_684)) :named @p_429))) :rule connective_def) +(step t28.t1.t2 (cl (= @p_392 (! (forall ((veriT_vr32 F$) (veriT_vr33 D$)) @p_429) :named @p_428))) :rule bind) +(step t28.t1.t3 (cl (= @p_418 (! (and @p_427 @p_428) :named @p_430))) :rule cong :premises (t28.t1.t1 t28.t1.t2)) +(step t28.t1.t4 (cl (= @p_419 (! (or @p_327 @p_355 @p_430) :named @p_431))) :rule cong :premises (t28.t1.t3)) +(step t28.t1 (cl (= @p_353 (! (forall ((veriT_vr25 D$)) @p_431) :named @p_426))) :rule bind) +(step t28.t2 (cl (= @p_420 (! (=> @p_330 @p_426) :named @p_432))) :rule cong :premises (t28.t1)) +(step t28 (cl (= @p_421 (! (forall ((veriT_vr19 A_b_c_M_state_fun$) (veriT_vr20 A_b_c_M_state_fun$)) @p_432) :named @p_433))) :rule bind) +(step t29 (cl (= @p_423 (! (not @p_433) :named @p_434))) :rule cong :premises (t28)) +(step t30 (cl (! (= @p_425 (! (and @p_159 @p_422 @p_434) :named @p_436)) :named @p_435)) :rule cong :premises (t29)) +(step t31 (cl (not @p_435) (not @p_425) @p_436) :rule equiv_pos2) +(step t32 (cl @p_436) :rule th_resolution :premises (t27 t30 t31)) +(anchor :step t33 :args ((:= (veriT_vr0 A_b_c_M_state_fun$) veriT_vr38) (:= (veriT_vr1 A_b_c_M_state_fun$) veriT_vr39))) +(anchor :step t33.t1 :args ((:= (veriT_vr2 A$) veriT_vr40) (:= (veriT_vr3 C$) veriT_vr41))) +(step t33.t1.t1 (cl (! (= veriT_vr0 veriT_vr38) :named @p_451)) :rule refl) +(step t33.t1.t2 (cl (! (= veriT_vr2 veriT_vr40) :named @p_447)) :rule refl) +(step t33.t1.t3 (cl (! (= @p_35 (! (fun_app$ veriT_vr38 veriT_vr40) :named @p_446)) :named @p_452)) :rule cong :premises (t33.t1.t1 t33.t1.t2)) +(step t33.t1.t4 (cl (! (= veriT_vr3 veriT_vr41) :named @p_449)) :rule refl) +(step t33.t1.t5 (cl (! (= @p_27 (! (run$ @p_446 veriT_vr41) :named @p_439)) :named @p_453)) :rule cong :premises (t33.t1.t3 t33.t1.t4)) +(step t33.t1.t6 (cl (! (= @p_25 (! (is_fail$ @p_439) :named @p_437)) :named @p_454)) :rule cong :premises (t33.t1.t5)) +(step t33.t1.t7 (cl (! (= veriT_vr1 veriT_vr39) :named @p_457)) :rule refl) +(step t33.t1.t8 (cl @p_447) :rule refl) +(step t33.t1.t9 (cl (! (= @p_38 (! (fun_app$ veriT_vr39 veriT_vr40) :named @p_448)) :named @p_458)) :rule cong :premises (t33.t1.t7 t33.t1.t8)) +(step t33.t1.t10 (cl @p_449) :rule refl) +(step t33.t1.t11 (cl (! (= @p_26 (! (run$ @p_448 veriT_vr41) :named @p_438)) :named @p_459)) :rule cong :premises (t33.t1.t9 t33.t1.t10)) +(step t33.t1.t12 (cl (= @p_41 (! (is_fail$ @p_438) :named @p_450))) :rule cong :premises (t33.t1.t11)) +(step t33.t1.t13 (cl @p_451) :rule refl) +(step t33.t1.t14 (cl @p_447) :rule refl) +(step t33.t1.t15 (cl @p_452) :rule cong :premises (t33.t1.t13 t33.t1.t14)) +(step t33.t1.t16 (cl @p_449) :rule refl) +(step t33.t1.t17 (cl @p_453) :rule cong :premises (t33.t1.t15 t33.t1.t16)) +(step t33.t1.t18 (cl @p_454) :rule cong :premises (t33.t1.t17)) +(step t33.t1.t19 (cl (= @p_47 (! (= @p_450 @p_437) :named @p_455))) :rule cong :premises (t33.t1.t12 t33.t1.t18)) +(anchor :step t33.t1.t20 :args ((:= (veriT_vr4 B$) veriT_vr42) (:= (veriT_vr5 C$) veriT_vr43))) +(step t33.t1.t20.t1 (cl @p_457) :rule refl) +(step t33.t1.t20.t2 (cl @p_447) :rule refl) +(step t33.t1.t20.t3 (cl @p_458) :rule cong :premises (t33.t1.t20.t1 t33.t1.t20.t2)) +(step t33.t1.t20.t4 (cl @p_449) :rule refl) +(step t33.t1.t20.t5 (cl @p_459) :rule cong :premises (t33.t1.t20.t3 t33.t1.t20.t4)) +(step t33.t1.t20.t6 (cl (! (= veriT_vr4 veriT_vr42) :named @p_461)) :rule refl) +(step t33.t1.t20.t7 (cl (! (= veriT_vr5 veriT_vr43) :named @p_462)) :rule refl) +(step t33.t1.t20.t8 (cl (! (= @p_28 (! (pair$ veriT_vr42 veriT_vr43) :named @p_440)) :named @p_463)) :rule cong :premises (t33.t1.t20.t6 t33.t1.t20.t7)) +(step t33.t1.t20.t9 (cl (= @p_54 (! (is_res$ @p_438 @p_440) :named @p_460))) :rule cong :premises (t33.t1.t20.t5 t33.t1.t20.t8)) +(step t33.t1.t20.t10 (cl @p_451) :rule refl) +(step t33.t1.t20.t11 (cl @p_447) :rule refl) +(step t33.t1.t20.t12 (cl @p_452) :rule cong :premises (t33.t1.t20.t10 t33.t1.t20.t11)) +(step t33.t1.t20.t13 (cl @p_449) :rule refl) +(step t33.t1.t20.t14 (cl @p_453) :rule cong :premises (t33.t1.t20.t12 t33.t1.t20.t13)) +(step t33.t1.t20.t15 (cl @p_461) :rule refl) +(step t33.t1.t20.t16 (cl @p_462) :rule refl) +(step t33.t1.t20.t17 (cl @p_463) :rule cong :premises (t33.t1.t20.t15 t33.t1.t20.t16)) +(step t33.t1.t20.t18 (cl (= @p_59 (! (is_res$ @p_439 @p_440) :named @p_464))) :rule cong :premises (t33.t1.t20.t14 t33.t1.t20.t17)) +(step t33.t1.t20.t19 (cl (= @p_61 (! (= @p_460 @p_464) :named @p_465))) :rule cong :premises (t33.t1.t20.t9 t33.t1.t20.t18)) +(step t33.t1.t20 (cl (= @p_49 (! (forall ((veriT_vr42 B$) (veriT_vr43 C$)) @p_465) :named @p_456))) :rule bind) +(step t33.t1.t21 (cl (= @p_63 (! (and @p_455 @p_456) :named @p_466))) :rule cong :premises (t33.t1.t19 t33.t1.t20)) +(step t33.t1.t22 (cl (= @p_65 (! (or @p_437 @p_466) :named @p_467))) :rule cong :premises (t33.t1.t6 t33.t1.t21)) +(step t33.t1 (cl (= @p_33 (! (forall ((veriT_vr40 A$) (veriT_vr41 C$)) @p_467) :named @p_445))) :rule bind) +(anchor :step t33.t2 :args ((:= (veriT_vr6 D$) veriT_vr44))) +(step t33.t2.t1 (cl @p_451) :rule refl) +(step t33.t2.t2 (cl (! (= @p_69 (! (b$ veriT_vr38) :named @p_469)) :named @p_473)) :rule cong :premises (t33.t2.t1)) +(step t33.t2.t3 (cl (! (= veriT_vr6 veriT_vr44) :named @p_471)) :rule refl) +(step t33.t2.t4 (cl (! (= @p_31 (! (run$a @p_469 veriT_vr44) :named @p_443)) :named @p_474)) :rule cong :premises (t33.t2.t2 t33.t2.t3)) +(step t33.t2.t5 (cl (! (= @p_29 (! (is_fail$a @p_443) :named @p_441)) :named @p_475)) :rule cong :premises (t33.t2.t4)) +(step t33.t2.t6 (cl @p_457) :rule refl) +(step t33.t2.t7 (cl (! (= @p_71 (! (b$ veriT_vr39) :named @p_470)) :named @p_478)) :rule cong :premises (t33.t2.t6)) +(step t33.t2.t8 (cl @p_471) :rule refl) +(step t33.t2.t9 (cl (! (= @p_30 (! (run$a @p_470 veriT_vr44) :named @p_442)) :named @p_479)) :rule cong :premises (t33.t2.t7 t33.t2.t8)) +(step t33.t2.t10 (cl (= @p_73 (! (is_fail$a @p_442) :named @p_472))) :rule cong :premises (t33.t2.t9)) +(step t33.t2.t11 (cl @p_451) :rule refl) +(step t33.t2.t12 (cl @p_473) :rule cong :premises (t33.t2.t11)) +(step t33.t2.t13 (cl @p_471) :rule refl) +(step t33.t2.t14 (cl @p_474) :rule cong :premises (t33.t2.t12 t33.t2.t13)) +(step t33.t2.t15 (cl @p_475) :rule cong :premises (t33.t2.t14)) +(step t33.t2.t16 (cl (= @p_78 (! (= @p_472 @p_441) :named @p_476))) :rule cong :premises (t33.t2.t10 t33.t2.t15)) +(anchor :step t33.t2.t17 :args ((:= (veriT_vr7 E$) veriT_vr45) (:= (veriT_vr8 D$) veriT_vr46))) +(step t33.t2.t17.t1 (cl @p_457) :rule refl) +(step t33.t2.t17.t2 (cl @p_478) :rule cong :premises (t33.t2.t17.t1)) +(step t33.t2.t17.t3 (cl @p_471) :rule refl) +(step t33.t2.t17.t4 (cl @p_479) :rule cong :premises (t33.t2.t17.t2 t33.t2.t17.t3)) +(step t33.t2.t17.t5 (cl (! (= veriT_vr7 veriT_vr45) :named @p_481)) :rule refl) +(step t33.t2.t17.t6 (cl (! (= veriT_vr8 veriT_vr46) :named @p_482)) :rule refl) +(step t33.t2.t17.t7 (cl (! (= @p_32 (! (pair$a veriT_vr45 veriT_vr46) :named @p_444)) :named @p_483)) :rule cong :premises (t33.t2.t17.t5 t33.t2.t17.t6)) +(step t33.t2.t17.t8 (cl (= @p_83 (! (is_res$a @p_442 @p_444) :named @p_480))) :rule cong :premises (t33.t2.t17.t4 t33.t2.t17.t7)) +(step t33.t2.t17.t9 (cl @p_451) :rule refl) +(step t33.t2.t17.t10 (cl @p_473) :rule cong :premises (t33.t2.t17.t9)) +(step t33.t2.t17.t11 (cl @p_471) :rule refl) +(step t33.t2.t17.t12 (cl @p_474) :rule cong :premises (t33.t2.t17.t10 t33.t2.t17.t11)) +(step t33.t2.t17.t13 (cl @p_481) :rule refl) +(step t33.t2.t17.t14 (cl @p_482) :rule refl) +(step t33.t2.t17.t15 (cl @p_483) :rule cong :premises (t33.t2.t17.t13 t33.t2.t17.t14)) +(step t33.t2.t17.t16 (cl (= @p_87 (! (is_res$a @p_443 @p_444) :named @p_484))) :rule cong :premises (t33.t2.t17.t12 t33.t2.t17.t15)) +(step t33.t2.t17.t17 (cl (= @p_89 (! (= @p_480 @p_484) :named @p_485))) :rule cong :premises (t33.t2.t17.t8 t33.t2.t17.t16)) +(step t33.t2.t17 (cl (= @p_80 (! (forall ((veriT_vr45 E$) (veriT_vr46 D$)) @p_485) :named @p_477))) :rule bind) +(step t33.t2.t18 (cl (= @p_91 (! (and @p_476 @p_477) :named @p_486))) :rule cong :premises (t33.t2.t16 t33.t2.t17)) +(step t33.t2.t19 (cl (= @p_93 (! (or @p_441 @p_486) :named @p_487))) :rule cong :premises (t33.t2.t5 t33.t2.t18)) +(step t33.t2 (cl (= @p_67 (! (forall ((veriT_vr44 D$)) @p_487) :named @p_468))) :rule bind) +(step t33.t3 (cl (= @p_95 (! (=> @p_445 @p_468) :named @p_488))) :rule cong :premises (t33.t1 t33.t2)) +(step t33 (cl (= @p_159 (! (forall ((veriT_vr38 A_b_c_M_state_fun$) (veriT_vr39 A_b_c_M_state_fun$)) @p_488) :named @p_691))) :rule bind) +(anchor :step t34 :args ((:= (veriT_vr11 E$) veriT_vr47) (:= (veriT_vr12 A_b_c_M_state_fun$) veriT_vr48) (:= (veriT_vr13 A_b_c_M_state_fun$) veriT_vr49))) +(anchor :step t34.t1 :args ((:= (veriT_vr14 A$) veriT_vr50) (:= (veriT_vr15 C$) veriT_vr51))) +(step t34.t1.t1 (cl (! (= veriT_vr12 veriT_vr48) :named @p_503)) :rule refl) +(step t34.t1.t2 (cl (! (= veriT_vr14 veriT_vr50) :named @p_499)) :rule refl) +(step t34.t1.t3 (cl (! (= @p_278 (! (fun_app$ veriT_vr48 veriT_vr50) :named @p_498)) :named @p_504)) :rule cong :premises (t34.t1.t1 t34.t1.t2)) +(step t34.t1.t4 (cl (! (= veriT_vr15 veriT_vr51) :named @p_501)) :rule refl) +(step t34.t1.t5 (cl (! (= @p_271 (! (run$ @p_498 veriT_vr51) :named @p_491)) :named @p_505)) :rule cong :premises (t34.t1.t3 t34.t1.t4)) +(step t34.t1.t6 (cl (! (= @p_269 (! (is_fail$ @p_491) :named @p_489)) :named @p_506)) :rule cong :premises (t34.t1.t5)) +(step t34.t1.t7 (cl (! (= veriT_vr13 veriT_vr49) :named @p_509)) :rule refl) +(step t34.t1.t8 (cl @p_499) :rule refl) +(step t34.t1.t9 (cl (! (= @p_280 (! (fun_app$ veriT_vr49 veriT_vr50) :named @p_500)) :named @p_510)) :rule cong :premises (t34.t1.t7 t34.t1.t8)) +(step t34.t1.t10 (cl @p_501) :rule refl) +(step t34.t1.t11 (cl (! (= @p_270 (! (run$ @p_500 veriT_vr51) :named @p_490)) :named @p_511)) :rule cong :premises (t34.t1.t9 t34.t1.t10)) +(step t34.t1.t12 (cl (= @p_282 (! (is_fail$ @p_490) :named @p_502))) :rule cong :premises (t34.t1.t11)) +(step t34.t1.t13 (cl @p_503) :rule refl) +(step t34.t1.t14 (cl @p_499) :rule refl) +(step t34.t1.t15 (cl @p_504) :rule cong :premises (t34.t1.t13 t34.t1.t14)) +(step t34.t1.t16 (cl @p_501) :rule refl) +(step t34.t1.t17 (cl @p_505) :rule cong :premises (t34.t1.t15 t34.t1.t16)) +(step t34.t1.t18 (cl @p_506) :rule cong :premises (t34.t1.t17)) +(step t34.t1.t19 (cl (= @p_287 (! (= @p_502 @p_489) :named @p_507))) :rule cong :premises (t34.t1.t12 t34.t1.t18)) +(anchor :step t34.t1.t20 :args ((:= (veriT_vr16 B$) veriT_vr52) (:= (veriT_vr17 C$) veriT_vr53))) +(step t34.t1.t20.t1 (cl @p_509) :rule refl) +(step t34.t1.t20.t2 (cl @p_499) :rule refl) +(step t34.t1.t20.t3 (cl @p_510) :rule cong :premises (t34.t1.t20.t1 t34.t1.t20.t2)) +(step t34.t1.t20.t4 (cl @p_501) :rule refl) +(step t34.t1.t20.t5 (cl @p_511) :rule cong :premises (t34.t1.t20.t3 t34.t1.t20.t4)) +(step t34.t1.t20.t6 (cl (! (= veriT_vr16 veriT_vr52) :named @p_513)) :rule refl) +(step t34.t1.t20.t7 (cl (! (= veriT_vr17 veriT_vr53) :named @p_514)) :rule refl) +(step t34.t1.t20.t8 (cl (! (= @p_272 (! (pair$ veriT_vr52 veriT_vr53) :named @p_492)) :named @p_515)) :rule cong :premises (t34.t1.t20.t6 t34.t1.t20.t7)) +(step t34.t1.t20.t9 (cl (= @p_292 (! (is_res$ @p_490 @p_492) :named @p_512))) :rule cong :premises (t34.t1.t20.t5 t34.t1.t20.t8)) +(step t34.t1.t20.t10 (cl @p_503) :rule refl) +(step t34.t1.t20.t11 (cl @p_499) :rule refl) +(step t34.t1.t20.t12 (cl @p_504) :rule cong :premises (t34.t1.t20.t10 t34.t1.t20.t11)) +(step t34.t1.t20.t13 (cl @p_501) :rule refl) +(step t34.t1.t20.t14 (cl @p_505) :rule cong :premises (t34.t1.t20.t12 t34.t1.t20.t13)) +(step t34.t1.t20.t15 (cl @p_513) :rule refl) +(step t34.t1.t20.t16 (cl @p_514) :rule refl) +(step t34.t1.t20.t17 (cl @p_515) :rule cong :premises (t34.t1.t20.t15 t34.t1.t20.t16)) +(step t34.t1.t20.t18 (cl (= @p_296 (! (is_res$ @p_491 @p_492) :named @p_516))) :rule cong :premises (t34.t1.t20.t14 t34.t1.t20.t17)) +(step t34.t1.t20.t19 (cl (= @p_297 (! (= @p_512 @p_516) :named @p_517))) :rule cong :premises (t34.t1.t20.t9 t34.t1.t20.t18)) +(step t34.t1.t20 (cl (= @p_288 (! (forall ((veriT_vr52 B$) (veriT_vr53 C$)) @p_517) :named @p_508))) :rule bind) +(step t34.t1.t21 (cl (= @p_298 (! (and @p_507 @p_508) :named @p_518))) :rule cong :premises (t34.t1.t19 t34.t1.t20)) +(step t34.t1.t22 (cl (= @p_299 (! (or @p_489 @p_518) :named @p_519))) :rule cong :premises (t34.t1.t6 t34.t1.t21)) +(step t34.t1 (cl (= @p_277 (! (forall ((veriT_vr50 A$) (veriT_vr51 C$)) @p_519) :named @p_497))) :rule bind) +(anchor :step t34.t2 :args ((:= (veriT_vr10 D$) veriT_vr54))) +(step t34.t2.t1 (cl (! (= veriT_vr11 veriT_vr47) :named @p_522)) :rule refl) +(step t34.t2.t2 (cl @p_503) :rule refl) +(step t34.t2.t3 (cl (! (= @p_301 (! (c$ veriT_vr47 veriT_vr48) :named @p_521)) :named @p_526)) :rule cong :premises (t34.t2.t1 t34.t2.t2)) +(step t34.t2.t4 (cl (! (= veriT_vr10 veriT_vr54) :named @p_524)) :rule refl) +(step t34.t2.t5 (cl (! (= @p_275 (! (run$b @p_521 veriT_vr54) :named @p_495)) :named @p_527)) :rule cong :premises (t34.t2.t3 t34.t2.t4)) +(step t34.t2.t6 (cl (! (= @p_273 (! (is_fail$b @p_495) :named @p_493)) :named @p_528)) :rule cong :premises (t34.t2.t5)) +(step t34.t2.t7 (cl @p_522) :rule refl) +(step t34.t2.t8 (cl @p_509) :rule refl) +(step t34.t2.t9 (cl (! (= @p_303 (! (c$ veriT_vr47 veriT_vr49) :named @p_523)) :named @p_531)) :rule cong :premises (t34.t2.t7 t34.t2.t8)) +(step t34.t2.t10 (cl @p_524) :rule refl) +(step t34.t2.t11 (cl (! (= @p_274 (! (run$b @p_523 veriT_vr54) :named @p_494)) :named @p_532)) :rule cong :premises (t34.t2.t9 t34.t2.t10)) +(step t34.t2.t12 (cl (= @p_305 (! (is_fail$b @p_494) :named @p_525))) :rule cong :premises (t34.t2.t11)) +(step t34.t2.t13 (cl @p_522) :rule refl) +(step t34.t2.t14 (cl @p_503) :rule refl) +(step t34.t2.t15 (cl @p_526) :rule cong :premises (t34.t2.t13 t34.t2.t14)) +(step t34.t2.t16 (cl @p_524) :rule refl) +(step t34.t2.t17 (cl @p_527) :rule cong :premises (t34.t2.t15 t34.t2.t16)) +(step t34.t2.t18 (cl @p_528) :rule cong :premises (t34.t2.t17)) +(step t34.t2.t19 (cl (= @p_309 (! (= @p_525 @p_493) :named @p_529))) :rule cong :premises (t34.t2.t12 t34.t2.t18)) +(anchor :step t34.t2.t20 :args ((:= (veriT_vr9 F$) veriT_vr55) (:= (veriT_vr18 D$) veriT_vr56))) +(step t34.t2.t20.t1 (cl @p_522) :rule refl) +(step t34.t2.t20.t2 (cl @p_509) :rule refl) +(step t34.t2.t20.t3 (cl @p_531) :rule cong :premises (t34.t2.t20.t1 t34.t2.t20.t2)) +(step t34.t2.t20.t4 (cl @p_524) :rule refl) +(step t34.t2.t20.t5 (cl @p_532) :rule cong :premises (t34.t2.t20.t3 t34.t2.t20.t4)) +(step t34.t2.t20.t6 (cl (! (= veriT_vr9 veriT_vr55) :named @p_534)) :rule refl) +(step t34.t2.t20.t7 (cl (! (= veriT_vr18 veriT_vr56) :named @p_535)) :rule refl) +(step t34.t2.t20.t8 (cl (! (= @p_276 (! (pair$b veriT_vr55 veriT_vr56) :named @p_496)) :named @p_536)) :rule cong :premises (t34.t2.t20.t6 t34.t2.t20.t7)) +(step t34.t2.t20.t9 (cl (= @p_313 (! (is_res$b @p_494 @p_496) :named @p_533))) :rule cong :premises (t34.t2.t20.t5 t34.t2.t20.t8)) +(step t34.t2.t20.t10 (cl @p_522) :rule refl) +(step t34.t2.t20.t11 (cl @p_503) :rule refl) +(step t34.t2.t20.t12 (cl @p_526) :rule cong :premises (t34.t2.t20.t10 t34.t2.t20.t11)) +(step t34.t2.t20.t13 (cl @p_524) :rule refl) +(step t34.t2.t20.t14 (cl @p_527) :rule cong :premises (t34.t2.t20.t12 t34.t2.t20.t13)) +(step t34.t2.t20.t15 (cl @p_534) :rule refl) +(step t34.t2.t20.t16 (cl @p_535) :rule refl) +(step t34.t2.t20.t17 (cl @p_536) :rule cong :premises (t34.t2.t20.t15 t34.t2.t20.t16)) +(step t34.t2.t20.t18 (cl (= @p_316 (! (is_res$b @p_495 @p_496) :named @p_537))) :rule cong :premises (t34.t2.t20.t14 t34.t2.t20.t17)) +(step t34.t2.t20.t19 (cl (= @p_317 (! (= @p_533 @p_537) :named @p_538))) :rule cong :premises (t34.t2.t20.t9 t34.t2.t20.t18)) +(step t34.t2.t20 (cl (= @p_310 (! (forall ((veriT_vr55 F$) (veriT_vr56 D$)) @p_538) :named @p_530))) :rule bind) +(step t34.t2.t21 (cl (= @p_318 (! (and @p_529 @p_530) :named @p_539))) :rule cong :premises (t34.t2.t19 t34.t2.t20)) +(step t34.t2.t22 (cl (= @p_319 (! (or @p_493 @p_539) :named @p_540))) :rule cong :premises (t34.t2.t6 t34.t2.t21)) +(step t34.t2 (cl (= @p_300 (! (forall ((veriT_vr54 D$)) @p_540) :named @p_520))) :rule bind) +(step t34.t3 (cl (= @p_320 (! (=> @p_497 @p_520) :named @p_541))) :rule cong :premises (t34.t1 t34.t2)) +(step t34 (cl (= @p_422 (! (forall ((veriT_vr47 E$) (veriT_vr48 A_b_c_M_state_fun$) (veriT_vr49 A_b_c_M_state_fun$)) @p_541) :named @p_692))) :rule bind) +(anchor :step t35 :args ((:= (veriT_vr19 A_b_c_M_state_fun$) veriT_vr57) (:= (veriT_vr20 A_b_c_M_state_fun$) veriT_vr58))) +(anchor :step t35.t1 :args ((:= (veriT_vr21 A$) veriT_vr59) (:= (veriT_vr22 C$) veriT_vr60))) +(step t35.t1.t1 (cl (! (= veriT_vr19 veriT_vr57) :named @p_557)) :rule refl) +(step t35.t1.t2 (cl (! (= veriT_vr21 veriT_vr59) :named @p_553)) :rule refl) +(step t35.t1.t3 (cl (! (= @p_331 @p_552) :named @p_558)) :rule cong :premises (t35.t1.t1 t35.t1.t2)) +(step t35.t1.t4 (cl (! (= veriT_vr22 veriT_vr60) :named @p_555)) :rule refl) +(step t35.t1.t5 (cl (! (= @p_323 @p_544) :named @p_559)) :rule cong :premises (t35.t1.t3 t35.t1.t4)) +(step t35.t1.t6 (cl (! (= @p_321 @p_542) :named @p_560)) :rule cong :premises (t35.t1.t5)) +(step t35.t1.t7 (cl (! (= veriT_vr20 veriT_vr58) :named @p_563)) :rule refl) +(step t35.t1.t8 (cl @p_553) :rule refl) +(step t35.t1.t9 (cl (! (= @p_333 @p_554) :named @p_564)) :rule cong :premises (t35.t1.t7 t35.t1.t8)) +(step t35.t1.t10 (cl @p_555) :rule refl) +(step t35.t1.t11 (cl (! (= @p_322 @p_543) :named @p_565)) :rule cong :premises (t35.t1.t9 t35.t1.t10)) +(step t35.t1.t12 (cl (= @p_335 @p_556)) :rule cong :premises (t35.t1.t11)) +(step t35.t1.t13 (cl @p_557) :rule refl) +(step t35.t1.t14 (cl @p_553) :rule refl) +(step t35.t1.t15 (cl @p_558) :rule cong :premises (t35.t1.t13 t35.t1.t14)) +(step t35.t1.t16 (cl @p_555) :rule refl) +(step t35.t1.t17 (cl @p_559) :rule cong :premises (t35.t1.t15 t35.t1.t16)) +(step t35.t1.t18 (cl @p_560) :rule cong :premises (t35.t1.t17)) +(step t35.t1.t19 (cl (= @p_340 @p_561)) :rule cong :premises (t35.t1.t12 t35.t1.t18)) +(anchor :step t35.t1.t20 :args ((:= (veriT_vr23 B$) veriT_vr61) (:= (veriT_vr24 C$) veriT_vr62))) +(step t35.t1.t20.t1 (cl @p_563) :rule refl) +(step t35.t1.t20.t2 (cl @p_553) :rule refl) +(step t35.t1.t20.t3 (cl @p_564) :rule cong :premises (t35.t1.t20.t1 t35.t1.t20.t2)) +(step t35.t1.t20.t4 (cl @p_555) :rule refl) +(step t35.t1.t20.t5 (cl @p_565) :rule cong :premises (t35.t1.t20.t3 t35.t1.t20.t4)) +(step t35.t1.t20.t6 (cl (! (= veriT_vr23 veriT_vr61) :named @p_567)) :rule refl) +(step t35.t1.t20.t7 (cl (! (= veriT_vr24 veriT_vr62) :named @p_568)) :rule refl) +(step t35.t1.t20.t8 (cl (! (= @p_324 @p_545) :named @p_569)) :rule cong :premises (t35.t1.t20.t6 t35.t1.t20.t7)) +(step t35.t1.t20.t9 (cl (= @p_345 @p_566)) :rule cong :premises (t35.t1.t20.t5 t35.t1.t20.t8)) +(step t35.t1.t20.t10 (cl @p_557) :rule refl) +(step t35.t1.t20.t11 (cl @p_553) :rule refl) +(step t35.t1.t20.t12 (cl @p_558) :rule cong :premises (t35.t1.t20.t10 t35.t1.t20.t11)) +(step t35.t1.t20.t13 (cl @p_555) :rule refl) +(step t35.t1.t20.t14 (cl @p_559) :rule cong :premises (t35.t1.t20.t12 t35.t1.t20.t13)) +(step t35.t1.t20.t15 (cl @p_567) :rule refl) +(step t35.t1.t20.t16 (cl @p_568) :rule refl) +(step t35.t1.t20.t17 (cl @p_569) :rule cong :premises (t35.t1.t20.t15 t35.t1.t20.t16)) +(step t35.t1.t20.t18 (cl (= @p_349 @p_570)) :rule cong :premises (t35.t1.t20.t14 t35.t1.t20.t17)) +(step t35.t1.t20.t19 (cl (= @p_350 @p_571)) :rule cong :premises (t35.t1.t20.t9 t35.t1.t20.t18)) +(step t35.t1.t20 (cl (= @p_341 @p_562)) :rule bind) +(step t35.t1.t21 (cl (= @p_351 @p_572)) :rule cong :premises (t35.t1.t19 t35.t1.t20)) +(step t35.t1.t22 (cl (= @p_352 @p_573)) :rule cong :premises (t35.t1.t6 t35.t1.t21)) +(step t35.t1 (cl (= @p_330 @p_551)) :rule bind) +(anchor :step t35.t2 :args ((:= (veriT_vr25 D$) veriT_vr63))) +(step t35.t2.t1 (cl @p_557) :rule refl) +(step t35.t2.t2 (cl (! (= @p_354 @p_575) :named @p_577)) :rule cong :premises (t35.t2.t1)) +(step t35.t2.t3 (cl (! (= veriT_vr25 veriT_vr63) :named @p_578)) :rule refl) +(step t35.t2.t4 (cl (! (= @p_325 @p_546) :named @p_579)) :rule cong :premises (t35.t2.t2 t35.t2.t3)) +(step t35.t2.t5 (cl (! (= @p_327 @p_548) :named @p_601)) :rule cong :premises (t35.t2.t4)) +(anchor :step t35.t2.t6 :args ((:= (veriT_vr26 E$) veriT_vr64) (:= (veriT_vr27 D$) veriT_vr65))) +(step t35.t2.t6.t1 (cl @p_557) :rule refl) +(step t35.t2.t6.t2 (cl @p_577) :rule cong :premises (t35.t2.t6.t1)) +(step t35.t2.t6.t3 (cl @p_578) :rule refl) +(step t35.t2.t6.t4 (cl @p_579) :rule cong :premises (t35.t2.t6.t2 t35.t2.t6.t3)) +(step t35.t2.t6.t5 (cl (! (= veriT_vr26 veriT_vr64) :named @p_582)) :rule refl) +(step t35.t2.t6.t6 (cl (! (= veriT_vr27 veriT_vr65) :named @p_584)) :rule refl) +(step t35.t2.t6.t7 (cl (= @p_359 @p_580)) :rule cong :premises (t35.t2.t6.t5 t35.t2.t6.t6)) +(step t35.t2.t6.t8 (cl (= @p_360 @p_581)) :rule cong :premises (t35.t2.t6.t4 t35.t2.t6.t7)) +(step t35.t2.t6.t9 (cl @p_582) :rule refl) +(step t35.t2.t6.t10 (cl @p_557) :rule refl) +(step t35.t2.t6.t11 (cl (= @p_362 @p_583)) :rule cong :premises (t35.t2.t6.t9 t35.t2.t6.t10)) +(step t35.t2.t6.t12 (cl @p_584) :rule refl) +(step t35.t2.t6.t13 (cl (= @p_364 @p_585)) :rule cong :premises (t35.t2.t6.t11 t35.t2.t6.t12)) +(step t35.t2.t6.t14 (cl (= @p_365 @p_586)) :rule cong :premises (t35.t2.t6.t13)) +(step t35.t2.t6.t15 (cl (= @p_366 @p_587)) :rule cong :premises (t35.t2.t6.t8 t35.t2.t6.t14)) +(step t35.t2.t6 (cl (= @p_355 @p_576)) :rule bind) +(step t35.t2.t7 (cl @p_563) :rule refl) +(step t35.t2.t8 (cl (! (= @p_367 @p_588) :named @p_590)) :rule cong :premises (t35.t2.t7)) +(step t35.t2.t9 (cl @p_578) :rule refl) +(step t35.t2.t10 (cl (! (= @p_326 @p_547) :named @p_591)) :rule cong :premises (t35.t2.t8 t35.t2.t9)) +(step t35.t2.t11 (cl (! (= @p_328 @p_549) :named @p_624)) :rule cong :premises (t35.t2.t10)) +(anchor :step t35.t2.t12 :args ((:= (veriT_vr28 E$) veriT_vr66) (:= (veriT_vr29 D$) veriT_vr67))) +(step t35.t2.t12.t1 (cl @p_563) :rule refl) +(step t35.t2.t12.t2 (cl @p_590) :rule cong :premises (t35.t2.t12.t1)) +(step t35.t2.t12.t3 (cl @p_578) :rule refl) +(step t35.t2.t12.t4 (cl @p_591) :rule cong :premises (t35.t2.t12.t2 t35.t2.t12.t3)) +(step t35.t2.t12.t5 (cl (! (= veriT_vr28 veriT_vr66) :named @p_594)) :rule refl) +(step t35.t2.t12.t6 (cl (! (= veriT_vr29 veriT_vr67) :named @p_596)) :rule refl) +(step t35.t2.t12.t7 (cl (= @p_371 @p_592)) :rule cong :premises (t35.t2.t12.t5 t35.t2.t12.t6)) +(step t35.t2.t12.t8 (cl (= @p_372 @p_593)) :rule cong :premises (t35.t2.t12.t4 t35.t2.t12.t7)) +(step t35.t2.t12.t9 (cl @p_594) :rule refl) +(step t35.t2.t12.t10 (cl @p_563) :rule refl) +(step t35.t2.t12.t11 (cl (= @p_374 @p_595)) :rule cong :premises (t35.t2.t12.t9 t35.t2.t12.t10)) +(step t35.t2.t12.t12 (cl @p_596) :rule refl) +(step t35.t2.t12.t13 (cl (= @p_376 @p_597)) :rule cong :premises (t35.t2.t12.t11 t35.t2.t12.t12)) +(step t35.t2.t12.t14 (cl (= @p_377 @p_598)) :rule cong :premises (t35.t2.t12.t13)) +(step t35.t2.t12.t15 (cl (= @p_378 @p_599)) :rule cong :premises (t35.t2.t12.t8 t35.t2.t12.t14)) +(step t35.t2.t12 (cl (= @p_368 @p_589)) :rule bind) +(step t35.t2.t13 (cl (= @p_379 @p_600)) :rule cong :premises (t35.t2.t11 t35.t2.t12)) +(step t35.t2.t14 (cl @p_557) :rule refl) +(step t35.t2.t15 (cl @p_577) :rule cong :premises (t35.t2.t14)) +(step t35.t2.t16 (cl @p_578) :rule refl) +(step t35.t2.t17 (cl @p_579) :rule cong :premises (t35.t2.t15 t35.t2.t16)) +(step t35.t2.t18 (cl @p_601) :rule cong :premises (t35.t2.t17)) +(anchor :step t35.t2.t19 :args ((:= (veriT_vr30 E$) veriT_vr68) (:= (veriT_vr31 D$) veriT_vr69))) +(step t35.t2.t19.t1 (cl @p_557) :rule refl) +(step t35.t2.t19.t2 (cl @p_577) :rule cong :premises (t35.t2.t19.t1)) +(step t35.t2.t19.t3 (cl @p_578) :rule refl) +(step t35.t2.t19.t4 (cl @p_579) :rule cong :premises (t35.t2.t19.t2 t35.t2.t19.t3)) +(step t35.t2.t19.t5 (cl (! (= veriT_vr30 veriT_vr68) :named @p_605)) :rule refl) +(step t35.t2.t19.t6 (cl (! (= veriT_vr31 veriT_vr69) :named @p_607)) :rule refl) +(step t35.t2.t19.t7 (cl (= @p_382 @p_603)) :rule cong :premises (t35.t2.t19.t5 t35.t2.t19.t6)) +(step t35.t2.t19.t8 (cl (= @p_383 @p_604)) :rule cong :premises (t35.t2.t19.t4 t35.t2.t19.t7)) +(step t35.t2.t19.t9 (cl @p_605) :rule refl) +(step t35.t2.t19.t10 (cl @p_557) :rule refl) +(step t35.t2.t19.t11 (cl (= @p_385 @p_606)) :rule cong :premises (t35.t2.t19.t9 t35.t2.t19.t10)) +(step t35.t2.t19.t12 (cl @p_607) :rule refl) +(step t35.t2.t19.t13 (cl (= @p_387 @p_608)) :rule cong :premises (t35.t2.t19.t11 t35.t2.t19.t12)) +(step t35.t2.t19.t14 (cl (= @p_388 @p_609)) :rule cong :premises (t35.t2.t19.t13)) +(step t35.t2.t19.t15 (cl (= @p_389 @p_610)) :rule cong :premises (t35.t2.t19.t8 t35.t2.t19.t14)) +(step t35.t2.t19 (cl (= @p_381 @p_602)) :rule bind) +(step t35.t2.t20 (cl (= @p_390 @p_611)) :rule cong :premises (t35.t2.t18 t35.t2.t19)) +(step t35.t2.t21 (cl (= @p_612 @p_613)) :rule cong :premises (t35.t2.t13 t35.t2.t20)) +(step t35.t2.t22 (cl @p_557) :rule refl) +(step t35.t2.t23 (cl @p_577) :rule cong :premises (t35.t2.t22)) +(step t35.t2.t24 (cl @p_578) :rule refl) +(step t35.t2.t25 (cl @p_579) :rule cong :premises (t35.t2.t23 t35.t2.t24)) +(step t35.t2.t26 (cl @p_601) :rule cong :premises (t35.t2.t25)) +(anchor :step t35.t2.t27 :args ((:= (veriT_vr30 E$) veriT_vr70) (:= (veriT_vr31 D$) veriT_vr71))) +(step t35.t2.t27.t1 (cl @p_557) :rule refl) +(step t35.t2.t27.t2 (cl @p_577) :rule cong :premises (t35.t2.t27.t1)) +(step t35.t2.t27.t3 (cl @p_578) :rule refl) +(step t35.t2.t27.t4 (cl @p_579) :rule cong :premises (t35.t2.t27.t2 t35.t2.t27.t3)) +(step t35.t2.t27.t5 (cl (! (= veriT_vr30 veriT_vr70) :named @p_617)) :rule refl) +(step t35.t2.t27.t6 (cl (! (= veriT_vr31 veriT_vr71) :named @p_619)) :rule refl) +(step t35.t2.t27.t7 (cl (= @p_382 @p_615)) :rule cong :premises (t35.t2.t27.t5 t35.t2.t27.t6)) +(step t35.t2.t27.t8 (cl (= @p_383 @p_616)) :rule cong :premises (t35.t2.t27.t4 t35.t2.t27.t7)) +(step t35.t2.t27.t9 (cl @p_617) :rule refl) +(step t35.t2.t27.t10 (cl @p_557) :rule refl) +(step t35.t2.t27.t11 (cl (= @p_385 @p_618)) :rule cong :premises (t35.t2.t27.t9 t35.t2.t27.t10)) +(step t35.t2.t27.t12 (cl @p_619) :rule refl) +(step t35.t2.t27.t13 (cl (= @p_387 @p_620)) :rule cong :premises (t35.t2.t27.t11 t35.t2.t27.t12)) +(step t35.t2.t27.t14 (cl (= @p_388 @p_621)) :rule cong :premises (t35.t2.t27.t13)) +(step t35.t2.t27.t15 (cl (= @p_389 @p_622)) :rule cong :premises (t35.t2.t27.t8 t35.t2.t27.t14)) +(step t35.t2.t27 (cl (= @p_381 @p_614)) :rule bind) +(step t35.t2.t28 (cl (= @p_390 @p_623)) :rule cong :premises (t35.t2.t26 t35.t2.t27)) +(step t35.t2.t29 (cl @p_563) :rule refl) +(step t35.t2.t30 (cl @p_590) :rule cong :premises (t35.t2.t29)) +(step t35.t2.t31 (cl @p_578) :rule refl) +(step t35.t2.t32 (cl @p_591) :rule cong :premises (t35.t2.t30 t35.t2.t31)) +(step t35.t2.t33 (cl @p_624) :rule cong :premises (t35.t2.t32)) +(anchor :step t35.t2.t34 :args ((:= (veriT_vr28 E$) veriT_vr72) (:= (veriT_vr29 D$) veriT_vr73))) +(step t35.t2.t34.t1 (cl @p_563) :rule refl) +(step t35.t2.t34.t2 (cl @p_590) :rule cong :premises (t35.t2.t34.t1)) +(step t35.t2.t34.t3 (cl @p_578) :rule refl) +(step t35.t2.t34.t4 (cl @p_591) :rule cong :premises (t35.t2.t34.t2 t35.t2.t34.t3)) +(step t35.t2.t34.t5 (cl (! (= veriT_vr28 veriT_vr72) :named @p_628)) :rule refl) +(step t35.t2.t34.t6 (cl (! (= veriT_vr29 veriT_vr73) :named @p_630)) :rule refl) +(step t35.t2.t34.t7 (cl (= @p_371 @p_626)) :rule cong :premises (t35.t2.t34.t5 t35.t2.t34.t6)) +(step t35.t2.t34.t8 (cl (= @p_372 @p_627)) :rule cong :premises (t35.t2.t34.t4 t35.t2.t34.t7)) +(step t35.t2.t34.t9 (cl @p_628) :rule refl) +(step t35.t2.t34.t10 (cl @p_563) :rule refl) +(step t35.t2.t34.t11 (cl (= @p_374 @p_629)) :rule cong :premises (t35.t2.t34.t9 t35.t2.t34.t10)) +(step t35.t2.t34.t12 (cl @p_630) :rule refl) +(step t35.t2.t34.t13 (cl (= @p_376 @p_631)) :rule cong :premises (t35.t2.t34.t11 t35.t2.t34.t12)) +(step t35.t2.t34.t14 (cl (= @p_377 @p_632)) :rule cong :premises (t35.t2.t34.t13)) +(step t35.t2.t34.t15 (cl (= @p_378 @p_633)) :rule cong :premises (t35.t2.t34.t8 t35.t2.t34.t14)) +(step t35.t2.t34 (cl (= @p_368 @p_625)) :rule bind) +(step t35.t2.t35 (cl (= @p_379 @p_634)) :rule cong :premises (t35.t2.t33 t35.t2.t34)) +(step t35.t2.t36 (cl (= @p_635 @p_636)) :rule cong :premises (t35.t2.t28 t35.t2.t35)) +(step t35.t2.t37 (cl (= @p_427 @p_637)) :rule cong :premises (t35.t2.t21 t35.t2.t36)) +(anchor :step t35.t2.t38 :args ((:= (veriT_vr32 F$) veriT_vr74) (:= (veriT_vr33 D$) veriT_vr75))) +(step t35.t2.t38.t1 (cl @p_563) :rule refl) +(step t35.t2.t38.t2 (cl @p_590) :rule cong :premises (t35.t2.t38.t1)) +(step t35.t2.t38.t3 (cl @p_578) :rule refl) +(step t35.t2.t38.t4 (cl @p_591) :rule cong :premises (t35.t2.t38.t2 t35.t2.t38.t3)) +(step t35.t2.t38.t5 (cl @p_624) :rule cong :premises (t35.t2.t38.t4)) +(anchor :step t35.t2.t38.t6 :args ((:= (veriT_vr34 E$) veriT_vr76) (:= (veriT_vr35 D$) veriT_vr77))) +(step t35.t2.t38.t6.t1 (cl @p_563) :rule refl) +(step t35.t2.t38.t6.t2 (cl @p_590) :rule cong :premises (t35.t2.t38.t6.t1)) +(step t35.t2.t38.t6.t3 (cl @p_578) :rule refl) +(step t35.t2.t38.t6.t4 (cl @p_591) :rule cong :premises (t35.t2.t38.t6.t2 t35.t2.t38.t6.t3)) +(step t35.t2.t38.t6.t5 (cl (! (= veriT_vr34 veriT_vr76) :named @p_642)) :rule refl) +(step t35.t2.t38.t6.t6 (cl (! (= veriT_vr35 veriT_vr77) :named @p_644)) :rule refl) +(step t35.t2.t38.t6.t7 (cl (= @p_395 @p_640)) :rule cong :premises (t35.t2.t38.t6.t5 t35.t2.t38.t6.t6)) +(step t35.t2.t38.t6.t8 (cl (= @p_396 @p_641)) :rule cong :premises (t35.t2.t38.t6.t4 t35.t2.t38.t6.t7)) +(step t35.t2.t38.t6.t9 (cl @p_642) :rule refl) +(step t35.t2.t38.t6.t10 (cl @p_563) :rule refl) +(step t35.t2.t38.t6.t11 (cl (= @p_398 @p_643)) :rule cong :premises (t35.t2.t38.t6.t9 t35.t2.t38.t6.t10)) +(step t35.t2.t38.t6.t12 (cl @p_644) :rule refl) +(step t35.t2.t38.t6.t13 (cl (= @p_400 @p_645)) :rule cong :premises (t35.t2.t38.t6.t11 t35.t2.t38.t6.t12)) +(step t35.t2.t38.t6.t14 (cl (! (= veriT_vr32 veriT_vr74) :named @p_656)) :rule refl) +(step t35.t2.t38.t6.t15 (cl (! (= veriT_vr33 veriT_vr75) :named @p_657)) :rule refl) +(step t35.t2.t38.t6.t16 (cl (! (= @p_329 @p_550) :named @p_658)) :rule cong :premises (t35.t2.t38.t6.t14 t35.t2.t38.t6.t15)) +(step t35.t2.t38.t6.t17 (cl (= @p_401 @p_646)) :rule cong :premises (t35.t2.t38.t6.t13 t35.t2.t38.t6.t16)) +(step t35.t2.t38.t6.t18 (cl (= @p_402 @p_647)) :rule cong :premises (t35.t2.t38.t6.t8 t35.t2.t38.t6.t17)) +(step t35.t2.t38.t6 (cl (= @p_394 @p_639)) :rule bind) +(step t35.t2.t38.t7 (cl (= @p_403 @p_648)) :rule cong :premises (t35.t2.t38.t5 t35.t2.t38.t6)) +(step t35.t2.t38.t8 (cl @p_557) :rule refl) +(step t35.t2.t38.t9 (cl @p_577) :rule cong :premises (t35.t2.t38.t8)) +(step t35.t2.t38.t10 (cl @p_578) :rule refl) +(step t35.t2.t38.t11 (cl @p_579) :rule cong :premises (t35.t2.t38.t9 t35.t2.t38.t10)) +(step t35.t2.t38.t12 (cl @p_601) :rule cong :premises (t35.t2.t38.t11)) +(anchor :step t35.t2.t38.t13 :args ((:= (veriT_vr36 E$) veriT_vr78) (:= (veriT_vr37 D$) veriT_vr79))) +(step t35.t2.t38.t13.t1 (cl @p_557) :rule refl) +(step t35.t2.t38.t13.t2 (cl @p_577) :rule cong :premises (t35.t2.t38.t13.t1)) +(step t35.t2.t38.t13.t3 (cl @p_578) :rule refl) +(step t35.t2.t38.t13.t4 (cl @p_579) :rule cong :premises (t35.t2.t38.t13.t2 t35.t2.t38.t13.t3)) +(step t35.t2.t38.t13.t5 (cl (! (= veriT_vr36 veriT_vr78) :named @p_652)) :rule refl) +(step t35.t2.t38.t13.t6 (cl (! (= veriT_vr37 veriT_vr79) :named @p_654)) :rule refl) +(step t35.t2.t38.t13.t7 (cl (= @p_405 @p_650)) :rule cong :premises (t35.t2.t38.t13.t5 t35.t2.t38.t13.t6)) +(step t35.t2.t38.t13.t8 (cl (= @p_406 @p_651)) :rule cong :premises (t35.t2.t38.t13.t4 t35.t2.t38.t13.t7)) +(step t35.t2.t38.t13.t9 (cl @p_652) :rule refl) +(step t35.t2.t38.t13.t10 (cl @p_557) :rule refl) +(step t35.t2.t38.t13.t11 (cl (= @p_408 @p_653)) :rule cong :premises (t35.t2.t38.t13.t9 t35.t2.t38.t13.t10)) +(step t35.t2.t38.t13.t12 (cl @p_654) :rule refl) +(step t35.t2.t38.t13.t13 (cl (= @p_410 @p_655)) :rule cong :premises (t35.t2.t38.t13.t11 t35.t2.t38.t13.t12)) +(step t35.t2.t38.t13.t14 (cl @p_656) :rule refl) +(step t35.t2.t38.t13.t15 (cl @p_657) :rule refl) +(step t35.t2.t38.t13.t16 (cl @p_658) :rule cong :premises (t35.t2.t38.t13.t14 t35.t2.t38.t13.t15)) +(step t35.t2.t38.t13.t17 (cl (= @p_414 @p_659)) :rule cong :premises (t35.t2.t38.t13.t13 t35.t2.t38.t13.t16)) +(step t35.t2.t38.t13.t18 (cl (= @p_415 @p_660)) :rule cong :premises (t35.t2.t38.t13.t8 t35.t2.t38.t13.t17)) +(step t35.t2.t38.t13 (cl (= @p_404 @p_649)) :rule bind) +(step t35.t2.t38.t14 (cl (= @p_416 @p_661)) :rule cong :premises (t35.t2.t38.t12 t35.t2.t38.t13)) +(step t35.t2.t38.t15 (cl (= @p_662 @p_663)) :rule cong :premises (t35.t2.t38.t7 t35.t2.t38.t14)) +(step t35.t2.t38.t16 (cl @p_557) :rule refl) +(step t35.t2.t38.t17 (cl @p_577) :rule cong :premises (t35.t2.t38.t16)) +(step t35.t2.t38.t18 (cl @p_578) :rule refl) +(step t35.t2.t38.t19 (cl @p_579) :rule cong :premises (t35.t2.t38.t17 t35.t2.t38.t18)) +(step t35.t2.t38.t20 (cl @p_601) :rule cong :premises (t35.t2.t38.t19)) +(anchor :step t35.t2.t38.t21 :args ((:= (veriT_vr36 E$) veriT_vr80) (:= (veriT_vr37 D$) veriT_vr81))) +(step t35.t2.t38.t21.t1 (cl @p_557) :rule refl) +(step t35.t2.t38.t21.t2 (cl @p_577) :rule cong :premises (t35.t2.t38.t21.t1)) +(step t35.t2.t38.t21.t3 (cl @p_578) :rule refl) +(step t35.t2.t38.t21.t4 (cl @p_579) :rule cong :premises (t35.t2.t38.t21.t2 t35.t2.t38.t21.t3)) +(step t35.t2.t38.t21.t5 (cl (! (= veriT_vr36 veriT_vr80) :named @p_667)) :rule refl) +(step t35.t2.t38.t21.t6 (cl (! (= veriT_vr37 veriT_vr81) :named @p_669)) :rule refl) +(step t35.t2.t38.t21.t7 (cl (= @p_405 @p_665)) :rule cong :premises (t35.t2.t38.t21.t5 t35.t2.t38.t21.t6)) +(step t35.t2.t38.t21.t8 (cl (= @p_406 @p_666)) :rule cong :premises (t35.t2.t38.t21.t4 t35.t2.t38.t21.t7)) +(step t35.t2.t38.t21.t9 (cl @p_667) :rule refl) +(step t35.t2.t38.t21.t10 (cl @p_557) :rule refl) +(step t35.t2.t38.t21.t11 (cl (= @p_408 @p_668)) :rule cong :premises (t35.t2.t38.t21.t9 t35.t2.t38.t21.t10)) +(step t35.t2.t38.t21.t12 (cl @p_669) :rule refl) +(step t35.t2.t38.t21.t13 (cl (= @p_410 @p_670)) :rule cong :premises (t35.t2.t38.t21.t11 t35.t2.t38.t21.t12)) +(step t35.t2.t38.t21.t14 (cl @p_656) :rule refl) +(step t35.t2.t38.t21.t15 (cl @p_657) :rule refl) +(step t35.t2.t38.t21.t16 (cl @p_658) :rule cong :premises (t35.t2.t38.t21.t14 t35.t2.t38.t21.t15)) +(step t35.t2.t38.t21.t17 (cl (= @p_414 @p_671)) :rule cong :premises (t35.t2.t38.t21.t13 t35.t2.t38.t21.t16)) +(step t35.t2.t38.t21.t18 (cl (= @p_415 @p_672)) :rule cong :premises (t35.t2.t38.t21.t8 t35.t2.t38.t21.t17)) +(step t35.t2.t38.t21 (cl (= @p_404 @p_664)) :rule bind) +(step t35.t2.t38.t22 (cl (= @p_416 @p_673)) :rule cong :premises (t35.t2.t38.t20 t35.t2.t38.t21)) +(step t35.t2.t38.t23 (cl @p_563) :rule refl) +(step t35.t2.t38.t24 (cl @p_590) :rule cong :premises (t35.t2.t38.t23)) +(step t35.t2.t38.t25 (cl @p_578) :rule refl) +(step t35.t2.t38.t26 (cl @p_591) :rule cong :premises (t35.t2.t38.t24 t35.t2.t38.t25)) +(step t35.t2.t38.t27 (cl @p_624) :rule cong :premises (t35.t2.t38.t26)) +(anchor :step t35.t2.t38.t28 :args ((:= (veriT_vr34 E$) veriT_vr82) (:= (veriT_vr35 D$) veriT_vr83))) +(step t35.t2.t38.t28.t1 (cl @p_563) :rule refl) +(step t35.t2.t38.t28.t2 (cl @p_590) :rule cong :premises (t35.t2.t38.t28.t1)) +(step t35.t2.t38.t28.t3 (cl @p_578) :rule refl) +(step t35.t2.t38.t28.t4 (cl @p_591) :rule cong :premises (t35.t2.t38.t28.t2 t35.t2.t38.t28.t3)) +(step t35.t2.t38.t28.t5 (cl (! (= veriT_vr34 veriT_vr82) :named @p_677)) :rule refl) +(step t35.t2.t38.t28.t6 (cl (! (= veriT_vr35 veriT_vr83) :named @p_679)) :rule refl) +(step t35.t2.t38.t28.t7 (cl (= @p_395 @p_675)) :rule cong :premises (t35.t2.t38.t28.t5 t35.t2.t38.t28.t6)) +(step t35.t2.t38.t28.t8 (cl (= @p_396 @p_676)) :rule cong :premises (t35.t2.t38.t28.t4 t35.t2.t38.t28.t7)) +(step t35.t2.t38.t28.t9 (cl @p_677) :rule refl) +(step t35.t2.t38.t28.t10 (cl @p_563) :rule refl) +(step t35.t2.t38.t28.t11 (cl (= @p_398 @p_678)) :rule cong :premises (t35.t2.t38.t28.t9 t35.t2.t38.t28.t10)) +(step t35.t2.t38.t28.t12 (cl @p_679) :rule refl) +(step t35.t2.t38.t28.t13 (cl (= @p_400 @p_680)) :rule cong :premises (t35.t2.t38.t28.t11 t35.t2.t38.t28.t12)) +(step t35.t2.t38.t28.t14 (cl @p_656) :rule refl) +(step t35.t2.t38.t28.t15 (cl @p_657) :rule refl) +(step t35.t2.t38.t28.t16 (cl @p_658) :rule cong :premises (t35.t2.t38.t28.t14 t35.t2.t38.t28.t15)) +(step t35.t2.t38.t28.t17 (cl (= @p_401 @p_681)) :rule cong :premises (t35.t2.t38.t28.t13 t35.t2.t38.t28.t16)) +(step t35.t2.t38.t28.t18 (cl (= @p_402 @p_682)) :rule cong :premises (t35.t2.t38.t28.t8 t35.t2.t38.t28.t17)) +(step t35.t2.t38.t28 (cl (= @p_394 @p_674)) :rule bind) +(step t35.t2.t38.t29 (cl (= @p_403 @p_683)) :rule cong :premises (t35.t2.t38.t27 t35.t2.t38.t28)) +(step t35.t2.t38.t30 (cl (= @p_684 @p_685)) :rule cong :premises (t35.t2.t38.t22 t35.t2.t38.t29)) +(step t35.t2.t38.t31 (cl (= @p_429 @p_686)) :rule cong :premises (t35.t2.t38.t15 t35.t2.t38.t30)) +(step t35.t2.t38 (cl (= @p_428 @p_638)) :rule bind) +(step t35.t2.t39 (cl (= @p_430 @p_687)) :rule cong :premises (t35.t2.t37 t35.t2.t38)) +(step t35.t2.t40 (cl (= @p_431 @p_688)) :rule cong :premises (t35.t2.t5 t35.t2.t6 t35.t2.t39)) +(step t35.t2 (cl (= @p_426 @p_574)) :rule bind) +(step t35.t3 (cl (= @p_432 @p_689)) :rule cong :premises (t35.t1 t35.t2)) +(step t35 (cl (= @p_433 (! (forall ((veriT_vr57 A_b_c_M_state_fun$) (veriT_vr58 A_b_c_M_state_fun$)) @p_689) :named @p_690))) :rule bind) +(step t36 (cl (= @p_434 (! (not @p_690) :named @p_693))) :rule cong :premises (t35)) +(step t37 (cl (! (= @p_436 (! (and @p_691 @p_692 @p_693) :named @p_695)) :named @p_694)) :rule cong :premises (t33 t34 t36)) +(step t38 (cl (not @p_694) (not @p_436) @p_695) :rule equiv_pos2) +(step t39 (cl @p_695) :rule th_resolution :premises (t32 t37 t38)) +(anchor :step t40 :args ((:= (veriT_vr57 A_b_c_M_state_fun$) veriT_sk0) (:= (veriT_vr58 A_b_c_M_state_fun$) veriT_sk1))) +(anchor :step t40.t1 :args ((veriT_vr59 A$) (veriT_vr60 C$))) +(step t40.t1.t1 (cl (! (= veriT_vr57 veriT_sk0) :named @p_754)) :rule refl) +(step t40.t1.t2 (cl (! (= @p_552 (! (fun_app$ veriT_sk0 veriT_vr59) :named @p_751)) :named @p_755)) :rule cong :premises (t40.t1.t1)) +(step t40.t1.t3 (cl (! (= @p_544 (! (run$ @p_751 veriT_vr60) :named @p_703)) :named @p_756)) :rule cong :premises (t40.t1.t2)) +(step t40.t1.t4 (cl (! (= @p_542 (! (is_fail$ @p_703) :named @p_701)) :named @p_757)) :rule cong :premises (t40.t1.t3)) +(step t40.t1.t5 (cl (! (= veriT_vr58 veriT_sk1) :named @p_760)) :rule refl) +(step t40.t1.t6 (cl (! (= @p_554 (! (fun_app$ veriT_sk1 veriT_vr59) :named @p_752)) :named @p_761)) :rule cong :premises (t40.t1.t5)) +(step t40.t1.t7 (cl (! (= @p_543 (! (run$ @p_752 veriT_vr60) :named @p_702)) :named @p_762)) :rule cong :premises (t40.t1.t6)) +(step t40.t1.t8 (cl (= @p_556 (! (is_fail$ @p_702) :named @p_753))) :rule cong :premises (t40.t1.t7)) +(step t40.t1.t9 (cl @p_754) :rule refl) +(step t40.t1.t10 (cl @p_755) :rule cong :premises (t40.t1.t9)) +(step t40.t1.t11 (cl @p_756) :rule cong :premises (t40.t1.t10)) +(step t40.t1.t12 (cl @p_757) :rule cong :premises (t40.t1.t11)) +(step t40.t1.t13 (cl (= @p_561 (! (= @p_753 @p_701) :named @p_758))) :rule cong :premises (t40.t1.t8 t40.t1.t12)) +(anchor :step t40.t1.t14 :args ((veriT_vr61 B$) (veriT_vr62 C$))) +(step t40.t1.t14.t1 (cl @p_760) :rule refl) +(step t40.t1.t14.t2 (cl @p_761) :rule cong :premises (t40.t1.t14.t1)) +(step t40.t1.t14.t3 (cl @p_762) :rule cong :premises (t40.t1.t14.t2)) +(step t40.t1.t14.t4 (cl (= @p_566 (! (is_res$ @p_702 @p_545) :named @p_763))) :rule cong :premises (t40.t1.t14.t3)) +(step t40.t1.t14.t5 (cl @p_754) :rule refl) +(step t40.t1.t14.t6 (cl @p_755) :rule cong :premises (t40.t1.t14.t5)) +(step t40.t1.t14.t7 (cl @p_756) :rule cong :premises (t40.t1.t14.t6)) +(step t40.t1.t14.t8 (cl (= @p_570 (! (is_res$ @p_703 @p_545) :named @p_764))) :rule cong :premises (t40.t1.t14.t7)) +(step t40.t1.t14.t9 (cl (= @p_571 (! (= @p_763 @p_764) :named @p_765))) :rule cong :premises (t40.t1.t14.t4 t40.t1.t14.t8)) +(step t40.t1.t14 (cl (= @p_562 (! (forall ((veriT_vr61 B$) (veriT_vr62 C$)) @p_765) :named @p_759))) :rule bind) +(step t40.t1.t15 (cl (= @p_572 (! (and @p_758 @p_759) :named @p_766))) :rule cong :premises (t40.t1.t13 t40.t1.t14)) +(step t40.t1.t16 (cl (= @p_573 (! (or @p_701 @p_766) :named @p_767))) :rule cong :premises (t40.t1.t4 t40.t1.t15)) +(step t40.t1 (cl (= @p_551 (! (forall ((veriT_vr59 A$) (veriT_vr60 C$)) @p_767) :named @p_750))) :rule bind) +(anchor :step t40.t2 :args ((:= (veriT_vr63 D$) veriT_sk2))) +(step t40.t2.t1 (cl @p_754) :rule refl) +(step t40.t2.t2 (cl (! (= @p_575 (! (b$ veriT_sk0) :named @p_769)) :named @p_771)) :rule cong :premises (t40.t2.t1)) +(step t40.t2.t3 (cl (! (= veriT_vr63 veriT_sk2) :named @p_772)) :rule refl) +(step t40.t2.t4 (cl (! (= @p_546 (! (run$a @p_769 veriT_sk2) :named @p_712)) :named @p_773)) :rule cong :premises (t40.t2.t2 t40.t2.t3)) +(step t40.t2.t5 (cl (! (= @p_548 (! (is_fail$a @p_712) :named @p_720)) :named @p_791)) :rule cong :premises (t40.t2.t4)) +(anchor :step t40.t2.t6 :args ((veriT_vr64 E$) (veriT_vr65 D$))) +(step t40.t2.t6.t1 (cl @p_754) :rule refl) +(step t40.t2.t6.t2 (cl @p_771) :rule cong :premises (t40.t2.t6.t1)) +(step t40.t2.t6.t3 (cl @p_772) :rule refl) +(step t40.t2.t6.t4 (cl @p_773) :rule cong :premises (t40.t2.t6.t2 t40.t2.t6.t3)) +(step t40.t2.t6.t5 (cl (= @p_581 (! (is_res$a @p_712 @p_580) :named @p_774))) :rule cong :premises (t40.t2.t6.t4)) +(step t40.t2.t6.t6 (cl @p_754) :rule refl) +(step t40.t2.t6.t7 (cl (= @p_583 (! (c$ veriT_vr64 veriT_sk0) :named @p_775))) :rule cong :premises (t40.t2.t6.t6)) +(step t40.t2.t6.t8 (cl (= @p_585 (! (run$b @p_775 veriT_vr65) :named @p_776))) :rule cong :premises (t40.t2.t6.t7)) +(step t40.t2.t6.t9 (cl (= @p_586 (! (is_fail$b @p_776) :named @p_777))) :rule cong :premises (t40.t2.t6.t8)) +(step t40.t2.t6.t10 (cl (= @p_587 (! (and @p_774 @p_777) :named @p_778))) :rule cong :premises (t40.t2.t6.t5 t40.t2.t6.t9)) +(step t40.t2.t6 (cl (= @p_576 (! (exists ((veriT_vr64 E$) (veriT_vr65 D$)) @p_778) :named @p_770))) :rule bind) +(step t40.t2.t7 (cl @p_760) :rule refl) +(step t40.t2.t8 (cl (! (= @p_588 (! (b$ veriT_sk1) :named @p_779)) :named @p_781)) :rule cong :premises (t40.t2.t7)) +(step t40.t2.t9 (cl @p_772) :rule refl) +(step t40.t2.t10 (cl (! (= @p_547 (! (run$a @p_779 veriT_sk2) :named @p_714)) :named @p_782)) :rule cong :premises (t40.t2.t8 t40.t2.t9)) +(step t40.t2.t11 (cl (! (= @p_549 (! (is_fail$a @p_714) :named @p_726)) :named @p_809)) :rule cong :premises (t40.t2.t10)) +(anchor :step t40.t2.t12 :args ((:= (veriT_vr66 E$) veriT_sk3) (:= (veriT_vr67 D$) veriT_sk4))) +(step t40.t2.t12.t1 (cl @p_760) :rule refl) +(step t40.t2.t12.t2 (cl @p_781) :rule cong :premises (t40.t2.t12.t1)) +(step t40.t2.t12.t3 (cl @p_772) :rule refl) +(step t40.t2.t12.t4 (cl @p_782) :rule cong :premises (t40.t2.t12.t2 t40.t2.t12.t3)) +(step t40.t2.t12.t5 (cl (! (= veriT_vr66 veriT_sk3) :named @p_785)) :rule refl) +(step t40.t2.t12.t6 (cl (! (= veriT_vr67 veriT_sk4) :named @p_787)) :rule refl) +(step t40.t2.t12.t7 (cl (= @p_592 (! (pair$a veriT_sk3 veriT_sk4) :named @p_783))) :rule cong :premises (t40.t2.t12.t5 t40.t2.t12.t6)) +(step t40.t2.t12.t8 (cl (= @p_593 (! (is_res$a @p_714 @p_783) :named @p_784))) :rule cong :premises (t40.t2.t12.t4 t40.t2.t12.t7)) +(step t40.t2.t12.t9 (cl @p_785) :rule refl) +(step t40.t2.t12.t10 (cl @p_760) :rule refl) +(step t40.t2.t12.t11 (cl (= @p_595 (! (c$ veriT_sk3 veriT_sk1) :named @p_786))) :rule cong :premises (t40.t2.t12.t9 t40.t2.t12.t10)) +(step t40.t2.t12.t12 (cl @p_787) :rule refl) +(step t40.t2.t12.t13 (cl (= @p_597 (! (run$b @p_786 veriT_sk4) :named @p_788))) :rule cong :premises (t40.t2.t12.t11 t40.t2.t12.t12)) +(step t40.t2.t12.t14 (cl (= @p_598 (! (is_fail$b @p_788) :named @p_789))) :rule cong :premises (t40.t2.t12.t13)) +(step t40.t2.t12.t15 (cl (= @p_599 (! (and @p_784 @p_789) :named @p_780))) :rule cong :premises (t40.t2.t12.t8 t40.t2.t12.t14)) +(step t40.t2.t12 (cl (= @p_589 @p_780)) :rule sko_ex) +(step t40.t2.t13 (cl (= @p_600 (! (or @p_726 @p_780) :named @p_790))) :rule cong :premises (t40.t2.t11 t40.t2.t12)) +(step t40.t2.t14 (cl @p_754) :rule refl) +(step t40.t2.t15 (cl @p_771) :rule cong :premises (t40.t2.t14)) +(step t40.t2.t16 (cl @p_772) :rule refl) +(step t40.t2.t17 (cl @p_773) :rule cong :premises (t40.t2.t15 t40.t2.t16)) +(step t40.t2.t18 (cl @p_791) :rule cong :premises (t40.t2.t17)) +(anchor :step t40.t2.t19 :args ((veriT_vr68 E$) (veriT_vr69 D$))) +(step t40.t2.t19.t1 (cl @p_754) :rule refl) +(step t40.t2.t19.t2 (cl @p_771) :rule cong :premises (t40.t2.t19.t1)) +(step t40.t2.t19.t3 (cl @p_772) :rule refl) +(step t40.t2.t19.t4 (cl @p_773) :rule cong :premises (t40.t2.t19.t2 t40.t2.t19.t3)) +(step t40.t2.t19.t5 (cl (= @p_604 (! (is_res$a @p_712 @p_603) :named @p_793))) :rule cong :premises (t40.t2.t19.t4)) +(step t40.t2.t19.t6 (cl @p_754) :rule refl) +(step t40.t2.t19.t7 (cl (= @p_606 (! (c$ veriT_vr68 veriT_sk0) :named @p_794))) :rule cong :premises (t40.t2.t19.t6)) +(step t40.t2.t19.t8 (cl (= @p_608 (! (run$b @p_794 veriT_vr69) :named @p_795))) :rule cong :premises (t40.t2.t19.t7)) +(step t40.t2.t19.t9 (cl (= @p_609 (! (is_fail$b @p_795) :named @p_796))) :rule cong :premises (t40.t2.t19.t8)) +(step t40.t2.t19.t10 (cl (= @p_610 (! (and @p_793 @p_796) :named @p_797))) :rule cong :premises (t40.t2.t19.t5 t40.t2.t19.t9)) +(step t40.t2.t19 (cl (= @p_602 (! (exists ((veriT_vr68 E$) (veriT_vr69 D$)) @p_797) :named @p_792))) :rule bind) +(step t40.t2.t20 (cl (= @p_611 (! (or @p_720 @p_792) :named @p_798))) :rule cong :premises (t40.t2.t18 t40.t2.t19)) +(step t40.t2.t21 (cl (= @p_613 (! (=> @p_790 @p_798) :named @p_799))) :rule cong :premises (t40.t2.t13 t40.t2.t20)) +(step t40.t2.t22 (cl @p_754) :rule refl) +(step t40.t2.t23 (cl @p_771) :rule cong :premises (t40.t2.t22)) +(step t40.t2.t24 (cl @p_772) :rule refl) +(step t40.t2.t25 (cl @p_773) :rule cong :premises (t40.t2.t23 t40.t2.t24)) +(step t40.t2.t26 (cl @p_791) :rule cong :premises (t40.t2.t25)) +(anchor :step t40.t2.t27 :args ((:= (veriT_vr70 E$) veriT_sk5) (:= (veriT_vr71 D$) veriT_sk6))) +(step t40.t2.t27.t1 (cl @p_754) :rule refl) +(step t40.t2.t27.t2 (cl @p_771) :rule cong :premises (t40.t2.t27.t1)) +(step t40.t2.t27.t3 (cl @p_772) :rule refl) +(step t40.t2.t27.t4 (cl @p_773) :rule cong :premises (t40.t2.t27.t2 t40.t2.t27.t3)) +(step t40.t2.t27.t5 (cl (! (= veriT_vr70 veriT_sk5) :named @p_803)) :rule refl) +(step t40.t2.t27.t6 (cl (! (= veriT_vr71 veriT_sk6) :named @p_805)) :rule refl) +(step t40.t2.t27.t7 (cl (= @p_615 (! (pair$a veriT_sk5 veriT_sk6) :named @p_801))) :rule cong :premises (t40.t2.t27.t5 t40.t2.t27.t6)) +(step t40.t2.t27.t8 (cl (= @p_616 (! (is_res$a @p_712 @p_801) :named @p_802))) :rule cong :premises (t40.t2.t27.t4 t40.t2.t27.t7)) +(step t40.t2.t27.t9 (cl @p_803) :rule refl) +(step t40.t2.t27.t10 (cl @p_754) :rule refl) +(step t40.t2.t27.t11 (cl (= @p_618 (! (c$ veriT_sk5 veriT_sk0) :named @p_804))) :rule cong :premises (t40.t2.t27.t9 t40.t2.t27.t10)) +(step t40.t2.t27.t12 (cl @p_805) :rule refl) +(step t40.t2.t27.t13 (cl (= @p_620 (! (run$b @p_804 veriT_sk6) :named @p_806))) :rule cong :premises (t40.t2.t27.t11 t40.t2.t27.t12)) +(step t40.t2.t27.t14 (cl (= @p_621 (! (is_fail$b @p_806) :named @p_807))) :rule cong :premises (t40.t2.t27.t13)) +(step t40.t2.t27.t15 (cl (= @p_622 (! (and @p_802 @p_807) :named @p_800))) :rule cong :premises (t40.t2.t27.t8 t40.t2.t27.t14)) +(step t40.t2.t27 (cl (= @p_614 @p_800)) :rule sko_ex) +(step t40.t2.t28 (cl (= @p_623 (! (or @p_720 @p_800) :named @p_808))) :rule cong :premises (t40.t2.t26 t40.t2.t27)) +(step t40.t2.t29 (cl @p_760) :rule refl) +(step t40.t2.t30 (cl @p_781) :rule cong :premises (t40.t2.t29)) +(step t40.t2.t31 (cl @p_772) :rule refl) +(step t40.t2.t32 (cl @p_782) :rule cong :premises (t40.t2.t30 t40.t2.t31)) +(step t40.t2.t33 (cl @p_809) :rule cong :premises (t40.t2.t32)) +(anchor :step t40.t2.t34 :args ((veriT_vr72 E$) (veriT_vr73 D$))) +(step t40.t2.t34.t1 (cl @p_760) :rule refl) +(step t40.t2.t34.t2 (cl @p_781) :rule cong :premises (t40.t2.t34.t1)) +(step t40.t2.t34.t3 (cl @p_772) :rule refl) +(step t40.t2.t34.t4 (cl @p_782) :rule cong :premises (t40.t2.t34.t2 t40.t2.t34.t3)) +(step t40.t2.t34.t5 (cl (= @p_627 (! (is_res$a @p_714 @p_626) :named @p_811))) :rule cong :premises (t40.t2.t34.t4)) +(step t40.t2.t34.t6 (cl @p_760) :rule refl) +(step t40.t2.t34.t7 (cl (= @p_629 (! (c$ veriT_vr72 veriT_sk1) :named @p_812))) :rule cong :premises (t40.t2.t34.t6)) +(step t40.t2.t34.t8 (cl (= @p_631 (! (run$b @p_812 veriT_vr73) :named @p_813))) :rule cong :premises (t40.t2.t34.t7)) +(step t40.t2.t34.t9 (cl (= @p_632 (! (is_fail$b @p_813) :named @p_814))) :rule cong :premises (t40.t2.t34.t8)) +(step t40.t2.t34.t10 (cl (= @p_633 (! (and @p_811 @p_814) :named @p_815))) :rule cong :premises (t40.t2.t34.t5 t40.t2.t34.t9)) +(step t40.t2.t34 (cl (= @p_625 (! (exists ((veriT_vr72 E$) (veriT_vr73 D$)) @p_815) :named @p_810))) :rule bind) +(step t40.t2.t35 (cl (= @p_634 (! (or @p_726 @p_810) :named @p_816))) :rule cong :premises (t40.t2.t33 t40.t2.t34)) +(step t40.t2.t36 (cl (= @p_636 (! (=> @p_808 @p_816) :named @p_817))) :rule cong :premises (t40.t2.t28 t40.t2.t35)) +(step t40.t2.t37 (cl (= @p_637 (! (and @p_799 @p_817) :named @p_818))) :rule cong :premises (t40.t2.t21 t40.t2.t36)) +(anchor :step t40.t2.t38 :args ((:= (veriT_vr74 F$) veriT_sk7) (:= (veriT_vr75 D$) veriT_sk8))) +(step t40.t2.t38.t1 (cl @p_760) :rule refl) +(step t40.t2.t38.t2 (cl @p_781) :rule cong :premises (t40.t2.t38.t1)) +(step t40.t2.t38.t3 (cl @p_772) :rule refl) +(step t40.t2.t38.t4 (cl @p_782) :rule cong :premises (t40.t2.t38.t2 t40.t2.t38.t3)) +(step t40.t2.t38.t5 (cl @p_809) :rule cong :premises (t40.t2.t38.t4)) +(anchor :step t40.t2.t38.t6 :args ((:= (veriT_vr76 E$) veriT_sk9) (:= (veriT_vr77 D$) veriT_sk10))) +(step t40.t2.t38.t6.t1 (cl @p_760) :rule refl) +(step t40.t2.t38.t6.t2 (cl @p_781) :rule cong :premises (t40.t2.t38.t6.t1)) +(step t40.t2.t38.t6.t3 (cl @p_772) :rule refl) +(step t40.t2.t38.t6.t4 (cl @p_782) :rule cong :premises (t40.t2.t38.t6.t2 t40.t2.t38.t6.t3)) +(step t40.t2.t38.t6.t5 (cl (! (= veriT_vr76 veriT_sk9) :named @p_823)) :rule refl) +(step t40.t2.t38.t6.t6 (cl (! (= veriT_vr77 veriT_sk10) :named @p_825)) :rule refl) +(step t40.t2.t38.t6.t7 (cl (= @p_640 (! (pair$a veriT_sk9 veriT_sk10) :named @p_821))) :rule cong :premises (t40.t2.t38.t6.t5 t40.t2.t38.t6.t6)) +(step t40.t2.t38.t6.t8 (cl (= @p_641 (! (is_res$a @p_714 @p_821) :named @p_822))) :rule cong :premises (t40.t2.t38.t6.t4 t40.t2.t38.t6.t7)) +(step t40.t2.t38.t6.t9 (cl @p_823) :rule refl) +(step t40.t2.t38.t6.t10 (cl @p_760) :rule refl) +(step t40.t2.t38.t6.t11 (cl (= @p_643 (! (c$ veriT_sk9 veriT_sk1) :named @p_824))) :rule cong :premises (t40.t2.t38.t6.t9 t40.t2.t38.t6.t10)) +(step t40.t2.t38.t6.t12 (cl @p_825) :rule refl) +(step t40.t2.t38.t6.t13 (cl (= @p_645 (! (run$b @p_824 veriT_sk10) :named @p_826))) :rule cong :premises (t40.t2.t38.t6.t11 t40.t2.t38.t6.t12)) +(step t40.t2.t38.t6.t14 (cl (! (= veriT_vr74 veriT_sk7) :named @p_833)) :rule refl) +(step t40.t2.t38.t6.t15 (cl (! (= veriT_vr75 veriT_sk8) :named @p_834)) :rule refl) +(step t40.t2.t38.t6.t16 (cl (! (= @p_550 (! (pair$b veriT_sk7 veriT_sk8) :named @p_747)) :named @p_835)) :rule cong :premises (t40.t2.t38.t6.t14 t40.t2.t38.t6.t15)) +(step t40.t2.t38.t6.t17 (cl (= @p_646 (! (is_res$b @p_826 @p_747) :named @p_827))) :rule cong :premises (t40.t2.t38.t6.t13 t40.t2.t38.t6.t16)) +(step t40.t2.t38.t6.t18 (cl (= @p_647 (! (and @p_822 @p_827) :named @p_820))) :rule cong :premises (t40.t2.t38.t6.t8 t40.t2.t38.t6.t17)) +(step t40.t2.t38.t6 (cl (= @p_639 @p_820)) :rule sko_ex) +(step t40.t2.t38.t7 (cl (= @p_648 (! (or @p_726 @p_820) :named @p_828))) :rule cong :premises (t40.t2.t38.t5 t40.t2.t38.t6)) +(step t40.t2.t38.t8 (cl @p_754) :rule refl) +(step t40.t2.t38.t9 (cl @p_771) :rule cong :premises (t40.t2.t38.t8)) +(step t40.t2.t38.t10 (cl @p_772) :rule refl) +(step t40.t2.t38.t11 (cl @p_773) :rule cong :premises (t40.t2.t38.t9 t40.t2.t38.t10)) +(step t40.t2.t38.t12 (cl @p_791) :rule cong :premises (t40.t2.t38.t11)) +(anchor :step t40.t2.t38.t13 :args ((veriT_vr78 E$) (veriT_vr79 D$))) +(step t40.t2.t38.t13.t1 (cl @p_754) :rule refl) +(step t40.t2.t38.t13.t2 (cl @p_771) :rule cong :premises (t40.t2.t38.t13.t1)) +(step t40.t2.t38.t13.t3 (cl @p_772) :rule refl) +(step t40.t2.t38.t13.t4 (cl @p_773) :rule cong :premises (t40.t2.t38.t13.t2 t40.t2.t38.t13.t3)) +(step t40.t2.t38.t13.t5 (cl (= @p_651 (! (is_res$a @p_712 @p_650) :named @p_830))) :rule cong :premises (t40.t2.t38.t13.t4)) +(step t40.t2.t38.t13.t6 (cl @p_754) :rule refl) +(step t40.t2.t38.t13.t7 (cl (= @p_653 (! (c$ veriT_vr78 veriT_sk0) :named @p_831))) :rule cong :premises (t40.t2.t38.t13.t6)) +(step t40.t2.t38.t13.t8 (cl (= @p_655 (! (run$b @p_831 veriT_vr79) :named @p_832))) :rule cong :premises (t40.t2.t38.t13.t7)) +(step t40.t2.t38.t13.t9 (cl @p_833) :rule refl) +(step t40.t2.t38.t13.t10 (cl @p_834) :rule refl) +(step t40.t2.t38.t13.t11 (cl @p_835) :rule cong :premises (t40.t2.t38.t13.t9 t40.t2.t38.t13.t10)) +(step t40.t2.t38.t13.t12 (cl (= @p_659 (! (is_res$b @p_832 @p_747) :named @p_836))) :rule cong :premises (t40.t2.t38.t13.t8 t40.t2.t38.t13.t11)) +(step t40.t2.t38.t13.t13 (cl (= @p_660 (! (and @p_830 @p_836) :named @p_837))) :rule cong :premises (t40.t2.t38.t13.t5 t40.t2.t38.t13.t12)) +(step t40.t2.t38.t13 (cl (= @p_649 (! (exists ((veriT_vr78 E$) (veriT_vr79 D$)) @p_837) :named @p_829))) :rule bind) +(step t40.t2.t38.t14 (cl (= @p_661 (! (or @p_720 @p_829) :named @p_838))) :rule cong :premises (t40.t2.t38.t12 t40.t2.t38.t13)) +(step t40.t2.t38.t15 (cl (= @p_663 (! (=> @p_828 @p_838) :named @p_839))) :rule cong :premises (t40.t2.t38.t7 t40.t2.t38.t14)) +(step t40.t2.t38.t16 (cl @p_754) :rule refl) +(step t40.t2.t38.t17 (cl @p_771) :rule cong :premises (t40.t2.t38.t16)) +(step t40.t2.t38.t18 (cl @p_772) :rule refl) +(step t40.t2.t38.t19 (cl @p_773) :rule cong :premises (t40.t2.t38.t17 t40.t2.t38.t18)) +(step t40.t2.t38.t20 (cl @p_791) :rule cong :premises (t40.t2.t38.t19)) +(anchor :step t40.t2.t38.t21 :args ((:= (veriT_vr80 E$) veriT_sk11) (:= (veriT_vr81 D$) veriT_sk12))) +(step t40.t2.t38.t21.t1 (cl @p_754) :rule refl) +(step t40.t2.t38.t21.t2 (cl @p_771) :rule cong :premises (t40.t2.t38.t21.t1)) +(step t40.t2.t38.t21.t3 (cl @p_772) :rule refl) +(step t40.t2.t38.t21.t4 (cl @p_773) :rule cong :premises (t40.t2.t38.t21.t2 t40.t2.t38.t21.t3)) +(step t40.t2.t38.t21.t5 (cl (! (= veriT_vr80 veriT_sk11) :named @p_843)) :rule refl) +(step t40.t2.t38.t21.t6 (cl (! (= veriT_vr81 veriT_sk12) :named @p_845)) :rule refl) +(step t40.t2.t38.t21.t7 (cl (= @p_665 (! (pair$a veriT_sk11 veriT_sk12) :named @p_841))) :rule cong :premises (t40.t2.t38.t21.t5 t40.t2.t38.t21.t6)) +(step t40.t2.t38.t21.t8 (cl (= @p_666 (! (is_res$a @p_712 @p_841) :named @p_842))) :rule cong :premises (t40.t2.t38.t21.t4 t40.t2.t38.t21.t7)) +(step t40.t2.t38.t21.t9 (cl @p_843) :rule refl) +(step t40.t2.t38.t21.t10 (cl @p_754) :rule refl) +(step t40.t2.t38.t21.t11 (cl (= @p_668 (! (c$ veriT_sk11 veriT_sk0) :named @p_844))) :rule cong :premises (t40.t2.t38.t21.t9 t40.t2.t38.t21.t10)) +(step t40.t2.t38.t21.t12 (cl @p_845) :rule refl) +(step t40.t2.t38.t21.t13 (cl (= @p_670 (! (run$b @p_844 veriT_sk12) :named @p_846))) :rule cong :premises (t40.t2.t38.t21.t11 t40.t2.t38.t21.t12)) +(step t40.t2.t38.t21.t14 (cl @p_833) :rule refl) +(step t40.t2.t38.t21.t15 (cl @p_834) :rule refl) +(step t40.t2.t38.t21.t16 (cl @p_835) :rule cong :premises (t40.t2.t38.t21.t14 t40.t2.t38.t21.t15)) +(step t40.t2.t38.t21.t17 (cl (= @p_671 (! (is_res$b @p_846 @p_747) :named @p_847))) :rule cong :premises (t40.t2.t38.t21.t13 t40.t2.t38.t21.t16)) +(step t40.t2.t38.t21.t18 (cl (= @p_672 (! (and @p_842 @p_847) :named @p_840))) :rule cong :premises (t40.t2.t38.t21.t8 t40.t2.t38.t21.t17)) +(step t40.t2.t38.t21 (cl (= @p_664 @p_840)) :rule sko_ex) +(step t40.t2.t38.t22 (cl (= @p_673 (! (or @p_720 @p_840) :named @p_848))) :rule cong :premises (t40.t2.t38.t20 t40.t2.t38.t21)) +(step t40.t2.t38.t23 (cl @p_760) :rule refl) +(step t40.t2.t38.t24 (cl @p_781) :rule cong :premises (t40.t2.t38.t23)) +(step t40.t2.t38.t25 (cl @p_772) :rule refl) +(step t40.t2.t38.t26 (cl @p_782) :rule cong :premises (t40.t2.t38.t24 t40.t2.t38.t25)) +(step t40.t2.t38.t27 (cl @p_809) :rule cong :premises (t40.t2.t38.t26)) +(anchor :step t40.t2.t38.t28 :args ((veriT_vr82 E$) (veriT_vr83 D$))) +(step t40.t2.t38.t28.t1 (cl @p_760) :rule refl) +(step t40.t2.t38.t28.t2 (cl @p_781) :rule cong :premises (t40.t2.t38.t28.t1)) +(step t40.t2.t38.t28.t3 (cl @p_772) :rule refl) +(step t40.t2.t38.t28.t4 (cl @p_782) :rule cong :premises (t40.t2.t38.t28.t2 t40.t2.t38.t28.t3)) +(step t40.t2.t38.t28.t5 (cl (= @p_676 (! (is_res$a @p_714 @p_675) :named @p_850))) :rule cong :premises (t40.t2.t38.t28.t4)) +(step t40.t2.t38.t28.t6 (cl @p_760) :rule refl) +(step t40.t2.t38.t28.t7 (cl (= @p_678 (! (c$ veriT_vr82 veriT_sk1) :named @p_851))) :rule cong :premises (t40.t2.t38.t28.t6)) +(step t40.t2.t38.t28.t8 (cl (= @p_680 (! (run$b @p_851 veriT_vr83) :named @p_852))) :rule cong :premises (t40.t2.t38.t28.t7)) +(step t40.t2.t38.t28.t9 (cl @p_833) :rule refl) +(step t40.t2.t38.t28.t10 (cl @p_834) :rule refl) +(step t40.t2.t38.t28.t11 (cl @p_835) :rule cong :premises (t40.t2.t38.t28.t9 t40.t2.t38.t28.t10)) +(step t40.t2.t38.t28.t12 (cl (= @p_681 (! (is_res$b @p_852 @p_747) :named @p_853))) :rule cong :premises (t40.t2.t38.t28.t8 t40.t2.t38.t28.t11)) +(step t40.t2.t38.t28.t13 (cl (= @p_682 (! (and @p_850 @p_853) :named @p_854))) :rule cong :premises (t40.t2.t38.t28.t5 t40.t2.t38.t28.t12)) +(step t40.t2.t38.t28 (cl (= @p_674 (! (exists ((veriT_vr82 E$) (veriT_vr83 D$)) @p_854) :named @p_849))) :rule bind) +(step t40.t2.t38.t29 (cl (= @p_683 (! (or @p_726 @p_849) :named @p_855))) :rule cong :premises (t40.t2.t38.t27 t40.t2.t38.t28)) +(step t40.t2.t38.t30 (cl (= @p_685 (! (=> @p_848 @p_855) :named @p_856))) :rule cong :premises (t40.t2.t38.t22 t40.t2.t38.t29)) +(step t40.t2.t38.t31 (cl (= @p_686 (! (and @p_839 @p_856) :named @p_819))) :rule cong :premises (t40.t2.t38.t15 t40.t2.t38.t30)) +(step t40.t2.t38 (cl (= @p_638 @p_819)) :rule sko_forall) +(step t40.t2.t39 (cl (= @p_687 (! (and @p_818 @p_819) :named @p_857))) :rule cong :premises (t40.t2.t37 t40.t2.t38)) +(step t40.t2.t40 (cl (= @p_688 (! (or @p_720 @p_770 @p_857) :named @p_768))) :rule cong :premises (t40.t2.t5 t40.t2.t6 t40.t2.t39)) +(step t40.t2 (cl (= @p_574 @p_768)) :rule sko_forall) +(step t40.t3 (cl (= @p_689 (! (=> @p_750 @p_768) :named @p_858))) :rule cong :premises (t40.t1 t40.t2)) +(step t40 (cl (= @p_690 @p_858)) :rule sko_forall) +(step t41 (cl (= @p_693 (! (not @p_858) :named @p_859))) :rule cong :premises (t40)) +(step t42 (cl (! (= @p_695 (! (and @p_691 @p_692 @p_859) :named @p_861)) :named @p_860)) :rule cong :premises (t41)) +(step t43 (cl (not @p_860) (not @p_695) @p_861) :rule equiv_pos2) +(step t44 (cl @p_861) :rule th_resolution :premises (t39 t42 t43)) +(anchor :step t45 :args ((:= (veriT_vr38 A_b_c_M_state_fun$) veriT_vr84) (:= (veriT_vr39 A_b_c_M_state_fun$) veriT_vr85))) +(anchor :step t45.t1 :args ((:= (veriT_vr40 A$) veriT_vr86) (:= (veriT_vr41 C$) veriT_vr87))) +(step t45.t1.t1 (cl (! (= veriT_vr38 veriT_vr84) :named @p_876)) :rule refl) +(step t45.t1.t2 (cl (! (= veriT_vr40 veriT_vr86) :named @p_872)) :rule refl) +(step t45.t1.t3 (cl (! (= @p_446 (! (fun_app$ veriT_vr84 veriT_vr86) :named @p_871)) :named @p_877)) :rule cong :premises (t45.t1.t1 t45.t1.t2)) +(step t45.t1.t4 (cl (! (= veriT_vr41 veriT_vr87) :named @p_874)) :rule refl) +(step t45.t1.t5 (cl (! (= @p_439 (! (run$ @p_871 veriT_vr87) :named @p_864)) :named @p_878)) :rule cong :premises (t45.t1.t3 t45.t1.t4)) +(step t45.t1.t6 (cl (! (= @p_437 (! (is_fail$ @p_864) :named @p_862)) :named @p_879)) :rule cong :premises (t45.t1.t5)) +(step t45.t1.t7 (cl (! (= veriT_vr39 veriT_vr85) :named @p_882)) :rule refl) +(step t45.t1.t8 (cl @p_872) :rule refl) +(step t45.t1.t9 (cl (! (= @p_448 (! (fun_app$ veriT_vr85 veriT_vr86) :named @p_873)) :named @p_883)) :rule cong :premises (t45.t1.t7 t45.t1.t8)) +(step t45.t1.t10 (cl @p_874) :rule refl) +(step t45.t1.t11 (cl (! (= @p_438 (! (run$ @p_873 veriT_vr87) :named @p_863)) :named @p_884)) :rule cong :premises (t45.t1.t9 t45.t1.t10)) +(step t45.t1.t12 (cl (= @p_450 (! (is_fail$ @p_863) :named @p_875))) :rule cong :premises (t45.t1.t11)) +(step t45.t1.t13 (cl @p_876) :rule refl) +(step t45.t1.t14 (cl @p_872) :rule refl) +(step t45.t1.t15 (cl @p_877) :rule cong :premises (t45.t1.t13 t45.t1.t14)) +(step t45.t1.t16 (cl @p_874) :rule refl) +(step t45.t1.t17 (cl @p_878) :rule cong :premises (t45.t1.t15 t45.t1.t16)) +(step t45.t1.t18 (cl @p_879) :rule cong :premises (t45.t1.t17)) +(step t45.t1.t19 (cl (= @p_455 (! (= @p_875 @p_862) :named @p_880))) :rule cong :premises (t45.t1.t12 t45.t1.t18)) +(anchor :step t45.t1.t20 :args ((:= (veriT_vr42 B$) veriT_vr88) (:= (veriT_vr43 C$) veriT_vr89))) +(step t45.t1.t20.t1 (cl @p_882) :rule refl) +(step t45.t1.t20.t2 (cl @p_872) :rule refl) +(step t45.t1.t20.t3 (cl @p_883) :rule cong :premises (t45.t1.t20.t1 t45.t1.t20.t2)) +(step t45.t1.t20.t4 (cl @p_874) :rule refl) +(step t45.t1.t20.t5 (cl @p_884) :rule cong :premises (t45.t1.t20.t3 t45.t1.t20.t4)) +(step t45.t1.t20.t6 (cl (! (= veriT_vr42 veriT_vr88) :named @p_886)) :rule refl) +(step t45.t1.t20.t7 (cl (! (= veriT_vr43 veriT_vr89) :named @p_887)) :rule refl) +(step t45.t1.t20.t8 (cl (! (= @p_440 (! (pair$ veriT_vr88 veriT_vr89) :named @p_865)) :named @p_888)) :rule cong :premises (t45.t1.t20.t6 t45.t1.t20.t7)) +(step t45.t1.t20.t9 (cl (= @p_460 (! (is_res$ @p_863 @p_865) :named @p_885))) :rule cong :premises (t45.t1.t20.t5 t45.t1.t20.t8)) +(step t45.t1.t20.t10 (cl @p_876) :rule refl) +(step t45.t1.t20.t11 (cl @p_872) :rule refl) +(step t45.t1.t20.t12 (cl @p_877) :rule cong :premises (t45.t1.t20.t10 t45.t1.t20.t11)) +(step t45.t1.t20.t13 (cl @p_874) :rule refl) +(step t45.t1.t20.t14 (cl @p_878) :rule cong :premises (t45.t1.t20.t12 t45.t1.t20.t13)) +(step t45.t1.t20.t15 (cl @p_886) :rule refl) +(step t45.t1.t20.t16 (cl @p_887) :rule refl) +(step t45.t1.t20.t17 (cl @p_888) :rule cong :premises (t45.t1.t20.t15 t45.t1.t20.t16)) +(step t45.t1.t20.t18 (cl (= @p_464 (! (is_res$ @p_864 @p_865) :named @p_889))) :rule cong :premises (t45.t1.t20.t14 t45.t1.t20.t17)) +(step t45.t1.t20.t19 (cl (= @p_465 (! (= @p_885 @p_889) :named @p_890))) :rule cong :premises (t45.t1.t20.t9 t45.t1.t20.t18)) +(step t45.t1.t20 (cl (= @p_456 (! (forall ((veriT_vr88 B$) (veriT_vr89 C$)) @p_890) :named @p_881))) :rule bind) +(step t45.t1.t21 (cl (= @p_466 (! (and @p_880 @p_881) :named @p_891))) :rule cong :premises (t45.t1.t19 t45.t1.t20)) +(step t45.t1.t22 (cl (= @p_467 (! (or @p_862 @p_891) :named @p_892))) :rule cong :premises (t45.t1.t6 t45.t1.t21)) +(step t45.t1 (cl (= @p_445 (! (forall ((veriT_vr86 A$) (veriT_vr87 C$)) @p_892) :named @p_870))) :rule bind) +(anchor :step t45.t2 :args ((:= (veriT_vr44 D$) veriT_vr90))) +(step t45.t2.t1 (cl @p_876) :rule refl) +(step t45.t2.t2 (cl (! (= @p_469 (! (b$ veriT_vr84) :named @p_894)) :named @p_898)) :rule cong :premises (t45.t2.t1)) +(step t45.t2.t3 (cl (! (= veriT_vr44 veriT_vr90) :named @p_896)) :rule refl) +(step t45.t2.t4 (cl (! (= @p_443 (! (run$a @p_894 veriT_vr90) :named @p_868)) :named @p_899)) :rule cong :premises (t45.t2.t2 t45.t2.t3)) +(step t45.t2.t5 (cl (! (= @p_441 (! (is_fail$a @p_868) :named @p_866)) :named @p_900)) :rule cong :premises (t45.t2.t4)) +(step t45.t2.t6 (cl @p_882) :rule refl) +(step t45.t2.t7 (cl (! (= @p_470 (! (b$ veriT_vr85) :named @p_895)) :named @p_903)) :rule cong :premises (t45.t2.t6)) +(step t45.t2.t8 (cl @p_896) :rule refl) +(step t45.t2.t9 (cl (! (= @p_442 (! (run$a @p_895 veriT_vr90) :named @p_867)) :named @p_904)) :rule cong :premises (t45.t2.t7 t45.t2.t8)) +(step t45.t2.t10 (cl (= @p_472 (! (is_fail$a @p_867) :named @p_897))) :rule cong :premises (t45.t2.t9)) +(step t45.t2.t11 (cl @p_876) :rule refl) +(step t45.t2.t12 (cl @p_898) :rule cong :premises (t45.t2.t11)) +(step t45.t2.t13 (cl @p_896) :rule refl) +(step t45.t2.t14 (cl @p_899) :rule cong :premises (t45.t2.t12 t45.t2.t13)) +(step t45.t2.t15 (cl @p_900) :rule cong :premises (t45.t2.t14)) +(step t45.t2.t16 (cl (= @p_476 (! (= @p_897 @p_866) :named @p_901))) :rule cong :premises (t45.t2.t10 t45.t2.t15)) +(anchor :step t45.t2.t17 :args ((:= (veriT_vr45 E$) veriT_vr91) (:= (veriT_vr46 D$) veriT_vr92))) +(step t45.t2.t17.t1 (cl @p_882) :rule refl) +(step t45.t2.t17.t2 (cl @p_903) :rule cong :premises (t45.t2.t17.t1)) +(step t45.t2.t17.t3 (cl @p_896) :rule refl) +(step t45.t2.t17.t4 (cl @p_904) :rule cong :premises (t45.t2.t17.t2 t45.t2.t17.t3)) +(step t45.t2.t17.t5 (cl (! (= veriT_vr45 veriT_vr91) :named @p_906)) :rule refl) +(step t45.t2.t17.t6 (cl (! (= veriT_vr46 veriT_vr92) :named @p_907)) :rule refl) +(step t45.t2.t17.t7 (cl (! (= @p_444 (! (pair$a veriT_vr91 veriT_vr92) :named @p_869)) :named @p_908)) :rule cong :premises (t45.t2.t17.t5 t45.t2.t17.t6)) +(step t45.t2.t17.t8 (cl (= @p_480 (! (is_res$a @p_867 @p_869) :named @p_905))) :rule cong :premises (t45.t2.t17.t4 t45.t2.t17.t7)) +(step t45.t2.t17.t9 (cl @p_876) :rule refl) +(step t45.t2.t17.t10 (cl @p_898) :rule cong :premises (t45.t2.t17.t9)) +(step t45.t2.t17.t11 (cl @p_896) :rule refl) +(step t45.t2.t17.t12 (cl @p_899) :rule cong :premises (t45.t2.t17.t10 t45.t2.t17.t11)) +(step t45.t2.t17.t13 (cl @p_906) :rule refl) +(step t45.t2.t17.t14 (cl @p_907) :rule refl) +(step t45.t2.t17.t15 (cl @p_908) :rule cong :premises (t45.t2.t17.t13 t45.t2.t17.t14)) +(step t45.t2.t17.t16 (cl (= @p_484 (! (is_res$a @p_868 @p_869) :named @p_909))) :rule cong :premises (t45.t2.t17.t12 t45.t2.t17.t15)) +(step t45.t2.t17.t17 (cl (= @p_485 (! (= @p_905 @p_909) :named @p_910))) :rule cong :premises (t45.t2.t17.t8 t45.t2.t17.t16)) +(step t45.t2.t17 (cl (= @p_477 (! (forall ((veriT_vr91 E$) (veriT_vr92 D$)) @p_910) :named @p_902))) :rule bind) +(step t45.t2.t18 (cl (= @p_486 (! (and @p_901 @p_902) :named @p_911))) :rule cong :premises (t45.t2.t16 t45.t2.t17)) +(step t45.t2.t19 (cl (= @p_487 (! (or @p_866 @p_911) :named @p_912))) :rule cong :premises (t45.t2.t5 t45.t2.t18)) +(step t45.t2 (cl (= @p_468 (! (forall ((veriT_vr90 D$)) @p_912) :named @p_893))) :rule bind) +(step t45.t3 (cl (= @p_488 (! (=> @p_870 @p_893) :named @p_913))) :rule cong :premises (t45.t1 t45.t2)) +(step t45 (cl (= @p_691 (! (forall ((veriT_vr84 A_b_c_M_state_fun$) (veriT_vr85 A_b_c_M_state_fun$)) @p_913) :named @p_1018))) :rule bind) +(anchor :step t46 :args ((:= (veriT_vr47 E$) veriT_vr91) (:= (veriT_vr48 A_b_c_M_state_fun$) veriT_vr84) (:= (veriT_vr49 A_b_c_M_state_fun$) veriT_vr85))) +(anchor :step t46.t1 :args ((:= (veriT_vr50 A$) veriT_vr86) (:= (veriT_vr51 C$) veriT_vr87))) +(step t46.t1.t1 (cl (! (= veriT_vr48 veriT_vr84) :named @p_920)) :rule refl) +(step t46.t1.t2 (cl (! (= veriT_vr50 veriT_vr86) :named @p_918)) :rule refl) +(step t46.t1.t3 (cl (! (= @p_498 @p_871) :named @p_921)) :rule cong :premises (t46.t1.t1 t46.t1.t2)) +(step t46.t1.t4 (cl (! (= veriT_vr51 veriT_vr87) :named @p_919)) :rule refl) +(step t46.t1.t5 (cl (! (= @p_491 @p_864) :named @p_922)) :rule cong :premises (t46.t1.t3 t46.t1.t4)) +(step t46.t1.t6 (cl (! (= @p_489 @p_862) :named @p_923)) :rule cong :premises (t46.t1.t5)) +(step t46.t1.t7 (cl (! (= veriT_vr49 veriT_vr85) :named @p_924)) :rule refl) +(step t46.t1.t8 (cl @p_918) :rule refl) +(step t46.t1.t9 (cl (! (= @p_500 @p_873) :named @p_925)) :rule cong :premises (t46.t1.t7 t46.t1.t8)) +(step t46.t1.t10 (cl @p_919) :rule refl) +(step t46.t1.t11 (cl (! (= @p_490 @p_863) :named @p_926)) :rule cong :premises (t46.t1.t9 t46.t1.t10)) +(step t46.t1.t12 (cl (= @p_502 @p_875)) :rule cong :premises (t46.t1.t11)) +(step t46.t1.t13 (cl @p_920) :rule refl) +(step t46.t1.t14 (cl @p_918) :rule refl) +(step t46.t1.t15 (cl @p_921) :rule cong :premises (t46.t1.t13 t46.t1.t14)) +(step t46.t1.t16 (cl @p_919) :rule refl) +(step t46.t1.t17 (cl @p_922) :rule cong :premises (t46.t1.t15 t46.t1.t16)) +(step t46.t1.t18 (cl @p_923) :rule cong :premises (t46.t1.t17)) +(step t46.t1.t19 (cl (= @p_507 @p_880)) :rule cong :premises (t46.t1.t12 t46.t1.t18)) +(anchor :step t46.t1.t20 :args ((:= (veriT_vr52 B$) veriT_vr88) (:= (veriT_vr53 C$) veriT_vr89))) +(step t46.t1.t20.t1 (cl @p_924) :rule refl) +(step t46.t1.t20.t2 (cl @p_918) :rule refl) +(step t46.t1.t20.t3 (cl @p_925) :rule cong :premises (t46.t1.t20.t1 t46.t1.t20.t2)) +(step t46.t1.t20.t4 (cl @p_919) :rule refl) +(step t46.t1.t20.t5 (cl @p_926) :rule cong :premises (t46.t1.t20.t3 t46.t1.t20.t4)) +(step t46.t1.t20.t6 (cl (! (= veriT_vr52 veriT_vr88) :named @p_927)) :rule refl) +(step t46.t1.t20.t7 (cl (! (= veriT_vr53 veriT_vr89) :named @p_928)) :rule refl) +(step t46.t1.t20.t8 (cl (! (= @p_492 @p_865) :named @p_929)) :rule cong :premises (t46.t1.t20.t6 t46.t1.t20.t7)) +(step t46.t1.t20.t9 (cl (= @p_512 @p_885)) :rule cong :premises (t46.t1.t20.t5 t46.t1.t20.t8)) +(step t46.t1.t20.t10 (cl @p_920) :rule refl) +(step t46.t1.t20.t11 (cl @p_918) :rule refl) +(step t46.t1.t20.t12 (cl @p_921) :rule cong :premises (t46.t1.t20.t10 t46.t1.t20.t11)) +(step t46.t1.t20.t13 (cl @p_919) :rule refl) +(step t46.t1.t20.t14 (cl @p_922) :rule cong :premises (t46.t1.t20.t12 t46.t1.t20.t13)) +(step t46.t1.t20.t15 (cl @p_927) :rule refl) +(step t46.t1.t20.t16 (cl @p_928) :rule refl) +(step t46.t1.t20.t17 (cl @p_929) :rule cong :premises (t46.t1.t20.t15 t46.t1.t20.t16)) +(step t46.t1.t20.t18 (cl (= @p_516 @p_889)) :rule cong :premises (t46.t1.t20.t14 t46.t1.t20.t17)) +(step t46.t1.t20.t19 (cl (= @p_517 @p_890)) :rule cong :premises (t46.t1.t20.t9 t46.t1.t20.t18)) +(step t46.t1.t20 (cl (= @p_508 @p_881)) :rule bind) +(step t46.t1.t21 (cl (= @p_518 @p_891)) :rule cong :premises (t46.t1.t19 t46.t1.t20)) +(step t46.t1.t22 (cl (= @p_519 @p_892)) :rule cong :premises (t46.t1.t6 t46.t1.t21)) +(step t46.t1 (cl (= @p_497 @p_870)) :rule bind) +(anchor :step t46.t2 :args ((:= (veriT_vr54 D$) veriT_vr90))) +(step t46.t2.t1 (cl (! (= veriT_vr47 veriT_vr91) :named @p_932)) :rule refl) +(step t46.t2.t2 (cl @p_920) :rule refl) +(step t46.t2.t3 (cl (! (= @p_521 (! (c$ veriT_vr91 veriT_vr84) :named @p_931)) :named @p_936)) :rule cong :premises (t46.t2.t1 t46.t2.t2)) +(step t46.t2.t4 (cl (! (= veriT_vr54 veriT_vr90) :named @p_934)) :rule refl) +(step t46.t2.t5 (cl (! (= @p_495 (! (run$b @p_931 veriT_vr90) :named @p_916)) :named @p_937)) :rule cong :premises (t46.t2.t3 t46.t2.t4)) +(step t46.t2.t6 (cl (! (= @p_493 (! (is_fail$b @p_916) :named @p_914)) :named @p_938)) :rule cong :premises (t46.t2.t5)) +(step t46.t2.t7 (cl @p_932) :rule refl) +(step t46.t2.t8 (cl @p_924) :rule refl) +(step t46.t2.t9 (cl (! (= @p_523 (! (c$ veriT_vr91 veriT_vr85) :named @p_933)) :named @p_941)) :rule cong :premises (t46.t2.t7 t46.t2.t8)) +(step t46.t2.t10 (cl @p_934) :rule refl) +(step t46.t2.t11 (cl (! (= @p_494 (! (run$b @p_933 veriT_vr90) :named @p_915)) :named @p_942)) :rule cong :premises (t46.t2.t9 t46.t2.t10)) +(step t46.t2.t12 (cl (= @p_525 (! (is_fail$b @p_915) :named @p_935))) :rule cong :premises (t46.t2.t11)) +(step t46.t2.t13 (cl @p_932) :rule refl) +(step t46.t2.t14 (cl @p_920) :rule refl) +(step t46.t2.t15 (cl @p_936) :rule cong :premises (t46.t2.t13 t46.t2.t14)) +(step t46.t2.t16 (cl @p_934) :rule refl) +(step t46.t2.t17 (cl @p_937) :rule cong :premises (t46.t2.t15 t46.t2.t16)) +(step t46.t2.t18 (cl @p_938) :rule cong :premises (t46.t2.t17)) +(step t46.t2.t19 (cl (= @p_529 (! (= @p_935 @p_914) :named @p_939))) :rule cong :premises (t46.t2.t12 t46.t2.t18)) +(anchor :step t46.t2.t20 :args ((:= (veriT_vr55 F$) veriT_vr93) (:= (veriT_vr56 D$) veriT_vr92))) +(step t46.t2.t20.t1 (cl @p_932) :rule refl) +(step t46.t2.t20.t2 (cl @p_924) :rule refl) +(step t46.t2.t20.t3 (cl @p_941) :rule cong :premises (t46.t2.t20.t1 t46.t2.t20.t2)) +(step t46.t2.t20.t4 (cl @p_934) :rule refl) +(step t46.t2.t20.t5 (cl @p_942) :rule cong :premises (t46.t2.t20.t3 t46.t2.t20.t4)) +(step t46.t2.t20.t6 (cl (! (= veriT_vr55 veriT_vr93) :named @p_944)) :rule refl) +(step t46.t2.t20.t7 (cl (! (= veriT_vr56 veriT_vr92) :named @p_945)) :rule refl) +(step t46.t2.t20.t8 (cl (! (= @p_496 (! (pair$b veriT_vr93 veriT_vr92) :named @p_917)) :named @p_946)) :rule cong :premises (t46.t2.t20.t6 t46.t2.t20.t7)) +(step t46.t2.t20.t9 (cl (= @p_533 (! (is_res$b @p_915 @p_917) :named @p_943))) :rule cong :premises (t46.t2.t20.t5 t46.t2.t20.t8)) +(step t46.t2.t20.t10 (cl @p_932) :rule refl) +(step t46.t2.t20.t11 (cl @p_920) :rule refl) +(step t46.t2.t20.t12 (cl @p_936) :rule cong :premises (t46.t2.t20.t10 t46.t2.t20.t11)) +(step t46.t2.t20.t13 (cl @p_934) :rule refl) +(step t46.t2.t20.t14 (cl @p_937) :rule cong :premises (t46.t2.t20.t12 t46.t2.t20.t13)) +(step t46.t2.t20.t15 (cl @p_944) :rule refl) +(step t46.t2.t20.t16 (cl @p_945) :rule refl) +(step t46.t2.t20.t17 (cl @p_946) :rule cong :premises (t46.t2.t20.t15 t46.t2.t20.t16)) +(step t46.t2.t20.t18 (cl (= @p_537 (! (is_res$b @p_916 @p_917) :named @p_947))) :rule cong :premises (t46.t2.t20.t14 t46.t2.t20.t17)) +(step t46.t2.t20.t19 (cl (= @p_538 (! (= @p_943 @p_947) :named @p_948))) :rule cong :premises (t46.t2.t20.t9 t46.t2.t20.t18)) +(step t46.t2.t20 (cl (= @p_530 (! (forall ((veriT_vr93 F$) (veriT_vr92 D$)) @p_948) :named @p_940))) :rule bind) +(step t46.t2.t21 (cl (= @p_539 (! (and @p_939 @p_940) :named @p_949))) :rule cong :premises (t46.t2.t19 t46.t2.t20)) +(step t46.t2.t22 (cl (= @p_540 (! (or @p_914 @p_949) :named @p_950))) :rule cong :premises (t46.t2.t6 t46.t2.t21)) +(step t46.t2 (cl (= @p_520 (! (forall ((veriT_vr90 D$)) @p_950) :named @p_930))) :rule bind) +(step t46.t3 (cl (= @p_541 (! (=> @p_870 @p_930) :named @p_951))) :rule cong :premises (t46.t1 t46.t2)) +(step t46 (cl (= @p_692 (! (forall ((veriT_vr91 E$) (veriT_vr84 A_b_c_M_state_fun$) (veriT_vr85 A_b_c_M_state_fun$)) @p_951) :named @p_1019))) :rule bind) +(anchor :step t47 :args ((:= (veriT_vr59 A$) veriT_vr86) (:= (veriT_vr60 C$) veriT_vr87))) +(step t47.t1 (cl (! (= veriT_vr59 veriT_vr86) :named @p_956)) :rule refl) +(step t47.t2 (cl (! (= @p_751 (! (fun_app$ veriT_sk0 veriT_vr86) :named @p_955)) :named @p_960)) :rule cong :premises (t47.t1)) +(step t47.t3 (cl (! (= veriT_vr60 veriT_vr87) :named @p_958)) :rule refl) +(step t47.t4 (cl (! (= @p_703 (! (run$ @p_955 veriT_vr87) :named @p_954)) :named @p_961)) :rule cong :premises (t47.t2 t47.t3)) +(step t47.t5 (cl (! (= @p_701 (! (is_fail$ @p_954) :named @p_952)) :named @p_962)) :rule cong :premises (t47.t4)) +(step t47.t6 (cl @p_956) :rule refl) +(step t47.t7 (cl (! (= @p_752 (! (fun_app$ veriT_sk1 veriT_vr86) :named @p_957)) :named @p_965)) :rule cong :premises (t47.t6)) +(step t47.t8 (cl @p_958) :rule refl) +(step t47.t9 (cl (! (= @p_702 (! (run$ @p_957 veriT_vr87) :named @p_953)) :named @p_966)) :rule cong :premises (t47.t7 t47.t8)) +(step t47.t10 (cl (= @p_753 (! (is_fail$ @p_953) :named @p_959))) :rule cong :premises (t47.t9)) +(step t47.t11 (cl @p_956) :rule refl) +(step t47.t12 (cl @p_960) :rule cong :premises (t47.t11)) +(step t47.t13 (cl @p_958) :rule refl) +(step t47.t14 (cl @p_961) :rule cong :premises (t47.t12 t47.t13)) +(step t47.t15 (cl @p_962) :rule cong :premises (t47.t14)) +(step t47.t16 (cl (= @p_758 (! (= @p_959 @p_952) :named @p_963))) :rule cong :premises (t47.t10 t47.t15)) +(anchor :step t47.t17 :args ((:= (veriT_vr61 B$) veriT_vr88) (:= (veriT_vr62 C$) veriT_vr89))) +(step t47.t17.t1 (cl @p_956) :rule refl) +(step t47.t17.t2 (cl @p_965) :rule cong :premises (t47.t17.t1)) +(step t47.t17.t3 (cl @p_958) :rule refl) +(step t47.t17.t4 (cl @p_966) :rule cong :premises (t47.t17.t2 t47.t17.t3)) +(step t47.t17.t5 (cl (! (= veriT_vr61 veriT_vr88) :named @p_968)) :rule refl) +(step t47.t17.t6 (cl (! (= veriT_vr62 veriT_vr89) :named @p_969)) :rule refl) +(step t47.t17.t7 (cl (! (= @p_545 @p_865) :named @p_970)) :rule cong :premises (t47.t17.t5 t47.t17.t6)) +(step t47.t17.t8 (cl (= @p_763 (! (is_res$ @p_953 @p_865) :named @p_967))) :rule cong :premises (t47.t17.t4 t47.t17.t7)) +(step t47.t17.t9 (cl @p_956) :rule refl) +(step t47.t17.t10 (cl @p_960) :rule cong :premises (t47.t17.t9)) +(step t47.t17.t11 (cl @p_958) :rule refl) +(step t47.t17.t12 (cl @p_961) :rule cong :premises (t47.t17.t10 t47.t17.t11)) +(step t47.t17.t13 (cl @p_968) :rule refl) +(step t47.t17.t14 (cl @p_969) :rule refl) +(step t47.t17.t15 (cl @p_970) :rule cong :premises (t47.t17.t13 t47.t17.t14)) +(step t47.t17.t16 (cl (= @p_764 (! (is_res$ @p_954 @p_865) :named @p_971))) :rule cong :premises (t47.t17.t12 t47.t17.t15)) +(step t47.t17.t17 (cl (= @p_765 (! (= @p_967 @p_971) :named @p_972))) :rule cong :premises (t47.t17.t8 t47.t17.t16)) +(step t47.t17 (cl (= @p_759 (! (forall ((veriT_vr88 B$) (veriT_vr89 C$)) @p_972) :named @p_964))) :rule bind) +(step t47.t18 (cl (= @p_766 (! (and @p_963 @p_964) :named @p_973))) :rule cong :premises (t47.t16 t47.t17)) +(step t47.t19 (cl (= @p_767 (! (or @p_952 @p_973) :named @p_974))) :rule cong :premises (t47.t5 t47.t18)) +(step t47 (cl (= @p_750 (! (forall ((veriT_vr86 A$) (veriT_vr87 C$)) @p_974) :named @p_1015))) :rule bind) +(anchor :step t48 :args ((:= (veriT_vr64 E$) veriT_vr91) (:= (veriT_vr65 D$) veriT_vr90))) +(step t48.t1 (cl (! (= veriT_vr64 veriT_vr91) :named @p_977)) :rule refl) +(step t48.t2 (cl (! (= veriT_vr65 veriT_vr90) :named @p_979)) :rule refl) +(step t48.t3 (cl (= @p_580 (! (pair$a veriT_vr91 veriT_vr90) :named @p_975))) :rule cong :premises (t48.t1 t48.t2)) +(step t48.t4 (cl (= @p_774 (! (is_res$a @p_712 @p_975) :named @p_976))) :rule cong :premises (t48.t3)) +(step t48.t5 (cl @p_977) :rule refl) +(step t48.t6 (cl (= @p_775 (! (c$ veriT_vr91 veriT_sk0) :named @p_978))) :rule cong :premises (t48.t5)) +(step t48.t7 (cl @p_979) :rule refl) +(step t48.t8 (cl (= @p_776 (! (run$b @p_978 veriT_vr90) :named @p_980))) :rule cong :premises (t48.t6 t48.t7)) +(step t48.t9 (cl (= @p_777 (! (is_fail$b @p_980) :named @p_981))) :rule cong :premises (t48.t8)) +(step t48.t10 (cl (= @p_778 (! (and @p_976 @p_981) :named @p_982))) :rule cong :premises (t48.t4 t48.t9)) +(step t48 (cl (= @p_770 (! (exists ((veriT_vr91 E$) (veriT_vr90 D$)) @p_982) :named @p_983))) :rule bind) +(anchor :step t49 :args ((:= (veriT_vr68 E$) veriT_vr91) (:= (veriT_vr69 D$) veriT_vr90))) +(step t49.t1 (cl (! (= veriT_vr68 veriT_vr91) :named @p_984)) :rule refl) +(step t49.t2 (cl (! (= veriT_vr69 veriT_vr90) :named @p_985)) :rule refl) +(step t49.t3 (cl (= @p_603 @p_975)) :rule cong :premises (t49.t1 t49.t2)) +(step t49.t4 (cl (= @p_793 @p_976)) :rule cong :premises (t49.t3)) +(step t49.t5 (cl @p_984) :rule refl) +(step t49.t6 (cl (= @p_794 @p_978)) :rule cong :premises (t49.t5)) +(step t49.t7 (cl @p_985) :rule refl) +(step t49.t8 (cl (= @p_795 @p_980)) :rule cong :premises (t49.t6 t49.t7)) +(step t49.t9 (cl (= @p_796 @p_981)) :rule cong :premises (t49.t8)) +(step t49.t10 (cl (= @p_797 @p_982)) :rule cong :premises (t49.t4 t49.t9)) +(step t49 (cl (= @p_792 @p_983)) :rule bind) +(step t50 (cl (= @p_798 (! (or @p_720 @p_983) :named @p_986))) :rule cong :premises (t49)) +(step t51 (cl (= @p_799 (! (=> @p_790 @p_986) :named @p_996))) :rule cong :premises (t50)) +(anchor :step t52 :args ((:= (veriT_vr72 E$) veriT_vr91) (:= (veriT_vr73 D$) veriT_vr90))) +(step t52.t1 (cl (! (= veriT_vr72 veriT_vr91) :named @p_988)) :rule refl) +(step t52.t2 (cl (! (= veriT_vr73 veriT_vr90) :named @p_990)) :rule refl) +(step t52.t3 (cl (= @p_626 @p_975)) :rule cong :premises (t52.t1 t52.t2)) +(step t52.t4 (cl (= @p_811 (! (is_res$a @p_714 @p_975) :named @p_987))) :rule cong :premises (t52.t3)) +(step t52.t5 (cl @p_988) :rule refl) +(step t52.t6 (cl (= @p_812 (! (c$ veriT_vr91 veriT_sk1) :named @p_989))) :rule cong :premises (t52.t5)) +(step t52.t7 (cl @p_990) :rule refl) +(step t52.t8 (cl (= @p_813 (! (run$b @p_989 veriT_vr90) :named @p_991))) :rule cong :premises (t52.t6 t52.t7)) +(step t52.t9 (cl (= @p_814 (! (is_fail$b @p_991) :named @p_992))) :rule cong :premises (t52.t8)) +(step t52.t10 (cl (= @p_815 (! (and @p_987 @p_992) :named @p_993))) :rule cong :premises (t52.t4 t52.t9)) +(step t52 (cl (= @p_810 (! (exists ((veriT_vr91 E$) (veriT_vr90 D$)) @p_993) :named @p_994))) :rule bind) +(step t53 (cl (= @p_816 (! (or @p_726 @p_994) :named @p_995))) :rule cong :premises (t52)) +(step t54 (cl (= @p_817 (! (=> @p_808 @p_995) :named @p_997))) :rule cong :premises (t53)) +(step t55 (cl (= @p_818 (! (and @p_996 @p_997) :named @p_1012))) :rule cong :premises (t51 t54)) +(anchor :step t56 :args ((:= (veriT_vr78 E$) veriT_vr91) (:= (veriT_vr79 D$) veriT_vr90))) +(step t56.t1 (cl (! (= veriT_vr78 veriT_vr91) :named @p_998)) :rule refl) +(step t56.t2 (cl (! (= veriT_vr79 veriT_vr90) :named @p_999)) :rule refl) +(step t56.t3 (cl (= @p_650 @p_975)) :rule cong :premises (t56.t1 t56.t2)) +(step t56.t4 (cl (= @p_830 @p_976)) :rule cong :premises (t56.t3)) +(step t56.t5 (cl @p_998) :rule refl) +(step t56.t6 (cl (= @p_831 @p_978)) :rule cong :premises (t56.t5)) +(step t56.t7 (cl @p_999) :rule refl) +(step t56.t8 (cl (= @p_832 @p_980)) :rule cong :premises (t56.t6 t56.t7)) +(step t56.t9 (cl (= @p_836 (! (is_res$b @p_980 @p_747) :named @p_1000))) :rule cong :premises (t56.t8)) +(step t56.t10 (cl (= @p_837 (! (and @p_976 @p_1000) :named @p_1001))) :rule cong :premises (t56.t4 t56.t9)) +(step t56 (cl (= @p_829 (! (exists ((veriT_vr91 E$) (veriT_vr90 D$)) @p_1001) :named @p_1002))) :rule bind) +(step t57 (cl (= @p_838 (! (or @p_720 @p_1002) :named @p_1003))) :rule cong :premises (t56)) +(step t58 (cl (= @p_839 (! (=> @p_828 @p_1003) :named @p_1010))) :rule cong :premises (t57)) +(anchor :step t59 :args ((:= (veriT_vr82 E$) veriT_vr91) (:= (veriT_vr83 D$) veriT_vr90))) +(step t59.t1 (cl (! (= veriT_vr82 veriT_vr91) :named @p_1004)) :rule refl) +(step t59.t2 (cl (! (= veriT_vr83 veriT_vr90) :named @p_1005)) :rule refl) +(step t59.t3 (cl (= @p_675 @p_975)) :rule cong :premises (t59.t1 t59.t2)) +(step t59.t4 (cl (= @p_850 @p_987)) :rule cong :premises (t59.t3)) +(step t59.t5 (cl @p_1004) :rule refl) +(step t59.t6 (cl (= @p_851 @p_989)) :rule cong :premises (t59.t5)) +(step t59.t7 (cl @p_1005) :rule refl) +(step t59.t8 (cl (= @p_852 @p_991)) :rule cong :premises (t59.t6 t59.t7)) +(step t59.t9 (cl (= @p_853 (! (is_res$b @p_991 @p_747) :named @p_1006))) :rule cong :premises (t59.t8)) +(step t59.t10 (cl (= @p_854 (! (and @p_987 @p_1006) :named @p_1007))) :rule cong :premises (t59.t4 t59.t9)) +(step t59 (cl (= @p_849 (! (exists ((veriT_vr91 E$) (veriT_vr90 D$)) @p_1007) :named @p_1008))) :rule bind) +(step t60 (cl (= @p_855 (! (or @p_726 @p_1008) :named @p_1009))) :rule cong :premises (t59)) +(step t61 (cl (= @p_856 (! (=> @p_848 @p_1009) :named @p_1011))) :rule cong :premises (t60)) +(step t62 (cl (= @p_819 (! (and @p_1010 @p_1011) :named @p_1013))) :rule cong :premises (t58 t61)) +(step t63 (cl (= @p_857 (! (and @p_1012 @p_1013) :named @p_1014))) :rule cong :premises (t55 t62)) +(step t64 (cl (= @p_768 (! (or @p_720 @p_983 @p_1014) :named @p_1016))) :rule cong :premises (t48 t63)) +(step t65 (cl (= @p_858 (! (=> @p_1015 @p_1016) :named @p_1017))) :rule cong :premises (t47 t64)) +(step t66 (cl (= @p_859 (! (not @p_1017) :named @p_1020))) :rule cong :premises (t65)) +(step t67 (cl (! (= @p_861 (! (and @p_1018 @p_1019 @p_1020) :named @p_1022)) :named @p_1021)) :rule cong :premises (t45 t46 t66)) +(step t68 (cl (not @p_1021) (not @p_861) @p_1022) :rule equiv_pos2) +(step t69 (cl @p_1022) :rule th_resolution :premises (t44 t67 t68)) +(step t70 (cl (= @p_1020 (! (and @p_1015 (! (not @p_1016) :named @p_1027)) :named @p_1023))) :rule bool_simplify) +(step t71 (cl (! (= @p_1022 (! (and @p_1018 @p_1019 @p_1023) :named @p_1025)) :named @p_1024)) :rule cong :premises (t70)) +(step t72 (cl (not @p_1024) (not @p_1022) @p_1025) :rule equiv_pos2) +(step t73 (cl @p_1025) :rule th_resolution :premises (t69 t71 t72)) +(step t74 (cl (= @p_1014 (! (and @p_996 @p_997 @p_1010 @p_1011) :named @p_1026))) :rule ac_simp) +(step t75 (cl (= @p_1016 (! (or @p_720 @p_983 @p_1026) :named @p_1028))) :rule ac_simp :premises (t74)) +(step t76 (cl (= @p_1027 (! (not @p_1028) :named @p_1029))) :rule cong :premises (t75)) +(step t77 (cl (! (= @p_1025 (! (and @p_1018 @p_1019 @p_1015 @p_1029) :named @p_1031)) :named @p_1030)) :rule ac_simp :premises (t76)) +(step t78 (cl (not @p_1030) (not @p_1025) @p_1031) :rule equiv_pos2) +(step t79 (cl @p_1031) :rule th_resolution :premises (t73 t77 t78)) +(anchor :step t80 :args ((:= (veriT_vr91 E$) veriT_vr94) (:= (veriT_vr84 A_b_c_M_state_fun$) veriT_vr95) (:= (veriT_vr85 A_b_c_M_state_fun$) veriT_vr96))) +(anchor :step t80.t1 :args ((:= (veriT_vr86 A$) veriT_vr97) (:= (veriT_vr87 C$) veriT_vr98))) +(step t80.t1.t1 (cl (! (= veriT_vr84 veriT_vr95) :named @p_1046)) :rule refl) +(step t80.t1.t2 (cl (! (= veriT_vr86 veriT_vr97) :named @p_1042)) :rule refl) +(step t80.t1.t3 (cl (! (= @p_871 (! (fun_app$ veriT_vr95 veriT_vr97) :named @p_1041)) :named @p_1047)) :rule cong :premises (t80.t1.t1 t80.t1.t2)) +(step t80.t1.t4 (cl (! (= veriT_vr87 veriT_vr98) :named @p_1044)) :rule refl) +(step t80.t1.t5 (cl (! (= @p_864 (! (run$ @p_1041 veriT_vr98) :named @p_1034)) :named @p_1048)) :rule cong :premises (t80.t1.t3 t80.t1.t4)) +(step t80.t1.t6 (cl (! (= @p_862 (! (is_fail$ @p_1034) :named @p_1032)) :named @p_1049)) :rule cong :premises (t80.t1.t5)) +(step t80.t1.t7 (cl (! (= veriT_vr85 veriT_vr96) :named @p_1052)) :rule refl) +(step t80.t1.t8 (cl @p_1042) :rule refl) +(step t80.t1.t9 (cl (! (= @p_873 (! (fun_app$ veriT_vr96 veriT_vr97) :named @p_1043)) :named @p_1053)) :rule cong :premises (t80.t1.t7 t80.t1.t8)) +(step t80.t1.t10 (cl @p_1044) :rule refl) +(step t80.t1.t11 (cl (! (= @p_863 (! (run$ @p_1043 veriT_vr98) :named @p_1033)) :named @p_1054)) :rule cong :premises (t80.t1.t9 t80.t1.t10)) +(step t80.t1.t12 (cl (= @p_875 (! (is_fail$ @p_1033) :named @p_1045))) :rule cong :premises (t80.t1.t11)) +(step t80.t1.t13 (cl @p_1046) :rule refl) +(step t80.t1.t14 (cl @p_1042) :rule refl) +(step t80.t1.t15 (cl @p_1047) :rule cong :premises (t80.t1.t13 t80.t1.t14)) +(step t80.t1.t16 (cl @p_1044) :rule refl) +(step t80.t1.t17 (cl @p_1048) :rule cong :premises (t80.t1.t15 t80.t1.t16)) +(step t80.t1.t18 (cl @p_1049) :rule cong :premises (t80.t1.t17)) +(step t80.t1.t19 (cl (= @p_880 (! (= @p_1045 @p_1032) :named @p_1050))) :rule cong :premises (t80.t1.t12 t80.t1.t18)) +(anchor :step t80.t1.t20 :args ((:= (veriT_vr88 B$) veriT_vr99) (:= (veriT_vr89 C$) veriT_vr100))) +(step t80.t1.t20.t1 (cl @p_1052) :rule refl) +(step t80.t1.t20.t2 (cl @p_1042) :rule refl) +(step t80.t1.t20.t3 (cl @p_1053) :rule cong :premises (t80.t1.t20.t1 t80.t1.t20.t2)) +(step t80.t1.t20.t4 (cl @p_1044) :rule refl) +(step t80.t1.t20.t5 (cl @p_1054) :rule cong :premises (t80.t1.t20.t3 t80.t1.t20.t4)) +(step t80.t1.t20.t6 (cl (! (= veriT_vr88 veriT_vr99) :named @p_1056)) :rule refl) +(step t80.t1.t20.t7 (cl (! (= veriT_vr89 veriT_vr100) :named @p_1057)) :rule refl) +(step t80.t1.t20.t8 (cl (! (= @p_865 (! (pair$ veriT_vr99 veriT_vr100) :named @p_1035)) :named @p_1058)) :rule cong :premises (t80.t1.t20.t6 t80.t1.t20.t7)) +(step t80.t1.t20.t9 (cl (= @p_885 (! (is_res$ @p_1033 @p_1035) :named @p_1055))) :rule cong :premises (t80.t1.t20.t5 t80.t1.t20.t8)) +(step t80.t1.t20.t10 (cl @p_1046) :rule refl) +(step t80.t1.t20.t11 (cl @p_1042) :rule refl) +(step t80.t1.t20.t12 (cl @p_1047) :rule cong :premises (t80.t1.t20.t10 t80.t1.t20.t11)) +(step t80.t1.t20.t13 (cl @p_1044) :rule refl) +(step t80.t1.t20.t14 (cl @p_1048) :rule cong :premises (t80.t1.t20.t12 t80.t1.t20.t13)) +(step t80.t1.t20.t15 (cl @p_1056) :rule refl) +(step t80.t1.t20.t16 (cl @p_1057) :rule refl) +(step t80.t1.t20.t17 (cl @p_1058) :rule cong :premises (t80.t1.t20.t15 t80.t1.t20.t16)) +(step t80.t1.t20.t18 (cl (= @p_889 (! (is_res$ @p_1034 @p_1035) :named @p_1059))) :rule cong :premises (t80.t1.t20.t14 t80.t1.t20.t17)) +(step t80.t1.t20.t19 (cl (= @p_890 (! (= @p_1055 @p_1059) :named @p_1060))) :rule cong :premises (t80.t1.t20.t9 t80.t1.t20.t18)) +(step t80.t1.t20 (cl (= @p_881 (! (forall ((veriT_vr99 B$) (veriT_vr100 C$)) @p_1060) :named @p_1051))) :rule bind) +(step t80.t1.t21 (cl (= @p_891 (! (and @p_1050 @p_1051) :named @p_1061))) :rule cong :premises (t80.t1.t19 t80.t1.t20)) +(step t80.t1.t22 (cl (= @p_892 (! (or @p_1032 @p_1061) :named @p_1062))) :rule cong :premises (t80.t1.t6 t80.t1.t21)) +(step t80.t1 (cl (= @p_870 (! (forall ((veriT_vr97 A$) (veriT_vr98 C$)) @p_1062) :named @p_1040))) :rule bind) +(anchor :step t80.t2 :args ((:= (veriT_vr90 D$) veriT_vr101))) +(step t80.t2.t1 (cl (! (= veriT_vr91 veriT_vr94) :named @p_1065)) :rule refl) +(step t80.t2.t2 (cl @p_1046) :rule refl) +(step t80.t2.t3 (cl (! (= @p_931 (! (c$ veriT_vr94 veriT_vr95) :named @p_1064)) :named @p_1069)) :rule cong :premises (t80.t2.t1 t80.t2.t2)) +(step t80.t2.t4 (cl (! (= veriT_vr90 veriT_vr101) :named @p_1067)) :rule refl) +(step t80.t2.t5 (cl (! (= @p_916 (! (run$b @p_1064 veriT_vr101) :named @p_1038)) :named @p_1070)) :rule cong :premises (t80.t2.t3 t80.t2.t4)) +(step t80.t2.t6 (cl (! (= @p_914 (! (is_fail$b @p_1038) :named @p_1036)) :named @p_1071)) :rule cong :premises (t80.t2.t5)) +(step t80.t2.t7 (cl @p_1065) :rule refl) +(step t80.t2.t8 (cl @p_1052) :rule refl) +(step t80.t2.t9 (cl (! (= @p_933 (! (c$ veriT_vr94 veriT_vr96) :named @p_1066)) :named @p_1074)) :rule cong :premises (t80.t2.t7 t80.t2.t8)) +(step t80.t2.t10 (cl @p_1067) :rule refl) +(step t80.t2.t11 (cl (! (= @p_915 (! (run$b @p_1066 veriT_vr101) :named @p_1037)) :named @p_1075)) :rule cong :premises (t80.t2.t9 t80.t2.t10)) +(step t80.t2.t12 (cl (= @p_935 (! (is_fail$b @p_1037) :named @p_1068))) :rule cong :premises (t80.t2.t11)) +(step t80.t2.t13 (cl @p_1065) :rule refl) +(step t80.t2.t14 (cl @p_1046) :rule refl) +(step t80.t2.t15 (cl @p_1069) :rule cong :premises (t80.t2.t13 t80.t2.t14)) +(step t80.t2.t16 (cl @p_1067) :rule refl) +(step t80.t2.t17 (cl @p_1070) :rule cong :premises (t80.t2.t15 t80.t2.t16)) +(step t80.t2.t18 (cl @p_1071) :rule cong :premises (t80.t2.t17)) +(step t80.t2.t19 (cl (= @p_939 (! (= @p_1068 @p_1036) :named @p_1072))) :rule cong :premises (t80.t2.t12 t80.t2.t18)) +(anchor :step t80.t2.t20 :args ((:= (veriT_vr93 F$) veriT_vr93) (:= (veriT_vr92 D$) veriT_vr102))) +(step t80.t2.t20.t1 (cl @p_1065) :rule refl) +(step t80.t2.t20.t2 (cl @p_1052) :rule refl) +(step t80.t2.t20.t3 (cl @p_1074) :rule cong :premises (t80.t2.t20.t1 t80.t2.t20.t2)) +(step t80.t2.t20.t4 (cl @p_1067) :rule refl) +(step t80.t2.t20.t5 (cl @p_1075) :rule cong :premises (t80.t2.t20.t3 t80.t2.t20.t4)) +(step t80.t2.t20.t6 (cl (! (= veriT_vr92 veriT_vr102) :named @p_1077)) :rule refl) +(step t80.t2.t20.t7 (cl (! (= @p_917 (! (pair$b veriT_vr93 veriT_vr102) :named @p_1039)) :named @p_1078)) :rule cong :premises (t80.t2.t20.t6)) +(step t80.t2.t20.t8 (cl (= @p_943 (! (is_res$b @p_1037 @p_1039) :named @p_1076))) :rule cong :premises (t80.t2.t20.t5 t80.t2.t20.t7)) +(step t80.t2.t20.t9 (cl @p_1065) :rule refl) +(step t80.t2.t20.t10 (cl @p_1046) :rule refl) +(step t80.t2.t20.t11 (cl @p_1069) :rule cong :premises (t80.t2.t20.t9 t80.t2.t20.t10)) +(step t80.t2.t20.t12 (cl @p_1067) :rule refl) +(step t80.t2.t20.t13 (cl @p_1070) :rule cong :premises (t80.t2.t20.t11 t80.t2.t20.t12)) +(step t80.t2.t20.t14 (cl @p_1077) :rule refl) +(step t80.t2.t20.t15 (cl @p_1078) :rule cong :premises (t80.t2.t20.t14)) +(step t80.t2.t20.t16 (cl (= @p_947 (! (is_res$b @p_1038 @p_1039) :named @p_1079))) :rule cong :premises (t80.t2.t20.t13 t80.t2.t20.t15)) +(step t80.t2.t20.t17 (cl (= @p_948 (! (= @p_1076 @p_1079) :named @p_1080))) :rule cong :premises (t80.t2.t20.t8 t80.t2.t20.t16)) +(step t80.t2.t20 (cl (= @p_940 (! (forall ((veriT_vr93 F$) (veriT_vr102 D$)) @p_1080) :named @p_1073))) :rule bind) +(step t80.t2.t21 (cl (= @p_949 (! (and @p_1072 @p_1073) :named @p_1081))) :rule cong :premises (t80.t2.t19 t80.t2.t20)) +(step t80.t2.t22 (cl (= @p_950 (! (or @p_1036 @p_1081) :named @p_1082))) :rule cong :premises (t80.t2.t6 t80.t2.t21)) +(step t80.t2 (cl (= @p_930 (! (forall ((veriT_vr101 D$)) @p_1082) :named @p_1063))) :rule bind) +(step t80.t3 (cl (= @p_951 (! (=> @p_1040 @p_1063) :named @p_1083))) :rule cong :premises (t80.t1 t80.t2)) +(step t80 (cl (= @p_1019 (! (forall ((veriT_vr94 E$) (veriT_vr95 A_b_c_M_state_fun$) (veriT_vr96 A_b_c_M_state_fun$)) @p_1083) :named @p_1154))) :rule bind) +(anchor :step t81 :args ((:= (veriT_vr86 A$) veriT_vr103) (:= (veriT_vr87 C$) veriT_vr104))) +(step t81.t1 (cl (! (= veriT_vr86 veriT_vr103) :named @p_1089)) :rule refl) +(step t81.t2 (cl (! (= @p_955 (! (fun_app$ veriT_sk0 veriT_vr103) :named @p_1088)) :named @p_1093)) :rule cong :premises (t81.t1)) +(step t81.t3 (cl (! (= veriT_vr87 veriT_vr104) :named @p_1091)) :rule refl) +(step t81.t4 (cl (! (= @p_954 (! (run$ @p_1088 veriT_vr104) :named @p_1086)) :named @p_1094)) :rule cong :premises (t81.t2 t81.t3)) +(step t81.t5 (cl (! (= @p_952 (! (is_fail$ @p_1086) :named @p_1084)) :named @p_1095)) :rule cong :premises (t81.t4)) +(step t81.t6 (cl @p_1089) :rule refl) +(step t81.t7 (cl (! (= @p_957 (! (fun_app$ veriT_sk1 veriT_vr103) :named @p_1090)) :named @p_1098)) :rule cong :premises (t81.t6)) +(step t81.t8 (cl @p_1091) :rule refl) +(step t81.t9 (cl (! (= @p_953 (! (run$ @p_1090 veriT_vr104) :named @p_1085)) :named @p_1099)) :rule cong :premises (t81.t7 t81.t8)) +(step t81.t10 (cl (= @p_959 (! (is_fail$ @p_1085) :named @p_1092))) :rule cong :premises (t81.t9)) +(step t81.t11 (cl @p_1089) :rule refl) +(step t81.t12 (cl @p_1093) :rule cong :premises (t81.t11)) +(step t81.t13 (cl @p_1091) :rule refl) +(step t81.t14 (cl @p_1094) :rule cong :premises (t81.t12 t81.t13)) +(step t81.t15 (cl @p_1095) :rule cong :premises (t81.t14)) +(step t81.t16 (cl (= @p_963 (! (= @p_1092 @p_1084) :named @p_1096))) :rule cong :premises (t81.t10 t81.t15)) +(anchor :step t81.t17 :args ((:= (veriT_vr88 B$) veriT_vr105) (:= (veriT_vr89 C$) veriT_vr106))) +(step t81.t17.t1 (cl @p_1089) :rule refl) +(step t81.t17.t2 (cl @p_1098) :rule cong :premises (t81.t17.t1)) +(step t81.t17.t3 (cl @p_1091) :rule refl) +(step t81.t17.t4 (cl @p_1099) :rule cong :premises (t81.t17.t2 t81.t17.t3)) +(step t81.t17.t5 (cl (! (= veriT_vr88 veriT_vr105) :named @p_1101)) :rule refl) +(step t81.t17.t6 (cl (! (= veriT_vr89 veriT_vr106) :named @p_1102)) :rule refl) +(step t81.t17.t7 (cl (! (= @p_865 (! (pair$ veriT_vr105 veriT_vr106) :named @p_1087)) :named @p_1103)) :rule cong :premises (t81.t17.t5 t81.t17.t6)) +(step t81.t17.t8 (cl (= @p_967 (! (is_res$ @p_1085 @p_1087) :named @p_1100))) :rule cong :premises (t81.t17.t4 t81.t17.t7)) +(step t81.t17.t9 (cl @p_1089) :rule refl) +(step t81.t17.t10 (cl @p_1093) :rule cong :premises (t81.t17.t9)) +(step t81.t17.t11 (cl @p_1091) :rule refl) +(step t81.t17.t12 (cl @p_1094) :rule cong :premises (t81.t17.t10 t81.t17.t11)) +(step t81.t17.t13 (cl @p_1101) :rule refl) +(step t81.t17.t14 (cl @p_1102) :rule refl) +(step t81.t17.t15 (cl @p_1103) :rule cong :premises (t81.t17.t13 t81.t17.t14)) +(step t81.t17.t16 (cl (= @p_971 (! (is_res$ @p_1086 @p_1087) :named @p_1104))) :rule cong :premises (t81.t17.t12 t81.t17.t15)) +(step t81.t17.t17 (cl (= @p_972 (! (= @p_1100 @p_1104) :named @p_1105))) :rule cong :premises (t81.t17.t8 t81.t17.t16)) +(step t81.t17 (cl (= @p_964 (! (forall ((veriT_vr105 B$) (veriT_vr106 C$)) @p_1105) :named @p_1097))) :rule bind) +(step t81.t18 (cl (= @p_973 (! (and @p_1096 @p_1097) :named @p_1106))) :rule cong :premises (t81.t16 t81.t17)) +(step t81.t19 (cl (= @p_974 (! (or @p_1084 @p_1106) :named @p_1107))) :rule cong :premises (t81.t5 t81.t18)) +(step t81 (cl (= @p_1015 (! (forall ((veriT_vr103 A$) (veriT_vr104 C$)) @p_1107) :named @p_1155))) :rule bind) +(anchor :step t82 :args ((:= (veriT_vr91 E$) veriT_vr107) (:= (veriT_vr90 D$) veriT_vr108))) +(step t82.t1 (cl (! (= veriT_vr91 veriT_vr107) :named @p_1110)) :rule refl) +(step t82.t2 (cl (! (= veriT_vr90 veriT_vr108) :named @p_1112)) :rule refl) +(step t82.t3 (cl (= @p_975 (! (pair$a veriT_vr107 veriT_vr108) :named @p_1108))) :rule cong :premises (t82.t1 t82.t2)) +(step t82.t4 (cl (= @p_976 (! (is_res$a @p_712 @p_1108) :named @p_1109))) :rule cong :premises (t82.t3)) +(step t82.t5 (cl @p_1110) :rule refl) +(step t82.t6 (cl (= @p_978 (! (c$ veriT_vr107 veriT_sk0) :named @p_1111))) :rule cong :premises (t82.t5)) +(step t82.t7 (cl @p_1112) :rule refl) +(step t82.t8 (cl (= @p_980 (! (run$b @p_1111 veriT_vr108) :named @p_1113))) :rule cong :premises (t82.t6 t82.t7)) +(step t82.t9 (cl (= @p_981 (! (is_fail$b @p_1113) :named @p_1114))) :rule cong :premises (t82.t8)) +(step t82.t10 (cl (= @p_982 (! (and @p_1109 @p_1114) :named @p_1115))) :rule cong :premises (t82.t4 t82.t9)) +(step t82 (cl (= @p_983 (! (exists ((veriT_vr107 E$) (veriT_vr108 D$)) @p_1115) :named @p_1116))) :rule bind) +(step t83 (cl (= @p_986 (! (or @p_720 @p_1116) :named @p_1117))) :rule cong :premises (t82)) +(step t84 (cl (= @p_996 (! (=> @p_790 @p_1117) :named @p_1148))) :rule cong :premises (t83)) +(anchor :step t85 :args ((:= (veriT_vr91 E$) veriT_vr109) (:= (veriT_vr90 D$) veriT_vr110))) +(step t85.t1 (cl (! (= veriT_vr91 veriT_vr109) :named @p_1120)) :rule refl) +(step t85.t2 (cl (! (= veriT_vr90 veriT_vr110) :named @p_1122)) :rule refl) +(step t85.t3 (cl (= @p_975 (! (pair$a veriT_vr109 veriT_vr110) :named @p_1118))) :rule cong :premises (t85.t1 t85.t2)) +(step t85.t4 (cl (= @p_987 (! (is_res$a @p_714 @p_1118) :named @p_1119))) :rule cong :premises (t85.t3)) +(step t85.t5 (cl @p_1120) :rule refl) +(step t85.t6 (cl (= @p_989 (! (c$ veriT_vr109 veriT_sk1) :named @p_1121))) :rule cong :premises (t85.t5)) +(step t85.t7 (cl @p_1122) :rule refl) +(step t85.t8 (cl (= @p_991 (! (run$b @p_1121 veriT_vr110) :named @p_1123))) :rule cong :premises (t85.t6 t85.t7)) +(step t85.t9 (cl (= @p_992 (! (is_fail$b @p_1123) :named @p_1124))) :rule cong :premises (t85.t8)) +(step t85.t10 (cl (= @p_993 (! (and @p_1119 @p_1124) :named @p_1125))) :rule cong :premises (t85.t4 t85.t9)) +(step t85 (cl (= @p_994 (! (exists ((veriT_vr109 E$) (veriT_vr110 D$)) @p_1125) :named @p_1126))) :rule bind) +(step t86 (cl (= @p_995 (! (or @p_726 @p_1126) :named @p_1127))) :rule cong :premises (t85)) +(step t87 (cl (= @p_997 (! (=> @p_808 @p_1127) :named @p_1149))) :rule cong :premises (t86)) +(anchor :step t88 :args ((:= (veriT_vr91 E$) veriT_vr111) (:= (veriT_vr90 D$) veriT_vr112))) +(step t88.t1 (cl (! (= veriT_vr91 veriT_vr111) :named @p_1130)) :rule refl) +(step t88.t2 (cl (! (= veriT_vr90 veriT_vr112) :named @p_1132)) :rule refl) +(step t88.t3 (cl (= @p_975 (! (pair$a veriT_vr111 veriT_vr112) :named @p_1128))) :rule cong :premises (t88.t1 t88.t2)) +(step t88.t4 (cl (= @p_976 (! (is_res$a @p_712 @p_1128) :named @p_1129))) :rule cong :premises (t88.t3)) +(step t88.t5 (cl @p_1130) :rule refl) +(step t88.t6 (cl (= @p_978 (! (c$ veriT_vr111 veriT_sk0) :named @p_1131))) :rule cong :premises (t88.t5)) +(step t88.t7 (cl @p_1132) :rule refl) +(step t88.t8 (cl (= @p_980 (! (run$b @p_1131 veriT_vr112) :named @p_1133))) :rule cong :premises (t88.t6 t88.t7)) +(step t88.t9 (cl (= @p_1000 (! (is_res$b @p_1133 @p_747) :named @p_1134))) :rule cong :premises (t88.t8)) +(step t88.t10 (cl (= @p_1001 (! (and @p_1129 @p_1134) :named @p_1135))) :rule cong :premises (t88.t4 t88.t9)) +(step t88 (cl (= @p_1002 (! (exists ((veriT_vr111 E$) (veriT_vr112 D$)) @p_1135) :named @p_1136))) :rule bind) +(step t89 (cl (= @p_1003 (! (or @p_720 @p_1136) :named @p_1137))) :rule cong :premises (t88)) +(step t90 (cl (= @p_1010 (! (=> @p_828 @p_1137) :named @p_1150))) :rule cong :premises (t89)) +(anchor :step t91 :args ((:= (veriT_vr91 E$) veriT_vr113) (:= (veriT_vr90 D$) veriT_vr114))) +(step t91.t1 (cl (! (= veriT_vr91 veriT_vr113) :named @p_1140)) :rule refl) +(step t91.t2 (cl (! (= veriT_vr90 veriT_vr114) :named @p_1142)) :rule refl) +(step t91.t3 (cl (= @p_975 (! (pair$a veriT_vr113 veriT_vr114) :named @p_1138))) :rule cong :premises (t91.t1 t91.t2)) +(step t91.t4 (cl (= @p_987 (! (is_res$a @p_714 @p_1138) :named @p_1139))) :rule cong :premises (t91.t3)) +(step t91.t5 (cl @p_1140) :rule refl) +(step t91.t6 (cl (= @p_989 (! (c$ veriT_vr113 veriT_sk1) :named @p_1141))) :rule cong :premises (t91.t5)) +(step t91.t7 (cl @p_1142) :rule refl) +(step t91.t8 (cl (= @p_991 (! (run$b @p_1141 veriT_vr114) :named @p_1143))) :rule cong :premises (t91.t6 t91.t7)) +(step t91.t9 (cl (= @p_1006 (! (is_res$b @p_1143 @p_747) :named @p_1144))) :rule cong :premises (t91.t8)) +(step t91.t10 (cl (= @p_1007 (! (and @p_1139 @p_1144) :named @p_1145))) :rule cong :premises (t91.t4 t91.t9)) +(step t91 (cl (= @p_1008 (! (exists ((veriT_vr113 E$) (veriT_vr114 D$)) @p_1145) :named @p_1146))) :rule bind) +(step t92 (cl (= @p_1009 (! (or @p_726 @p_1146) :named @p_1147))) :rule cong :premises (t91)) +(step t93 (cl (= @p_1011 (! (=> @p_848 @p_1147) :named @p_1151))) :rule cong :premises (t92)) +(step t94 (cl (= @p_1026 (! (and @p_1148 @p_1149 @p_1150 @p_1151) :named @p_1152))) :rule cong :premises (t84 t87 t90 t93)) +(step t95 (cl (= @p_1028 (! (or @p_720 @p_1116 @p_1152) :named @p_1153))) :rule cong :premises (t82 t94)) +(step t96 (cl (= @p_1029 (! (not @p_1153) :named @p_1156))) :rule cong :premises (t95)) +(step t97 (cl (! (= @p_1031 (! (and @p_1018 @p_1154 @p_1155 @p_1156) :named @p_1158)) :named @p_1157)) :rule cong :premises (t80 t81 t96)) +(step t98 (cl (not @p_1157) (not @p_1031) @p_1158) :rule equiv_pos2) +(step t99 (cl @p_1158) :rule th_resolution :premises (t79 t97 t98)) +(step t100 (cl (= @p_1116 (! (not (! (forall ((veriT_vr107 E$) (veriT_vr108 D$)) (not @p_1115)) :named @p_1177)) :named @p_1159))) :rule connective_def) +(step t101 (cl (= @p_1117 (! (or @p_720 @p_1159) :named @p_1160))) :rule cong :premises (t100)) +(step t102 (cl (= @p_1148 (! (=> @p_790 @p_1160) :named @p_1167))) :rule cong :premises (t101)) +(step t103 (cl (= @p_1126 (! (not (forall ((veriT_vr109 E$) (veriT_vr110 D$)) (not @p_1125))) :named @p_1161))) :rule connective_def) +(step t104 (cl (= @p_1127 (! (or @p_726 @p_1161) :named @p_1162))) :rule cong :premises (t103)) +(step t105 (cl (= @p_1149 (! (=> @p_808 @p_1162) :named @p_1168))) :rule cong :premises (t104)) +(step t106 (cl (= @p_1136 (! (not (! (forall ((veriT_vr111 E$) (veriT_vr112 D$)) (not @p_1135)) :named @p_1181)) :named @p_1163))) :rule connective_def) +(step t107 (cl (= @p_1137 (! (or @p_720 @p_1163) :named @p_1164))) :rule cong :premises (t106)) +(step t108 (cl (= @p_1150 (! (=> @p_828 @p_1164) :named @p_1169))) :rule cong :premises (t107)) +(step t109 (cl (= @p_1146 (! (not (! (forall ((veriT_vr113 E$) (veriT_vr114 D$)) (not @p_1145)) :named @p_1184)) :named @p_1165))) :rule connective_def) +(step t110 (cl (= @p_1147 (! (or @p_726 @p_1165) :named @p_1166))) :rule cong :premises (t109)) +(step t111 (cl (= @p_1151 (! (=> @p_848 @p_1166) :named @p_1170))) :rule cong :premises (t110)) +(step t112 (cl (= @p_1152 (! (and @p_1167 @p_1168 @p_1169 @p_1170) :named @p_1171))) :rule cong :premises (t102 t105 t108 t111)) +(step t113 (cl (= @p_1153 (! (or @p_720 @p_1159 @p_1171) :named @p_1172))) :rule cong :premises (t100 t112)) +(step t114 (cl (= @p_1156 (! (not @p_1172) :named @p_1173))) :rule cong :premises (t113)) +(step t115 (cl (! (= @p_1158 (! (and @p_1018 @p_1154 @p_1155 @p_1173) :named @p_1175)) :named @p_1174)) :rule cong :premises (t114)) +(step t116 (cl (not @p_1174) (not @p_1158) @p_1175) :rule equiv_pos2) +(step t117 (cl @p_1175) :rule th_resolution :premises (t99 t115 t116)) +(step t118 (cl @p_1018) :rule and :premises (t117)) +(step t119 (cl @p_1154) :rule and :premises (t117)) +(step t120 (cl @p_1155) :rule and :premises (t117)) +(step t121 (cl @p_1173) :rule and :premises (t117)) +(step t122 (cl (not @p_720)) :rule not_or :premises (t121)) +(step t123 (cl (! (not @p_1159) :named @p_1176)) :rule not_or :premises (t121)) +(step t124 (cl (not @p_1176) @p_1177) :rule not_not) +(step t125 (cl @p_1177) :rule th_resolution :premises (t124 t123)) +(step t126 (cl (not @p_1171)) :rule not_or :premises (t121)) +(step t127 (cl (! (not @p_780) :named @p_1178) @p_784) :rule and_pos) +(step t128 (cl @p_1178 @p_789) :rule and_pos) +(step t129 (cl (! (not @p_790) :named @p_1925) @p_726 @p_780) :rule or_pos) +(step t130 (cl @p_1167 @p_790) :rule implies_neg1) +(step t131 (cl (! (not @p_808) :named @p_1185) @p_720 @p_800) :rule or_pos) +(step t132 (cl @p_1168 @p_808) :rule implies_neg1) +(step t133 (cl (! (not @p_820) :named @p_1179) @p_822) :rule and_pos) +(step t134 (cl @p_1179 @p_827) :rule and_pos) +(step t135 (cl (not @p_828) @p_726 @p_820) :rule or_pos) +(step t136 (cl @p_1169 @p_828) :rule implies_neg1) +(step t137 (cl @p_1164 (! (not @p_1163) :named @p_1180)) :rule or_neg) +(step t138 (cl (not @p_1180) @p_1181) :rule not_not) +(step t139 (cl @p_1164 @p_1181) :rule th_resolution :premises (t138 t137)) +(step t140 (cl @p_1169 (! (not @p_1164) :named @p_1927)) :rule implies_neg2) +(step t141 (cl (! (not @p_840) :named @p_1182) @p_842) :rule and_pos) +(step t142 (cl @p_1182 @p_847) :rule and_pos) +(step t143 (cl (! (not @p_848) :named @p_1186) @p_720 @p_840) :rule or_pos) +(step t144 (cl @p_1170 @p_848) :rule implies_neg1) +(step t145 (cl @p_1166 (! (not @p_1165) :named @p_1183)) :rule or_neg) +(step t146 (cl (not @p_1183) @p_1184) :rule not_not) +(step t147 (cl @p_1166 @p_1184) :rule th_resolution :premises (t146 t145)) +(step t148 (cl @p_1170 (not @p_1166)) :rule implies_neg2) +(step t149 (cl (not @p_1167) (not @p_1168) (! (not @p_1169) :named @p_1926) (not @p_1170)) :rule not_and :premises (t126)) +(step t150 (cl @p_1185 @p_800) :rule resolution :premises (t131 t122)) +(step t151 (cl @p_1186 @p_840) :rule resolution :premises (t143 t122)) +(step t152 (cl (or (! (not @p_1155) :named @p_1187) (! (forall ((veriT_vr103 A$) (veriT_vr104 C$)) (or @p_1084 (not @p_1092) @p_1084)) :named @p_1340))) :rule qnt_cnf) +(step t153 (cl (or @p_1187 (! (forall ((veriT_vr103 A$) (veriT_vr104 C$) (veriT_vr105 B$) (veriT_vr106 C$)) (or @p_1084 (not @p_1100) @p_1104)) :named @p_1351))) :rule qnt_cnf) +(step t154 (cl (or @p_1187 (! (forall ((veriT_vr103 A$) (veriT_vr104 C$) (veriT_vr105 B$) (veriT_vr106 C$)) (or @p_1084 (not @p_1104) @p_1100)) :named @p_1352))) :rule qnt_cnf) +(step t155 (cl (or @p_1159 (! (not @p_800) :named @p_1323))) :rule forall_inst :args ((:= veriT_vr107 veriT_sk5) (:= veriT_vr108 veriT_sk6))) +(step t156 (cl (or @p_1159 (! (not (! (and @p_842 (! (is_fail$b @p_846) :named @p_1325)) :named @p_1324)) :named @p_1326))) :rule forall_inst :args ((:= veriT_vr107 veriT_sk11) (:= veriT_vr108 veriT_sk12))) +(step t157 (cl (or @p_1159 (! (not (! (and (! (is_res$a @p_712 @p_821) :named @p_1328) (! (is_fail$b (! (run$b (c$ veriT_sk9 veriT_sk0) veriT_sk10) :named @p_1353)) :named @p_1329)) :named @p_1327)) :named @p_1330))) :rule forall_inst :args ((:= veriT_vr107 veriT_sk9) (:= veriT_vr108 veriT_sk10))) +(step t158 (cl (or @p_1159 (! (not (! (and (! (is_res$a @p_712 @p_783) :named @p_1332) (! (is_fail$b (! (run$b (c$ veriT_sk3 veriT_sk0) veriT_sk4) :named @p_1389)) :named @p_1333)) :named @p_1331)) :named @p_1334))) :rule forall_inst :args ((:= veriT_vr107 veriT_sk3) (:= veriT_vr108 veriT_sk4))) +(step t159 (cl (not (! (not (! (not @p_1154) :named @p_1385)) :named @p_1533)) @p_1154) :rule not_not) +(step t160 (cl (or (! (not @p_1018) :named @p_1320) (! (=> @p_1015 (! (or @p_720 (! (and (! (= @p_726 @p_720) :named @p_1208) (! (forall ((veriT_vr91 E$) (veriT_vr92 D$)) (! (= (! (is_res$a @p_714 @p_869) :named @p_1236) (! (is_res$a @p_712 @p_869) :named @p_1241)) :named @p_1243)) :named @p_1234)) :named @p_1245)) :named @p_1247)) :named @p_1188))) :rule forall_inst :args ((:= veriT_vr84 veriT_sk0) (:= veriT_vr85 veriT_sk1) (:= veriT_vr90 veriT_sk2))) +(anchor :step t161) +(assume t161.h1 @p_1188) +(anchor :step t161.t2 :args ((:= (veriT_vr86 A$) veriT_vr115) (:= (veriT_vr87 C$) veriT_vr116))) +(step t161.t2.t1 (cl (! (= veriT_vr86 veriT_vr115) :named @p_1215)) :rule refl) +(step t161.t2.t2 (cl (! (= @p_955 (! (fun_app$ veriT_sk0 veriT_vr115) :named @p_1214)) :named @p_1219)) :rule cong :premises (t161.t2.t1)) +(step t161.t2.t3 (cl (! (= veriT_vr87 veriT_vr116) :named @p_1217)) :rule refl) +(step t161.t2.t4 (cl (! (= @p_954 (! (run$ @p_1214 veriT_vr116) :named @p_1212)) :named @p_1220)) :rule cong :premises (t161.t2.t2 t161.t2.t3)) +(step t161.t2.t5 (cl (! (= @p_952 (! (is_fail$ @p_1212) :named @p_1210)) :named @p_1221)) :rule cong :premises (t161.t2.t4)) +(step t161.t2.t6 (cl @p_1215) :rule refl) +(step t161.t2.t7 (cl (! (= @p_957 (! (fun_app$ veriT_sk1 veriT_vr115) :named @p_1216)) :named @p_1224)) :rule cong :premises (t161.t2.t6)) +(step t161.t2.t8 (cl @p_1217) :rule refl) +(step t161.t2.t9 (cl (! (= @p_953 (! (run$ @p_1216 veriT_vr116) :named @p_1211)) :named @p_1225)) :rule cong :premises (t161.t2.t7 t161.t2.t8)) +(step t161.t2.t10 (cl (= @p_959 (! (is_fail$ @p_1211) :named @p_1218))) :rule cong :premises (t161.t2.t9)) +(step t161.t2.t11 (cl @p_1215) :rule refl) +(step t161.t2.t12 (cl @p_1219) :rule cong :premises (t161.t2.t11)) +(step t161.t2.t13 (cl @p_1217) :rule refl) +(step t161.t2.t14 (cl @p_1220) :rule cong :premises (t161.t2.t12 t161.t2.t13)) +(step t161.t2.t15 (cl @p_1221) :rule cong :premises (t161.t2.t14)) +(step t161.t2.t16 (cl (= @p_963 (! (= @p_1218 @p_1210) :named @p_1222))) :rule cong :premises (t161.t2.t10 t161.t2.t15)) +(anchor :step t161.t2.t17 :args ((:= (veriT_vr88 B$) veriT_vr117) (:= (veriT_vr89 C$) veriT_vr118))) +(step t161.t2.t17.t1 (cl @p_1215) :rule refl) +(step t161.t2.t17.t2 (cl @p_1224) :rule cong :premises (t161.t2.t17.t1)) +(step t161.t2.t17.t3 (cl @p_1217) :rule refl) +(step t161.t2.t17.t4 (cl @p_1225) :rule cong :premises (t161.t2.t17.t2 t161.t2.t17.t3)) +(step t161.t2.t17.t5 (cl (! (= veriT_vr88 veriT_vr117) :named @p_1227)) :rule refl) +(step t161.t2.t17.t6 (cl (! (= veriT_vr89 veriT_vr118) :named @p_1228)) :rule refl) +(step t161.t2.t17.t7 (cl (! (= @p_865 (! (pair$ veriT_vr117 veriT_vr118) :named @p_1213)) :named @p_1229)) :rule cong :premises (t161.t2.t17.t5 t161.t2.t17.t6)) +(step t161.t2.t17.t8 (cl (= @p_967 (! (is_res$ @p_1211 @p_1213) :named @p_1226))) :rule cong :premises (t161.t2.t17.t4 t161.t2.t17.t7)) +(step t161.t2.t17.t9 (cl @p_1215) :rule refl) +(step t161.t2.t17.t10 (cl @p_1219) :rule cong :premises (t161.t2.t17.t9)) +(step t161.t2.t17.t11 (cl @p_1217) :rule refl) +(step t161.t2.t17.t12 (cl @p_1220) :rule cong :premises (t161.t2.t17.t10 t161.t2.t17.t11)) +(step t161.t2.t17.t13 (cl @p_1227) :rule refl) +(step t161.t2.t17.t14 (cl @p_1228) :rule refl) +(step t161.t2.t17.t15 (cl @p_1229) :rule cong :premises (t161.t2.t17.t13 t161.t2.t17.t14)) +(step t161.t2.t17.t16 (cl (= @p_971 (! (is_res$ @p_1212 @p_1213) :named @p_1230))) :rule cong :premises (t161.t2.t17.t12 t161.t2.t17.t15)) +(step t161.t2.t17.t17 (cl (= @p_972 (! (= @p_1226 @p_1230) :named @p_1231))) :rule cong :premises (t161.t2.t17.t8 t161.t2.t17.t16)) +(step t161.t2.t17 (cl (= @p_964 (! (forall ((veriT_vr117 B$) (veriT_vr118 C$)) @p_1231) :named @p_1223))) :rule bind) +(step t161.t2.t18 (cl (= @p_973 (! (and @p_1222 @p_1223) :named @p_1232))) :rule cong :premises (t161.t2.t16 t161.t2.t17)) +(step t161.t2.t19 (cl (= @p_974 (! (or @p_1210 @p_1232) :named @p_1233))) :rule cong :premises (t161.t2.t5 t161.t2.t18)) +(step t161.t2 (cl (= @p_1015 (! (forall ((veriT_vr115 A$) (veriT_vr116 C$)) @p_1233) :named @p_1249))) :rule bind) +(anchor :step t161.t3 :args ((:= (veriT_vr91 E$) veriT_vr121) (:= (veriT_vr92 D$) veriT_vr120))) +(step t161.t3.t1 (cl (! (= veriT_vr91 veriT_vr121) :named @p_1238)) :rule refl) +(step t161.t3.t2 (cl (! (= veriT_vr92 veriT_vr120) :named @p_1239)) :rule refl) +(step t161.t3.t3 (cl (! (= @p_869 (! (pair$a veriT_vr121 veriT_vr120) :named @p_1235)) :named @p_1240)) :rule cong :premises (t161.t3.t1 t161.t3.t2)) +(step t161.t3.t4 (cl (= @p_1236 (! (is_res$a @p_714 @p_1235) :named @p_1237))) :rule cong :premises (t161.t3.t3)) +(step t161.t3.t5 (cl @p_1238) :rule refl) +(step t161.t3.t6 (cl @p_1239) :rule refl) +(step t161.t3.t7 (cl @p_1240) :rule cong :premises (t161.t3.t5 t161.t3.t6)) +(step t161.t3.t8 (cl (= @p_1241 (! (is_res$a @p_712 @p_1235) :named @p_1242))) :rule cong :premises (t161.t3.t7)) +(step t161.t3.t9 (cl (= @p_1243 (! (= @p_1237 @p_1242) :named @p_1244))) :rule cong :premises (t161.t3.t4 t161.t3.t8)) +(step t161.t3 (cl (= @p_1234 (! (forall ((veriT_vr121 E$) (veriT_vr120 D$)) @p_1244) :named @p_1246))) :rule bind) +(step t161.t4 (cl (= @p_1245 (! (and @p_1208 @p_1246) :named @p_1248))) :rule cong :premises (t161.t3)) +(step t161.t5 (cl (= @p_1247 (! (or @p_720 @p_1248) :named @p_1250))) :rule cong :premises (t161.t4)) +(step t161.t6 (cl (! (= @p_1188 (! (=> @p_1249 @p_1250) :named @p_1253)) :named @p_1251)) :rule cong :premises (t161.t2 t161.t5)) +(step t161.t7 (cl (not @p_1251) (! (not @p_1188) :named @p_1252) @p_1253) :rule equiv_pos2) +(step t161.t8 (cl @p_1253) :rule th_resolution :premises (t161.h1 t161.t6 t161.t7)) +(anchor :step t161.t9 :args ((:= (veriT_vr115 A$) veriT_vr122) (:= (veriT_vr116 C$) veriT_vr123))) +(step t161.t9.t1 (cl (! (= veriT_vr115 veriT_vr122) :named @p_1256)) :rule refl) +(step t161.t9.t2 (cl (! (= @p_1214 @p_1255) :named @p_1260)) :rule cong :premises (t161.t9.t1)) +(step t161.t9.t3 (cl (! (= veriT_vr116 veriT_vr123) :named @p_1258)) :rule refl) +(step t161.t9.t4 (cl (! (= @p_1212 @p_1191) :named @p_1261)) :rule cong :premises (t161.t9.t2 t161.t9.t3)) +(step t161.t9.t5 (cl (! (= @p_1210 @p_1189) :named @p_1262)) :rule cong :premises (t161.t9.t4)) +(step t161.t9.t6 (cl @p_1256) :rule refl) +(step t161.t9.t7 (cl (! (= @p_1216 @p_1257) :named @p_1265)) :rule cong :premises (t161.t9.t6)) +(step t161.t9.t8 (cl @p_1258) :rule refl) +(step t161.t9.t9 (cl (! (= @p_1211 @p_1190) :named @p_1266)) :rule cong :premises (t161.t9.t7 t161.t9.t8)) +(step t161.t9.t10 (cl (= @p_1218 @p_1259)) :rule cong :premises (t161.t9.t9)) +(step t161.t9.t11 (cl @p_1256) :rule refl) +(step t161.t9.t12 (cl @p_1260) :rule cong :premises (t161.t9.t11)) +(step t161.t9.t13 (cl @p_1258) :rule refl) +(step t161.t9.t14 (cl @p_1261) :rule cong :premises (t161.t9.t12 t161.t9.t13)) +(step t161.t9.t15 (cl @p_1262) :rule cong :premises (t161.t9.t14)) +(step t161.t9.t16 (cl (= @p_1222 @p_1263)) :rule cong :premises (t161.t9.t10 t161.t9.t15)) +(anchor :step t161.t9.t17 :args ((:= (veriT_vr117 B$) veriT_vr124) (:= (veriT_vr118 C$) veriT_vr125))) +(step t161.t9.t17.t1 (cl @p_1256) :rule refl) +(step t161.t9.t17.t2 (cl @p_1265) :rule cong :premises (t161.t9.t17.t1)) +(step t161.t9.t17.t3 (cl @p_1258) :rule refl) +(step t161.t9.t17.t4 (cl @p_1266) :rule cong :premises (t161.t9.t17.t2 t161.t9.t17.t3)) +(step t161.t9.t17.t5 (cl (! (= veriT_vr117 veriT_vr124) :named @p_1268)) :rule refl) +(step t161.t9.t17.t6 (cl (! (= veriT_vr118 veriT_vr125) :named @p_1269)) :rule refl) +(step t161.t9.t17.t7 (cl (! (= @p_1213 @p_1192) :named @p_1270)) :rule cong :premises (t161.t9.t17.t5 t161.t9.t17.t6)) +(step t161.t9.t17.t8 (cl (= @p_1226 @p_1267)) :rule cong :premises (t161.t9.t17.t4 t161.t9.t17.t7)) +(step t161.t9.t17.t9 (cl @p_1256) :rule refl) +(step t161.t9.t17.t10 (cl @p_1260) :rule cong :premises (t161.t9.t17.t9)) +(step t161.t9.t17.t11 (cl @p_1258) :rule refl) +(step t161.t9.t17.t12 (cl @p_1261) :rule cong :premises (t161.t9.t17.t10 t161.t9.t17.t11)) +(step t161.t9.t17.t13 (cl @p_1268) :rule refl) +(step t161.t9.t17.t14 (cl @p_1269) :rule refl) +(step t161.t9.t17.t15 (cl @p_1270) :rule cong :premises (t161.t9.t17.t13 t161.t9.t17.t14)) +(step t161.t9.t17.t16 (cl (= @p_1230 @p_1271)) :rule cong :premises (t161.t9.t17.t12 t161.t9.t17.t15)) +(step t161.t9.t17.t17 (cl (= @p_1231 @p_1272)) :rule cong :premises (t161.t9.t17.t8 t161.t9.t17.t16)) +(step t161.t9.t17 (cl (= @p_1223 @p_1264)) :rule bind) +(step t161.t9.t18 (cl (= @p_1232 @p_1273)) :rule cong :premises (t161.t9.t16 t161.t9.t17)) +(step t161.t9.t19 (cl (= @p_1233 @p_1254)) :rule cong :premises (t161.t9.t5 t161.t9.t18)) +(step t161.t9 (cl (= @p_1249 (! (forall ((veriT_vr122 A$) (veriT_vr123 C$)) @p_1254) :named @p_1283))) :rule bind) +(anchor :step t161.t10 :args ((:= (veriT_vr121 E$) veriT_vr126) (:= (veriT_vr120 D$) veriT_vr127))) +(step t161.t10.t1 (cl (! (= veriT_vr121 veriT_vr126) :named @p_1276)) :rule refl) +(step t161.t10.t2 (cl (! (= veriT_vr120 veriT_vr127) :named @p_1277)) :rule refl) +(step t161.t10.t3 (cl (! (= @p_1235 (! (pair$a veriT_vr126 veriT_vr127) :named @p_1274)) :named @p_1278)) :rule cong :premises (t161.t10.t1 t161.t10.t2)) +(step t161.t10.t4 (cl (= @p_1237 (! (is_res$a @p_714 @p_1274) :named @p_1275))) :rule cong :premises (t161.t10.t3)) +(step t161.t10.t5 (cl @p_1276) :rule refl) +(step t161.t10.t6 (cl @p_1277) :rule refl) +(step t161.t10.t7 (cl @p_1278) :rule cong :premises (t161.t10.t5 t161.t10.t6)) +(step t161.t10.t8 (cl (= @p_1242 (! (is_res$a @p_712 @p_1274) :named @p_1279))) :rule cong :premises (t161.t10.t7)) +(step t161.t10.t9 (cl (= @p_1244 (! (= @p_1275 @p_1279) :named @p_1280))) :rule cong :premises (t161.t10.t4 t161.t10.t8)) +(step t161.t10 (cl (= @p_1246 (! (forall ((veriT_vr126 E$) (veriT_vr127 D$)) @p_1280) :named @p_1281))) :rule bind) +(step t161.t11 (cl (= @p_1248 (! (and @p_1208 @p_1281) :named @p_1282))) :rule cong :premises (t161.t10)) +(step t161.t12 (cl (= @p_1250 (! (or @p_720 @p_1282) :named @p_1284))) :rule cong :premises (t161.t11)) +(step t161.t13 (cl (! (= @p_1253 (! (=> @p_1283 @p_1284) :named @p_1286)) :named @p_1285)) :rule cong :premises (t161.t9 t161.t12)) +(step t161.t14 (cl (not @p_1285) (not @p_1253) @p_1286) :rule equiv_pos2) +(step t161.t15 (cl @p_1286) :rule th_resolution :premises (t161.t8 t161.t13 t161.t14)) +(anchor :step t161.t16 :args ((:= (veriT_vr122 A$) veriT_sk13) (:= (veriT_vr123 C$) veriT_sk14))) +(step t161.t16.t1 (cl (! (= veriT_vr122 veriT_sk13) :named @p_1289)) :rule refl) +(step t161.t16.t2 (cl (! (= @p_1255 (! (fun_app$ veriT_sk0 veriT_sk13) :named @p_1288)) :named @p_1293)) :rule cong :premises (t161.t16.t1)) +(step t161.t16.t3 (cl (! (= veriT_vr123 veriT_sk14) :named @p_1291)) :rule refl) +(step t161.t16.t4 (cl (! (= @p_1191 (! (run$ @p_1288 veriT_sk14) :named @p_1206)) :named @p_1294)) :rule cong :premises (t161.t16.t2 t161.t16.t3)) +(step t161.t16.t5 (cl (! (= @p_1189 (! (is_fail$ @p_1206) :named @p_1198)) :named @p_1295)) :rule cong :premises (t161.t16.t4)) +(step t161.t16.t6 (cl @p_1289) :rule refl) +(step t161.t16.t7 (cl (! (= @p_1257 (! (fun_app$ veriT_sk1 veriT_sk13) :named @p_1290)) :named @p_1298)) :rule cong :premises (t161.t16.t6)) +(step t161.t16.t8 (cl @p_1291) :rule refl) +(step t161.t16.t9 (cl (! (= @p_1190 (! (run$ @p_1290 veriT_sk14) :named @p_1199)) :named @p_1299)) :rule cong :premises (t161.t16.t7 t161.t16.t8)) +(step t161.t16.t10 (cl (= @p_1259 (! (is_fail$ @p_1199) :named @p_1292))) :rule cong :premises (t161.t16.t9)) +(step t161.t16.t11 (cl @p_1289) :rule refl) +(step t161.t16.t12 (cl @p_1293) :rule cong :premises (t161.t16.t11)) +(step t161.t16.t13 (cl @p_1291) :rule refl) +(step t161.t16.t14 (cl @p_1294) :rule cong :premises (t161.t16.t12 t161.t16.t13)) +(step t161.t16.t15 (cl @p_1295) :rule cong :premises (t161.t16.t14)) +(step t161.t16.t16 (cl (= @p_1263 (! (= @p_1292 @p_1198) :named @p_1296))) :rule cong :premises (t161.t16.t10 t161.t16.t15)) +(anchor :step t161.t16.t17 :args ((:= (veriT_vr124 B$) veriT_sk15) (:= (veriT_vr125 C$) veriT_sk16))) +(step t161.t16.t17.t1 (cl @p_1289) :rule refl) +(step t161.t16.t17.t2 (cl @p_1298) :rule cong :premises (t161.t16.t17.t1)) +(step t161.t16.t17.t3 (cl @p_1291) :rule refl) +(step t161.t16.t17.t4 (cl @p_1299) :rule cong :premises (t161.t16.t17.t2 t161.t16.t17.t3)) +(step t161.t16.t17.t5 (cl (! (= veriT_vr124 veriT_sk15) :named @p_1302)) :rule refl) +(step t161.t16.t17.t6 (cl (! (= veriT_vr125 veriT_sk16) :named @p_1303)) :rule refl) +(step t161.t16.t17.t7 (cl (! (= @p_1192 (! (pair$ veriT_sk15 veriT_sk16) :named @p_1207)) :named @p_1304)) :rule cong :premises (t161.t16.t17.t5 t161.t16.t17.t6)) +(step t161.t16.t17.t8 (cl (= @p_1267 (! (is_res$ @p_1199 @p_1207) :named @p_1301))) :rule cong :premises (t161.t16.t17.t4 t161.t16.t17.t7)) +(step t161.t16.t17.t9 (cl @p_1289) :rule refl) +(step t161.t16.t17.t10 (cl @p_1293) :rule cong :premises (t161.t16.t17.t9)) +(step t161.t16.t17.t11 (cl @p_1291) :rule refl) +(step t161.t16.t17.t12 (cl @p_1294) :rule cong :premises (t161.t16.t17.t10 t161.t16.t17.t11)) +(step t161.t16.t17.t13 (cl @p_1302) :rule refl) +(step t161.t16.t17.t14 (cl @p_1303) :rule refl) +(step t161.t16.t17.t15 (cl @p_1304) :rule cong :premises (t161.t16.t17.t13 t161.t16.t17.t14)) +(step t161.t16.t17.t16 (cl (= @p_1271 (! (is_res$ @p_1206 @p_1207) :named @p_1305))) :rule cong :premises (t161.t16.t17.t12 t161.t16.t17.t15)) +(step t161.t16.t17.t17 (cl (= @p_1272 (! (= @p_1301 @p_1305) :named @p_1297))) :rule cong :premises (t161.t16.t17.t8 t161.t16.t17.t16)) +(step t161.t16.t17 (cl (= @p_1264 @p_1297)) :rule sko_forall) +(step t161.t16.t18 (cl (= @p_1273 (! (and @p_1296 @p_1297) :named @p_1306))) :rule cong :premises (t161.t16.t16 t161.t16.t17)) +(step t161.t16.t19 (cl (= @p_1254 (! (or @p_1198 @p_1306) :named @p_1287))) :rule cong :premises (t161.t16.t5 t161.t16.t18)) +(step t161.t16 (cl (= @p_1283 @p_1287)) :rule sko_forall) +(step t161.t17 (cl (! (= @p_1286 (! (=> @p_1287 @p_1284) :named @p_1308)) :named @p_1307)) :rule cong :premises (t161.t16)) +(step t161.t18 (cl (not @p_1307) (not @p_1286) @p_1308) :rule equiv_pos2) +(step t161.t19 (cl @p_1308) :rule th_resolution :premises (t161.t15 t161.t17 t161.t18)) +(anchor :step t161.t20 :args ((:= (veriT_vr126 E$) veriT_vr128) (:= (veriT_vr127 D$) veriT_vr129))) +(step t161.t20.t1 (cl (! (= veriT_vr126 veriT_vr128) :named @p_1311)) :rule refl) +(step t161.t20.t2 (cl (! (= veriT_vr127 veriT_vr129) :named @p_1312)) :rule refl) +(step t161.t20.t3 (cl (! (= @p_1274 (! (pair$a veriT_vr128 veriT_vr129) :named @p_1209)) :named @p_1313)) :rule cong :premises (t161.t20.t1 t161.t20.t2)) +(step t161.t20.t4 (cl (= @p_1275 (! (is_res$a @p_714 @p_1209) :named @p_1310))) :rule cong :premises (t161.t20.t3)) +(step t161.t20.t5 (cl @p_1311) :rule refl) +(step t161.t20.t6 (cl @p_1312) :rule refl) +(step t161.t20.t7 (cl @p_1313) :rule cong :premises (t161.t20.t5 t161.t20.t6)) +(step t161.t20.t8 (cl (= @p_1279 (! (is_res$a @p_712 @p_1209) :named @p_1314))) :rule cong :premises (t161.t20.t7)) +(step t161.t20.t9 (cl (= @p_1280 (! (= @p_1310 @p_1314) :named @p_1315))) :rule cong :premises (t161.t20.t4 t161.t20.t8)) +(step t161.t20 (cl (= @p_1281 (! (forall ((veriT_vr128 E$) (veriT_vr129 D$)) @p_1315) :named @p_1309))) :rule bind) +(step t161.t21 (cl (= @p_1282 (! (and @p_1208 @p_1309) :named @p_1316))) :rule cong :premises (t161.t20)) +(step t161.t22 (cl (= @p_1284 (! (or @p_720 @p_1316) :named @p_1317))) :rule cong :premises (t161.t21)) +(step t161.t23 (cl (! (= @p_1308 (! (=> @p_1287 @p_1317) :named @p_1318)) :named @p_1319)) :rule cong :premises (t161.t22)) +(step t161.t24 (cl (not @p_1319) (not @p_1308) @p_1318) :rule equiv_pos2) +(step t161.t25 (cl @p_1318) :rule th_resolution :premises (t161.t19 t161.t23 t161.t24)) +(step t161 (cl @p_1252 @p_1318) :rule subproof :discharge (h1)) +(step t162 (cl @p_1320 @p_1188) :rule or :premises (t160)) +(step t163 (cl (! (or @p_1320 @p_1318) :named @p_1322) (! (not @p_1320) :named @p_1321)) :rule or_neg) +(step t164 (cl (not @p_1321) @p_1018) :rule not_not) +(step t165 (cl @p_1322 @p_1018) :rule th_resolution :premises (t164 t163)) +(step t166 (cl @p_1322 (! (not @p_1318) :named @p_1336)) :rule or_neg) +(step t167 (cl @p_1322) :rule th_resolution :premises (t162 t161 t165 t166)) +(step t168 (cl @p_1159 @p_1323) :rule or :premises (t155)) +(step t169 (cl @p_1323) :rule resolution :premises (t168 t125)) +(step t170 (cl @p_1185) :rule resolution :premises (t150 t169)) +(step t171 (cl @p_1168) :rule resolution :premises (t132 t170)) +(step t172 (cl @p_1324 (! (not @p_842) :named @p_1676) (not @p_1325)) :rule and_neg) +(step t173 (cl @p_1159 @p_1326) :rule or :premises (t156)) +(step t174 (cl @p_1326) :rule resolution :premises (t173 t125)) +(step t175 (cl @p_1327 (! (not @p_1328) :named @p_1669) (not @p_1329)) :rule and_neg) +(step t176 (cl @p_1159 @p_1330) :rule or :premises (t157)) +(step t177 (cl @p_1330) :rule resolution :premises (t176 t125)) +(step t178 (cl @p_1331 (not @p_1332) (not @p_1333)) :rule and_neg) +(step t179 (cl @p_1159 @p_1334) :rule or :premises (t158)) +(step t180 (cl @p_1334) :rule resolution :premises (t179 t125)) +(step t181 (cl @p_1287 (not @p_1198)) :rule or_neg) +(step t182 (cl @p_1296 @p_1292 @p_1198) :rule equiv_neg2) +(step t183 (cl @p_1306 (not @p_1296) (not @p_1297)) :rule and_neg) +(step t184 (cl @p_1287 (not @p_1306)) :rule or_neg) +(step t185 (cl (! (not @p_1208) :named @p_1337) (! (not @p_726) :named @p_1338) @p_720) :rule equiv_pos2) +(step t186 (cl (! (not @p_1316) :named @p_1335) @p_1208) :rule and_pos) +(step t187 (cl @p_1335 @p_1309) :rule and_pos) +(step t188 (cl (! (not @p_1317) :named @p_1339) @p_720 @p_1316) :rule or_pos) +(step t189 (cl @p_1336 (not @p_1287) @p_1317) :rule implies_pos) +(step t190 (cl @p_1320 @p_1318) :rule or :premises (t167)) +(step t191 (cl @p_1337 @p_1338) :rule resolution :premises (t185 t122)) +(step t192 (cl @p_1339 @p_1316) :rule resolution :premises (t188 t122)) +(step t193 (cl @p_1318) :rule resolution :premises (t190 t118)) +(step t194 (cl @p_1187 @p_1340) :rule or :premises (t152)) +(step t195 (cl (or (! (not @p_1340) :named @p_1341) (! (or @p_1198 (! (not @p_1292) :named @p_1346) @p_1198) :named @p_1342))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk13) (:= veriT_vr104 veriT_sk14))) +(step t196 (cl @p_1341 @p_1342) :rule or :premises (t195)) +(step t197 (cl (! (or @p_1187 @p_1342) :named @p_1344) (! (not @p_1187) :named @p_1343)) :rule or_neg) +(step t198 (cl (not @p_1343) @p_1155) :rule not_not) +(step t199 (cl @p_1344 @p_1155) :rule th_resolution :premises (t198 t197)) +(step t200 (cl @p_1344 (! (not @p_1342) :named @p_1345)) :rule or_neg) +(step t201 (cl @p_1344) :rule th_resolution :premises (t194 t196 t199 t200)) +(anchor :step t202) +(assume t202.h1 @p_1342) +(step t202.t2 (cl (! (= @p_1342 (! (or @p_1198 @p_1346) :named @p_1347)) :named @p_1348)) :rule ac_simp) +(step t202.t3 (cl (not @p_1348) @p_1345 @p_1347) :rule equiv_pos2) +(step t202.t4 (cl @p_1347) :rule th_resolution :premises (t202.h1 t202.t2 t202.t3)) +(step t202 (cl @p_1345 @p_1347) :rule subproof :discharge (h1)) +(step t203 (cl @p_1187 @p_1342) :rule or :premises (t201)) +(step t204 (cl (! (or @p_1187 @p_1347) :named @p_1349) @p_1343) :rule or_neg) +(step t205 (cl @p_1349 @p_1155) :rule th_resolution :premises (t198 t204)) +(step t206 (cl @p_1349 (! (not @p_1347) :named @p_1350)) :rule or_neg) +(step t207 (cl @p_1349) :rule th_resolution :premises (t203 t202 t205 t206)) +(step t208 (cl @p_1350 @p_1198 @p_1346) :rule or_pos) +(step t209 (cl @p_1187 @p_1347) :rule or :premises (t207)) +(step t210 (cl @p_1347) :rule resolution :premises (t209 t120)) +(step t211 (cl @p_1187 @p_1351) :rule or :premises (t153)) +(step t212 (cl @p_1187 @p_1352) :rule or :premises (t154)) +(step t213 (cl (or @p_1163 (! (not (! (and @p_1328 (! (is_res$b @p_1353 @p_747) :named @p_1670)) :named @p_1668)) :named @p_1671))) :rule forall_inst :args ((:= veriT_vr111 veriT_sk9) (:= veriT_vr112 veriT_sk10))) +(step t214 (cl (or (! (not @p_1309) :named @p_1354) (! (= @p_822 @p_1328) :named @p_1672))) :rule forall_inst :args ((:= veriT_vr128 veriT_sk9) (:= veriT_vr129 veriT_sk10))) +(step t215 (cl (or @p_1354 (! (= @p_784 @p_1332) :named @p_1673))) :rule forall_inst :args ((:= veriT_vr128 veriT_sk3) (:= veriT_vr129 veriT_sk4))) +(step t216 (cl (or @p_1354 (! (= (! (is_res$a @p_714 @p_841) :named @p_1675) @p_842) :named @p_1674))) :rule forall_inst :args ((:= veriT_vr128 veriT_sk11) (:= veriT_vr129 veriT_sk12))) +(step t217 (cl (or @p_1187 (! (or @p_1198 (! (and @p_1296 (! (forall ((veriT_vr105 B$) (veriT_vr106 C$)) (! (= (! (is_res$ @p_1199 @p_1087) :named @p_1359) (! (is_res$ @p_1206 @p_1087) :named @p_1364)) :named @p_1366)) :named @p_1357)) :named @p_1368)) :named @p_1355))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk13) (:= veriT_vr104 veriT_sk14))) +(anchor :step t218) +(assume t218.h1 @p_1355) +(anchor :step t218.t2 :args ((:= (veriT_vr105 B$) veriT_vr148) (:= (veriT_vr106 C$) veriT_vr149))) +(step t218.t2.t1 (cl (! (= veriT_vr105 veriT_vr148) :named @p_1361)) :rule refl) +(step t218.t2.t2 (cl (! (= veriT_vr106 veriT_vr149) :named @p_1362)) :rule refl) +(step t218.t2.t3 (cl (! (= @p_1087 (! (pair$ veriT_vr148 veriT_vr149) :named @p_1358)) :named @p_1363)) :rule cong :premises (t218.t2.t1 t218.t2.t2)) +(step t218.t2.t4 (cl (= @p_1359 (! (is_res$ @p_1199 @p_1358) :named @p_1360))) :rule cong :premises (t218.t2.t3)) +(step t218.t2.t5 (cl @p_1361) :rule refl) +(step t218.t2.t6 (cl @p_1362) :rule refl) +(step t218.t2.t7 (cl @p_1363) :rule cong :premises (t218.t2.t5 t218.t2.t6)) +(step t218.t2.t8 (cl (= @p_1364 (! (is_res$ @p_1206 @p_1358) :named @p_1365))) :rule cong :premises (t218.t2.t7)) +(step t218.t2.t9 (cl (= @p_1366 (! (= @p_1360 @p_1365) :named @p_1367))) :rule cong :premises (t218.t2.t4 t218.t2.t8)) +(step t218.t2 (cl (= @p_1357 (! (forall ((veriT_vr148 B$) (veriT_vr149 C$)) @p_1367) :named @p_1369))) :rule bind) +(step t218.t3 (cl (= @p_1368 (! (and @p_1296 @p_1369) :named @p_1370))) :rule cong :premises (t218.t2)) +(step t218.t4 (cl (! (= @p_1355 (! (or @p_1198 @p_1370) :named @p_1373)) :named @p_1371)) :rule cong :premises (t218.t3)) +(step t218.t5 (cl (not @p_1371) (! (not @p_1355) :named @p_1372) @p_1373) :rule equiv_pos2) +(step t218.t6 (cl @p_1373) :rule th_resolution :premises (t218.h1 t218.t4 t218.t5)) +(anchor :step t218.t7 :args ((:= (veriT_vr148 B$) veriT_vr150) (:= (veriT_vr149 C$) veriT_vr151))) +(step t218.t7.t1 (cl (! (= veriT_vr148 veriT_vr150) :named @p_1376)) :rule refl) +(step t218.t7.t2 (cl (! (= veriT_vr149 veriT_vr151) :named @p_1377)) :rule refl) +(step t218.t7.t3 (cl (! (= @p_1358 (! (pair$ veriT_vr150 veriT_vr151) :named @p_1356)) :named @p_1378)) :rule cong :premises (t218.t7.t1 t218.t7.t2)) +(step t218.t7.t4 (cl (= @p_1360 (! (is_res$ @p_1199 @p_1356) :named @p_1375))) :rule cong :premises (t218.t7.t3)) +(step t218.t7.t5 (cl @p_1376) :rule refl) +(step t218.t7.t6 (cl @p_1377) :rule refl) +(step t218.t7.t7 (cl @p_1378) :rule cong :premises (t218.t7.t5 t218.t7.t6)) +(step t218.t7.t8 (cl (= @p_1365 (! (is_res$ @p_1206 @p_1356) :named @p_1379))) :rule cong :premises (t218.t7.t7)) +(step t218.t7.t9 (cl (= @p_1367 (! (= @p_1375 @p_1379) :named @p_1380))) :rule cong :premises (t218.t7.t4 t218.t7.t8)) +(step t218.t7 (cl (= @p_1369 (! (forall ((veriT_vr150 B$) (veriT_vr151 C$)) @p_1380) :named @p_1374))) :rule bind) +(step t218.t8 (cl (= @p_1370 (! (and @p_1296 @p_1374) :named @p_1381))) :rule cong :premises (t218.t7)) +(step t218.t9 (cl (! (= @p_1373 (! (or @p_1198 @p_1381) :named @p_1382)) :named @p_1383)) :rule cong :premises (t218.t8)) +(step t218.t10 (cl (not @p_1383) (not @p_1373) @p_1382) :rule equiv_pos2) +(step t218.t11 (cl @p_1382) :rule th_resolution :premises (t218.t6 t218.t9 t218.t10)) +(step t218 (cl @p_1372 @p_1382) :rule subproof :discharge (h1)) +(step t219 (cl @p_1187 @p_1355) :rule or :premises (t217)) +(step t220 (cl (! (or @p_1187 @p_1382) :named @p_1384) @p_1343) :rule or_neg) +(step t221 (cl @p_1384 @p_1155) :rule th_resolution :premises (t198 t220)) +(step t222 (cl @p_1384 (! (not @p_1382) :named @p_1677)) :rule or_neg) +(step t223 (cl @p_1384) :rule th_resolution :premises (t219 t218 t221 t222)) +(step t224 (cl (or @p_1385 (! (=> (! (forall ((veriT_vr97 A$) (veriT_vr98 C$)) (! (or (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk0 veriT_vr97) :named @p_1417) veriT_vr98) :named @p_1388)) :named @p_1386) (! (and (! (= (! (is_fail$ (! (run$ (! (fun_app$ veriT_sk1 veriT_vr97) :named @p_1420) veriT_vr98) :named @p_1387)) :named @p_1423) @p_1386) :named @p_1428) (! (forall ((veriT_vr99 B$) (veriT_vr100 C$)) (! (= (! (is_res$ @p_1387 @p_1035) :named @p_1434) (! (is_res$ @p_1388 @p_1035) :named @p_1439)) :named @p_1441)) :named @p_1430)) :named @p_1443)) :named @p_1445)) :named @p_1412) (! (or @p_1333 (! (and (! (= @p_789 @p_1333) :named @p_1410) (! (forall ((veriT_vr93 F$) (veriT_vr102 D$)) (! (= (! (is_res$b @p_788 @p_1039) :named @p_1449) (! (is_res$b @p_1389 @p_1039) :named @p_1454)) :named @p_1456)) :named @p_1447)) :named @p_1458)) :named @p_1460)) :named @p_1390))) :rule forall_inst :args ((:= veriT_vr94 veriT_sk3) (:= veriT_vr95 veriT_sk0) (:= veriT_vr96 veriT_sk1) (:= veriT_vr101 veriT_sk4))) +(anchor :step t225) +(assume t225.h1 @p_1390) +(anchor :step t225.t2 :args ((:= (veriT_vr97 A$) veriT_vr165) (:= (veriT_vr98 C$) veriT_vr166))) +(step t225.t2.t1 (cl (! (= veriT_vr97 veriT_vr165) :named @p_1419)) :rule refl) +(step t225.t2.t2 (cl (! (= @p_1417 (! (fun_app$ veriT_sk0 veriT_vr165) :named @p_1418)) :named @p_1425)) :rule cong :premises (t225.t2.t1)) +(step t225.t2.t3 (cl (! (= veriT_vr98 veriT_vr166) :named @p_1422)) :rule refl) +(step t225.t2.t4 (cl (! (= @p_1388 (! (run$ @p_1418 veriT_vr166) :named @p_1415)) :named @p_1426)) :rule cong :premises (t225.t2.t2 t225.t2.t3)) +(step t225.t2.t5 (cl (! (= @p_1386 (! (is_fail$ @p_1415) :named @p_1413)) :named @p_1427)) :rule cong :premises (t225.t2.t4)) +(step t225.t2.t6 (cl @p_1419) :rule refl) +(step t225.t2.t7 (cl (! (= @p_1420 (! (fun_app$ veriT_sk1 veriT_vr165) :named @p_1421)) :named @p_1432)) :rule cong :premises (t225.t2.t6)) +(step t225.t2.t8 (cl @p_1422) :rule refl) +(step t225.t2.t9 (cl (! (= @p_1387 (! (run$ @p_1421 veriT_vr166) :named @p_1414)) :named @p_1433)) :rule cong :premises (t225.t2.t7 t225.t2.t8)) +(step t225.t2.t10 (cl (= @p_1423 (! (is_fail$ @p_1414) :named @p_1424))) :rule cong :premises (t225.t2.t9)) +(step t225.t2.t11 (cl @p_1419) :rule refl) +(step t225.t2.t12 (cl @p_1425) :rule cong :premises (t225.t2.t11)) +(step t225.t2.t13 (cl @p_1422) :rule refl) +(step t225.t2.t14 (cl @p_1426) :rule cong :premises (t225.t2.t12 t225.t2.t13)) +(step t225.t2.t15 (cl @p_1427) :rule cong :premises (t225.t2.t14)) +(step t225.t2.t16 (cl (= @p_1428 (! (= @p_1424 @p_1413) :named @p_1429))) :rule cong :premises (t225.t2.t10 t225.t2.t15)) +(anchor :step t225.t2.t17 :args ((:= (veriT_vr99 B$) veriT_vr167) (:= (veriT_vr100 C$) veriT_vr168))) +(step t225.t2.t17.t1 (cl @p_1419) :rule refl) +(step t225.t2.t17.t2 (cl @p_1432) :rule cong :premises (t225.t2.t17.t1)) +(step t225.t2.t17.t3 (cl @p_1422) :rule refl) +(step t225.t2.t17.t4 (cl @p_1433) :rule cong :premises (t225.t2.t17.t2 t225.t2.t17.t3)) +(step t225.t2.t17.t5 (cl (! (= veriT_vr99 veriT_vr167) :named @p_1436)) :rule refl) +(step t225.t2.t17.t6 (cl (! (= veriT_vr100 veriT_vr168) :named @p_1437)) :rule refl) +(step t225.t2.t17.t7 (cl (! (= @p_1035 (! (pair$ veriT_vr167 veriT_vr168) :named @p_1416)) :named @p_1438)) :rule cong :premises (t225.t2.t17.t5 t225.t2.t17.t6)) +(step t225.t2.t17.t8 (cl (= @p_1434 (! (is_res$ @p_1414 @p_1416) :named @p_1435))) :rule cong :premises (t225.t2.t17.t4 t225.t2.t17.t7)) +(step t225.t2.t17.t9 (cl @p_1419) :rule refl) +(step t225.t2.t17.t10 (cl @p_1425) :rule cong :premises (t225.t2.t17.t9)) +(step t225.t2.t17.t11 (cl @p_1422) :rule refl) +(step t225.t2.t17.t12 (cl @p_1426) :rule cong :premises (t225.t2.t17.t10 t225.t2.t17.t11)) +(step t225.t2.t17.t13 (cl @p_1436) :rule refl) +(step t225.t2.t17.t14 (cl @p_1437) :rule refl) +(step t225.t2.t17.t15 (cl @p_1438) :rule cong :premises (t225.t2.t17.t13 t225.t2.t17.t14)) +(step t225.t2.t17.t16 (cl (= @p_1439 (! (is_res$ @p_1415 @p_1416) :named @p_1440))) :rule cong :premises (t225.t2.t17.t12 t225.t2.t17.t15)) +(step t225.t2.t17.t17 (cl (= @p_1441 (! (= @p_1435 @p_1440) :named @p_1442))) :rule cong :premises (t225.t2.t17.t8 t225.t2.t17.t16)) +(step t225.t2.t17 (cl (= @p_1430 (! (forall ((veriT_vr167 B$) (veriT_vr168 C$)) @p_1442) :named @p_1431))) :rule bind) +(step t225.t2.t18 (cl (= @p_1443 (! (and @p_1429 @p_1431) :named @p_1444))) :rule cong :premises (t225.t2.t16 t225.t2.t17)) +(step t225.t2.t19 (cl (= @p_1445 (! (or @p_1413 @p_1444) :named @p_1446))) :rule cong :premises (t225.t2.t5 t225.t2.t18)) +(step t225.t2 (cl (= @p_1412 (! (forall ((veriT_vr165 A$) (veriT_vr166 C$)) @p_1446) :named @p_1462))) :rule bind) +(anchor :step t225.t3 :args ((:= (veriT_vr93 F$) veriT_vr169) (:= (veriT_vr102 D$) veriT_vr170))) +(step t225.t3.t1 (cl (! (= veriT_vr93 veriT_vr169) :named @p_1451)) :rule refl) +(step t225.t3.t2 (cl (! (= veriT_vr102 veriT_vr170) :named @p_1452)) :rule refl) +(step t225.t3.t3 (cl (! (= @p_1039 (! (pair$b veriT_vr169 veriT_vr170) :named @p_1448)) :named @p_1453)) :rule cong :premises (t225.t3.t1 t225.t3.t2)) +(step t225.t3.t4 (cl (= @p_1449 (! (is_res$b @p_788 @p_1448) :named @p_1450))) :rule cong :premises (t225.t3.t3)) +(step t225.t3.t5 (cl @p_1451) :rule refl) +(step t225.t3.t6 (cl @p_1452) :rule refl) +(step t225.t3.t7 (cl @p_1453) :rule cong :premises (t225.t3.t5 t225.t3.t6)) +(step t225.t3.t8 (cl (= @p_1454 (! (is_res$b @p_1389 @p_1448) :named @p_1455))) :rule cong :premises (t225.t3.t7)) +(step t225.t3.t9 (cl (= @p_1456 (! (= @p_1450 @p_1455) :named @p_1457))) :rule cong :premises (t225.t3.t4 t225.t3.t8)) +(step t225.t3 (cl (= @p_1447 (! (forall ((veriT_vr169 F$) (veriT_vr170 D$)) @p_1457) :named @p_1459))) :rule bind) +(step t225.t4 (cl (= @p_1458 (! (and @p_1410 @p_1459) :named @p_1461))) :rule cong :premises (t225.t3)) +(step t225.t5 (cl (= @p_1460 (! (or @p_1333 @p_1461) :named @p_1463))) :rule cong :premises (t225.t4)) +(step t225.t6 (cl (! (= @p_1390 (! (=> @p_1462 @p_1463) :named @p_1466)) :named @p_1464)) :rule cong :premises (t225.t2 t225.t5)) +(step t225.t7 (cl (not @p_1464) (! (not @p_1390) :named @p_1465) @p_1466) :rule equiv_pos2) +(step t225.t8 (cl @p_1466) :rule th_resolution :premises (t225.h1 t225.t6 t225.t7)) +(anchor :step t225.t9 :args ((:= (veriT_vr165 A$) veriT_vr171) (:= (veriT_vr166 C$) veriT_vr172))) +(step t225.t9.t1 (cl (! (= veriT_vr165 veriT_vr171) :named @p_1469)) :rule refl) +(step t225.t9.t2 (cl (! (= @p_1418 @p_1468) :named @p_1473)) :rule cong :premises (t225.t9.t1)) +(step t225.t9.t3 (cl (! (= veriT_vr166 veriT_vr172) :named @p_1471)) :rule refl) +(step t225.t9.t4 (cl (! (= @p_1415 @p_1393) :named @p_1474)) :rule cong :premises (t225.t9.t2 t225.t9.t3)) +(step t225.t9.t5 (cl (! (= @p_1413 @p_1391) :named @p_1475)) :rule cong :premises (t225.t9.t4)) +(step t225.t9.t6 (cl @p_1469) :rule refl) +(step t225.t9.t7 (cl (! (= @p_1421 @p_1470) :named @p_1478)) :rule cong :premises (t225.t9.t6)) +(step t225.t9.t8 (cl @p_1471) :rule refl) +(step t225.t9.t9 (cl (! (= @p_1414 @p_1392) :named @p_1479)) :rule cong :premises (t225.t9.t7 t225.t9.t8)) +(step t225.t9.t10 (cl (= @p_1424 @p_1472)) :rule cong :premises (t225.t9.t9)) +(step t225.t9.t11 (cl @p_1469) :rule refl) +(step t225.t9.t12 (cl @p_1473) :rule cong :premises (t225.t9.t11)) +(step t225.t9.t13 (cl @p_1471) :rule refl) +(step t225.t9.t14 (cl @p_1474) :rule cong :premises (t225.t9.t12 t225.t9.t13)) +(step t225.t9.t15 (cl @p_1475) :rule cong :premises (t225.t9.t14)) +(step t225.t9.t16 (cl (= @p_1429 @p_1476)) :rule cong :premises (t225.t9.t10 t225.t9.t15)) +(anchor :step t225.t9.t17 :args ((:= (veriT_vr167 B$) veriT_vr173) (:= (veriT_vr168 C$) veriT_vr174))) +(step t225.t9.t17.t1 (cl @p_1469) :rule refl) +(step t225.t9.t17.t2 (cl @p_1478) :rule cong :premises (t225.t9.t17.t1)) +(step t225.t9.t17.t3 (cl @p_1471) :rule refl) +(step t225.t9.t17.t4 (cl @p_1479) :rule cong :premises (t225.t9.t17.t2 t225.t9.t17.t3)) +(step t225.t9.t17.t5 (cl (! (= veriT_vr167 veriT_vr173) :named @p_1481)) :rule refl) +(step t225.t9.t17.t6 (cl (! (= veriT_vr168 veriT_vr174) :named @p_1482)) :rule refl) +(step t225.t9.t17.t7 (cl (! (= @p_1416 @p_1394) :named @p_1483)) :rule cong :premises (t225.t9.t17.t5 t225.t9.t17.t6)) +(step t225.t9.t17.t8 (cl (= @p_1435 @p_1480)) :rule cong :premises (t225.t9.t17.t4 t225.t9.t17.t7)) +(step t225.t9.t17.t9 (cl @p_1469) :rule refl) +(step t225.t9.t17.t10 (cl @p_1473) :rule cong :premises (t225.t9.t17.t9)) +(step t225.t9.t17.t11 (cl @p_1471) :rule refl) +(step t225.t9.t17.t12 (cl @p_1474) :rule cong :premises (t225.t9.t17.t10 t225.t9.t17.t11)) +(step t225.t9.t17.t13 (cl @p_1481) :rule refl) +(step t225.t9.t17.t14 (cl @p_1482) :rule refl) +(step t225.t9.t17.t15 (cl @p_1483) :rule cong :premises (t225.t9.t17.t13 t225.t9.t17.t14)) +(step t225.t9.t17.t16 (cl (= @p_1440 @p_1484)) :rule cong :premises (t225.t9.t17.t12 t225.t9.t17.t15)) +(step t225.t9.t17.t17 (cl (= @p_1442 @p_1485)) :rule cong :premises (t225.t9.t17.t8 t225.t9.t17.t16)) +(step t225.t9.t17 (cl (= @p_1431 @p_1477)) :rule bind) +(step t225.t9.t18 (cl (= @p_1444 @p_1486)) :rule cong :premises (t225.t9.t16 t225.t9.t17)) +(step t225.t9.t19 (cl (= @p_1446 @p_1467)) :rule cong :premises (t225.t9.t5 t225.t9.t18)) +(step t225.t9 (cl (= @p_1462 (! (forall ((veriT_vr171 A$) (veriT_vr172 C$)) @p_1467) :named @p_1496))) :rule bind) +(anchor :step t225.t10 :args ((:= (veriT_vr169 F$) veriT_vr175) (:= (veriT_vr170 D$) veriT_vr176))) +(step t225.t10.t1 (cl (! (= veriT_vr169 veriT_vr175) :named @p_1489)) :rule refl) +(step t225.t10.t2 (cl (! (= veriT_vr170 veriT_vr176) :named @p_1490)) :rule refl) +(step t225.t10.t3 (cl (! (= @p_1448 (! (pair$b veriT_vr175 veriT_vr176) :named @p_1487)) :named @p_1491)) :rule cong :premises (t225.t10.t1 t225.t10.t2)) +(step t225.t10.t4 (cl (= @p_1450 (! (is_res$b @p_788 @p_1487) :named @p_1488))) :rule cong :premises (t225.t10.t3)) +(step t225.t10.t5 (cl @p_1489) :rule refl) +(step t225.t10.t6 (cl @p_1490) :rule refl) +(step t225.t10.t7 (cl @p_1491) :rule cong :premises (t225.t10.t5 t225.t10.t6)) +(step t225.t10.t8 (cl (= @p_1455 (! (is_res$b @p_1389 @p_1487) :named @p_1492))) :rule cong :premises (t225.t10.t7)) +(step t225.t10.t9 (cl (= @p_1457 (! (= @p_1488 @p_1492) :named @p_1493))) :rule cong :premises (t225.t10.t4 t225.t10.t8)) +(step t225.t10 (cl (= @p_1459 (! (forall ((veriT_vr175 F$) (veriT_vr176 D$)) @p_1493) :named @p_1494))) :rule bind) +(step t225.t11 (cl (= @p_1461 (! (and @p_1410 @p_1494) :named @p_1495))) :rule cong :premises (t225.t10)) +(step t225.t12 (cl (= @p_1463 (! (or @p_1333 @p_1495) :named @p_1497))) :rule cong :premises (t225.t11)) +(step t225.t13 (cl (! (= @p_1466 (! (=> @p_1496 @p_1497) :named @p_1499)) :named @p_1498)) :rule cong :premises (t225.t9 t225.t12)) +(step t225.t14 (cl (not @p_1498) (not @p_1466) @p_1499) :rule equiv_pos2) +(step t225.t15 (cl @p_1499) :rule th_resolution :premises (t225.t8 t225.t13 t225.t14)) +(anchor :step t225.t16 :args ((:= (veriT_vr171 A$) veriT_sk25) (:= (veriT_vr172 C$) veriT_sk26))) +(step t225.t16.t1 (cl (! (= veriT_vr171 veriT_sk25) :named @p_1502)) :rule refl) +(step t225.t16.t2 (cl (! (= @p_1468 (! (fun_app$ veriT_sk0 veriT_sk25) :named @p_1501)) :named @p_1506)) :rule cong :premises (t225.t16.t1)) +(step t225.t16.t3 (cl (! (= veriT_vr172 veriT_sk26) :named @p_1504)) :rule refl) +(step t225.t16.t4 (cl (! (= @p_1393 (! (run$ @p_1501 veriT_sk26) :named @p_1408)) :named @p_1507)) :rule cong :premises (t225.t16.t2 t225.t16.t3)) +(step t225.t16.t5 (cl (! (= @p_1391 (! (is_fail$ @p_1408) :named @p_1400)) :named @p_1508)) :rule cong :premises (t225.t16.t4)) +(step t225.t16.t6 (cl @p_1502) :rule refl) +(step t225.t16.t7 (cl (! (= @p_1470 (! (fun_app$ veriT_sk1 veriT_sk25) :named @p_1503)) :named @p_1511)) :rule cong :premises (t225.t16.t6)) +(step t225.t16.t8 (cl @p_1504) :rule refl) +(step t225.t16.t9 (cl (! (= @p_1392 (! (run$ @p_1503 veriT_sk26) :named @p_1401)) :named @p_1512)) :rule cong :premises (t225.t16.t7 t225.t16.t8)) +(step t225.t16.t10 (cl (= @p_1472 (! (is_fail$ @p_1401) :named @p_1505))) :rule cong :premises (t225.t16.t9)) +(step t225.t16.t11 (cl @p_1502) :rule refl) +(step t225.t16.t12 (cl @p_1506) :rule cong :premises (t225.t16.t11)) +(step t225.t16.t13 (cl @p_1504) :rule refl) +(step t225.t16.t14 (cl @p_1507) :rule cong :premises (t225.t16.t12 t225.t16.t13)) +(step t225.t16.t15 (cl @p_1508) :rule cong :premises (t225.t16.t14)) +(step t225.t16.t16 (cl (= @p_1476 (! (= @p_1505 @p_1400) :named @p_1509))) :rule cong :premises (t225.t16.t10 t225.t16.t15)) +(anchor :step t225.t16.t17 :args ((:= (veriT_vr173 B$) veriT_sk27) (:= (veriT_vr174 C$) veriT_sk28))) +(step t225.t16.t17.t1 (cl @p_1502) :rule refl) +(step t225.t16.t17.t2 (cl @p_1511) :rule cong :premises (t225.t16.t17.t1)) +(step t225.t16.t17.t3 (cl @p_1504) :rule refl) +(step t225.t16.t17.t4 (cl @p_1512) :rule cong :premises (t225.t16.t17.t2 t225.t16.t17.t3)) +(step t225.t16.t17.t5 (cl (! (= veriT_vr173 veriT_sk27) :named @p_1515)) :rule refl) +(step t225.t16.t17.t6 (cl (! (= veriT_vr174 veriT_sk28) :named @p_1516)) :rule refl) +(step t225.t16.t17.t7 (cl (! (= @p_1394 (! (pair$ veriT_sk27 veriT_sk28) :named @p_1409)) :named @p_1517)) :rule cong :premises (t225.t16.t17.t5 t225.t16.t17.t6)) +(step t225.t16.t17.t8 (cl (= @p_1480 (! (is_res$ @p_1401 @p_1409) :named @p_1514))) :rule cong :premises (t225.t16.t17.t4 t225.t16.t17.t7)) +(step t225.t16.t17.t9 (cl @p_1502) :rule refl) +(step t225.t16.t17.t10 (cl @p_1506) :rule cong :premises (t225.t16.t17.t9)) +(step t225.t16.t17.t11 (cl @p_1504) :rule refl) +(step t225.t16.t17.t12 (cl @p_1507) :rule cong :premises (t225.t16.t17.t10 t225.t16.t17.t11)) +(step t225.t16.t17.t13 (cl @p_1515) :rule refl) +(step t225.t16.t17.t14 (cl @p_1516) :rule refl) +(step t225.t16.t17.t15 (cl @p_1517) :rule cong :premises (t225.t16.t17.t13 t225.t16.t17.t14)) +(step t225.t16.t17.t16 (cl (= @p_1484 (! (is_res$ @p_1408 @p_1409) :named @p_1518))) :rule cong :premises (t225.t16.t17.t12 t225.t16.t17.t15)) +(step t225.t16.t17.t17 (cl (= @p_1485 (! (= @p_1514 @p_1518) :named @p_1510))) :rule cong :premises (t225.t16.t17.t8 t225.t16.t17.t16)) +(step t225.t16.t17 (cl (= @p_1477 @p_1510)) :rule sko_forall) +(step t225.t16.t18 (cl (= @p_1486 (! (and @p_1509 @p_1510) :named @p_1519))) :rule cong :premises (t225.t16.t16 t225.t16.t17)) +(step t225.t16.t19 (cl (= @p_1467 (! (or @p_1400 @p_1519) :named @p_1500))) :rule cong :premises (t225.t16.t5 t225.t16.t18)) +(step t225.t16 (cl (= @p_1496 @p_1500)) :rule sko_forall) +(step t225.t17 (cl (! (= @p_1499 (! (=> @p_1500 @p_1497) :named @p_1521)) :named @p_1520)) :rule cong :premises (t225.t16)) +(step t225.t18 (cl (not @p_1520) (not @p_1499) @p_1521) :rule equiv_pos2) +(step t225.t19 (cl @p_1521) :rule th_resolution :premises (t225.t15 t225.t17 t225.t18)) +(anchor :step t225.t20 :args ((:= (veriT_vr175 F$) veriT_vr177) (:= (veriT_vr176 D$) veriT_vr178))) +(step t225.t20.t1 (cl (! (= veriT_vr175 veriT_vr177) :named @p_1524)) :rule refl) +(step t225.t20.t2 (cl (! (= veriT_vr176 veriT_vr178) :named @p_1525)) :rule refl) +(step t225.t20.t3 (cl (! (= @p_1487 (! (pair$b veriT_vr177 veriT_vr178) :named @p_1411)) :named @p_1526)) :rule cong :premises (t225.t20.t1 t225.t20.t2)) +(step t225.t20.t4 (cl (= @p_1488 (! (is_res$b @p_788 @p_1411) :named @p_1523))) :rule cong :premises (t225.t20.t3)) +(step t225.t20.t5 (cl @p_1524) :rule refl) +(step t225.t20.t6 (cl @p_1525) :rule refl) +(step t225.t20.t7 (cl @p_1526) :rule cong :premises (t225.t20.t5 t225.t20.t6)) +(step t225.t20.t8 (cl (= @p_1492 (! (is_res$b @p_1389 @p_1411) :named @p_1527))) :rule cong :premises (t225.t20.t7)) +(step t225.t20.t9 (cl (= @p_1493 (! (= @p_1523 @p_1527) :named @p_1528))) :rule cong :premises (t225.t20.t4 t225.t20.t8)) +(step t225.t20 (cl (= @p_1494 (! (forall ((veriT_vr177 F$) (veriT_vr178 D$)) @p_1528) :named @p_1522))) :rule bind) +(step t225.t21 (cl (= @p_1495 (! (and @p_1410 @p_1522) :named @p_1529))) :rule cong :premises (t225.t20)) +(step t225.t22 (cl (= @p_1497 (! (or @p_1333 @p_1529) :named @p_1530))) :rule cong :premises (t225.t21)) +(step t225.t23 (cl (! (= @p_1521 (! (=> @p_1500 @p_1530) :named @p_1531)) :named @p_1532)) :rule cong :premises (t225.t22)) +(step t225.t24 (cl (not @p_1532) (not @p_1521) @p_1531) :rule equiv_pos2) +(step t225.t25 (cl @p_1531) :rule th_resolution :premises (t225.t19 t225.t23 t225.t24)) +(step t225 (cl @p_1465 @p_1531) :rule subproof :discharge (h1)) +(step t226 (cl @p_1385 @p_1390) :rule or :premises (t224)) +(step t227 (cl (! (or @p_1385 @p_1531) :named @p_1534) @p_1533) :rule or_neg) +(step t228 (cl @p_1534 @p_1154) :rule th_resolution :premises (t159 t227)) +(step t229 (cl @p_1534 (! (not @p_1531) :named @p_1678)) :rule or_neg) +(step t230 (cl @p_1534) :rule th_resolution :premises (t226 t225 t228 t229)) +(step t231 (cl (or @p_1385 (! (=> @p_1412 (! (or @p_1329 (! (and (! (= (is_fail$b @p_826) @p_1329) :named @p_1555) (! (forall ((veriT_vr93 F$) (veriT_vr102 D$)) (! (= (! (is_res$b @p_826 @p_1039) :named @p_1583) (! (is_res$b @p_1353 @p_1039) :named @p_1588)) :named @p_1590)) :named @p_1581)) :named @p_1592)) :named @p_1594)) :named @p_1535))) :rule forall_inst :args ((:= veriT_vr94 veriT_sk9) (:= veriT_vr95 veriT_sk0) (:= veriT_vr96 veriT_sk1) (:= veriT_vr101 veriT_sk10))) +(anchor :step t232) +(assume t232.h1 @p_1535) +(anchor :step t232.t2 :args ((:= (veriT_vr97 A$) veriT_vr179) (:= (veriT_vr98 C$) veriT_vr180))) +(step t232.t2.t1 (cl (! (= veriT_vr97 veriT_vr179) :named @p_1562)) :rule refl) +(step t232.t2.t2 (cl (! (= @p_1417 (! (fun_app$ veriT_sk0 veriT_vr179) :named @p_1561)) :named @p_1566)) :rule cong :premises (t232.t2.t1)) +(step t232.t2.t3 (cl (! (= veriT_vr98 veriT_vr180) :named @p_1564)) :rule refl) +(step t232.t2.t4 (cl (! (= @p_1388 (! (run$ @p_1561 veriT_vr180) :named @p_1559)) :named @p_1567)) :rule cong :premises (t232.t2.t2 t232.t2.t3)) +(step t232.t2.t5 (cl (! (= @p_1386 (! (is_fail$ @p_1559) :named @p_1557)) :named @p_1568)) :rule cong :premises (t232.t2.t4)) +(step t232.t2.t6 (cl @p_1562) :rule refl) +(step t232.t2.t7 (cl (! (= @p_1420 (! (fun_app$ veriT_sk1 veriT_vr179) :named @p_1563)) :named @p_1571)) :rule cong :premises (t232.t2.t6)) +(step t232.t2.t8 (cl @p_1564) :rule refl) +(step t232.t2.t9 (cl (! (= @p_1387 (! (run$ @p_1563 veriT_vr180) :named @p_1558)) :named @p_1572)) :rule cong :premises (t232.t2.t7 t232.t2.t8)) +(step t232.t2.t10 (cl (= @p_1423 (! (is_fail$ @p_1558) :named @p_1565))) :rule cong :premises (t232.t2.t9)) +(step t232.t2.t11 (cl @p_1562) :rule refl) +(step t232.t2.t12 (cl @p_1566) :rule cong :premises (t232.t2.t11)) +(step t232.t2.t13 (cl @p_1564) :rule refl) +(step t232.t2.t14 (cl @p_1567) :rule cong :premises (t232.t2.t12 t232.t2.t13)) +(step t232.t2.t15 (cl @p_1568) :rule cong :premises (t232.t2.t14)) +(step t232.t2.t16 (cl (= @p_1428 (! (= @p_1565 @p_1557) :named @p_1569))) :rule cong :premises (t232.t2.t10 t232.t2.t15)) +(anchor :step t232.t2.t17 :args ((:= (veriT_vr99 B$) veriT_vr181) (:= (veriT_vr100 C$) veriT_vr182))) +(step t232.t2.t17.t1 (cl @p_1562) :rule refl) +(step t232.t2.t17.t2 (cl @p_1571) :rule cong :premises (t232.t2.t17.t1)) +(step t232.t2.t17.t3 (cl @p_1564) :rule refl) +(step t232.t2.t17.t4 (cl @p_1572) :rule cong :premises (t232.t2.t17.t2 t232.t2.t17.t3)) +(step t232.t2.t17.t5 (cl (! (= veriT_vr99 veriT_vr181) :named @p_1574)) :rule refl) +(step t232.t2.t17.t6 (cl (! (= veriT_vr100 veriT_vr182) :named @p_1575)) :rule refl) +(step t232.t2.t17.t7 (cl (! (= @p_1035 (! (pair$ veriT_vr181 veriT_vr182) :named @p_1560)) :named @p_1576)) :rule cong :premises (t232.t2.t17.t5 t232.t2.t17.t6)) +(step t232.t2.t17.t8 (cl (= @p_1434 (! (is_res$ @p_1558 @p_1560) :named @p_1573))) :rule cong :premises (t232.t2.t17.t4 t232.t2.t17.t7)) +(step t232.t2.t17.t9 (cl @p_1562) :rule refl) +(step t232.t2.t17.t10 (cl @p_1566) :rule cong :premises (t232.t2.t17.t9)) +(step t232.t2.t17.t11 (cl @p_1564) :rule refl) +(step t232.t2.t17.t12 (cl @p_1567) :rule cong :premises (t232.t2.t17.t10 t232.t2.t17.t11)) +(step t232.t2.t17.t13 (cl @p_1574) :rule refl) +(step t232.t2.t17.t14 (cl @p_1575) :rule refl) +(step t232.t2.t17.t15 (cl @p_1576) :rule cong :premises (t232.t2.t17.t13 t232.t2.t17.t14)) +(step t232.t2.t17.t16 (cl (= @p_1439 (! (is_res$ @p_1559 @p_1560) :named @p_1577))) :rule cong :premises (t232.t2.t17.t12 t232.t2.t17.t15)) +(step t232.t2.t17.t17 (cl (= @p_1441 (! (= @p_1573 @p_1577) :named @p_1578))) :rule cong :premises (t232.t2.t17.t8 t232.t2.t17.t16)) +(step t232.t2.t17 (cl (= @p_1430 (! (forall ((veriT_vr181 B$) (veriT_vr182 C$)) @p_1578) :named @p_1570))) :rule bind) +(step t232.t2.t18 (cl (= @p_1443 (! (and @p_1569 @p_1570) :named @p_1579))) :rule cong :premises (t232.t2.t16 t232.t2.t17)) +(step t232.t2.t19 (cl (= @p_1445 (! (or @p_1557 @p_1579) :named @p_1580))) :rule cong :premises (t232.t2.t5 t232.t2.t18)) +(step t232.t2 (cl (= @p_1412 (! (forall ((veriT_vr179 A$) (veriT_vr180 C$)) @p_1580) :named @p_1596))) :rule bind) +(anchor :step t232.t3 :args ((:= (veriT_vr93 F$) veriT_vr183) (:= (veriT_vr102 D$) veriT_vr184))) +(step t232.t3.t1 (cl (! (= veriT_vr93 veriT_vr183) :named @p_1585)) :rule refl) +(step t232.t3.t2 (cl (! (= veriT_vr102 veriT_vr184) :named @p_1586)) :rule refl) +(step t232.t3.t3 (cl (! (= @p_1039 (! (pair$b veriT_vr183 veriT_vr184) :named @p_1582)) :named @p_1587)) :rule cong :premises (t232.t3.t1 t232.t3.t2)) +(step t232.t3.t4 (cl (= @p_1583 (! (is_res$b @p_826 @p_1582) :named @p_1584))) :rule cong :premises (t232.t3.t3)) +(step t232.t3.t5 (cl @p_1585) :rule refl) +(step t232.t3.t6 (cl @p_1586) :rule refl) +(step t232.t3.t7 (cl @p_1587) :rule cong :premises (t232.t3.t5 t232.t3.t6)) +(step t232.t3.t8 (cl (= @p_1588 (! (is_res$b @p_1353 @p_1582) :named @p_1589))) :rule cong :premises (t232.t3.t7)) +(step t232.t3.t9 (cl (= @p_1590 (! (= @p_1584 @p_1589) :named @p_1591))) :rule cong :premises (t232.t3.t4 t232.t3.t8)) +(step t232.t3 (cl (= @p_1581 (! (forall ((veriT_vr183 F$) (veriT_vr184 D$)) @p_1591) :named @p_1593))) :rule bind) +(step t232.t4 (cl (= @p_1592 (! (and @p_1555 @p_1593) :named @p_1595))) :rule cong :premises (t232.t3)) +(step t232.t5 (cl (= @p_1594 (! (or @p_1329 @p_1595) :named @p_1597))) :rule cong :premises (t232.t4)) +(step t232.t6 (cl (! (= @p_1535 (! (=> @p_1596 @p_1597) :named @p_1600)) :named @p_1598)) :rule cong :premises (t232.t2 t232.t5)) +(step t232.t7 (cl (not @p_1598) (! (not @p_1535) :named @p_1599) @p_1600) :rule equiv_pos2) +(step t232.t8 (cl @p_1600) :rule th_resolution :premises (t232.h1 t232.t6 t232.t7)) +(anchor :step t232.t9 :args ((:= (veriT_vr179 A$) veriT_vr185) (:= (veriT_vr180 C$) veriT_vr186))) +(step t232.t9.t1 (cl (! (= veriT_vr179 veriT_vr185) :named @p_1603)) :rule refl) +(step t232.t9.t2 (cl (! (= @p_1561 @p_1602) :named @p_1607)) :rule cong :premises (t232.t9.t1)) +(step t232.t9.t3 (cl (! (= veriT_vr180 veriT_vr186) :named @p_1605)) :rule refl) +(step t232.t9.t4 (cl (! (= @p_1559 @p_1538) :named @p_1608)) :rule cong :premises (t232.t9.t2 t232.t9.t3)) +(step t232.t9.t5 (cl (! (= @p_1557 @p_1536) :named @p_1609)) :rule cong :premises (t232.t9.t4)) +(step t232.t9.t6 (cl @p_1603) :rule refl) +(step t232.t9.t7 (cl (! (= @p_1563 @p_1604) :named @p_1612)) :rule cong :premises (t232.t9.t6)) +(step t232.t9.t8 (cl @p_1605) :rule refl) +(step t232.t9.t9 (cl (! (= @p_1558 @p_1537) :named @p_1613)) :rule cong :premises (t232.t9.t7 t232.t9.t8)) +(step t232.t9.t10 (cl (= @p_1565 @p_1606)) :rule cong :premises (t232.t9.t9)) +(step t232.t9.t11 (cl @p_1603) :rule refl) +(step t232.t9.t12 (cl @p_1607) :rule cong :premises (t232.t9.t11)) +(step t232.t9.t13 (cl @p_1605) :rule refl) +(step t232.t9.t14 (cl @p_1608) :rule cong :premises (t232.t9.t12 t232.t9.t13)) +(step t232.t9.t15 (cl @p_1609) :rule cong :premises (t232.t9.t14)) +(step t232.t9.t16 (cl (= @p_1569 @p_1610)) :rule cong :premises (t232.t9.t10 t232.t9.t15)) +(anchor :step t232.t9.t17 :args ((:= (veriT_vr181 B$) veriT_vr187) (:= (veriT_vr182 C$) veriT_vr188))) +(step t232.t9.t17.t1 (cl @p_1603) :rule refl) +(step t232.t9.t17.t2 (cl @p_1612) :rule cong :premises (t232.t9.t17.t1)) +(step t232.t9.t17.t3 (cl @p_1605) :rule refl) +(step t232.t9.t17.t4 (cl @p_1613) :rule cong :premises (t232.t9.t17.t2 t232.t9.t17.t3)) +(step t232.t9.t17.t5 (cl (! (= veriT_vr181 veriT_vr187) :named @p_1615)) :rule refl) +(step t232.t9.t17.t6 (cl (! (= veriT_vr182 veriT_vr188) :named @p_1616)) :rule refl) +(step t232.t9.t17.t7 (cl (! (= @p_1560 @p_1539) :named @p_1617)) :rule cong :premises (t232.t9.t17.t5 t232.t9.t17.t6)) +(step t232.t9.t17.t8 (cl (= @p_1573 @p_1614)) :rule cong :premises (t232.t9.t17.t4 t232.t9.t17.t7)) +(step t232.t9.t17.t9 (cl @p_1603) :rule refl) +(step t232.t9.t17.t10 (cl @p_1607) :rule cong :premises (t232.t9.t17.t9)) +(step t232.t9.t17.t11 (cl @p_1605) :rule refl) +(step t232.t9.t17.t12 (cl @p_1608) :rule cong :premises (t232.t9.t17.t10 t232.t9.t17.t11)) +(step t232.t9.t17.t13 (cl @p_1615) :rule refl) +(step t232.t9.t17.t14 (cl @p_1616) :rule refl) +(step t232.t9.t17.t15 (cl @p_1617) :rule cong :premises (t232.t9.t17.t13 t232.t9.t17.t14)) +(step t232.t9.t17.t16 (cl (= @p_1577 @p_1618)) :rule cong :premises (t232.t9.t17.t12 t232.t9.t17.t15)) +(step t232.t9.t17.t17 (cl (= @p_1578 @p_1619)) :rule cong :premises (t232.t9.t17.t8 t232.t9.t17.t16)) +(step t232.t9.t17 (cl (= @p_1570 @p_1611)) :rule bind) +(step t232.t9.t18 (cl (= @p_1579 @p_1620)) :rule cong :premises (t232.t9.t16 t232.t9.t17)) +(step t232.t9.t19 (cl (= @p_1580 @p_1601)) :rule cong :premises (t232.t9.t5 t232.t9.t18)) +(step t232.t9 (cl (= @p_1596 (! (forall ((veriT_vr185 A$) (veriT_vr186 C$)) @p_1601) :named @p_1630))) :rule bind) +(anchor :step t232.t10 :args ((:= (veriT_vr183 F$) veriT_vr189) (:= (veriT_vr184 D$) veriT_vr190))) +(step t232.t10.t1 (cl (! (= veriT_vr183 veriT_vr189) :named @p_1623)) :rule refl) +(step t232.t10.t2 (cl (! (= veriT_vr184 veriT_vr190) :named @p_1624)) :rule refl) +(step t232.t10.t3 (cl (! (= @p_1582 (! (pair$b veriT_vr189 veriT_vr190) :named @p_1621)) :named @p_1625)) :rule cong :premises (t232.t10.t1 t232.t10.t2)) +(step t232.t10.t4 (cl (= @p_1584 (! (is_res$b @p_826 @p_1621) :named @p_1622))) :rule cong :premises (t232.t10.t3)) +(step t232.t10.t5 (cl @p_1623) :rule refl) +(step t232.t10.t6 (cl @p_1624) :rule refl) +(step t232.t10.t7 (cl @p_1625) :rule cong :premises (t232.t10.t5 t232.t10.t6)) +(step t232.t10.t8 (cl (= @p_1589 (! (is_res$b @p_1353 @p_1621) :named @p_1626))) :rule cong :premises (t232.t10.t7)) +(step t232.t10.t9 (cl (= @p_1591 (! (= @p_1622 @p_1626) :named @p_1627))) :rule cong :premises (t232.t10.t4 t232.t10.t8)) +(step t232.t10 (cl (= @p_1593 (! (forall ((veriT_vr189 F$) (veriT_vr190 D$)) @p_1627) :named @p_1628))) :rule bind) +(step t232.t11 (cl (= @p_1595 (! (and @p_1555 @p_1628) :named @p_1629))) :rule cong :premises (t232.t10)) +(step t232.t12 (cl (= @p_1597 (! (or @p_1329 @p_1629) :named @p_1631))) :rule cong :premises (t232.t11)) +(step t232.t13 (cl (! (= @p_1600 (! (=> @p_1630 @p_1631) :named @p_1633)) :named @p_1632)) :rule cong :premises (t232.t9 t232.t12)) +(step t232.t14 (cl (not @p_1632) (not @p_1600) @p_1633) :rule equiv_pos2) +(step t232.t15 (cl @p_1633) :rule th_resolution :premises (t232.t8 t232.t13 t232.t14)) +(anchor :step t232.t16 :args ((:= (veriT_vr185 A$) veriT_sk29) (:= (veriT_vr186 C$) veriT_sk30))) +(step t232.t16.t1 (cl (! (= veriT_vr185 veriT_sk29) :named @p_1636)) :rule refl) +(step t232.t16.t2 (cl (! (= @p_1602 (! (fun_app$ veriT_sk0 veriT_sk29) :named @p_1635)) :named @p_1640)) :rule cong :premises (t232.t16.t1)) +(step t232.t16.t3 (cl (! (= veriT_vr186 veriT_sk30) :named @p_1638)) :rule refl) +(step t232.t16.t4 (cl (! (= @p_1538 (! (run$ @p_1635 veriT_sk30) :named @p_1553)) :named @p_1641)) :rule cong :premises (t232.t16.t2 t232.t16.t3)) +(step t232.t16.t5 (cl (! (= @p_1536 (! (is_fail$ @p_1553) :named @p_1545)) :named @p_1642)) :rule cong :premises (t232.t16.t4)) +(step t232.t16.t6 (cl @p_1636) :rule refl) +(step t232.t16.t7 (cl (! (= @p_1604 (! (fun_app$ veriT_sk1 veriT_sk29) :named @p_1637)) :named @p_1645)) :rule cong :premises (t232.t16.t6)) +(step t232.t16.t8 (cl @p_1638) :rule refl) +(step t232.t16.t9 (cl (! (= @p_1537 (! (run$ @p_1637 veriT_sk30) :named @p_1546)) :named @p_1646)) :rule cong :premises (t232.t16.t7 t232.t16.t8)) +(step t232.t16.t10 (cl (= @p_1606 (! (is_fail$ @p_1546) :named @p_1639))) :rule cong :premises (t232.t16.t9)) +(step t232.t16.t11 (cl @p_1636) :rule refl) +(step t232.t16.t12 (cl @p_1640) :rule cong :premises (t232.t16.t11)) +(step t232.t16.t13 (cl @p_1638) :rule refl) +(step t232.t16.t14 (cl @p_1641) :rule cong :premises (t232.t16.t12 t232.t16.t13)) +(step t232.t16.t15 (cl @p_1642) :rule cong :premises (t232.t16.t14)) +(step t232.t16.t16 (cl (= @p_1610 (! (= @p_1639 @p_1545) :named @p_1643))) :rule cong :premises (t232.t16.t10 t232.t16.t15)) +(anchor :step t232.t16.t17 :args ((:= (veriT_vr187 B$) veriT_sk31) (:= (veriT_vr188 C$) veriT_sk32))) +(step t232.t16.t17.t1 (cl @p_1636) :rule refl) +(step t232.t16.t17.t2 (cl @p_1645) :rule cong :premises (t232.t16.t17.t1)) +(step t232.t16.t17.t3 (cl @p_1638) :rule refl) +(step t232.t16.t17.t4 (cl @p_1646) :rule cong :premises (t232.t16.t17.t2 t232.t16.t17.t3)) +(step t232.t16.t17.t5 (cl (! (= veriT_vr187 veriT_sk31) :named @p_1649)) :rule refl) +(step t232.t16.t17.t6 (cl (! (= veriT_vr188 veriT_sk32) :named @p_1650)) :rule refl) +(step t232.t16.t17.t7 (cl (! (= @p_1539 (! (pair$ veriT_sk31 veriT_sk32) :named @p_1554)) :named @p_1651)) :rule cong :premises (t232.t16.t17.t5 t232.t16.t17.t6)) +(step t232.t16.t17.t8 (cl (= @p_1614 (! (is_res$ @p_1546 @p_1554) :named @p_1648))) :rule cong :premises (t232.t16.t17.t4 t232.t16.t17.t7)) +(step t232.t16.t17.t9 (cl @p_1636) :rule refl) +(step t232.t16.t17.t10 (cl @p_1640) :rule cong :premises (t232.t16.t17.t9)) +(step t232.t16.t17.t11 (cl @p_1638) :rule refl) +(step t232.t16.t17.t12 (cl @p_1641) :rule cong :premises (t232.t16.t17.t10 t232.t16.t17.t11)) +(step t232.t16.t17.t13 (cl @p_1649) :rule refl) +(step t232.t16.t17.t14 (cl @p_1650) :rule refl) +(step t232.t16.t17.t15 (cl @p_1651) :rule cong :premises (t232.t16.t17.t13 t232.t16.t17.t14)) +(step t232.t16.t17.t16 (cl (= @p_1618 (! (is_res$ @p_1553 @p_1554) :named @p_1652))) :rule cong :premises (t232.t16.t17.t12 t232.t16.t17.t15)) +(step t232.t16.t17.t17 (cl (= @p_1619 (! (= @p_1648 @p_1652) :named @p_1644))) :rule cong :premises (t232.t16.t17.t8 t232.t16.t17.t16)) +(step t232.t16.t17 (cl (= @p_1611 @p_1644)) :rule sko_forall) +(step t232.t16.t18 (cl (= @p_1620 (! (and @p_1643 @p_1644) :named @p_1653))) :rule cong :premises (t232.t16.t16 t232.t16.t17)) +(step t232.t16.t19 (cl (= @p_1601 (! (or @p_1545 @p_1653) :named @p_1634))) :rule cong :premises (t232.t16.t5 t232.t16.t18)) +(step t232.t16 (cl (= @p_1630 @p_1634)) :rule sko_forall) +(step t232.t17 (cl (! (= @p_1633 (! (=> @p_1634 @p_1631) :named @p_1655)) :named @p_1654)) :rule cong :premises (t232.t16)) +(step t232.t18 (cl (not @p_1654) (not @p_1633) @p_1655) :rule equiv_pos2) +(step t232.t19 (cl @p_1655) :rule th_resolution :premises (t232.t15 t232.t17 t232.t18)) +(anchor :step t232.t20 :args ((:= (veriT_vr189 F$) veriT_vr191) (:= (veriT_vr190 D$) veriT_vr192))) +(step t232.t20.t1 (cl (! (= veriT_vr189 veriT_vr191) :named @p_1658)) :rule refl) +(step t232.t20.t2 (cl (! (= veriT_vr190 veriT_vr192) :named @p_1659)) :rule refl) +(step t232.t20.t3 (cl (! (= @p_1621 (! (pair$b veriT_vr191 veriT_vr192) :named @p_1556)) :named @p_1660)) :rule cong :premises (t232.t20.t1 t232.t20.t2)) +(step t232.t20.t4 (cl (= @p_1622 (! (is_res$b @p_826 @p_1556) :named @p_1657))) :rule cong :premises (t232.t20.t3)) +(step t232.t20.t5 (cl @p_1658) :rule refl) +(step t232.t20.t6 (cl @p_1659) :rule refl) +(step t232.t20.t7 (cl @p_1660) :rule cong :premises (t232.t20.t5 t232.t20.t6)) +(step t232.t20.t8 (cl (= @p_1626 (! (is_res$b @p_1353 @p_1556) :named @p_1661))) :rule cong :premises (t232.t20.t7)) +(step t232.t20.t9 (cl (= @p_1627 (! (= @p_1657 @p_1661) :named @p_1662))) :rule cong :premises (t232.t20.t4 t232.t20.t8)) +(step t232.t20 (cl (= @p_1628 (! (forall ((veriT_vr191 F$) (veriT_vr192 D$)) @p_1662) :named @p_1656))) :rule bind) +(step t232.t21 (cl (= @p_1629 (! (and @p_1555 @p_1656) :named @p_1663))) :rule cong :premises (t232.t20)) +(step t232.t22 (cl (= @p_1631 (! (or @p_1329 @p_1663) :named @p_1664))) :rule cong :premises (t232.t21)) +(step t232.t23 (cl (! (= @p_1655 (! (=> @p_1634 @p_1664) :named @p_1665)) :named @p_1666)) :rule cong :premises (t232.t22)) +(step t232.t24 (cl (not @p_1666) (not @p_1655) @p_1665) :rule equiv_pos2) +(step t232.t25 (cl @p_1665) :rule th_resolution :premises (t232.t19 t232.t23 t232.t24)) +(step t232 (cl @p_1599 @p_1665) :rule subproof :discharge (h1)) +(step t233 (cl @p_1385 @p_1535) :rule or :premises (t231)) +(step t234 (cl (! (or @p_1385 @p_1665) :named @p_1667) @p_1533) :rule or_neg) +(step t235 (cl @p_1667 @p_1154) :rule th_resolution :premises (t159 t234)) +(step t236 (cl @p_1667 (! (not @p_1665) :named @p_1679)) :rule or_neg) +(step t237 (cl @p_1667) :rule th_resolution :premises (t233 t232 t235 t236)) +(step t238 (cl @p_1668 @p_1669 (not @p_1670)) :rule and_neg) +(step t239 (cl @p_1163 @p_1671) :rule or :premises (t213)) +(step t240 (cl (! (not @p_1672) :named @p_1721) (not @p_822) @p_1328) :rule equiv_pos2) +(step t241 (cl @p_1354 @p_1672) :rule or :premises (t214)) +(step t242 (cl (! (not @p_1673) :named @p_1711) (not @p_784) @p_1332) :rule equiv_pos2) +(step t243 (cl @p_1354 @p_1673) :rule or :premises (t215)) +(step t244 (cl (not @p_1674) @p_1675 @p_1676) :rule equiv_pos1) +(step t245 (cl @p_1354 @p_1674) :rule or :premises (t216)) +(step t246 (cl (not @p_1381) @p_1374) :rule and_pos) +(step t247 (cl @p_1677 @p_1198 @p_1381) :rule or_pos) +(step t248 (cl @p_1187 @p_1382) :rule or :premises (t223)) +(step t249 (cl @p_1382) :rule resolution :premises (t248 t120)) +(step t250 (cl @p_1500 (not @p_1400)) :rule or_neg) +(step t251 (cl @p_1509 @p_1505 @p_1400) :rule equiv_neg2) +(step t252 (cl @p_1510 (! (not @p_1514) :named @p_1696) (! (not @p_1518) :named @p_1919)) :rule equiv_neg1) +(step t253 (cl @p_1510 @p_1514 @p_1518) :rule equiv_neg2) +(step t254 (cl @p_1519 (not @p_1509) (not @p_1510)) :rule and_neg) +(step t255 (cl @p_1500 (not @p_1519)) :rule or_neg) +(step t256 (cl (not @p_1410) (not @p_789) @p_1333) :rule equiv_pos2) +(step t257 (cl (not @p_1529) @p_1410) :rule and_pos) +(step t258 (cl (! (not @p_1530) :named @p_1710) @p_1333 @p_1529) :rule or_pos) +(step t259 (cl @p_1678 (not @p_1500) @p_1530) :rule implies_pos) +(step t260 (cl @p_1385 @p_1531) :rule or :premises (t230)) +(step t261 (cl @p_1531) :rule resolution :premises (t260 t119)) +(step t262 (cl @p_1634 (! (not @p_1545) :named @p_1929)) :rule or_neg) +(step t263 (cl @p_1643 @p_1639 @p_1545) :rule equiv_neg2) +(step t264 (cl @p_1644 (! (not @p_1648) :named @p_1700) (! (not @p_1652) :named @p_1705)) :rule equiv_neg1) +(step t265 (cl @p_1644 @p_1648 @p_1652) :rule equiv_neg2) +(step t266 (cl @p_1653 (not @p_1643) (! (not @p_1644) :named @p_1931)) :rule and_neg) +(step t267 (cl @p_1634 (! (not @p_1653) :named @p_1930)) :rule or_neg) +(step t268 (cl (not @p_1663) @p_1656) :rule and_pos) +(step t269 (cl (! (not @p_1664) :named @p_1720) @p_1329 @p_1663) :rule or_pos) +(step t270 (cl @p_1679 (! (not @p_1634) :named @p_1928) @p_1664) :rule implies_pos) +(step t271 (cl @p_1385 @p_1665) :rule or :premises (t237)) +(step t272 (cl @p_1665) :rule resolution :premises (t271 t119)) +(step t273 (cl (or @p_1341 (! (or @p_1400 (! (not @p_1505) :named @p_1683) @p_1400) :named @p_1680))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk25) (:= veriT_vr104 veriT_sk26))) +(step t274 (cl @p_1341 @p_1680) :rule or :premises (t273)) +(step t275 (cl (! (or @p_1187 @p_1680) :named @p_1681) @p_1343) :rule or_neg) +(step t276 (cl @p_1681 @p_1155) :rule th_resolution :premises (t198 t275)) +(step t277 (cl @p_1681 (! (not @p_1680) :named @p_1682)) :rule or_neg) +(step t278 (cl @p_1681) :rule th_resolution :premises (t194 t274 t276 t277)) +(anchor :step t279) +(assume t279.h1 @p_1680) +(step t279.t2 (cl (! (= @p_1680 (! (or @p_1400 @p_1683) :named @p_1684)) :named @p_1685)) :rule ac_simp) +(step t279.t3 (cl (not @p_1685) @p_1682 @p_1684) :rule equiv_pos2) +(step t279.t4 (cl @p_1684) :rule th_resolution :premises (t279.h1 t279.t2 t279.t3)) +(step t279 (cl @p_1682 @p_1684) :rule subproof :discharge (h1)) +(step t280 (cl @p_1187 @p_1680) :rule or :premises (t278)) +(step t281 (cl (! (or @p_1187 @p_1684) :named @p_1686) @p_1343) :rule or_neg) +(step t282 (cl @p_1686 @p_1155) :rule th_resolution :premises (t198 t281)) +(step t283 (cl @p_1686 (! (not @p_1684) :named @p_1694)) :rule or_neg) +(step t284 (cl @p_1686) :rule th_resolution :premises (t280 t279 t282 t283)) +(step t285 (cl (or @p_1341 (! (or @p_1545 (! (not @p_1639) :named @p_1690) @p_1545) :named @p_1687))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk29) (:= veriT_vr104 veriT_sk30))) +(step t286 (cl @p_1341 @p_1687) :rule or :premises (t285)) +(step t287 (cl (! (or @p_1187 @p_1687) :named @p_1688) @p_1343) :rule or_neg) +(step t288 (cl @p_1688 @p_1155) :rule th_resolution :premises (t198 t287)) +(step t289 (cl @p_1688 (! (not @p_1687) :named @p_1689)) :rule or_neg) +(step t290 (cl @p_1688) :rule th_resolution :premises (t194 t286 t288 t289)) +(anchor :step t291) +(assume t291.h1 @p_1687) +(step t291.t2 (cl (! (= @p_1687 (! (or @p_1545 @p_1690) :named @p_1691)) :named @p_1692)) :rule ac_simp) +(step t291.t3 (cl (not @p_1692) @p_1689 @p_1691) :rule equiv_pos2) +(step t291.t4 (cl @p_1691) :rule th_resolution :premises (t291.h1 t291.t2 t291.t3)) +(step t291 (cl @p_1689 @p_1691) :rule subproof :discharge (h1)) +(step t292 (cl @p_1187 @p_1687) :rule or :premises (t290)) +(step t293 (cl (! (or @p_1187 @p_1691) :named @p_1693) @p_1343) :rule or_neg) +(step t294 (cl @p_1693 @p_1155) :rule th_resolution :premises (t198 t293)) +(step t295 (cl @p_1693 (! (not @p_1691) :named @p_1695)) :rule or_neg) +(step t296 (cl @p_1693) :rule th_resolution :premises (t292 t291 t294 t295)) +(step t297 (cl @p_1694 @p_1400 @p_1683) :rule or_pos) +(step t298 (cl @p_1187 @p_1684) :rule or :premises (t284)) +(step t299 (cl @p_1684) :rule resolution :premises (t298 t120)) +(step t300 (cl @p_1695 @p_1545 @p_1690) :rule or_pos) +(step t301 (cl @p_1187 @p_1691) :rule or :premises (t296)) +(step t302 (cl @p_1691) :rule resolution :premises (t301 t120)) +(step t303 (cl (or (! (not @p_1351) :named @p_1697) (! (or @p_1400 @p_1696 @p_1518) :named @p_1698))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk25) (:= veriT_vr104 veriT_sk26) (:= veriT_vr105 veriT_sk27) (:= veriT_vr106 veriT_sk28))) +(step t304 (cl @p_1697 @p_1698) :rule or :premises (t303)) +(step t305 (cl (! (or @p_1187 @p_1698) :named @p_1699) @p_1343) :rule or_neg) +(step t306 (cl @p_1699 @p_1155) :rule th_resolution :premises (t198 t305)) +(step t307 (cl @p_1699 (! (not @p_1698) :named @p_1703)) :rule or_neg) +(step t308 (cl @p_1699) :rule th_resolution :premises (t211 t304 t306 t307)) +(step t309 (cl (or @p_1697 (! (or @p_1545 @p_1700 @p_1652) :named @p_1701))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk29) (:= veriT_vr104 veriT_sk30) (:= veriT_vr105 veriT_sk31) (:= veriT_vr106 veriT_sk32))) +(step t310 (cl @p_1697 @p_1701) :rule or :premises (t309)) +(step t311 (cl (! (or @p_1187 @p_1701) :named @p_1702) @p_1343) :rule or_neg) +(step t312 (cl @p_1702 @p_1155) :rule th_resolution :premises (t198 t311)) +(step t313 (cl @p_1702 (! (not @p_1701) :named @p_1704)) :rule or_neg) +(step t314 (cl @p_1702) :rule th_resolution :premises (t211 t310 t312 t313)) +(step t315 (cl @p_1703 @p_1400 @p_1696 @p_1518) :rule or_pos) +(step t316 (cl @p_1187 @p_1698) :rule or :premises (t308)) +(step t317 (cl @p_1698) :rule resolution :premises (t316 t120)) +(step t318 (cl @p_1704 @p_1545 @p_1700 @p_1652) :rule or_pos) +(step t319 (cl @p_1187 @p_1701) :rule or :premises (t314)) +(step t320 (cl @p_1701) :rule resolution :premises (t319 t120)) +(step t321 (cl @p_1545 @p_1652 @p_1644) :rule resolution :premises (t318 t265 t320)) +(step t322 (cl (or (! (not @p_1352) :named @p_1706) (! (or @p_1545 @p_1705 @p_1648) :named @p_1707))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk29) (:= veriT_vr104 veriT_sk30) (:= veriT_vr105 veriT_sk31) (:= veriT_vr106 veriT_sk32))) +(step t323 (cl @p_1706 @p_1707) :rule or :premises (t322)) +(step t324 (cl (! (or @p_1187 @p_1707) :named @p_1708) @p_1343) :rule or_neg) +(step t325 (cl @p_1708 @p_1155) :rule th_resolution :premises (t198 t324)) +(step t326 (cl @p_1708 (! (not @p_1707) :named @p_1709)) :rule or_neg) +(step t327 (cl @p_1708) :rule th_resolution :premises (t212 t323 t325 t326)) +(step t328 (cl @p_1709 @p_1545 @p_1705 @p_1648) :rule or_pos) +(step t329 (cl @p_1187 @p_1707) :rule or :premises (t327)) +(step t330 (cl @p_1707) :rule resolution :premises (t329 t120)) +(step t331 (cl @p_1710 @p_1711 @p_1178) :rule resolution :premises (t257 t258 t256 t178 t242 t128 t127 t180)) +(step t332 (cl (or (! (not @p_1656) :named @p_1712) (! (forall ((veriT_vr191 F$) (veriT_vr192 D$)) (or (not @p_1657) @p_1661)) :named @p_1713))) :rule qnt_cnf) +(step t333 (cl @p_1712 @p_1713) :rule or :premises (t332)) +(step t334 (cl (or (! (not @p_1713) :named @p_1714) (! (or (! (not @p_827) :named @p_1719) @p_1670) :named @p_1715))) :rule forall_inst :args ((:= veriT_vr191 veriT_sk7) (:= veriT_vr192 veriT_sk8))) +(step t335 (cl @p_1714 @p_1715) :rule or :premises (t334)) +(step t336 (cl (! (or @p_1712 @p_1715) :named @p_1717) (! (not @p_1712) :named @p_1716)) :rule or_neg) +(step t337 (cl (not @p_1716) @p_1656) :rule not_not) +(step t338 (cl @p_1717 @p_1656) :rule th_resolution :premises (t337 t336)) +(step t339 (cl @p_1717 (! (not @p_1715) :named @p_1718)) :rule or_neg) +(step t340 (cl @p_1717) :rule th_resolution :premises (t333 t335 t338 t339)) +(step t341 (cl @p_1718 @p_1719 @p_1670) :rule or_pos) +(step t342 (cl @p_1712 @p_1715) :rule or :premises (t340)) +(step t343 (cl @p_1720 @p_1668 @p_1721 @p_1179) :rule resolution :premises (t342 t268 t341 t269 t238 t175 t240 t134 t133 t177)) +(step t344 (cl (or @p_1165 @p_1179)) :rule forall_inst :args ((:= veriT_vr113 veriT_sk9) (:= veriT_vr114 veriT_sk10))) +(step t345 (cl (or @p_1165 (! (not (! (and @p_1675 (! (is_res$b (! (run$b (c$ veriT_sk11 veriT_sk1) veriT_sk12) :named @p_1757) @p_747) :named @p_1753)) :named @p_1752)) :named @p_1754))) :rule forall_inst :args ((:= veriT_vr113 veriT_sk11) (:= veriT_vr114 veriT_sk12))) +(step t346 (cl (or (! (not @p_1374) :named @p_1755) @p_1297)) :rule forall_inst :args ((:= veriT_vr150 veriT_sk15) (:= veriT_vr151 veriT_sk16))) +(step t347 (cl (or @p_1187 (! (or @p_1400 (! (and @p_1509 (! (forall ((veriT_vr105 B$) (veriT_vr106 C$)) (! (= (! (is_res$ @p_1401 @p_1087) :named @p_1726) (! (is_res$ @p_1408 @p_1087) :named @p_1731)) :named @p_1733)) :named @p_1724)) :named @p_1735)) :named @p_1722))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk25) (:= veriT_vr104 veriT_sk26))) +(anchor :step t348) +(assume t348.h1 @p_1722) +(anchor :step t348.t2 :args ((:= (veriT_vr105 B$) veriT_vr207) (:= (veriT_vr106 C$) veriT_vr208))) +(step t348.t2.t1 (cl (! (= veriT_vr105 veriT_vr207) :named @p_1728)) :rule refl) +(step t348.t2.t2 (cl (! (= veriT_vr106 veriT_vr208) :named @p_1729)) :rule refl) +(step t348.t2.t3 (cl (! (= @p_1087 (! (pair$ veriT_vr207 veriT_vr208) :named @p_1725)) :named @p_1730)) :rule cong :premises (t348.t2.t1 t348.t2.t2)) +(step t348.t2.t4 (cl (= @p_1726 (! (is_res$ @p_1401 @p_1725) :named @p_1727))) :rule cong :premises (t348.t2.t3)) +(step t348.t2.t5 (cl @p_1728) :rule refl) +(step t348.t2.t6 (cl @p_1729) :rule refl) +(step t348.t2.t7 (cl @p_1730) :rule cong :premises (t348.t2.t5 t348.t2.t6)) +(step t348.t2.t8 (cl (= @p_1731 (! (is_res$ @p_1408 @p_1725) :named @p_1732))) :rule cong :premises (t348.t2.t7)) +(step t348.t2.t9 (cl (= @p_1733 (! (= @p_1727 @p_1732) :named @p_1734))) :rule cong :premises (t348.t2.t4 t348.t2.t8)) +(step t348.t2 (cl (= @p_1724 (! (forall ((veriT_vr207 B$) (veriT_vr208 C$)) @p_1734) :named @p_1736))) :rule bind) +(step t348.t3 (cl (= @p_1735 (! (and @p_1509 @p_1736) :named @p_1737))) :rule cong :premises (t348.t2)) +(step t348.t4 (cl (! (= @p_1722 (! (or @p_1400 @p_1737) :named @p_1740)) :named @p_1738)) :rule cong :premises (t348.t3)) +(step t348.t5 (cl (not @p_1738) (! (not @p_1722) :named @p_1739) @p_1740) :rule equiv_pos2) +(step t348.t6 (cl @p_1740) :rule th_resolution :premises (t348.h1 t348.t4 t348.t5)) +(anchor :step t348.t7 :args ((:= (veriT_vr207 B$) veriT_vr209) (:= (veriT_vr208 C$) veriT_vr210))) +(step t348.t7.t1 (cl (! (= veriT_vr207 veriT_vr209) :named @p_1743)) :rule refl) +(step t348.t7.t2 (cl (! (= veriT_vr208 veriT_vr210) :named @p_1744)) :rule refl) +(step t348.t7.t3 (cl (! (= @p_1725 (! (pair$ veriT_vr209 veriT_vr210) :named @p_1723)) :named @p_1745)) :rule cong :premises (t348.t7.t1 t348.t7.t2)) +(step t348.t7.t4 (cl (= @p_1727 (! (is_res$ @p_1401 @p_1723) :named @p_1742))) :rule cong :premises (t348.t7.t3)) +(step t348.t7.t5 (cl @p_1743) :rule refl) +(step t348.t7.t6 (cl @p_1744) :rule refl) +(step t348.t7.t7 (cl @p_1745) :rule cong :premises (t348.t7.t5 t348.t7.t6)) +(step t348.t7.t8 (cl (= @p_1732 (! (is_res$ @p_1408 @p_1723) :named @p_1746))) :rule cong :premises (t348.t7.t7)) +(step t348.t7.t9 (cl (= @p_1734 (! (= @p_1742 @p_1746) :named @p_1747))) :rule cong :premises (t348.t7.t4 t348.t7.t8)) +(step t348.t7 (cl (= @p_1736 (! (forall ((veriT_vr209 B$) (veriT_vr210 C$)) @p_1747) :named @p_1741))) :rule bind) +(step t348.t8 (cl (= @p_1737 (! (and @p_1509 @p_1741) :named @p_1748))) :rule cong :premises (t348.t7)) +(step t348.t9 (cl (! (= @p_1740 (! (or @p_1400 @p_1748) :named @p_1749)) :named @p_1750)) :rule cong :premises (t348.t8)) +(step t348.t10 (cl (not @p_1750) (not @p_1740) @p_1749) :rule equiv_pos2) +(step t348.t11 (cl @p_1749) :rule th_resolution :premises (t348.t6 t348.t9 t348.t10)) +(step t348 (cl @p_1739 @p_1749) :rule subproof :discharge (h1)) +(step t349 (cl @p_1187 @p_1722) :rule or :premises (t347)) +(step t350 (cl (! (or @p_1187 @p_1749) :named @p_1751) @p_1343) :rule or_neg) +(step t351 (cl @p_1751 @p_1155) :rule th_resolution :premises (t198 t350)) +(step t352 (cl @p_1751 (! (not @p_1749) :named @p_1756)) :rule or_neg) +(step t353 (cl @p_1751) :rule th_resolution :premises (t349 t348 t351 t352)) +(step t354 (cl @p_1165 @p_1179) :rule or :premises (t344)) +(step t355 (cl @p_1752 (not @p_1675) (not @p_1753)) :rule and_neg) +(step t356 (cl @p_1165 @p_1754) :rule or :premises (t345)) +(step t357 (cl @p_1755 @p_1297) :rule or :premises (t346)) +(step t358 (cl (not @p_1748) @p_1741) :rule and_pos) +(step t359 (cl @p_1756 @p_1400 @p_1748) :rule or_pos) +(step t360 (cl @p_1187 @p_1749) :rule or :premises (t353)) +(step t361 (cl @p_1749) :rule resolution :premises (t360 t120)) +(step t362 (cl (or @p_1385 (! (=> @p_1412 (! (or @p_1325 (! (and (! (= (is_fail$b @p_1757) @p_1325) :named @p_1778) (! (forall ((veriT_vr93 F$) (veriT_vr102 D$)) (! (= (! (is_res$b @p_1757 @p_1039) :named @p_1806) (! (is_res$b @p_846 @p_1039) :named @p_1811)) :named @p_1813)) :named @p_1804)) :named @p_1815)) :named @p_1817)) :named @p_1758))) :rule forall_inst :args ((:= veriT_vr94 veriT_sk11) (:= veriT_vr95 veriT_sk0) (:= veriT_vr96 veriT_sk1) (:= veriT_vr101 veriT_sk12))) +(anchor :step t363) +(assume t363.h1 @p_1758) +(anchor :step t363.t2 :args ((:= (veriT_vr97 A$) veriT_vr237) (:= (veriT_vr98 C$) veriT_vr238))) +(step t363.t2.t1 (cl (! (= veriT_vr97 veriT_vr237) :named @p_1785)) :rule refl) +(step t363.t2.t2 (cl (! (= @p_1417 (! (fun_app$ veriT_sk0 veriT_vr237) :named @p_1784)) :named @p_1789)) :rule cong :premises (t363.t2.t1)) +(step t363.t2.t3 (cl (! (= veriT_vr98 veriT_vr238) :named @p_1787)) :rule refl) +(step t363.t2.t4 (cl (! (= @p_1388 (! (run$ @p_1784 veriT_vr238) :named @p_1782)) :named @p_1790)) :rule cong :premises (t363.t2.t2 t363.t2.t3)) +(step t363.t2.t5 (cl (! (= @p_1386 (! (is_fail$ @p_1782) :named @p_1780)) :named @p_1791)) :rule cong :premises (t363.t2.t4)) +(step t363.t2.t6 (cl @p_1785) :rule refl) +(step t363.t2.t7 (cl (! (= @p_1420 (! (fun_app$ veriT_sk1 veriT_vr237) :named @p_1786)) :named @p_1794)) :rule cong :premises (t363.t2.t6)) +(step t363.t2.t8 (cl @p_1787) :rule refl) +(step t363.t2.t9 (cl (! (= @p_1387 (! (run$ @p_1786 veriT_vr238) :named @p_1781)) :named @p_1795)) :rule cong :premises (t363.t2.t7 t363.t2.t8)) +(step t363.t2.t10 (cl (= @p_1423 (! (is_fail$ @p_1781) :named @p_1788))) :rule cong :premises (t363.t2.t9)) +(step t363.t2.t11 (cl @p_1785) :rule refl) +(step t363.t2.t12 (cl @p_1789) :rule cong :premises (t363.t2.t11)) +(step t363.t2.t13 (cl @p_1787) :rule refl) +(step t363.t2.t14 (cl @p_1790) :rule cong :premises (t363.t2.t12 t363.t2.t13)) +(step t363.t2.t15 (cl @p_1791) :rule cong :premises (t363.t2.t14)) +(step t363.t2.t16 (cl (= @p_1428 (! (= @p_1788 @p_1780) :named @p_1792))) :rule cong :premises (t363.t2.t10 t363.t2.t15)) +(anchor :step t363.t2.t17 :args ((:= (veriT_vr99 B$) veriT_vr239) (:= (veriT_vr100 C$) veriT_vr240))) +(step t363.t2.t17.t1 (cl @p_1785) :rule refl) +(step t363.t2.t17.t2 (cl @p_1794) :rule cong :premises (t363.t2.t17.t1)) +(step t363.t2.t17.t3 (cl @p_1787) :rule refl) +(step t363.t2.t17.t4 (cl @p_1795) :rule cong :premises (t363.t2.t17.t2 t363.t2.t17.t3)) +(step t363.t2.t17.t5 (cl (! (= veriT_vr99 veriT_vr239) :named @p_1797)) :rule refl) +(step t363.t2.t17.t6 (cl (! (= veriT_vr100 veriT_vr240) :named @p_1798)) :rule refl) +(step t363.t2.t17.t7 (cl (! (= @p_1035 (! (pair$ veriT_vr239 veriT_vr240) :named @p_1783)) :named @p_1799)) :rule cong :premises (t363.t2.t17.t5 t363.t2.t17.t6)) +(step t363.t2.t17.t8 (cl (= @p_1434 (! (is_res$ @p_1781 @p_1783) :named @p_1796))) :rule cong :premises (t363.t2.t17.t4 t363.t2.t17.t7)) +(step t363.t2.t17.t9 (cl @p_1785) :rule refl) +(step t363.t2.t17.t10 (cl @p_1789) :rule cong :premises (t363.t2.t17.t9)) +(step t363.t2.t17.t11 (cl @p_1787) :rule refl) +(step t363.t2.t17.t12 (cl @p_1790) :rule cong :premises (t363.t2.t17.t10 t363.t2.t17.t11)) +(step t363.t2.t17.t13 (cl @p_1797) :rule refl) +(step t363.t2.t17.t14 (cl @p_1798) :rule refl) +(step t363.t2.t17.t15 (cl @p_1799) :rule cong :premises (t363.t2.t17.t13 t363.t2.t17.t14)) +(step t363.t2.t17.t16 (cl (= @p_1439 (! (is_res$ @p_1782 @p_1783) :named @p_1800))) :rule cong :premises (t363.t2.t17.t12 t363.t2.t17.t15)) +(step t363.t2.t17.t17 (cl (= @p_1441 (! (= @p_1796 @p_1800) :named @p_1801))) :rule cong :premises (t363.t2.t17.t8 t363.t2.t17.t16)) +(step t363.t2.t17 (cl (= @p_1430 (! (forall ((veriT_vr239 B$) (veriT_vr240 C$)) @p_1801) :named @p_1793))) :rule bind) +(step t363.t2.t18 (cl (= @p_1443 (! (and @p_1792 @p_1793) :named @p_1802))) :rule cong :premises (t363.t2.t16 t363.t2.t17)) +(step t363.t2.t19 (cl (= @p_1445 (! (or @p_1780 @p_1802) :named @p_1803))) :rule cong :premises (t363.t2.t5 t363.t2.t18)) +(step t363.t2 (cl (= @p_1412 (! (forall ((veriT_vr237 A$) (veriT_vr238 C$)) @p_1803) :named @p_1819))) :rule bind) +(anchor :step t363.t3 :args ((:= (veriT_vr93 F$) veriT_vr241) (:= (veriT_vr102 D$) veriT_vr242))) +(step t363.t3.t1 (cl (! (= veriT_vr93 veriT_vr241) :named @p_1808)) :rule refl) +(step t363.t3.t2 (cl (! (= veriT_vr102 veriT_vr242) :named @p_1809)) :rule refl) +(step t363.t3.t3 (cl (! (= @p_1039 (! (pair$b veriT_vr241 veriT_vr242) :named @p_1805)) :named @p_1810)) :rule cong :premises (t363.t3.t1 t363.t3.t2)) +(step t363.t3.t4 (cl (= @p_1806 (! (is_res$b @p_1757 @p_1805) :named @p_1807))) :rule cong :premises (t363.t3.t3)) +(step t363.t3.t5 (cl @p_1808) :rule refl) +(step t363.t3.t6 (cl @p_1809) :rule refl) +(step t363.t3.t7 (cl @p_1810) :rule cong :premises (t363.t3.t5 t363.t3.t6)) +(step t363.t3.t8 (cl (= @p_1811 (! (is_res$b @p_846 @p_1805) :named @p_1812))) :rule cong :premises (t363.t3.t7)) +(step t363.t3.t9 (cl (= @p_1813 (! (= @p_1807 @p_1812) :named @p_1814))) :rule cong :premises (t363.t3.t4 t363.t3.t8)) +(step t363.t3 (cl (= @p_1804 (! (forall ((veriT_vr241 F$) (veriT_vr242 D$)) @p_1814) :named @p_1816))) :rule bind) +(step t363.t4 (cl (= @p_1815 (! (and @p_1778 @p_1816) :named @p_1818))) :rule cong :premises (t363.t3)) +(step t363.t5 (cl (= @p_1817 (! (or @p_1325 @p_1818) :named @p_1820))) :rule cong :premises (t363.t4)) +(step t363.t6 (cl (! (= @p_1758 (! (=> @p_1819 @p_1820) :named @p_1823)) :named @p_1821)) :rule cong :premises (t363.t2 t363.t5)) +(step t363.t7 (cl (not @p_1821) (! (not @p_1758) :named @p_1822) @p_1823) :rule equiv_pos2) +(step t363.t8 (cl @p_1823) :rule th_resolution :premises (t363.h1 t363.t6 t363.t7)) +(anchor :step t363.t9 :args ((:= (veriT_vr237 A$) veriT_vr243) (:= (veriT_vr238 C$) veriT_vr244))) +(step t363.t9.t1 (cl (! (= veriT_vr237 veriT_vr243) :named @p_1826)) :rule refl) +(step t363.t9.t2 (cl (! (= @p_1784 @p_1825) :named @p_1830)) :rule cong :premises (t363.t9.t1)) +(step t363.t9.t3 (cl (! (= veriT_vr238 veriT_vr244) :named @p_1828)) :rule refl) +(step t363.t9.t4 (cl (! (= @p_1782 @p_1761) :named @p_1831)) :rule cong :premises (t363.t9.t2 t363.t9.t3)) +(step t363.t9.t5 (cl (! (= @p_1780 @p_1759) :named @p_1832)) :rule cong :premises (t363.t9.t4)) +(step t363.t9.t6 (cl @p_1826) :rule refl) +(step t363.t9.t7 (cl (! (= @p_1786 @p_1827) :named @p_1835)) :rule cong :premises (t363.t9.t6)) +(step t363.t9.t8 (cl @p_1828) :rule refl) +(step t363.t9.t9 (cl (! (= @p_1781 @p_1760) :named @p_1836)) :rule cong :premises (t363.t9.t7 t363.t9.t8)) +(step t363.t9.t10 (cl (= @p_1788 @p_1829)) :rule cong :premises (t363.t9.t9)) +(step t363.t9.t11 (cl @p_1826) :rule refl) +(step t363.t9.t12 (cl @p_1830) :rule cong :premises (t363.t9.t11)) +(step t363.t9.t13 (cl @p_1828) :rule refl) +(step t363.t9.t14 (cl @p_1831) :rule cong :premises (t363.t9.t12 t363.t9.t13)) +(step t363.t9.t15 (cl @p_1832) :rule cong :premises (t363.t9.t14)) +(step t363.t9.t16 (cl (= @p_1792 @p_1833)) :rule cong :premises (t363.t9.t10 t363.t9.t15)) +(anchor :step t363.t9.t17 :args ((:= (veriT_vr239 B$) veriT_vr245) (:= (veriT_vr240 C$) veriT_vr246))) +(step t363.t9.t17.t1 (cl @p_1826) :rule refl) +(step t363.t9.t17.t2 (cl @p_1835) :rule cong :premises (t363.t9.t17.t1)) +(step t363.t9.t17.t3 (cl @p_1828) :rule refl) +(step t363.t9.t17.t4 (cl @p_1836) :rule cong :premises (t363.t9.t17.t2 t363.t9.t17.t3)) +(step t363.t9.t17.t5 (cl (! (= veriT_vr239 veriT_vr245) :named @p_1838)) :rule refl) +(step t363.t9.t17.t6 (cl (! (= veriT_vr240 veriT_vr246) :named @p_1839)) :rule refl) +(step t363.t9.t17.t7 (cl (! (= @p_1783 @p_1762) :named @p_1840)) :rule cong :premises (t363.t9.t17.t5 t363.t9.t17.t6)) +(step t363.t9.t17.t8 (cl (= @p_1796 @p_1837)) :rule cong :premises (t363.t9.t17.t4 t363.t9.t17.t7)) +(step t363.t9.t17.t9 (cl @p_1826) :rule refl) +(step t363.t9.t17.t10 (cl @p_1830) :rule cong :premises (t363.t9.t17.t9)) +(step t363.t9.t17.t11 (cl @p_1828) :rule refl) +(step t363.t9.t17.t12 (cl @p_1831) :rule cong :premises (t363.t9.t17.t10 t363.t9.t17.t11)) +(step t363.t9.t17.t13 (cl @p_1838) :rule refl) +(step t363.t9.t17.t14 (cl @p_1839) :rule refl) +(step t363.t9.t17.t15 (cl @p_1840) :rule cong :premises (t363.t9.t17.t13 t363.t9.t17.t14)) +(step t363.t9.t17.t16 (cl (= @p_1800 @p_1841)) :rule cong :premises (t363.t9.t17.t12 t363.t9.t17.t15)) +(step t363.t9.t17.t17 (cl (= @p_1801 @p_1842)) :rule cong :premises (t363.t9.t17.t8 t363.t9.t17.t16)) +(step t363.t9.t17 (cl (= @p_1793 @p_1834)) :rule bind) +(step t363.t9.t18 (cl (= @p_1802 @p_1843)) :rule cong :premises (t363.t9.t16 t363.t9.t17)) +(step t363.t9.t19 (cl (= @p_1803 @p_1824)) :rule cong :premises (t363.t9.t5 t363.t9.t18)) +(step t363.t9 (cl (= @p_1819 (! (forall ((veriT_vr243 A$) (veriT_vr244 C$)) @p_1824) :named @p_1853))) :rule bind) +(anchor :step t363.t10 :args ((:= (veriT_vr241 F$) veriT_vr247) (:= (veriT_vr242 D$) veriT_vr248))) +(step t363.t10.t1 (cl (! (= veriT_vr241 veriT_vr247) :named @p_1846)) :rule refl) +(step t363.t10.t2 (cl (! (= veriT_vr242 veriT_vr248) :named @p_1847)) :rule refl) +(step t363.t10.t3 (cl (! (= @p_1805 (! (pair$b veriT_vr247 veriT_vr248) :named @p_1844)) :named @p_1848)) :rule cong :premises (t363.t10.t1 t363.t10.t2)) +(step t363.t10.t4 (cl (= @p_1807 (! (is_res$b @p_1757 @p_1844) :named @p_1845))) :rule cong :premises (t363.t10.t3)) +(step t363.t10.t5 (cl @p_1846) :rule refl) +(step t363.t10.t6 (cl @p_1847) :rule refl) +(step t363.t10.t7 (cl @p_1848) :rule cong :premises (t363.t10.t5 t363.t10.t6)) +(step t363.t10.t8 (cl (= @p_1812 (! (is_res$b @p_846 @p_1844) :named @p_1849))) :rule cong :premises (t363.t10.t7)) +(step t363.t10.t9 (cl (= @p_1814 (! (= @p_1845 @p_1849) :named @p_1850))) :rule cong :premises (t363.t10.t4 t363.t10.t8)) +(step t363.t10 (cl (= @p_1816 (! (forall ((veriT_vr247 F$) (veriT_vr248 D$)) @p_1850) :named @p_1851))) :rule bind) +(step t363.t11 (cl (= @p_1818 (! (and @p_1778 @p_1851) :named @p_1852))) :rule cong :premises (t363.t10)) +(step t363.t12 (cl (= @p_1820 (! (or @p_1325 @p_1852) :named @p_1854))) :rule cong :premises (t363.t11)) +(step t363.t13 (cl (! (= @p_1823 (! (=> @p_1853 @p_1854) :named @p_1856)) :named @p_1855)) :rule cong :premises (t363.t9 t363.t12)) +(step t363.t14 (cl (not @p_1855) (not @p_1823) @p_1856) :rule equiv_pos2) +(step t363.t15 (cl @p_1856) :rule th_resolution :premises (t363.t8 t363.t13 t363.t14)) +(anchor :step t363.t16 :args ((:= (veriT_vr243 A$) veriT_sk41) (:= (veriT_vr244 C$) veriT_sk42))) +(step t363.t16.t1 (cl (! (= veriT_vr243 veriT_sk41) :named @p_1859)) :rule refl) +(step t363.t16.t2 (cl (! (= @p_1825 (! (fun_app$ veriT_sk0 veriT_sk41) :named @p_1858)) :named @p_1863)) :rule cong :premises (t363.t16.t1)) +(step t363.t16.t3 (cl (! (= veriT_vr244 veriT_sk42) :named @p_1861)) :rule refl) +(step t363.t16.t4 (cl (! (= @p_1761 (! (run$ @p_1858 veriT_sk42) :named @p_1776)) :named @p_1864)) :rule cong :premises (t363.t16.t2 t363.t16.t3)) +(step t363.t16.t5 (cl (! (= @p_1759 (! (is_fail$ @p_1776) :named @p_1768)) :named @p_1865)) :rule cong :premises (t363.t16.t4)) +(step t363.t16.t6 (cl @p_1859) :rule refl) +(step t363.t16.t7 (cl (! (= @p_1827 (! (fun_app$ veriT_sk1 veriT_sk41) :named @p_1860)) :named @p_1868)) :rule cong :premises (t363.t16.t6)) +(step t363.t16.t8 (cl @p_1861) :rule refl) +(step t363.t16.t9 (cl (! (= @p_1760 (! (run$ @p_1860 veriT_sk42) :named @p_1769)) :named @p_1869)) :rule cong :premises (t363.t16.t7 t363.t16.t8)) +(step t363.t16.t10 (cl (= @p_1829 (! (is_fail$ @p_1769) :named @p_1862))) :rule cong :premises (t363.t16.t9)) +(step t363.t16.t11 (cl @p_1859) :rule refl) +(step t363.t16.t12 (cl @p_1863) :rule cong :premises (t363.t16.t11)) +(step t363.t16.t13 (cl @p_1861) :rule refl) +(step t363.t16.t14 (cl @p_1864) :rule cong :premises (t363.t16.t12 t363.t16.t13)) +(step t363.t16.t15 (cl @p_1865) :rule cong :premises (t363.t16.t14)) +(step t363.t16.t16 (cl (= @p_1833 (! (= @p_1862 @p_1768) :named @p_1866))) :rule cong :premises (t363.t16.t10 t363.t16.t15)) +(anchor :step t363.t16.t17 :args ((:= (veriT_vr245 B$) veriT_sk43) (:= (veriT_vr246 C$) veriT_sk44))) +(step t363.t16.t17.t1 (cl @p_1859) :rule refl) +(step t363.t16.t17.t2 (cl @p_1868) :rule cong :premises (t363.t16.t17.t1)) +(step t363.t16.t17.t3 (cl @p_1861) :rule refl) +(step t363.t16.t17.t4 (cl @p_1869) :rule cong :premises (t363.t16.t17.t2 t363.t16.t17.t3)) +(step t363.t16.t17.t5 (cl (! (= veriT_vr245 veriT_sk43) :named @p_1872)) :rule refl) +(step t363.t16.t17.t6 (cl (! (= veriT_vr246 veriT_sk44) :named @p_1873)) :rule refl) +(step t363.t16.t17.t7 (cl (! (= @p_1762 (! (pair$ veriT_sk43 veriT_sk44) :named @p_1777)) :named @p_1874)) :rule cong :premises (t363.t16.t17.t5 t363.t16.t17.t6)) +(step t363.t16.t17.t8 (cl (= @p_1837 (! (is_res$ @p_1769 @p_1777) :named @p_1871))) :rule cong :premises (t363.t16.t17.t4 t363.t16.t17.t7)) +(step t363.t16.t17.t9 (cl @p_1859) :rule refl) +(step t363.t16.t17.t10 (cl @p_1863) :rule cong :premises (t363.t16.t17.t9)) +(step t363.t16.t17.t11 (cl @p_1861) :rule refl) +(step t363.t16.t17.t12 (cl @p_1864) :rule cong :premises (t363.t16.t17.t10 t363.t16.t17.t11)) +(step t363.t16.t17.t13 (cl @p_1872) :rule refl) +(step t363.t16.t17.t14 (cl @p_1873) :rule refl) +(step t363.t16.t17.t15 (cl @p_1874) :rule cong :premises (t363.t16.t17.t13 t363.t16.t17.t14)) +(step t363.t16.t17.t16 (cl (= @p_1841 (! (is_res$ @p_1776 @p_1777) :named @p_1875))) :rule cong :premises (t363.t16.t17.t12 t363.t16.t17.t15)) +(step t363.t16.t17.t17 (cl (= @p_1842 (! (= @p_1871 @p_1875) :named @p_1867))) :rule cong :premises (t363.t16.t17.t8 t363.t16.t17.t16)) +(step t363.t16.t17 (cl (= @p_1834 @p_1867)) :rule sko_forall) +(step t363.t16.t18 (cl (= @p_1843 (! (and @p_1866 @p_1867) :named @p_1876))) :rule cong :premises (t363.t16.t16 t363.t16.t17)) +(step t363.t16.t19 (cl (= @p_1824 (! (or @p_1768 @p_1876) :named @p_1857))) :rule cong :premises (t363.t16.t5 t363.t16.t18)) +(step t363.t16 (cl (= @p_1853 @p_1857)) :rule sko_forall) +(step t363.t17 (cl (! (= @p_1856 (! (=> @p_1857 @p_1854) :named @p_1878)) :named @p_1877)) :rule cong :premises (t363.t16)) +(step t363.t18 (cl (not @p_1877) (not @p_1856) @p_1878) :rule equiv_pos2) +(step t363.t19 (cl @p_1878) :rule th_resolution :premises (t363.t15 t363.t17 t363.t18)) +(anchor :step t363.t20 :args ((:= (veriT_vr247 F$) veriT_vr249) (:= (veriT_vr248 D$) veriT_vr250))) +(step t363.t20.t1 (cl (! (= veriT_vr247 veriT_vr249) :named @p_1881)) :rule refl) +(step t363.t20.t2 (cl (! (= veriT_vr248 veriT_vr250) :named @p_1882)) :rule refl) +(step t363.t20.t3 (cl (! (= @p_1844 (! (pair$b veriT_vr249 veriT_vr250) :named @p_1779)) :named @p_1883)) :rule cong :premises (t363.t20.t1 t363.t20.t2)) +(step t363.t20.t4 (cl (= @p_1845 (! (is_res$b @p_1757 @p_1779) :named @p_1880))) :rule cong :premises (t363.t20.t3)) +(step t363.t20.t5 (cl @p_1881) :rule refl) +(step t363.t20.t6 (cl @p_1882) :rule refl) +(step t363.t20.t7 (cl @p_1883) :rule cong :premises (t363.t20.t5 t363.t20.t6)) +(step t363.t20.t8 (cl (= @p_1849 (! (is_res$b @p_846 @p_1779) :named @p_1884))) :rule cong :premises (t363.t20.t7)) +(step t363.t20.t9 (cl (= @p_1850 (! (= @p_1880 @p_1884) :named @p_1885))) :rule cong :premises (t363.t20.t4 t363.t20.t8)) +(step t363.t20 (cl (= @p_1851 (! (forall ((veriT_vr249 F$) (veriT_vr250 D$)) @p_1885) :named @p_1879))) :rule bind) +(step t363.t21 (cl (= @p_1852 (! (and @p_1778 @p_1879) :named @p_1886))) :rule cong :premises (t363.t20)) +(step t363.t22 (cl (= @p_1854 (! (or @p_1325 @p_1886) :named @p_1887))) :rule cong :premises (t363.t21)) +(step t363.t23 (cl (! (= @p_1878 (! (=> @p_1857 @p_1887) :named @p_1888)) :named @p_1889)) :rule cong :premises (t363.t22)) +(step t363.t24 (cl (not @p_1889) (not @p_1878) @p_1888) :rule equiv_pos2) +(step t363.t25 (cl @p_1888) :rule th_resolution :premises (t363.t19 t363.t23 t363.t24)) +(step t363 (cl @p_1822 @p_1888) :rule subproof :discharge (h1)) +(step t364 (cl @p_1385 @p_1758) :rule or :premises (t362)) +(step t365 (cl (! (or @p_1385 @p_1888) :named @p_1890) @p_1533) :rule or_neg) +(step t366 (cl @p_1890 @p_1154) :rule th_resolution :premises (t159 t365)) +(step t367 (cl @p_1890 (! (not @p_1888) :named @p_1891)) :rule or_neg) +(step t368 (cl @p_1890) :rule th_resolution :premises (t364 t363 t366 t367)) +(step t369 (cl @p_1857 (not @p_1768)) :rule or_neg) +(step t370 (cl @p_1866 @p_1862 @p_1768) :rule equiv_neg2) +(step t371 (cl @p_1867 (! (not @p_1871) :named @p_1900) (! (not @p_1875) :named @p_1904)) :rule equiv_neg1) +(step t372 (cl @p_1867 @p_1871 @p_1875) :rule equiv_neg2) +(step t373 (cl @p_1876 (not @p_1866) (not @p_1867)) :rule and_neg) +(step t374 (cl @p_1857 (not @p_1876)) :rule or_neg) +(step t375 (cl (not @p_1886) @p_1879) :rule and_pos) +(step t376 (cl (! (not @p_1887) :named @p_1916) @p_1325 @p_1886) :rule or_pos) +(step t377 (cl @p_1891 (not @p_1857) @p_1887) :rule implies_pos) +(step t378 (cl @p_1385 @p_1888) :rule or :premises (t368)) +(step t379 (cl @p_1888) :rule resolution :premises (t378 t119)) +(step t380 (cl (or @p_1341 (! (or @p_1768 (! (not @p_1862) :named @p_1895) @p_1768) :named @p_1892))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk41) (:= veriT_vr104 veriT_sk42))) +(step t381 (cl @p_1341 @p_1892) :rule or :premises (t380)) +(step t382 (cl (! (or @p_1187 @p_1892) :named @p_1893) @p_1343) :rule or_neg) +(step t383 (cl @p_1893 @p_1155) :rule th_resolution :premises (t198 t382)) +(step t384 (cl @p_1893 (! (not @p_1892) :named @p_1894)) :rule or_neg) +(step t385 (cl @p_1893) :rule th_resolution :premises (t194 t381 t383 t384)) +(anchor :step t386) +(assume t386.h1 @p_1892) +(step t386.t2 (cl (! (= @p_1892 (! (or @p_1768 @p_1895) :named @p_1896)) :named @p_1897)) :rule ac_simp) +(step t386.t3 (cl (not @p_1897) @p_1894 @p_1896) :rule equiv_pos2) +(step t386.t4 (cl @p_1896) :rule th_resolution :premises (t386.h1 t386.t2 t386.t3)) +(step t386 (cl @p_1894 @p_1896) :rule subproof :discharge (h1)) +(step t387 (cl @p_1187 @p_1892) :rule or :premises (t385)) +(step t388 (cl (! (or @p_1187 @p_1896) :named @p_1898) @p_1343) :rule or_neg) +(step t389 (cl @p_1898 @p_1155) :rule th_resolution :premises (t198 t388)) +(step t390 (cl @p_1898 (! (not @p_1896) :named @p_1899)) :rule or_neg) +(step t391 (cl @p_1898) :rule th_resolution :premises (t387 t386 t389 t390)) +(step t392 (cl @p_1899 @p_1768 @p_1895) :rule or_pos) +(step t393 (cl @p_1187 @p_1896) :rule or :premises (t391)) +(step t394 (cl @p_1896) :rule resolution :premises (t393 t120)) +(step t395 (cl (or @p_1697 (! (or @p_1768 @p_1900 @p_1875) :named @p_1901))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk41) (:= veriT_vr104 veriT_sk42) (:= veriT_vr105 veriT_sk43) (:= veriT_vr106 veriT_sk44))) +(step t396 (cl @p_1697 @p_1901) :rule or :premises (t395)) +(step t397 (cl (! (or @p_1187 @p_1901) :named @p_1902) @p_1343) :rule or_neg) +(step t398 (cl @p_1902 @p_1155) :rule th_resolution :premises (t198 t397)) +(step t399 (cl @p_1902 (! (not @p_1901) :named @p_1903)) :rule or_neg) +(step t400 (cl @p_1902) :rule th_resolution :premises (t211 t396 t398 t399)) +(step t401 (cl @p_1903 @p_1768 @p_1900 @p_1875) :rule or_pos) +(step t402 (cl @p_1187 @p_1901) :rule or :premises (t400)) +(step t403 (cl @p_1901) :rule resolution :premises (t402 t120)) +(step t404 (cl @p_1768 @p_1875 @p_1867) :rule resolution :premises (t401 t372 t403)) +(step t405 (cl (or @p_1706 (! (or @p_1768 @p_1904 @p_1871) :named @p_1905))) :rule forall_inst :args ((:= veriT_vr103 veriT_sk41) (:= veriT_vr104 veriT_sk42) (:= veriT_vr105 veriT_sk43) (:= veriT_vr106 veriT_sk44))) +(step t406 (cl @p_1706 @p_1905) :rule or :premises (t405)) +(step t407 (cl (! (or @p_1187 @p_1905) :named @p_1906) @p_1343) :rule or_neg) +(step t408 (cl @p_1906 @p_1155) :rule th_resolution :premises (t198 t407)) +(step t409 (cl @p_1906 (! (not @p_1905) :named @p_1907)) :rule or_neg) +(step t410 (cl @p_1906) :rule th_resolution :premises (t212 t406 t408 t409)) +(step t411 (cl @p_1907 @p_1768 @p_1904 @p_1871) :rule or_pos) +(step t412 (cl @p_1187 @p_1905) :rule or :premises (t410)) +(step t413 (cl @p_1905) :rule resolution :premises (t412 t120)) +(step t414 (cl (or (! (not @p_1879) :named @p_1908) (! (forall ((veriT_vr249 F$) (veriT_vr250 D$)) (or (not @p_1884) @p_1880)) :named @p_1909))) :rule qnt_cnf) +(step t415 (cl @p_1908 @p_1909) :rule or :premises (t414)) +(step t416 (cl (or (! (not @p_1909) :named @p_1910) (! (or (! (not @p_847) :named @p_1915) @p_1753) :named @p_1911))) :rule forall_inst :args ((:= veriT_vr249 veriT_sk7) (:= veriT_vr250 veriT_sk8))) +(step t417 (cl @p_1910 @p_1911) :rule or :premises (t416)) +(step t418 (cl (! (or @p_1908 @p_1911) :named @p_1913) (! (not @p_1908) :named @p_1912)) :rule or_neg) +(step t419 (cl (not @p_1912) @p_1879) :rule not_not) +(step t420 (cl @p_1913 @p_1879) :rule th_resolution :premises (t419 t418)) +(step t421 (cl @p_1913 (! (not @p_1911) :named @p_1914)) :rule or_neg) +(step t422 (cl @p_1913) :rule th_resolution :premises (t415 t417 t420 t421)) +(step t423 (cl @p_1914 @p_1915 @p_1753) :rule or_pos) +(step t424 (cl @p_1908 @p_1911) :rule or :premises (t422)) +(step t425 (cl @p_1916 @p_820 @p_780 @p_1335) :rule resolution :premises (t424 t423 t375 t355 t376 t244 t172 t356 t142 t141 t147 t151 t148 t144 t149 t136 t130 t135 t129 t245 t191 t187 t186 t174 t171)) +(step t426 (cl @p_1287) :rule resolution :premises (t357 t183 t246 t182 t247 t208 t184 t181 t249 t210)) +(step t427 (cl @p_1317) :rule resolution :premises (t189 t426 t193)) +(step t428 (cl @p_1316) :rule resolution :premises (t192 t427)) +(step t429 (cl @p_1208) :rule resolution :premises (t186 t428)) +(step t430 (cl @p_1309) :rule resolution :premises (t187 t428)) +(step t431 (cl @p_1338) :rule resolution :premises (t191 t429)) +(step t432 (cl @p_1672) :rule resolution :premises (t241 t430)) +(step t433 (cl @p_1673) :rule resolution :premises (t243 t430)) +(step t434 (cl @p_1857) :rule resolution :premises (t411 t371 t404 t373 t370 t392 t374 t369 t413 t394)) +(step t435 (cl @p_1887) :rule resolution :premises (t377 t434 t379)) +(step t436 (cl @p_1400 @p_1518 @p_1510) :rule resolution :premises (t315 t253 t317)) +(step t437 (cl (or (! (not @p_1741) :named @p_1917) (! (forall ((veriT_vr209 B$) (veriT_vr210 C$)) (or (not @p_1746) @p_1742)) :named @p_1918))) :rule qnt_cnf) +(step t438 (cl @p_1917 @p_1918) :rule or :premises (t437)) +(step t439 (cl (or (! (not @p_1918) :named @p_1920) (! (or @p_1919 @p_1514) :named @p_1921))) :rule forall_inst :args ((:= veriT_vr209 veriT_sk27) (:= veriT_vr210 veriT_sk28))) +(step t440 (cl @p_1920 @p_1921) :rule or :premises (t439)) +(step t441 (cl (! (or @p_1917 @p_1921) :named @p_1923) (! (not @p_1917) :named @p_1922)) :rule or_neg) +(step t442 (cl (not @p_1922) @p_1741) :rule not_not) +(step t443 (cl @p_1923 @p_1741) :rule th_resolution :premises (t442 t441)) +(step t444 (cl @p_1923 (! (not @p_1921) :named @p_1924)) :rule or_neg) +(step t445 (cl @p_1923) :rule th_resolution :premises (t438 t440 t443 t444)) +(step t446 (cl @p_1924 @p_1919 @p_1514) :rule or_pos) +(step t447 (cl @p_1917 @p_1921) :rule or :premises (t445)) +(step t448 (cl @p_1500) :rule resolution :premises (t446 t252 t436 t447 t254 t358 t251 t359 t297 t255 t250 t361 t299)) +(step t449 (cl @p_1530) :rule resolution :premises (t259 t448 t261)) +(step t450 (cl @p_1178) :rule resolution :premises (t331 t449 t433)) +(step t451 (cl @p_820) :rule resolution :premises (t425 t450 t428 t435)) +(step t452 (cl @p_1925) :rule resolution :premises (t129 t450 t431)) +(step t453 (cl @p_1165) :rule resolution :premises (t354 t451)) +(step t454 (cl @p_1167) :rule resolution :premises (t130 t452)) +(step t455 (cl @p_1166) :rule resolution :premises (t147 t453)) +(step t456 (cl @p_1170) :rule resolution :premises (t148 t455)) +(step t457 (cl @p_1926) :rule resolution :premises (t149 t456 t171 t454)) +(step t458 (cl @p_1927) :rule resolution :premises (t140 t457)) +(step t459 (cl @p_1181) :rule resolution :premises (t139 t458)) +(step t460 (cl @p_1671) :rule resolution :premises (t239 t459)) +(step t461 (cl @p_1720) :rule resolution :premises (t343 t460 t432 t451)) +(step t462 (cl @p_1928) :rule resolution :premises (t270 t461 t272)) +(step t463 (cl @p_1929) :rule resolution :premises (t262 t462)) +(step t464 (cl @p_1930) :rule resolution :premises (t267 t462)) +(step t465 (cl @p_1690) :rule resolution :premises (t300 t463 t302)) +(step t466 (cl @p_1643) :rule resolution :premises (t263 t465 t463)) +(step t467 (cl @p_1931) :rule resolution :premises (t266 t466 t464)) +(step t468 (cl @p_1652) :rule resolution :premises (t321 t467 t463)) +(step t469 (cl @p_1700) :rule resolution :premises (t264 t468 t467)) +(step t470 (cl) :rule resolution :premises (t328 t468 t463 t330 t469)) +a352c3d2d258129c9c0fa30de525ad6ea4644748 543 0 +unsat +(define-fun veriT_sk0 () Exp$ (! (choice ((veriT_vr40 Exp$)) (not (! (=> (! (member$ veriT_vr40 (! (myset$ z$) :named @p_199)) :named @p_278) (! (not (! (forall ((veriT_vr41 FreeExp$)) (! (not (! (= veriT_vr40 (! (fun_app$ uu$ veriT_vr41) :named @p_281)) :named @p_282)) :named @p_283)) :named @p_279)) :named @p_284)) :named @p_277))) :named @p_201)) +(define-fun veriT_sk1 () FreeExp_list$ (! (choice ((veriT_vr42 FreeExp_list$)) (! (= z$ (! (map2$ uu$ veriT_vr42) :named @p_286)) :named @p_285)) :named @p_301)) +(define-fun veriT_sk2 () FreeExp$ (! (choice ((veriT_vr48 FreeExp$)) (not (! (not (! (= veriT_sk0 (! (abs_Exp$ (! (myImage$ exprel$ (! (insert$ veriT_vr48 bot$) :named @p_356)) :named @p_357)) :named @p_358)) :named @p_359)) :named @p_355))) :named @p_366)) +(assume axiom0 (! (forall ((?v0 FreeExp$)) (! (= (! (fun_app$ uu$ ?v0) :named @p_3) (! (abs_Exp$ (! (myImage$ exprel$ (! (insert$ ?v0 bot$) :named @p_6)) :named @p_8)) :named @p_10)) :named @p_12)) :named @p_2)) +(assume axiom1 (! (forall ((?v0 FreeExp_list$)) (! (= (! (abs_ExpList$ ?v0) :named @p_1) (! (map2$ uu$ ?v0) :named @p_27)) :named @p_29)) :named @p_24)) +(assume axiom2 (! (forall ((?v0 Exp$)) (! (=> (! (forall ((?v1 FreeExp$)) (! (=> (! (= ?v0 (! (abs_Exp$ (! (myImage$ exprel$ (! (insert$ ?v1 bot$) :named @p_42)) :named @p_44)) :named @p_46)) :named @p_48) false) :named @p_50)) :named @p_40) false) :named @p_52)) :named @p_39)) +(assume axiom3 (! (forall ((?v0 Exp_list$) (?v1 FreeExp_exp_fun$)) (! (= (! (exists ((?v2 FreeExp_list$)) (! (= ?v0 (! (map2$ ?v1 ?v2) :named @p_74)) :named @p_76)) :named @p_72) (! (forall ((?v2 Exp$)) (! (=> (! (member$ ?v2 (! (myset$ ?v0) :named @p_81)) :named @p_83) (! (exists ((?v3 FreeExp$)) (! (= ?v2 (! (fun_app$ ?v1 ?v3) :named @p_89)) :named @p_91)) :named @p_85)) :named @p_93)) :named @p_78)) :named @p_95)) :named @p_71)) +(assume axiom4 (! (not (! (exists ((?v0 FreeExp_list$)) (! (= @p_1 z$) :named @p_178)) :named @p_176)) :named @p_180)) +(anchor :step t6 :args ((:= (?v0 FreeExp$) veriT_vr0))) +(step t6.t1 (cl (! (= ?v0 veriT_vr0) :named @p_5)) :rule refl) +(step t6.t2 (cl (= @p_3 (! (fun_app$ uu$ veriT_vr0) :named @p_4))) :rule cong :premises (t6.t1)) +(step t6.t3 (cl @p_5) :rule refl) +(step t6.t4 (cl (= @p_6 (! (insert$ veriT_vr0 bot$) :named @p_7))) :rule cong :premises (t6.t3)) +(step t6.t5 (cl (= @p_8 (! (myImage$ exprel$ @p_7) :named @p_9))) :rule cong :premises (t6.t4)) +(step t6.t6 (cl (= @p_10 (! (abs_Exp$ @p_9) :named @p_11))) :rule cong :premises (t6.t5)) +(step t6.t7 (cl (= @p_12 (! (= @p_4 @p_11) :named @p_13))) :rule cong :premises (t6.t2 t6.t6)) +(step t6 (cl (! (= @p_2 (! (forall ((veriT_vr0 FreeExp$)) @p_13) :named @p_15)) :named @p_14)) :rule bind) +(step t7 (cl (not @p_14) (not @p_2) @p_15) :rule equiv_pos2) +(step t8 (cl @p_15) :rule th_resolution :premises (axiom0 t6 t7)) +(anchor :step t9 :args ((:= (veriT_vr0 FreeExp$) veriT_vr1))) +(step t9.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_17)) :rule refl) +(step t9.t2 (cl (= @p_4 (! (fun_app$ uu$ veriT_vr1) :named @p_16))) :rule cong :premises (t9.t1)) +(step t9.t3 (cl @p_17) :rule refl) +(step t9.t4 (cl (= @p_7 (! (insert$ veriT_vr1 bot$) :named @p_18))) :rule cong :premises (t9.t3)) +(step t9.t5 (cl (= @p_9 (! (myImage$ exprel$ @p_18) :named @p_19))) :rule cong :premises (t9.t4)) +(step t9.t6 (cl (= @p_11 (! (abs_Exp$ @p_19) :named @p_20))) :rule cong :premises (t9.t5)) +(step t9.t7 (cl (= @p_13 (! (= @p_16 @p_20) :named @p_21))) :rule cong :premises (t9.t2 t9.t6)) +(step t9 (cl (! (= @p_15 (! (forall ((veriT_vr1 FreeExp$)) @p_21) :named @p_23)) :named @p_22)) :rule bind) +(step t10 (cl (not @p_22) (not @p_15) @p_23) :rule equiv_pos2) +(step t11 (cl @p_23) :rule th_resolution :premises (t8 t9 t10)) +(anchor :step t12 :args ((:= (?v0 FreeExp_list$) veriT_vr2))) +(step t12.t1 (cl (! (= ?v0 veriT_vr2) :named @p_26)) :rule refl) +(step t12.t2 (cl (= @p_1 (! (abs_ExpList$ veriT_vr2) :named @p_25))) :rule cong :premises (t12.t1)) +(step t12.t3 (cl @p_26) :rule refl) +(step t12.t4 (cl (= @p_27 (! (map2$ uu$ veriT_vr2) :named @p_28))) :rule cong :premises (t12.t3)) +(step t12.t5 (cl (= @p_29 (! (= @p_25 @p_28) :named @p_30))) :rule cong :premises (t12.t2 t12.t4)) +(step t12 (cl (! (= @p_24 (! (forall ((veriT_vr2 FreeExp_list$)) @p_30) :named @p_32)) :named @p_31)) :rule bind) +(step t13 (cl (not @p_31) (not @p_24) @p_32) :rule equiv_pos2) +(step t14 (cl @p_32) :rule th_resolution :premises (axiom1 t12 t13)) +(anchor :step t15 :args ((:= (veriT_vr2 FreeExp_list$) veriT_vr3))) +(step t15.t1 (cl (! (= veriT_vr2 veriT_vr3) :named @p_34)) :rule refl) +(step t15.t2 (cl (= @p_25 (! (abs_ExpList$ veriT_vr3) :named @p_33))) :rule cong :premises (t15.t1)) +(step t15.t3 (cl @p_34) :rule refl) +(step t15.t4 (cl (= @p_28 (! (map2$ uu$ veriT_vr3) :named @p_35))) :rule cong :premises (t15.t3)) +(step t15.t5 (cl (= @p_30 (! (= @p_33 @p_35) :named @p_36))) :rule cong :premises (t15.t2 t15.t4)) +(step t15 (cl (! (= @p_32 (! (forall ((veriT_vr3 FreeExp_list$)) @p_36) :named @p_38)) :named @p_37)) :rule bind) +(step t16 (cl (not @p_37) (not @p_32) @p_38) :rule equiv_pos2) +(step t17 (cl @p_38) :rule th_resolution :premises (t14 t15 t16)) +(anchor :step t18 :args ((:= (?v0 Exp$) veriT_vr4))) +(anchor :step t18.t1 :args ((:= (?v1 FreeExp$) veriT_vr5))) +(step t18.t1.t1 (cl (= ?v0 veriT_vr4)) :rule refl) +(step t18.t1.t2 (cl (= ?v1 veriT_vr5)) :rule refl) +(step t18.t1.t3 (cl (= @p_42 (! (insert$ veriT_vr5 bot$) :named @p_43))) :rule cong :premises (t18.t1.t2)) +(step t18.t1.t4 (cl (= @p_44 (! (myImage$ exprel$ @p_43) :named @p_45))) :rule cong :premises (t18.t1.t3)) +(step t18.t1.t5 (cl (= @p_46 (! (abs_Exp$ @p_45) :named @p_47))) :rule cong :premises (t18.t1.t4)) +(step t18.t1.t6 (cl (= @p_48 (! (= veriT_vr4 @p_47) :named @p_49))) :rule cong :premises (t18.t1.t1 t18.t1.t5)) +(step t18.t1.t7 (cl (= @p_50 (! (=> @p_49 false) :named @p_51))) :rule cong :premises (t18.t1.t6)) +(step t18.t1 (cl (= @p_40 (! (forall ((veriT_vr5 FreeExp$)) @p_51) :named @p_41))) :rule bind) +(step t18.t2 (cl (= @p_52 (! (=> @p_41 false) :named @p_53))) :rule cong :premises (t18.t1)) +(step t18 (cl (! (= @p_39 (! (forall ((veriT_vr4 Exp$)) @p_53) :named @p_55)) :named @p_54)) :rule bind) +(step t19 (cl (not @p_54) (not @p_39) @p_55) :rule equiv_pos2) +(step t20 (cl @p_55) :rule th_resolution :premises (axiom2 t18 t19)) +(anchor :step t21 :args ((veriT_vr4 Exp$))) +(anchor :step t21.t1 :args ((veriT_vr5 FreeExp$))) +(step t21.t1.t1 (cl (= @p_51 (! (not @p_49) :named @p_57))) :rule implies_simplify) +(step t21.t1 (cl (= @p_41 (! (forall ((veriT_vr5 FreeExp$)) @p_57) :named @p_56))) :rule bind) +(step t21.t2 (cl (= @p_53 (! (=> @p_56 false) :named @p_58))) :rule cong :premises (t21.t1)) +(step t21.t3 (cl (= @p_58 (! (not @p_56) :named @p_59))) :rule implies_simplify) +(step t21.t4 (cl (= @p_53 @p_59)) :rule trans :premises (t21.t2 t21.t3)) +(step t21 (cl (! (= @p_55 (! (forall ((veriT_vr4 Exp$)) @p_59) :named @p_61)) :named @p_60)) :rule bind) +(step t22 (cl (not @p_60) (not @p_55) @p_61) :rule equiv_pos2) +(step t23 (cl @p_61) :rule th_resolution :premises (t20 t21 t22)) +(anchor :step t24 :args ((:= (veriT_vr4 Exp$) veriT_vr6))) +(anchor :step t24.t1 :args ((:= (veriT_vr5 FreeExp$) veriT_vr7))) +(step t24.t1.t1 (cl (= veriT_vr4 veriT_vr6)) :rule refl) +(step t24.t1.t2 (cl (= veriT_vr5 veriT_vr7)) :rule refl) +(step t24.t1.t3 (cl (= @p_43 (! (insert$ veriT_vr7 bot$) :named @p_63))) :rule cong :premises (t24.t1.t2)) +(step t24.t1.t4 (cl (= @p_45 (! (myImage$ exprel$ @p_63) :named @p_64))) :rule cong :premises (t24.t1.t3)) +(step t24.t1.t5 (cl (= @p_47 (! (abs_Exp$ @p_64) :named @p_65))) :rule cong :premises (t24.t1.t4)) +(step t24.t1.t6 (cl (= @p_49 (! (= veriT_vr6 @p_65) :named @p_66))) :rule cong :premises (t24.t1.t1 t24.t1.t5)) +(step t24.t1.t7 (cl (= @p_57 (! (not @p_66) :named @p_67))) :rule cong :premises (t24.t1.t6)) +(step t24.t1 (cl (= @p_56 (! (forall ((veriT_vr7 FreeExp$)) @p_67) :named @p_62))) :rule bind) +(step t24.t2 (cl (= @p_59 (! (not @p_62) :named @p_68))) :rule cong :premises (t24.t1)) +(step t24 (cl (! (= @p_61 (! (forall ((veriT_vr6 Exp$)) @p_68) :named @p_70)) :named @p_69)) :rule bind) +(step t25 (cl (not @p_69) (not @p_61) @p_70) :rule equiv_pos2) +(step t26 (cl @p_70) :rule th_resolution :premises (t23 t24 t25)) +(anchor :step t27 :args ((:= (?v0 Exp_list$) veriT_vr8) (:= (?v1 FreeExp_exp_fun$) veriT_vr9))) +(anchor :step t27.t1 :args ((:= (?v2 FreeExp_list$) veriT_vr10))) +(step t27.t1.t1 (cl (! (= ?v0 veriT_vr8) :named @p_80)) :rule refl) +(step t27.t1.t2 (cl (! (= ?v1 veriT_vr9) :named @p_88)) :rule refl) +(step t27.t1.t3 (cl (= ?v2 veriT_vr10)) :rule refl) +(step t27.t1.t4 (cl (= @p_74 (! (map2$ veriT_vr9 veriT_vr10) :named @p_75))) :rule cong :premises (t27.t1.t2 t27.t1.t3)) +(step t27.t1.t5 (cl (= @p_76 (! (= veriT_vr8 @p_75) :named @p_77))) :rule cong :premises (t27.t1.t1 t27.t1.t4)) +(step t27.t1 (cl (= @p_72 (! (exists ((veriT_vr10 FreeExp_list$)) @p_77) :named @p_73))) :rule bind) +(anchor :step t27.t2 :args ((:= (?v2 Exp$) veriT_vr11))) +(step t27.t2.t1 (cl (! (= ?v2 veriT_vr11) :named @p_87)) :rule refl) +(step t27.t2.t2 (cl @p_80) :rule refl) +(step t27.t2.t3 (cl (= @p_81 (! (myset$ veriT_vr8) :named @p_82))) :rule cong :premises (t27.t2.t2)) +(step t27.t2.t4 (cl (= @p_83 (! (member$ veriT_vr11 @p_82) :named @p_84))) :rule cong :premises (t27.t2.t1 t27.t2.t3)) +(anchor :step t27.t2.t5 :args ((:= (?v3 FreeExp$) veriT_vr12))) +(step t27.t2.t5.t1 (cl @p_87) :rule refl) +(step t27.t2.t5.t2 (cl @p_88) :rule refl) +(step t27.t2.t5.t3 (cl (= ?v3 veriT_vr12)) :rule refl) +(step t27.t2.t5.t4 (cl (= @p_89 (! (fun_app$ veriT_vr9 veriT_vr12) :named @p_90))) :rule cong :premises (t27.t2.t5.t2 t27.t2.t5.t3)) +(step t27.t2.t5.t5 (cl (= @p_91 (! (= veriT_vr11 @p_90) :named @p_92))) :rule cong :premises (t27.t2.t5.t1 t27.t2.t5.t4)) +(step t27.t2.t5 (cl (= @p_85 (! (exists ((veriT_vr12 FreeExp$)) @p_92) :named @p_86))) :rule bind) +(step t27.t2.t6 (cl (= @p_93 (! (=> @p_84 @p_86) :named @p_94))) :rule cong :premises (t27.t2.t4 t27.t2.t5)) +(step t27.t2 (cl (= @p_78 (! (forall ((veriT_vr11 Exp$)) @p_94) :named @p_79))) :rule bind) +(step t27.t3 (cl (= @p_95 (! (= @p_73 @p_79) :named @p_96))) :rule cong :premises (t27.t1 t27.t2)) +(step t27 (cl (! (= @p_71 (! (forall ((veriT_vr8 Exp_list$) (veriT_vr9 FreeExp_exp_fun$)) @p_96) :named @p_98)) :named @p_97)) :rule bind) +(step t28 (cl (not @p_97) (not @p_71) @p_98) :rule equiv_pos2) +(step t29 (cl @p_98) :rule th_resolution :premises (axiom3 t27 t28)) +(anchor :step t30 :args ((veriT_vr8 Exp_list$) (veriT_vr9 FreeExp_exp_fun$))) +(step t30.t1 (cl (= @p_96 (! (and (! (=> @p_73 @p_79) :named @p_115) (! (=> @p_79 @p_73) :named @p_128)) :named @p_99))) :rule connective_def) +(step t30 (cl (! (= @p_98 (! (forall ((veriT_vr8 Exp_list$) (veriT_vr9 FreeExp_exp_fun$)) @p_99) :named @p_101)) :named @p_100)) :rule bind) +(step t31 (cl (not @p_100) (not @p_98) @p_101) :rule equiv_pos2) +(step t32 (cl @p_101) :rule th_resolution :premises (t29 t30 t31)) +(anchor :step t33 :args ((:= (veriT_vr8 Exp_list$) veriT_vr13) (:= (veriT_vr9 FreeExp_exp_fun$) veriT_vr14))) +(anchor :step t33.t1 :args ((:= (veriT_vr10 FreeExp_list$) veriT_vr15))) +(step t33.t1.t1 (cl (! (= veriT_vr8 veriT_vr13) :named @p_107)) :rule refl) +(step t33.t1.t2 (cl (! (= veriT_vr9 veriT_vr14) :named @p_111)) :rule refl) +(step t33.t1.t3 (cl (= veriT_vr10 veriT_vr15)) :rule refl) +(step t33.t1.t4 (cl (= @p_75 (! (map2$ veriT_vr14 veriT_vr15) :named @p_104))) :rule cong :premises (t33.t1.t2 t33.t1.t3)) +(step t33.t1.t5 (cl (= @p_77 (! (= veriT_vr13 @p_104) :named @p_105))) :rule cong :premises (t33.t1.t1 t33.t1.t4)) +(step t33.t1 (cl (= @p_73 (! (exists ((veriT_vr15 FreeExp_list$)) @p_105) :named @p_103))) :rule bind) +(anchor :step t33.t2 :args ((:= (veriT_vr11 Exp$) veriT_vr16))) +(step t33.t2.t1 (cl (! (= veriT_vr11 veriT_vr16) :named @p_110)) :rule refl) +(step t33.t2.t2 (cl @p_107) :rule refl) +(step t33.t2.t3 (cl (! (= @p_82 (! (myset$ veriT_vr13) :named @p_102)) :named @p_118)) :rule cong :premises (t33.t2.t2)) +(step t33.t2.t4 (cl (= @p_84 (! (member$ veriT_vr16 @p_102) :named @p_108))) :rule cong :premises (t33.t2.t1 t33.t2.t3)) +(anchor :step t33.t2.t5 :args ((:= (veriT_vr12 FreeExp$) veriT_vr17))) +(step t33.t2.t5.t1 (cl @p_110) :rule refl) +(step t33.t2.t5.t2 (cl @p_111) :rule refl) +(step t33.t2.t5.t3 (cl (= veriT_vr12 veriT_vr17)) :rule refl) +(step t33.t2.t5.t4 (cl (= @p_90 (! (fun_app$ veriT_vr14 veriT_vr17) :named @p_112))) :rule cong :premises (t33.t2.t5.t2 t33.t2.t5.t3)) +(step t33.t2.t5.t5 (cl (= @p_92 (! (= veriT_vr16 @p_112) :named @p_113))) :rule cong :premises (t33.t2.t5.t1 t33.t2.t5.t4)) +(step t33.t2.t5 (cl (= @p_86 (! (exists ((veriT_vr17 FreeExp$)) @p_113) :named @p_109))) :rule bind) +(step t33.t2.t6 (cl (= @p_94 (! (=> @p_108 @p_109) :named @p_114))) :rule cong :premises (t33.t2.t4 t33.t2.t5)) +(step t33.t2 (cl (= @p_79 (! (forall ((veriT_vr16 Exp$)) @p_114) :named @p_106))) :rule bind) +(step t33.t3 (cl (= @p_115 (! (=> @p_103 @p_106) :named @p_116))) :rule cong :premises (t33.t1 t33.t2)) +(anchor :step t33.t4 :args ((:= (veriT_vr11 Exp$) veriT_vr18))) +(step t33.t4.t1 (cl (! (= veriT_vr11 veriT_vr18) :named @p_121)) :rule refl) +(step t33.t4.t2 (cl @p_107) :rule refl) +(step t33.t4.t3 (cl @p_118) :rule cong :premises (t33.t4.t2)) +(step t33.t4.t4 (cl (= @p_84 (! (member$ veriT_vr18 @p_102) :named @p_119))) :rule cong :premises (t33.t4.t1 t33.t4.t3)) +(anchor :step t33.t4.t5 :args ((:= (veriT_vr12 FreeExp$) veriT_vr19))) +(step t33.t4.t5.t1 (cl @p_121) :rule refl) +(step t33.t4.t5.t2 (cl @p_111) :rule refl) +(step t33.t4.t5.t3 (cl (= veriT_vr12 veriT_vr19)) :rule refl) +(step t33.t4.t5.t4 (cl (= @p_90 (! (fun_app$ veriT_vr14 veriT_vr19) :named @p_122))) :rule cong :premises (t33.t4.t5.t2 t33.t4.t5.t3)) +(step t33.t4.t5.t5 (cl (= @p_92 (! (= veriT_vr18 @p_122) :named @p_123))) :rule cong :premises (t33.t4.t5.t1 t33.t4.t5.t4)) +(step t33.t4.t5 (cl (= @p_86 (! (exists ((veriT_vr19 FreeExp$)) @p_123) :named @p_120))) :rule bind) +(step t33.t4.t6 (cl (= @p_94 (! (=> @p_119 @p_120) :named @p_124))) :rule cong :premises (t33.t4.t4 t33.t4.t5)) +(step t33.t4 (cl (= @p_79 (! (forall ((veriT_vr18 Exp$)) @p_124) :named @p_117))) :rule bind) +(anchor :step t33.t5 :args ((:= (veriT_vr10 FreeExp_list$) veriT_vr20))) +(step t33.t5.t1 (cl @p_107) :rule refl) +(step t33.t5.t2 (cl @p_111) :rule refl) +(step t33.t5.t3 (cl (= veriT_vr10 veriT_vr20)) :rule refl) +(step t33.t5.t4 (cl (= @p_75 (! (map2$ veriT_vr14 veriT_vr20) :named @p_126))) :rule cong :premises (t33.t5.t2 t33.t5.t3)) +(step t33.t5.t5 (cl (= @p_77 (! (= veriT_vr13 @p_126) :named @p_127))) :rule cong :premises (t33.t5.t1 t33.t5.t4)) +(step t33.t5 (cl (= @p_73 (! (exists ((veriT_vr20 FreeExp_list$)) @p_127) :named @p_125))) :rule bind) +(step t33.t6 (cl (= @p_128 (! (=> @p_117 @p_125) :named @p_129))) :rule cong :premises (t33.t4 t33.t5)) +(step t33.t7 (cl (= @p_99 (! (and @p_116 @p_129) :named @p_130))) :rule cong :premises (t33.t3 t33.t6)) +(step t33 (cl (! (= @p_101 (! (forall ((veriT_vr13 Exp_list$) (veriT_vr14 FreeExp_exp_fun$)) @p_130) :named @p_132)) :named @p_131)) :rule bind) +(step t34 (cl (not @p_131) (not @p_101) @p_132) :rule equiv_pos2) +(step t35 (cl @p_132) :rule th_resolution :premises (t32 t33 t34)) +(anchor :step t36 :args ((:= (veriT_vr13 Exp_list$) veriT_vr21) (:= (veriT_vr14 FreeExp_exp_fun$) veriT_vr22))) +(anchor :step t36.t1 :args ((:= (veriT_vr15 FreeExp_list$) veriT_vr23))) +(step t36.t1.t1 (cl (! (= veriT_vr13 veriT_vr21) :named @p_137)) :rule refl) +(step t36.t1.t2 (cl (! (= veriT_vr14 veriT_vr22) :named @p_142)) :rule refl) +(step t36.t1.t3 (cl (= veriT_vr15 veriT_vr23)) :rule refl) +(step t36.t1.t4 (cl (= @p_104 (! (map2$ veriT_vr22 veriT_vr23) :named @p_135))) :rule cong :premises (t36.t1.t2 t36.t1.t3)) +(step t36.t1.t5 (cl (= @p_105 (! (= veriT_vr21 @p_135) :named @p_136))) :rule cong :premises (t36.t1.t1 t36.t1.t4)) +(step t36.t1 (cl (= @p_103 (! (exists ((veriT_vr23 FreeExp_list$)) @p_136) :named @p_134))) :rule bind) +(anchor :step t36.t2 :args ((:= (veriT_vr16 Exp$) veriT_vr24))) +(step t36.t2.t1 (cl (! (= veriT_vr16 veriT_vr24) :named @p_141)) :rule refl) +(step t36.t2.t2 (cl @p_137) :rule refl) +(step t36.t2.t3 (cl (! (= @p_102 (! (myset$ veriT_vr21) :named @p_138)) :named @p_147)) :rule cong :premises (t36.t2.t2)) +(step t36.t2.t4 (cl (= @p_108 (! (member$ veriT_vr24 @p_138) :named @p_139))) :rule cong :premises (t36.t2.t1 t36.t2.t3)) +(anchor :step t36.t2.t5 :args ((:= (veriT_vr17 FreeExp$) veriT_vr25))) +(step t36.t2.t5.t1 (cl @p_141) :rule refl) +(step t36.t2.t5.t2 (cl @p_142) :rule refl) +(step t36.t2.t5.t3 (cl (= veriT_vr17 veriT_vr25)) :rule refl) +(step t36.t2.t5.t4 (cl (= @p_112 (! (fun_app$ veriT_vr22 veriT_vr25) :named @p_143))) :rule cong :premises (t36.t2.t5.t2 t36.t2.t5.t3)) +(step t36.t2.t5.t5 (cl (= @p_113 (! (= veriT_vr24 @p_143) :named @p_144))) :rule cong :premises (t36.t2.t5.t1 t36.t2.t5.t4)) +(step t36.t2.t5 (cl (= @p_109 (! (exists ((veriT_vr25 FreeExp$)) @p_144) :named @p_140))) :rule bind) +(step t36.t2.t6 (cl (= @p_114 (! (=> @p_139 @p_140) :named @p_145))) :rule cong :premises (t36.t2.t4 t36.t2.t5)) +(step t36.t2 (cl (= @p_106 (! (forall ((veriT_vr24 Exp$)) @p_145) :named @p_133))) :rule bind) +(step t36.t3 (cl (= @p_116 (! (=> @p_134 @p_133) :named @p_146))) :rule cong :premises (t36.t1 t36.t2)) +(anchor :step t36.t4 :args ((:= (veriT_vr18 Exp$) veriT_vr24))) +(step t36.t4.t1 (cl (! (= veriT_vr18 veriT_vr24) :named @p_148)) :rule refl) +(step t36.t4.t2 (cl @p_137) :rule refl) +(step t36.t4.t3 (cl @p_147) :rule cong :premises (t36.t4.t2)) +(step t36.t4.t4 (cl (= @p_119 @p_139)) :rule cong :premises (t36.t4.t1 t36.t4.t3)) +(anchor :step t36.t4.t5 :args ((:= (veriT_vr19 FreeExp$) veriT_vr25))) +(step t36.t4.t5.t1 (cl @p_148) :rule refl) +(step t36.t4.t5.t2 (cl @p_142) :rule refl) +(step t36.t4.t5.t3 (cl (= veriT_vr19 veriT_vr25)) :rule refl) +(step t36.t4.t5.t4 (cl (= @p_122 @p_143)) :rule cong :premises (t36.t4.t5.t2 t36.t4.t5.t3)) +(step t36.t4.t5.t5 (cl (= @p_123 @p_144)) :rule cong :premises (t36.t4.t5.t1 t36.t4.t5.t4)) +(step t36.t4.t5 (cl (= @p_120 @p_140)) :rule bind) +(step t36.t4.t6 (cl (= @p_124 @p_145)) :rule cong :premises (t36.t4.t4 t36.t4.t5)) +(step t36.t4 (cl (= @p_117 @p_133)) :rule bind) +(anchor :step t36.t5 :args ((:= (veriT_vr20 FreeExp_list$) veriT_vr23))) +(step t36.t5.t1 (cl @p_137) :rule refl) +(step t36.t5.t2 (cl @p_142) :rule refl) +(step t36.t5.t3 (cl (= veriT_vr20 veriT_vr23)) :rule refl) +(step t36.t5.t4 (cl (= @p_126 @p_135)) :rule cong :premises (t36.t5.t2 t36.t5.t3)) +(step t36.t5.t5 (cl (= @p_127 @p_136)) :rule cong :premises (t36.t5.t1 t36.t5.t4)) +(step t36.t5 (cl (= @p_125 @p_134)) :rule bind) +(step t36.t6 (cl (= @p_129 (! (=> @p_133 @p_134) :named @p_149))) :rule cong :premises (t36.t4 t36.t5)) +(step t36.t7 (cl (= @p_130 (! (and @p_146 @p_149) :named @p_150))) :rule cong :premises (t36.t3 t36.t6)) +(step t36 (cl (! (= @p_132 (! (forall ((veriT_vr21 Exp_list$) (veriT_vr22 FreeExp_exp_fun$)) @p_150) :named @p_152)) :named @p_151)) :rule bind) +(step t37 (cl (not @p_151) (not @p_132) @p_152) :rule equiv_pos2) +(step t38 (cl @p_152) :rule th_resolution :premises (t35 t36 t37)) +(anchor :step t39 :args ((:= (veriT_vr21 Exp_list$) veriT_vr21) (:= (veriT_vr22 FreeExp_exp_fun$) veriT_vr22))) +(anchor :step t39.t1 :args ((:= (veriT_vr24 Exp$) veriT_vr26))) +(step t39.t1.t1 (cl (! (= veriT_vr24 veriT_vr26) :named @p_156)) :rule refl) +(step t39.t1.t2 (cl (= @p_139 (! (member$ veriT_vr26 @p_138) :named @p_154))) :rule cong :premises (t39.t1.t1)) +(anchor :step t39.t1.t3 :args ((:= (veriT_vr25 FreeExp$) veriT_vr27))) +(step t39.t1.t3.t1 (cl @p_156) :rule refl) +(step t39.t1.t3.t2 (cl (= veriT_vr25 veriT_vr27)) :rule refl) +(step t39.t1.t3.t3 (cl (= @p_143 (! (fun_app$ veriT_vr22 veriT_vr27) :named @p_157))) :rule cong :premises (t39.t1.t3.t2)) +(step t39.t1.t3.t4 (cl (= @p_144 (! (= veriT_vr26 @p_157) :named @p_158))) :rule cong :premises (t39.t1.t3.t1 t39.t1.t3.t3)) +(step t39.t1.t3 (cl (= @p_140 (! (exists ((veriT_vr27 FreeExp$)) @p_158) :named @p_155))) :rule bind) +(step t39.t1.t4 (cl (= @p_145 (! (=> @p_154 @p_155) :named @p_159))) :rule cong :premises (t39.t1.t2 t39.t1.t3)) +(step t39.t1 (cl (= @p_133 (! (forall ((veriT_vr26 Exp$)) @p_159) :named @p_153))) :rule bind) +(anchor :step t39.t2 :args ((:= (veriT_vr23 FreeExp_list$) veriT_vr28))) +(step t39.t2.t1 (cl (= veriT_vr23 veriT_vr28)) :rule refl) +(step t39.t2.t2 (cl (= @p_135 (! (map2$ veriT_vr22 veriT_vr28) :named @p_161))) :rule cong :premises (t39.t2.t1)) +(step t39.t2.t3 (cl (= @p_136 (! (= veriT_vr21 @p_161) :named @p_162))) :rule cong :premises (t39.t2.t2)) +(step t39.t2 (cl (= @p_134 (! (exists ((veriT_vr28 FreeExp_list$)) @p_162) :named @p_160))) :rule bind) +(step t39.t3 (cl (= @p_149 (! (=> @p_153 @p_160) :named @p_163))) :rule cong :premises (t39.t1 t39.t2)) +(step t39.t4 (cl (= @p_150 (! (and @p_146 @p_163) :named @p_164))) :rule cong :premises (t39.t3)) +(step t39 (cl (! (= @p_152 (! (forall ((veriT_vr21 Exp_list$) (veriT_vr22 FreeExp_exp_fun$)) @p_164) :named @p_166)) :named @p_165)) :rule bind) +(step t40 (cl (not @p_165) (not @p_152) @p_166) :rule equiv_pos2) +(step t41 (cl @p_166) :rule th_resolution :premises (t38 t39 t40)) +(anchor :step t42 :args ((veriT_vr21 Exp_list$) (veriT_vr22 FreeExp_exp_fun$))) +(step t42.t1 (cl (= @p_134 (! (not (forall ((veriT_vr23 FreeExp_list$)) (not @p_136))) :named @p_167))) :rule connective_def) +(step t42.t2 (cl (= @p_146 (! (=> @p_167 @p_133) :named @p_168))) :rule cong :premises (t42.t1)) +(anchor :step t42.t3 :args ((veriT_vr26 Exp$))) +(step t42.t3.t1 (cl (= @p_155 (! (not (forall ((veriT_vr27 FreeExp$)) (not @p_158))) :named @p_170))) :rule connective_def) +(step t42.t3.t2 (cl (= @p_159 (! (=> @p_154 @p_170) :named @p_171))) :rule cong :premises (t42.t3.t1)) +(step t42.t3 (cl (= @p_153 (! (forall ((veriT_vr26 Exp$)) @p_171) :named @p_169))) :rule bind) +(step t42.t4 (cl (= @p_163 (! (=> @p_169 @p_160) :named @p_172))) :rule cong :premises (t42.t3)) +(step t42.t5 (cl (= @p_164 (! (and @p_168 @p_172) :named @p_173))) :rule cong :premises (t42.t2 t42.t4)) +(step t42 (cl (! (= @p_166 (! (forall ((veriT_vr21 Exp_list$) (veriT_vr22 FreeExp_exp_fun$)) @p_173) :named @p_175)) :named @p_174)) :rule bind) +(step t43 (cl (not @p_174) (not @p_166) @p_175) :rule equiv_pos2) +(step t44 (cl @p_175) :rule th_resolution :premises (t41 t42 t43)) +(anchor :step t45 :args ((:= (?v0 FreeExp_list$) veriT_vr29))) +(step t45.t1 (cl (= ?v0 veriT_vr29)) :rule refl) +(step t45.t2 (cl (= @p_1 (! (abs_ExpList$ veriT_vr29) :named @p_177))) :rule cong :premises (t45.t1)) +(step t45.t3 (cl (= @p_178 (! (= z$ @p_177) :named @p_179))) :rule cong :premises (t45.t2)) +(step t45 (cl (= @p_176 (! (exists ((veriT_vr29 FreeExp_list$)) @p_179) :named @p_181))) :rule bind) +(step t46 (cl (! (= @p_180 (! (not @p_181) :named @p_183)) :named @p_182)) :rule cong :premises (t45)) +(step t47 (cl (! (not @p_182) :named @p_185) (! (not @p_180) :named @p_184) @p_183) :rule equiv_pos2) +(step t48 (cl (not @p_184) @p_176) :rule not_not) +(step t49 (cl @p_185 @p_176 @p_183) :rule th_resolution :premises (t48 t47)) +(step t50 (cl @p_183) :rule th_resolution :premises (axiom4 t46 t49)) +(anchor :step t51 :args ((:= (veriT_vr29 FreeExp_list$) veriT_vr30))) +(step t51.t1 (cl (= veriT_vr29 veriT_vr30)) :rule refl) +(step t51.t2 (cl (= @p_177 (! (abs_ExpList$ veriT_vr30) :named @p_186))) :rule cong :premises (t51.t1)) +(step t51.t3 (cl (= @p_179 (! (= z$ @p_186) :named @p_187))) :rule cong :premises (t51.t2)) +(step t51 (cl (= @p_181 (! (exists ((veriT_vr30 FreeExp_list$)) @p_187) :named @p_188))) :rule bind) +(step t52 (cl (! (= @p_183 (! (not @p_188) :named @p_190)) :named @p_189)) :rule cong :premises (t51)) +(step t53 (cl (! (not @p_189) :named @p_192) (! (not @p_183) :named @p_191) @p_190) :rule equiv_pos2) +(step t54 (cl (not @p_191) @p_181) :rule not_not) +(step t55 (cl @p_192 @p_181 @p_190) :rule th_resolution :premises (t54 t53)) +(step t56 (cl @p_190) :rule th_resolution :premises (t50 t52 t55)) +(step t57 (cl (= @p_188 (! (not (! (forall ((veriT_vr30 FreeExp_list$)) (not @p_187)) :named @p_198)) :named @p_193))) :rule connective_def) +(step t58 (cl (! (= @p_190 (! (not @p_193) :named @p_195)) :named @p_194)) :rule cong :premises (t57)) +(step t59 (cl (! (not @p_194) :named @p_197) (! (not @p_190) :named @p_196) @p_195) :rule equiv_pos2) +(step t60 (cl (not @p_196) @p_188) :rule not_not) +(step t61 (cl @p_197 @p_188 @p_195) :rule th_resolution :premises (t60 t59)) +(step t62 (cl (not @p_195) @p_198) :rule not_not) +(step t63 (cl @p_197 @p_188 @p_198) :rule th_resolution :premises (t62 t61)) +(step t64 (cl @p_195) :rule th_resolution :premises (t56 t58 t63)) +(step t65 (cl @p_198) :rule th_resolution :premises (t62 t64)) +(step t66 (cl (or (! (not @p_175) :named @p_336) (! (and (! (=> (! (not (! (forall ((veriT_vr23 FreeExp_list$)) (! (not (! (= z$ (! (map2$ uu$ veriT_vr23) :named @p_203)) :named @p_205)) :named @p_207)) :named @p_202)) :named @p_209) (! (forall ((veriT_vr24 Exp$)) (! (=> (! (member$ veriT_vr24 @p_199) :named @p_212) (! (exists ((veriT_vr25 FreeExp$)) (! (= veriT_vr24 (! (fun_app$ uu$ veriT_vr25) :named @p_217)) :named @p_219)) :named @p_214)) :named @p_221)) :named @p_211)) :named @p_223) (! (=> (! (forall ((veriT_vr26 Exp$)) (! (=> (! (member$ veriT_vr26 @p_199) :named @p_227) (! (not (! (forall ((veriT_vr27 FreeExp$)) (! (not (! (= veriT_vr26 (! (fun_app$ uu$ veriT_vr27) :named @p_231)) :named @p_232)) :named @p_233)) :named @p_228)) :named @p_235)) :named @p_237)) :named @p_226) (! (exists ((veriT_vr28 FreeExp_list$)) (! (= z$ (! (map2$ uu$ veriT_vr28) :named @p_240)) :named @p_241)) :named @p_239)) :named @p_242)) :named @p_200))) :rule forall_inst :args ((:= veriT_vr21 z$) (:= veriT_vr22 uu$))) +(anchor :step t67) +(assume t67.h1 @p_200) +(anchor :step t67.t2 :args ((:= (veriT_vr23 FreeExp_list$) veriT_vr31))) +(step t67.t2.t1 (cl (= veriT_vr23 veriT_vr31)) :rule refl) +(step t67.t2.t2 (cl (= @p_203 (! (map2$ uu$ veriT_vr31) :named @p_204))) :rule cong :premises (t67.t2.t1)) +(step t67.t2.t3 (cl (= @p_205 (! (= z$ @p_204) :named @p_206))) :rule cong :premises (t67.t2.t2)) +(step t67.t2.t4 (cl (= @p_207 (! (not @p_206) :named @p_208))) :rule cong :premises (t67.t2.t3)) +(step t67.t2 (cl (= @p_202 (! (forall ((veriT_vr31 FreeExp_list$)) @p_208) :named @p_210))) :rule bind) +(step t67.t3 (cl (= @p_209 (! (not @p_210) :named @p_224))) :rule cong :premises (t67.t2)) +(anchor :step t67.t4 :args ((:= (veriT_vr24 Exp$) veriT_vr32))) +(step t67.t4.t1 (cl (! (= veriT_vr24 veriT_vr32) :named @p_216)) :rule refl) +(step t67.t4.t2 (cl (= @p_212 (! (member$ veriT_vr32 @p_199) :named @p_213))) :rule cong :premises (t67.t4.t1)) +(anchor :step t67.t4.t3 :args ((:= (veriT_vr25 FreeExp$) veriT_vr33))) +(step t67.t4.t3.t1 (cl @p_216) :rule refl) +(step t67.t4.t3.t2 (cl (= veriT_vr25 veriT_vr33)) :rule refl) +(step t67.t4.t3.t3 (cl (= @p_217 (! (fun_app$ uu$ veriT_vr33) :named @p_218))) :rule cong :premises (t67.t4.t3.t2)) +(step t67.t4.t3.t4 (cl (= @p_219 (! (= veriT_vr32 @p_218) :named @p_220))) :rule cong :premises (t67.t4.t3.t1 t67.t4.t3.t3)) +(step t67.t4.t3 (cl (= @p_214 (! (exists ((veriT_vr33 FreeExp$)) @p_220) :named @p_215))) :rule bind) +(step t67.t4.t4 (cl (= @p_221 (! (=> @p_213 @p_215) :named @p_222))) :rule cong :premises (t67.t4.t2 t67.t4.t3)) +(step t67.t4 (cl (= @p_211 (! (forall ((veriT_vr32 Exp$)) @p_222) :named @p_225))) :rule bind) +(step t67.t5 (cl (= @p_223 (! (=> @p_224 @p_225) :named @p_245))) :rule cong :premises (t67.t3 t67.t4)) +(anchor :step t67.t6 :args ((:= (veriT_vr26 Exp$) veriT_vr32))) +(step t67.t6.t1 (cl (! (= veriT_vr26 veriT_vr32) :named @p_230)) :rule refl) +(step t67.t6.t2 (cl (= @p_227 @p_213)) :rule cong :premises (t67.t6.t1)) +(anchor :step t67.t6.t3 :args ((:= (veriT_vr27 FreeExp$) veriT_vr33))) +(step t67.t6.t3.t1 (cl @p_230) :rule refl) +(step t67.t6.t3.t2 (cl (= veriT_vr27 veriT_vr33)) :rule refl) +(step t67.t6.t3.t3 (cl (= @p_231 @p_218)) :rule cong :premises (t67.t6.t3.t2)) +(step t67.t6.t3.t4 (cl (= @p_232 @p_220)) :rule cong :premises (t67.t6.t3.t1 t67.t6.t3.t3)) +(step t67.t6.t3.t5 (cl (= @p_233 (! (not @p_220) :named @p_234))) :rule cong :premises (t67.t6.t3.t4)) +(step t67.t6.t3 (cl (= @p_228 (! (forall ((veriT_vr33 FreeExp$)) @p_234) :named @p_229))) :rule bind) +(step t67.t6.t4 (cl (= @p_235 (! (not @p_229) :named @p_236))) :rule cong :premises (t67.t6.t3)) +(step t67.t6.t5 (cl (= @p_237 (! (=> @p_213 @p_236) :named @p_238))) :rule cong :premises (t67.t6.t2 t67.t6.t4)) +(step t67.t6 (cl (= @p_226 (! (forall ((veriT_vr32 Exp$)) @p_238) :named @p_243))) :rule bind) +(anchor :step t67.t7 :args ((:= (veriT_vr28 FreeExp_list$) veriT_vr31))) +(step t67.t7.t1 (cl (= veriT_vr28 veriT_vr31)) :rule refl) +(step t67.t7.t2 (cl (= @p_240 @p_204)) :rule cong :premises (t67.t7.t1)) +(step t67.t7.t3 (cl (= @p_241 @p_206)) :rule cong :premises (t67.t7.t2)) +(step t67.t7 (cl (= @p_239 (! (exists ((veriT_vr31 FreeExp_list$)) @p_206) :named @p_244))) :rule bind) +(step t67.t8 (cl (= @p_242 (! (=> @p_243 @p_244) :named @p_246))) :rule cong :premises (t67.t6 t67.t7)) +(step t67.t9 (cl (! (= @p_200 (! (and @p_245 @p_246) :named @p_249)) :named @p_247)) :rule cong :premises (t67.t5 t67.t8)) +(step t67.t10 (cl (not @p_247) (! (not @p_200) :named @p_248) @p_249) :rule equiv_pos2) +(step t67.t11 (cl @p_249) :rule th_resolution :premises (t67.h1 t67.t9 t67.t10)) +(anchor :step t67.t12 :args ((:= (veriT_vr32 Exp$) veriT_vr34))) +(step t67.t12.t1 (cl (! (= veriT_vr32 veriT_vr34) :named @p_252)) :rule refl) +(step t67.t12.t2 (cl (= @p_213 (! (member$ veriT_vr34 @p_199) :named @p_250))) :rule cong :premises (t67.t12.t1)) +(anchor :step t67.t12.t3 :args ((:= (veriT_vr33 FreeExp$) veriT_vr35))) +(step t67.t12.t3.t1 (cl @p_252) :rule refl) +(step t67.t12.t3.t2 (cl (= veriT_vr33 veriT_vr35)) :rule refl) +(step t67.t12.t3.t3 (cl (= @p_218 (! (fun_app$ uu$ veriT_vr35) :named @p_253))) :rule cong :premises (t67.t12.t3.t2)) +(step t67.t12.t3.t4 (cl (= @p_220 (! (= veriT_vr34 @p_253) :named @p_254))) :rule cong :premises (t67.t12.t3.t1 t67.t12.t3.t3)) +(step t67.t12.t3.t5 (cl (= @p_234 (! (not @p_254) :named @p_255))) :rule cong :premises (t67.t12.t3.t4)) +(step t67.t12.t3 (cl (= @p_229 (! (forall ((veriT_vr35 FreeExp$)) @p_255) :named @p_251))) :rule bind) +(step t67.t12.t4 (cl (= @p_236 (! (not @p_251) :named @p_256))) :rule cong :premises (t67.t12.t3)) +(step t67.t12.t5 (cl (= @p_238 (! (=> @p_250 @p_256) :named @p_257))) :rule cong :premises (t67.t12.t2 t67.t12.t4)) +(step t67.t12 (cl (= @p_243 (! (forall ((veriT_vr34 Exp$)) @p_257) :named @p_260))) :rule bind) +(anchor :step t67.t13 :args ((:= (veriT_vr31 FreeExp_list$) veriT_vr36))) +(step t67.t13.t1 (cl (= veriT_vr31 veriT_vr36)) :rule refl) +(step t67.t13.t2 (cl (= @p_204 (! (map2$ uu$ veriT_vr36) :named @p_258))) :rule cong :premises (t67.t13.t1)) +(step t67.t13.t3 (cl (= @p_206 (! (= z$ @p_258) :named @p_259))) :rule cong :premises (t67.t13.t2)) +(step t67.t13 (cl (= @p_244 (! (exists ((veriT_vr36 FreeExp_list$)) @p_259) :named @p_261))) :rule bind) +(step t67.t14 (cl (= @p_246 (! (=> @p_260 @p_261) :named @p_262))) :rule cong :premises (t67.t12 t67.t13)) +(step t67.t15 (cl (! (= @p_249 (! (and @p_245 @p_262) :named @p_264)) :named @p_263)) :rule cong :premises (t67.t14)) +(step t67.t16 (cl (not @p_263) (not @p_249) @p_264) :rule equiv_pos2) +(step t67.t17 (cl @p_264) :rule th_resolution :premises (t67.t11 t67.t15 t67.t16)) +(anchor :step t67.t18 :args ((:= (veriT_vr31 FreeExp_list$) veriT_vr37))) +(step t67.t18.t1 (cl (= veriT_vr31 veriT_vr37)) :rule refl) +(step t67.t18.t2 (cl (= @p_204 (! (map2$ uu$ veriT_vr37) :named @p_265))) :rule cong :premises (t67.t18.t1)) +(step t67.t18.t3 (cl (= @p_206 (! (= z$ @p_265) :named @p_266))) :rule cong :premises (t67.t18.t2)) +(step t67.t18.t4 (cl (= @p_208 (! (not @p_266) :named @p_267))) :rule cong :premises (t67.t18.t3)) +(step t67.t18 (cl (= @p_210 (! (forall ((veriT_vr37 FreeExp_list$)) @p_267) :named @p_268))) :rule bind) +(step t67.t19 (cl (= @p_224 (! (not @p_268) :named @p_275))) :rule cong :premises (t67.t18)) +(anchor :step t67.t20 :args ((:= (veriT_vr32 Exp$) veriT_vr38))) +(step t67.t20.t1 (cl (! (= veriT_vr32 veriT_vr38) :named @p_271)) :rule refl) +(step t67.t20.t2 (cl (= @p_213 (! (member$ veriT_vr38 @p_199) :named @p_269))) :rule cong :premises (t67.t20.t1)) +(anchor :step t67.t20.t3 :args ((:= (veriT_vr33 FreeExp$) veriT_vr39))) +(step t67.t20.t3.t1 (cl @p_271) :rule refl) +(step t67.t20.t3.t2 (cl (= veriT_vr33 veriT_vr39)) :rule refl) +(step t67.t20.t3.t3 (cl (= @p_218 (! (fun_app$ uu$ veriT_vr39) :named @p_272))) :rule cong :premises (t67.t20.t3.t2)) +(step t67.t20.t3.t4 (cl (= @p_220 (! (= veriT_vr38 @p_272) :named @p_273))) :rule cong :premises (t67.t20.t3.t1 t67.t20.t3.t3)) +(step t67.t20.t3 (cl (= @p_215 (! (exists ((veriT_vr39 FreeExp$)) @p_273) :named @p_270))) :rule bind) +(step t67.t20.t4 (cl (= @p_222 (! (=> @p_269 @p_270) :named @p_274))) :rule cong :premises (t67.t20.t2 t67.t20.t3)) +(step t67.t20 (cl (= @p_225 (! (forall ((veriT_vr38 Exp$)) @p_274) :named @p_276))) :rule bind) +(step t67.t21 (cl (= @p_245 (! (=> @p_275 @p_276) :named @p_289))) :rule cong :premises (t67.t19 t67.t20)) +(anchor :step t67.t22 :args ((:= (veriT_vr34 Exp$) veriT_vr40))) +(step t67.t22.t1 (cl (! (= veriT_vr34 veriT_vr40) :named @p_280)) :rule refl) +(step t67.t22.t2 (cl (= @p_250 @p_278)) :rule cong :premises (t67.t22.t1)) +(anchor :step t67.t22.t3 :args ((:= (veriT_vr35 FreeExp$) veriT_vr41))) +(step t67.t22.t3.t1 (cl @p_280) :rule refl) +(step t67.t22.t3.t2 (cl (= veriT_vr35 veriT_vr41)) :rule refl) +(step t67.t22.t3.t3 (cl (= @p_253 @p_281)) :rule cong :premises (t67.t22.t3.t2)) +(step t67.t22.t3.t4 (cl (= @p_254 @p_282)) :rule cong :premises (t67.t22.t3.t1 t67.t22.t3.t3)) +(step t67.t22.t3.t5 (cl (= @p_255 @p_283)) :rule cong :premises (t67.t22.t3.t4)) +(step t67.t22.t3 (cl (= @p_251 @p_279)) :rule bind) +(step t67.t22.t4 (cl (= @p_256 @p_284)) :rule cong :premises (t67.t22.t3)) +(step t67.t22.t5 (cl (= @p_257 @p_277)) :rule cong :premises (t67.t22.t2 t67.t22.t4)) +(step t67.t22 (cl (= @p_260 (! (forall ((veriT_vr40 Exp$)) @p_277) :named @p_287))) :rule bind) +(anchor :step t67.t23 :args ((:= (veriT_vr36 FreeExp_list$) veriT_vr42))) +(step t67.t23.t1 (cl (= veriT_vr36 veriT_vr42)) :rule refl) +(step t67.t23.t2 (cl (= @p_258 @p_286)) :rule cong :premises (t67.t23.t1)) +(step t67.t23.t3 (cl (= @p_259 @p_285)) :rule cong :premises (t67.t23.t2)) +(step t67.t23 (cl (= @p_261 (! (exists ((veriT_vr42 FreeExp_list$)) @p_285) :named @p_288))) :rule bind) +(step t67.t24 (cl (= @p_262 (! (=> @p_287 @p_288) :named @p_290))) :rule cong :premises (t67.t22 t67.t23)) +(step t67.t25 (cl (! (= @p_264 (! (and @p_289 @p_290) :named @p_292)) :named @p_291)) :rule cong :premises (t67.t21 t67.t24)) +(step t67.t26 (cl (not @p_291) (not @p_264) @p_292) :rule equiv_pos2) +(step t67.t27 (cl @p_292) :rule th_resolution :premises (t67.t17 t67.t25 t67.t26)) +(anchor :step t67.t28 :args ((:= (veriT_vr40 Exp$) veriT_sk0))) +(step t67.t28.t1 (cl (! (= veriT_vr40 veriT_sk0) :named @p_295)) :rule refl) +(step t67.t28.t2 (cl (= @p_278 (! (member$ veriT_sk0 @p_199) :named @p_293))) :rule cong :premises (t67.t28.t1)) +(anchor :step t67.t28.t3 :args ((veriT_vr41 FreeExp$))) +(step t67.t28.t3.t1 (cl @p_295) :rule refl) +(step t67.t28.t3.t2 (cl (= @p_282 (! (= @p_281 veriT_sk0) :named @p_296))) :rule cong :premises (t67.t28.t3.t1)) +(step t67.t28.t3.t3 (cl (= @p_283 (! (not @p_296) :named @p_297))) :rule cong :premises (t67.t28.t3.t2)) +(step t67.t28.t3 (cl (= @p_279 (! (forall ((veriT_vr41 FreeExp$)) @p_297) :named @p_294))) :rule bind) +(step t67.t28.t4 (cl (= @p_284 (! (not @p_294) :named @p_298))) :rule cong :premises (t67.t28.t3)) +(step t67.t28.t5 (cl (= @p_277 (! (=> @p_293 @p_298) :named @p_299))) :rule cong :premises (t67.t28.t2 t67.t28.t4)) +(step t67.t28 (cl (= @p_287 @p_299)) :rule sko_forall) +(anchor :step t67.t29 :args ((:= (veriT_vr42 FreeExp_list$) veriT_sk1))) +(step t67.t29.t1 (cl (= veriT_vr42 veriT_sk1)) :rule refl) +(step t67.t29.t2 (cl (= @p_286 (! (map2$ uu$ veriT_sk1) :named @p_302))) :rule cong :premises (t67.t29.t1)) +(step t67.t29.t3 (cl (= @p_285 (! (= z$ @p_302) :named @p_300))) :rule cong :premises (t67.t29.t2)) +(step t67.t29 (cl (= @p_288 @p_300)) :rule sko_ex) +(step t67.t30 (cl (= @p_290 (! (=> @p_299 @p_300) :named @p_303))) :rule cong :premises (t67.t28 t67.t29)) +(step t67.t31 (cl (! (= @p_292 (! (and @p_289 @p_303) :named @p_305)) :named @p_304)) :rule cong :premises (t67.t30)) +(step t67.t32 (cl (not @p_304) (not @p_292) @p_305) :rule equiv_pos2) +(step t67.t33 (cl @p_305) :rule th_resolution :premises (t67.t27 t67.t31 t67.t32)) +(anchor :step t67.t34 :args ((:= (veriT_vr37 FreeExp_list$) veriT_vr43))) +(step t67.t34.t1 (cl (= veriT_vr37 veriT_vr43)) :rule refl) +(step t67.t34.t2 (cl (= @p_265 (! (map2$ uu$ veriT_vr43) :named @p_307))) :rule cong :premises (t67.t34.t1)) +(step t67.t34.t3 (cl (= @p_266 (! (= z$ @p_307) :named @p_308))) :rule cong :premises (t67.t34.t2)) +(step t67.t34.t4 (cl (= @p_267 (! (not @p_308) :named @p_309))) :rule cong :premises (t67.t34.t3)) +(step t67.t34 (cl (= @p_268 (! (forall ((veriT_vr43 FreeExp_list$)) @p_309) :named @p_306))) :rule bind) +(step t67.t35 (cl (= @p_275 (! (not @p_306) :named @p_310))) :rule cong :premises (t67.t34)) +(anchor :step t67.t36 :args ((:= (veriT_vr38 Exp$) veriT_vr44))) +(step t67.t36.t1 (cl (! (= veriT_vr38 veriT_vr44) :named @p_314)) :rule refl) +(step t67.t36.t2 (cl (= @p_269 (! (member$ veriT_vr44 @p_199) :named @p_312))) :rule cong :premises (t67.t36.t1)) +(anchor :step t67.t36.t3 :args ((:= (veriT_vr39 FreeExp$) veriT_vr45))) +(step t67.t36.t3.t1 (cl @p_314) :rule refl) +(step t67.t36.t3.t2 (cl (= veriT_vr39 veriT_vr45)) :rule refl) +(step t67.t36.t3.t3 (cl (= @p_272 (! (fun_app$ uu$ veriT_vr45) :named @p_315))) :rule cong :premises (t67.t36.t3.t2)) +(step t67.t36.t3.t4 (cl (= @p_273 (! (= veriT_vr44 @p_315) :named @p_316))) :rule cong :premises (t67.t36.t3.t1 t67.t36.t3.t3)) +(step t67.t36.t3 (cl (= @p_270 (! (exists ((veriT_vr45 FreeExp$)) @p_316) :named @p_313))) :rule bind) +(step t67.t36.t4 (cl (= @p_274 (! (=> @p_312 @p_313) :named @p_317))) :rule cong :premises (t67.t36.t2 t67.t36.t3)) +(step t67.t36 (cl (= @p_276 (! (forall ((veriT_vr44 Exp$)) @p_317) :named @p_311))) :rule bind) +(step t67.t37 (cl (= @p_289 (! (=> @p_310 @p_311) :named @p_318))) :rule cong :premises (t67.t35 t67.t36)) +(anchor :step t67.t38 :args ((:= (veriT_vr41 FreeExp$) veriT_vr45))) +(step t67.t38.t1 (cl (= veriT_vr41 veriT_vr45)) :rule refl) +(step t67.t38.t2 (cl (= @p_281 @p_315)) :rule cong :premises (t67.t38.t1)) +(step t67.t38.t3 (cl (= @p_296 (! (= veriT_sk0 @p_315) :named @p_319))) :rule cong :premises (t67.t38.t2)) +(step t67.t38.t4 (cl (= @p_297 (! (not @p_319) :named @p_320))) :rule cong :premises (t67.t38.t3)) +(step t67.t38 (cl (= @p_294 (! (forall ((veriT_vr45 FreeExp$)) @p_320) :named @p_321))) :rule bind) +(step t67.t39 (cl (= @p_298 (! (not @p_321) :named @p_322))) :rule cong :premises (t67.t38)) +(step t67.t40 (cl (= @p_299 (! (=> @p_293 @p_322) :named @p_323))) :rule cong :premises (t67.t39)) +(step t67.t41 (cl (= @p_303 (! (=> @p_323 @p_300) :named @p_324))) :rule cong :premises (t67.t40)) +(step t67.t42 (cl (! (= @p_305 (! (and @p_318 @p_324) :named @p_326)) :named @p_325)) :rule cong :premises (t67.t37 t67.t41)) +(step t67.t43 (cl (not @p_325) (not @p_305) @p_326) :rule equiv_pos2) +(step t67.t44 (cl @p_326) :rule th_resolution :premises (t67.t33 t67.t42 t67.t43)) +(anchor :step t67.t45 :args ((:= (veriT_vr45 FreeExp$) veriT_vr46))) +(step t67.t45.t1 (cl (= veriT_vr45 veriT_vr46)) :rule refl) +(step t67.t45.t2 (cl (= @p_315 (! (fun_app$ uu$ veriT_vr46) :named @p_328))) :rule cong :premises (t67.t45.t1)) +(step t67.t45.t3 (cl (= @p_319 (! (= veriT_sk0 @p_328) :named @p_329))) :rule cong :premises (t67.t45.t2)) +(step t67.t45.t4 (cl (= @p_320 (! (not @p_329) :named @p_330))) :rule cong :premises (t67.t45.t3)) +(step t67.t45 (cl (= @p_321 (! (forall ((veriT_vr46 FreeExp$)) @p_330) :named @p_327))) :rule bind) +(step t67.t46 (cl (= @p_322 (! (not @p_327) :named @p_331))) :rule cong :premises (t67.t45)) +(step t67.t47 (cl (= @p_323 (! (=> @p_293 @p_331) :named @p_332))) :rule cong :premises (t67.t46)) +(step t67.t48 (cl (= @p_324 (! (=> @p_332 @p_300) :named @p_333))) :rule cong :premises (t67.t47)) +(step t67.t49 (cl (! (= @p_326 (! (and @p_318 @p_333) :named @p_334)) :named @p_335)) :rule cong :premises (t67.t48)) +(step t67.t50 (cl (not @p_335) (not @p_326) @p_334) :rule equiv_pos2) +(step t67.t51 (cl @p_334) :rule th_resolution :premises (t67.t44 t67.t49 t67.t50)) +(step t67 (cl @p_248 @p_334) :rule subproof :discharge (h1)) +(step t68 (cl @p_336 @p_200) :rule or :premises (t66)) +(step t69 (cl (! (or @p_336 @p_334) :named @p_338) (! (not @p_336) :named @p_337)) :rule or_neg) +(step t70 (cl (not @p_337) @p_175) :rule not_not) +(step t71 (cl @p_338 @p_175) :rule th_resolution :premises (t70 t69)) +(step t72 (cl @p_338 (! (not @p_334) :named @p_340)) :rule or_neg) +(step t73 (cl @p_338) :rule th_resolution :premises (t68 t67 t71 t72)) +(step t74 (cl @p_332 (! (not @p_331) :named @p_339)) :rule implies_neg2) +(step t75 (cl (not @p_339) @p_327) :rule not_not) +(step t76 (cl @p_332 @p_327) :rule th_resolution :premises (t75 t74)) +(step t77 (cl (not @p_333) (! (not @p_332) :named @p_393) @p_300) :rule implies_pos) +(step t78 (cl @p_340 @p_333) :rule and_pos) +(step t79 (cl @p_336 @p_334) :rule or :premises (t73)) +(step t80 (cl @p_334) :rule resolution :premises (t79 t44)) +(step t81 (cl @p_333) :rule resolution :premises (t78 t80)) +(step t82 (cl (or (! (not @p_70) :named @p_377) (! (not (! (forall ((veriT_vr7 FreeExp$)) (! (not (! (= @p_65 veriT_sk0) :named @p_346)) :named @p_348)) :named @p_342)) :named @p_341))) :rule forall_inst :args ((:= veriT_vr6 veriT_sk0))) +(anchor :step t83) +(assume t83.h1 @p_341) +(anchor :step t83.t2 :args ((:= (veriT_vr7 FreeExp$) veriT_vr47))) +(step t83.t2.t1 (cl (= veriT_vr7 veriT_vr47)) :rule refl) +(step t83.t2.t2 (cl (= @p_63 (! (insert$ veriT_vr47 bot$) :named @p_343))) :rule cong :premises (t83.t2.t1)) +(step t83.t2.t3 (cl (= @p_64 (! (myImage$ exprel$ @p_343) :named @p_344))) :rule cong :premises (t83.t2.t2)) +(step t83.t2.t4 (cl (= @p_65 (! (abs_Exp$ @p_344) :named @p_345))) :rule cong :premises (t83.t2.t3)) +(step t83.t2.t5 (cl (= @p_346 (! (= veriT_sk0 @p_345) :named @p_347))) :rule cong :premises (t83.t2.t4)) +(step t83.t2.t6 (cl (= @p_348 (! (not @p_347) :named @p_349))) :rule cong :premises (t83.t2.t5)) +(step t83.t2 (cl (= @p_342 (! (forall ((veriT_vr47 FreeExp$)) @p_349) :named @p_350))) :rule bind) +(step t83.t3 (cl (! (= @p_341 (! (not @p_350) :named @p_353)) :named @p_351)) :rule cong :premises (t83.t2)) +(step t83.t4 (cl (! (not @p_351) :named @p_354) (! (not @p_341) :named @p_352) @p_353) :rule equiv_pos2) +(step t83.t5 (cl (! (not @p_352) :named @p_376) @p_342) :rule not_not) +(step t83.t6 (cl @p_354 @p_342 @p_353) :rule th_resolution :premises (t83.t5 t83.t4)) +(step t83.t7 (cl @p_353) :rule th_resolution :premises (t83.h1 t83.t3 t83.t6)) +(anchor :step t83.t8 :args ((:= (veriT_vr47 FreeExp$) veriT_vr48))) +(step t83.t8.t1 (cl (= veriT_vr47 veriT_vr48)) :rule refl) +(step t83.t8.t2 (cl (= @p_343 @p_356)) :rule cong :premises (t83.t8.t1)) +(step t83.t8.t3 (cl (= @p_344 @p_357)) :rule cong :premises (t83.t8.t2)) +(step t83.t8.t4 (cl (= @p_345 @p_358)) :rule cong :premises (t83.t8.t3)) +(step t83.t8.t5 (cl (= @p_347 @p_359)) :rule cong :premises (t83.t8.t4)) +(step t83.t8.t6 (cl (= @p_349 @p_355)) :rule cong :premises (t83.t8.t5)) +(step t83.t8 (cl (= @p_350 (! (forall ((veriT_vr48 FreeExp$)) @p_355) :named @p_360))) :rule bind) +(step t83.t9 (cl (! (= @p_353 (! (not @p_360) :named @p_362)) :named @p_361)) :rule cong :premises (t83.t8)) +(step t83.t10 (cl (! (not @p_361) :named @p_364) (! (not @p_353) :named @p_363) @p_362) :rule equiv_pos2) +(step t83.t11 (cl (not @p_363) @p_350) :rule not_not) +(step t83.t12 (cl @p_364 @p_350 @p_362) :rule th_resolution :premises (t83.t11 t83.t10)) +(step t83.t13 (cl @p_362) :rule th_resolution :premises (t83.t7 t83.t9 t83.t12)) +(anchor :step t83.t14 :args ((:= (veriT_vr48 FreeExp$) veriT_sk2))) +(step t83.t14.t1 (cl (= veriT_vr48 veriT_sk2)) :rule refl) +(step t83.t14.t2 (cl (= @p_356 (! (insert$ veriT_sk2 bot$) :named @p_367))) :rule cong :premises (t83.t14.t1)) +(step t83.t14.t3 (cl (= @p_357 (! (myImage$ exprel$ @p_367) :named @p_368))) :rule cong :premises (t83.t14.t2)) +(step t83.t14.t4 (cl (= @p_358 (! (abs_Exp$ @p_368) :named @p_369))) :rule cong :premises (t83.t14.t3)) +(step t83.t14.t5 (cl (= @p_359 (! (= veriT_sk0 @p_369) :named @p_365))) :rule cong :premises (t83.t14.t4)) +(step t83.t14.t6 (cl (= @p_355 (! (not @p_365) :named @p_370))) :rule cong :premises (t83.t14.t5)) +(step t83.t14 (cl (= @p_360 @p_370)) :rule sko_forall) +(step t83.t15 (cl (! (= @p_362 (! (not @p_370) :named @p_372)) :named @p_371)) :rule cong :premises (t83.t14)) +(step t83.t16 (cl (! (not @p_371) :named @p_374) (! (not @p_362) :named @p_373) @p_372) :rule equiv_pos2) +(step t83.t17 (cl (not @p_373) @p_360) :rule not_not) +(step t83.t18 (cl @p_374 @p_360 @p_372) :rule th_resolution :premises (t83.t17 t83.t16)) +(step t83.t19 (cl (! (not @p_372) :named @p_375) @p_365) :rule not_not) +(step t83.t20 (cl @p_374 @p_360 @p_365) :rule th_resolution :premises (t83.t19 t83.t18)) +(step t83.t21 (cl @p_372) :rule th_resolution :premises (t83.t13 t83.t15 t83.t20)) +(step t83.t22 (cl @p_375 @p_365) :rule not_not) +(step t83.t23 (cl @p_365) :rule th_resolution :premises (t83.t22 t83.t21)) +(step t83.t24 (cl (= @p_372 @p_365)) :rule not_simplify) +(step t83.t25 (cl @p_365) :rule th_resolution :premises (t83.t23 t83.t24)) +(step t83 (cl @p_352 @p_365) :rule subproof :discharge (h1)) +(step t84 (cl @p_376 @p_342) :rule not_not) +(step t85 (cl @p_342 @p_365) :rule th_resolution :premises (t84 t83)) +(step t86 (cl @p_377 @p_341) :rule or :premises (t82)) +(step t87 (cl (! (or @p_377 @p_365) :named @p_379) (! (not @p_377) :named @p_378)) :rule or_neg) +(step t88 (cl (not @p_378) @p_70) :rule not_not) +(step t89 (cl @p_379 @p_70) :rule th_resolution :premises (t88 t87)) +(step t90 (cl @p_379 @p_370) :rule or_neg) +(step t91 (cl @p_379) :rule th_resolution :premises (t86 t85 t89 t90)) +(step t92 (cl @p_377 @p_365) :rule or :premises (t91)) +(step t93 (cl @p_365) :rule resolution :premises (t92 t26)) +(step t94 (cl (or @p_331 (! (not (! (= veriT_sk0 (! (fun_app$ uu$ veriT_sk2) :named @p_381)) :named @p_392)) :named @p_382))) :rule forall_inst :args ((:= veriT_vr46 veriT_sk2))) +(step t95 (cl (or @p_193 (! (not (! (= z$ (! (abs_ExpList$ veriT_sk1) :named @p_380)) :named @p_389)) :named @p_383))) :rule forall_inst :args ((:= veriT_vr30 veriT_sk1))) +(step t96 (cl (or (! (not @p_38) :named @p_384) (! (= @p_302 @p_380) :named @p_385))) :rule forall_inst :args ((:= veriT_vr3 veriT_sk1))) +(step t97 (cl (or (! (not @p_23) :named @p_386) (! (= @p_369 @p_381) :named @p_387))) :rule forall_inst :args ((:= veriT_vr1 veriT_sk2))) +(step t98 (cl @p_331 @p_382) :rule or :premises (t94)) +(step t99 (cl @p_193 @p_383) :rule or :premises (t95)) +(step t100 (cl @p_383) :rule resolution :premises (t99 t65)) +(step t101 (cl @p_384 @p_385) :rule or :premises (t96)) +(step t102 (cl @p_385) :rule resolution :premises (t101 t17)) +(step t103 (cl @p_386 @p_387) :rule or :premises (t97)) +(step t104 (cl @p_387) :rule resolution :premises (t103 t11)) +(step t105 (cl (! (= z$ z$) :named @p_388)) :rule eq_reflexive) +(step t106 (cl (not @p_388) (! (not @p_300) :named @p_390) (! (not @p_385) :named @p_391) @p_389) :rule eq_transitive) +(step t107 (cl @p_390 @p_391 @p_389) :rule th_resolution :premises (t106 t105)) +(step t108 (cl @p_390) :rule resolution :premises (t107 t100 t102)) +(step t109 (cl @p_370 (not @p_387) @p_392) :rule eq_transitive) +(step t110 (cl @p_392) :rule resolution :premises (t109 t93 t104)) +(step t111 (cl @p_393) :rule resolution :premises (t77 t108 t81)) +(step t112 (cl @p_331) :rule resolution :premises (t98 t110)) +(step t113 (cl) :rule resolution :premises (t76 t111 t112)) +c24fc06f55d92aed7783d8234aedb7ced3e99be7 2 0 +(error "status is not unsat.") +unknown diff --git a/src/HOL/SMT_Examples/SMT_Examples_Verit.thy b/src/HOL/SMT_Examples/SMT_Examples_Verit.thy --- a/src/HOL/SMT_Examples/SMT_Examples_Verit.thy +++ b/src/HOL/SMT_Examples/SMT_Examples_Verit.thy @@ -1,740 +1,818 @@ (* Title: HOL/SMT_Examples/SMT_Examples_Verit.thy Author: Sascha Boehme, TU Muenchen Author: Mathias Fleury, JKU Half of the examples come from the corresponding file for z3, the others come from the Isabelle distribution or the AFP. *) section \Examples for the (smt (verit)) binding\ theory SMT_Examples_Verit imports Complex_Main begin external_file \SMT_Examples_Verit.certs\ declare [[smt_certificates = "SMT_Examples_Verit.certs"]] -declare [[smt_read_only_certificates = true]] +declare [[smt_read_only_certificates = false]] section \Propositional and first-order logic\ lemma "True" by (smt (verit)) lemma "p \ \p" by (smt (verit)) lemma "(p \ True) = p" by (smt (verit)) lemma "(p \ q) \ \p \ q" by (smt (verit)) lemma "(a \ b) \ (c \ d) \ (a \ b) \ (c \ d)" by (smt (verit)) lemma "(p1 \ p2) \ p3 \ (p1 \ (p3 \ p2) \ (p1 \ p3)) \ p1" by (smt (verit)) lemma "P = P = P = P = P = P = P = P = P = P" by (smt (verit)) lemma assumes "a \ b \ c \ d" and "e \ f \ (a \ d)" and "\ (a \ (c \ ~c)) \ b" and "\ (b \ (x \ \ x)) \ c" and "\ (d \ False) \ c" and "\ (c \ (\ p \ (p \ (q \ \ q))))" shows False using assms by (smt (verit)) axiomatization symm_f :: "'a \ 'a \ 'a" where symm_f: "symm_f x y = symm_f y x" lemma "a = a \ symm_f a b = symm_f b a" by (smt (verit) symm_f) (* Taken from ~~/src/HOL/ex/SAT_Examples.thy. Translated from TPTP problem library: PUZ015-2.006.dimacs *) lemma assumes "~x0" and "~x30" and "~x29" and "~x59" and "x1 \ x31 \ x0" and "x2 \ x32 \ x1" and "x3 \ x33 \ x2" and "x4 \ x34 \ x3" and "x35 \ x4" and "x5 \ x36 \ x30" and "x6 \ x37 \ x5 \ x31" and "x7 \ x38 \ x6 \ x32" and "x8 \ x39 \ x7 \ x33" and "x9 \ x40 \ x8 \ x34" and "x41 \ x9 \ x35" and "x10 \ x42 \ x36" and "x11 \ x43 \ x10 \ x37" and "x12 \ x44 \ x11 \ x38" and "x13 \ x45 \ x12 \ x39" and "x14 \ x46 \ x13 \ x40" and "x47 \ x14 \ x41" and "x15 \ x48 \ x42" and "x16 \ x49 \ x15 \ x43" and "x17 \ x50 \ x16 \ x44" and "x18 \ x51 \ x17 \ x45" and "x19 \ x52 \ x18 \ x46" and "x53 \ x19 \ x47" and "x20 \ x54 \ x48" and "x21 \ x55 \ x20 \ x49" and "x22 \ x56 \ x21 \ x50" and "x23 \ x57 \ x22 \ x51" and "x24 \ x58 \ x23 \ x52" and "x59 \ x24 \ x53" and "x25 \ x54" and "x26 \ x25 \ x55" and "x27 \ x26 \ x56" and "x28 \ x27 \ x57" and "x29 \ x28 \ x58" and "~x1 \ ~x31" and "~x1 \ ~x0" and "~x31 \ ~x0" and "~x2 \ ~x32" and "~x2 \ ~x1" and "~x32 \ ~x1" and "~x3 \ ~x33" and "~x3 \ ~x2" and "~x33 \ ~x2" and "~x4 \ ~x34" and "~x4 \ ~x3" and "~x34 \ ~x3" and "~x35 \ ~x4" and "~x5 \ ~x36" and "~x5 \ ~x30" and "~x36 \ ~x30" and "~x6 \ ~x37" and "~x6 \ ~x5" and "~x6 \ ~x31" and "~x37 \ ~x5" and "~x37 \ ~x31" and "~x5 \ ~x31" and "~x7 \ ~x38" and "~x7 \ ~x6" and "~x7 \ ~x32" and "~x38 \ ~x6" and "~x38 \ ~x32" and "~x6 \ ~x32" and "~x8 \ ~x39" and "~x8 \ ~x7" and "~x8 \ ~x33" and "~x39 \ ~x7" and "~x39 \ ~x33" and "~x7 \ ~x33" and "~x9 \ ~x40" and "~x9 \ ~x8" and "~x9 \ ~x34" and "~x40 \ ~x8" and "~x40 \ ~x34" and "~x8 \ ~x34" and "~x41 \ ~x9" and "~x41 \ ~x35" and "~x9 \ ~x35" and "~x10 \ ~x42" and "~x10 \ ~x36" and "~x42 \ ~x36" and "~x11 \ ~x43" and "~x11 \ ~x10" and "~x11 \ ~x37" and "~x43 \ ~x10" and "~x43 \ ~x37" and "~x10 \ ~x37" and "~x12 \ ~x44" and "~x12 \ ~x11" and "~x12 \ ~x38" and "~x44 \ ~x11" and "~x44 \ ~x38" and "~x11 \ ~x38" and "~x13 \ ~x45" and "~x13 \ ~x12" and "~x13 \ ~x39" and "~x45 \ ~x12" and "~x45 \ ~x39" and "~x12 \ ~x39" and "~x14 \ ~x46" and "~x14 \ ~x13" and "~x14 \ ~x40" and "~x46 \ ~x13" and "~x46 \ ~x40" and "~x13 \ ~x40" and "~x47 \ ~x14" and "~x47 \ ~x41" and "~x14 \ ~x41" and "~x15 \ ~x48" and "~x15 \ ~x42" and "~x48 \ ~x42" and "~x16 \ ~x49" and "~x16 \ ~x15" and "~x16 \ ~x43" and "~x49 \ ~x15" and "~x49 \ ~x43" and "~x15 \ ~x43" and "~x17 \ ~x50" and "~x17 \ ~x16" and "~x17 \ ~x44" and "~x50 \ ~x16" and "~x50 \ ~x44" and "~x16 \ ~x44" and "~x18 \ ~x51" and "~x18 \ ~x17" and "~x18 \ ~x45" and "~x51 \ ~x17" and "~x51 \ ~x45" and "~x17 \ ~x45" and "~x19 \ ~x52" and "~x19 \ ~x18" and "~x19 \ ~x46" and "~x52 \ ~x18" and "~x52 \ ~x46" and "~x18 \ ~x46" and "~x53 \ ~x19" and "~x53 \ ~x47" and "~x19 \ ~x47" and "~x20 \ ~x54" and "~x20 \ ~x48" and "~x54 \ ~x48" and "~x21 \ ~x55" and "~x21 \ ~x20" and "~x21 \ ~x49" and "~x55 \ ~x20" and "~x55 \ ~x49" and "~x20 \ ~x49" and "~x22 \ ~x56" and "~x22 \ ~x21" and "~x22 \ ~x50" and "~x56 \ ~x21" and "~x56 \ ~x50" and "~x21 \ ~x50" and "~x23 \ ~x57" and "~x23 \ ~x22" and "~x23 \ ~x51" and "~x57 \ ~x22" and "~x57 \ ~x51" and "~x22 \ ~x51" and "~x24 \ ~x58" and "~x24 \ ~x23" and "~x24 \ ~x52" and "~x58 \ ~x23" and "~x58 \ ~x52" and "~x23 \ ~x52" and "~x59 \ ~x24" and "~x59 \ ~x53" and "~x24 \ ~x53" and "~x25 \ ~x54" and "~x26 \ ~x25" and "~x26 \ ~x55" and "~x25 \ ~x55" and "~x27 \ ~x26" and "~x27 \ ~x56" and "~x26 \ ~x56" and "~x28 \ ~x27" and "~x28 \ ~x57" and "~x27 \ ~x57" and "~x29 \ ~x28" and "~x29 \ ~x58" and "~x28 \ ~x58" shows False using assms by (smt (verit)) lemma "\x::int. P x \ (\y::int. P x \ P y)" by (smt (verit)) lemma assumes "(\x y. P x y = x)" shows "(\y. P x y) = P x c" using assms by (smt (verit)) lemma assumes "(\x y. P x y = x)" and "(\x. \y. P x y) = (\x. P x c)" shows "(\y. P x y) = P x c" using assms by (smt (verit)) lemma assumes "if P x then \(\y. P y) else (\y. \P y)" shows "P x \ P y" using assms by (smt (verit)) section \Arithmetic\ subsection \Linear arithmetic over integers and reals\ lemma "(3::int) = 3" by (smt (verit)) lemma "(3::real) = 3" by (smt (verit)) lemma "(3 :: int) + 1 = 4" by (smt (verit)) lemma "x + (y + z) = y + (z + (x::int))" by (smt (verit)) lemma "max (3::int) 8 > 5" by (smt (verit)) lemma "\x :: real\ + \y\ \ \x + y\" by (smt (verit)) lemma "P ((2::int) < 3) = P True" supply[[smt_trace]] by (smt (verit)) lemma "x + 3 \ 4 \ x < (1::int)" by (smt (verit)) lemma assumes "x \ (3::int)" and "y = x + 4" shows "y - x > 0" using assms by (smt (verit)) lemma "let x = (2 :: int) in x + x \ 5" by (smt (verit)) lemma fixes x :: int assumes "3 * x + 7 * a < 4" and "3 < 2 * x" shows "a < 0" using assms by (smt (verit)) lemma "(0 \ y + -1 * x \ \ 0 \ x \ 0 \ (x::int)) = (\ False)" by (smt (verit)) lemma " (n < m \ m < n') \ (n < m \ m = n') \ (n < n' \ n' < m) \ (n = n' \ n' < m) \ (n = m \ m < n') \ (n' < m \ m < n) \ (n' < m \ m = n) \ (n' < n \ n < m) \ (n' = n \ n < m) \ (n' = m \ m < n) \ (m < n \ n < n') \ (m < n \ n' = n) \ (m < n' \ n' < n) \ (m = n \ n < n') \ (m = n' \ n' < n) \ (n' = m \ m = (n::int))" by (smt (verit)) text\ The following example was taken from HOL/ex/PresburgerEx.thy, where it says: This following theorem proves that all solutions to the recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with period 9. The example was brought to our attention by John Harrison. It does does not require Presburger arithmetic but merely quantifier-free linear arithmetic and holds for the rationals as well. Warning: it takes (in 2006) over 4.2 minutes! There, it is proved by "arith". (smt (verit)) is able to prove this within a fraction of one second. With proof reconstruction, it takes about 13 seconds on a Core2 processor. \ lemma "\ x3 = \x2\ - x1; x4 = \x3\ - x2; x5 = \x4\ - x3; x6 = \x5\ - x4; x7 = \x6\ - x5; x8 = \x7\ - x6; x9 = \x8\ - x7; x10 = \x9\ - x8; x11 = \x10\ - x9 \ \ x1 = x10 \ x2 = (x11::int)" by (smt (verit)) lemma "let P = 2 * x + 1 > x + (x::real) in P \ False \ P" by (smt (verit)) subsection \Linear arithmetic with quantifiers\ lemma "~ (\x::int. False)" by (smt (verit)) lemma "~ (\x::real. False)" by (smt (verit)) lemma "\x y::int. (x = 0 \ y = 1) \ x \ y" by (smt (verit)) lemma "\x y::int. x < y \ (2 * x + 1) < (2 * y)" by (smt (verit)) lemma "\x y::int. x + y > 2 \ x + y = 2 \ x + y < 2" by (smt (verit)) lemma "\x::int. if x > 0 then x + 1 > 0 else 1 > x" by (smt (verit)) lemma "(if (\x::int. x < 0 \ x > 0) then -1 else 3) > (0::int)" by (smt (verit)) lemma "\x::int. \x y. 0 < x \ 0 < y \ (0::int) < x + y" by (smt (verit)) lemma "\u::int. \(x::int) y::real. 0 < x \ 0 < y \ -1 < x" by (smt (verit)) lemma "\(a::int) b::int. 0 < b \ b < 1" by (smt (verit)) subsection \Linear arithmetic for natural numbers\ declare [[smt_nat_as_int]] lemma "2 * (x::nat) \ 1" by (smt (verit)) lemma "a < 3 \ (7::nat) > 2 * a" by (smt (verit)) lemma "let x = (1::nat) + y in x - y > 0 * x" by (smt (verit)) lemma "let x = (1::nat) + y in let P = (if x > 0 then True else False) in False \ P = (x - 1 = y) \ (\P \ False)" by (smt (verit)) lemma "int (nat \x::int\) = \x\" by (smt (verit) int_nat_eq) definition prime_nat :: "nat \ bool" where "prime_nat p = (1 < p \ (\m. m dvd p --> m = 1 \ m = p))" lemma "prime_nat (4*m + 1) \ m \ (1::nat)" by (smt (verit) prime_nat_def) lemma "2 * (x::nat) \ 1" by (smt (verit)) lemma \2*(x :: int) \ 1\ by (smt (verit)) declare [[smt_nat_as_int = false]] section \Pairs\ lemma "fst (x, y) = a \ x = a" using fst_conv by (smt (verit)) lemma "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" using fst_conv snd_conv by (smt (verit)) section \Higher-order problems and recursion\ lemma "i \ i1 \ i \ i2 \ (f (i1 := v1, i2 := v2)) i = f i" using fun_upd_same fun_upd_apply by (smt (verit)) lemma "(f g (x::'a::type) = (g x \ True)) \ (f g x = True) \ (g x = True)" by (smt (verit)) lemma "id x = x \ id True = True" by (smt (verit) id_def) lemma "i \ i1 \ i \ i2 \ ((f (i1 := v1)) (i2 := v2)) i = f i" using fun_upd_same fun_upd_apply by (smt (verit)) lemma "f (\x. g x) \ True" "f (\x. g x) \ True" by (smt (verit))+ lemma True using let_rsp by (smt (verit)) lemma "le = (\) \ le (3::int) 42" by (smt (verit)) lemma "map (\i::int. i + 1) [0, 1] = [1, 2]" by (smt (verit) list.map) lemma "(\x. P x) \ \ All P" by (smt (verit)) fun dec_10 :: "int \ int" where "dec_10 n = (if n < 10 then n else dec_10 (n - 10))" lemma "dec_10 (4 * dec_10 4) = 6" by (smt (verit) dec_10.simps) context complete_lattice begin lemma assumes "Sup {a | i::bool. True} \ Sup {b | i::bool. True}" and "Sup {b | i::bool. True} \ Sup {a | i::bool. True}" shows "Sup {a | i::bool. True} \ Sup {a | i::bool. True}" using assms by (smt (verit) order_trans) end lemma "eq_set (List.coset xs) (set ys) = rhs" if "\ys. subset' (List.coset xs) (set ys) = (let n = card (UNIV::'a set) in 0 < n \ card (set (xs @ ys)) = n)" and "\uu A. (uu::'a) \ - A \ uu \ A" and "\uu. card (set (uu::'a list)) = length (remdups uu)" and "\uu. finite (set (uu::'a list))" and "\uu. (uu::'a) \ UNIV" and "(UNIV::'a set) \ {}" and "\c A B P. \(c::'a) \ A \ B; c \ A \ P; c \ B \ P\ \ P" and "\a b. (a::nat) + b = b + a" and "\a b. ((a::nat) = a + b) = (b = 0)" and "card' (set xs) = length (remdups xs)" and "card' = (card :: 'a set \ nat)" and "\A B. \finite (A::'a set); finite B\ \ card A + card B = card (A \ B) + card (A \ B)" and "\A. (card (A::'a set) = 0) = (A = {} \ infinite A)" and "\A. \finite (UNIV::'a set); card (A::'a set) = card (UNIV::'a set)\ \ A = UNIV" and "\xs. - List.coset (xs::'a list) = set xs" and "\xs. - set (xs::'a list) = List.coset xs" and "\A B. (A \ B = {}) = (\x. (x::'a) \ A \ x \ B)" and "eq_set = (=)" and "\A. finite (A::'a set) \ finite (- A) = finite (UNIV::'a set)" and "rhs \ let n = card (UNIV::'a set) in if n = 0 then False else let xs' = remdups xs; ys' = remdups ys in length xs' + length ys' = n \ (\x\set xs'. x \ set ys') \ (\y\set ys'. y \ set xs')" and "\xs ys. set ((xs::'a list) @ ys) = set xs \ set ys" and "\A B. ((A::'a set) = B) = (A \ B \ B \ A)" and "\xs. set (remdups (xs::'a list)) = set xs" and "subset' = (\)" and "\A B. (\x. (x::'a) \ A \ x \ B) \ A \ B" and "\A B. \(A::'a set) \ B; B \ A\ \ A = B" and "\A ys. (A \ List.coset ys) = (\y\set ys. (y::'a) \ A)" using that by (smt (verit, default)) notepad begin have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" if \(k, g) \ one_chain_typeI\ \\A b B. ({} = (A::(real \ real) set) \ insert (b::real \ real) (B::(real \ real) set)) = (b \ A \ {} = A \ B)\ \finite ({} :: (real \ real) set)\ \\a A. finite (A::(real \ real) set) \ finite (insert (a::real \ real) A)\ \(i::real \ real) = (1::real, 0::real)\ \ \a A. (a::real \ real) \ (A::(real \ real) set) \ insert a A = A\ \j = (0, 1)\ \\x. (x::(real \ real) set) \ {} = {}\ \\y x A. insert (x::real \ real) (insert (y::real \ real) (A::(real \ real) set)) = insert y (insert x A)\ \\a A. insert (a::real \ real) (A::(real \ real) set) = {a} \ A\ \\F u basis2 basis1 \. finite (u :: (real \ real) set) \ line_integral_exists F basis1 \ \ line_integral_exists F basis2 \ \ basis1 \ basis2 = u \ basis1 \ basis2 = {} \ line_integral F u \ = line_integral F basis1 \ + line_integral F basis2 \\ \one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k, \)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k, \)\one_chain_typeII. line_integral_exists F {i} \)\ \ one_chain_line_integral (F::real \ real \ real \ real) {j::real \ real} (one_chain_typeII::(int \ (real \ real \ real)) set) = one_chain_line_integral F {j} (one_chain_typeI::(int \ (real \ real \ real)) set) \ (\(k::int, \::real \ real \ real)\one_chain_typeII. line_integral_exists F {j} \) \ (\(k::int, \::real \ real \ real)\one_chain_typeI. line_integral_exists F {j} \)\ for F i j g one_chain_typeI one_chain_typeII and line_integral :: \(real \ real \ real \ real) \ (real \ real) set \ (real \ real \ real) \ real\ and line_integral_exists :: \(real \ real \ real \ real) \ (real \ real) set \ (real \ real \ real) \ bool\ and one_chain_line_integral :: \(real \ real \ real \ real) \ (real \ real) set \ (int \ (real \ real \ real)) set \ real\ and k using prod.case_eq_if singleton_inject snd_conv that by (smt (verit)) end lemma fixes x y z :: real assumes \x + 2 * y > 0\ and \x - 2 * y > 0\ and \x < 0\ shows False using assms by (smt (verit)) (*test for arith reconstruction*) lemma fixes d :: real assumes \0 < d\ \diamond_y \ \t. d / 2 - \t\\ \\a b c :: real. (a / c < b / c) = ((0 < c \ a < b) \ (c < 0 \ b < a) \ c \ 0)\ \\a b c :: real. (a / c < b / c) = ((0 < c \ a < b) \ (c < 0 \ b < a) \ c \ 0)\ \\a b :: real. - a / b = - (a / b)\ \\a b :: real. - a * b = - (a * b)\ \\(x1 :: real) x2 y1 y2 :: real. ((x1, x2) = (y1, y2)) = (x1 = y1 \ x2 = y2)\ shows \(\y. (d / 2, (2 * y - 1) * diamond_y (d / 2))) \ (\x. ((x - 1 / 2) * d, - diamond_y ((x - 1 / 2) * d))) \ (\y. (- (d / 2), (2 * y - 1) * diamond_y (- (d / 2)))) = (\x. ((x - 1 / 2) * d, diamond_y ((x - 1 / 2) * d))) \ False\ using assms by (smt (verit,del_insts)) lemma fixes d :: real assumes \0 < d\ \diamond_y \ \t. d / 2 - \t\\ \\a b c :: real. (a / c < b / c) = ((0 < c \ a < b) \ (c < 0 \ b < a) \ c \ 0)\ \\a b c :: real. (a / c < b / c) = ((0 < c \ a < b) \ (c < 0 \ b < a) \ c \ 0)\ \\a b :: real. - a / b = - (a / b)\ \\a b :: real. - a * b = - (a * b)\ \\(x1 :: real) x2 y1 y2 :: real. ((x1, x2) = (y1, y2)) = (x1 = y1 \ x2 = y2)\ shows \(\y. (d / 2, (2 * y - 1) * diamond_y (d / 2))) \ (\x. ((x - 1 / 2) * d, - diamond_y ((x - 1 / 2) * d))) \ (\y. (- (d / 2), (2 * y - 1) * diamond_y (- (d / 2)))) = (\x. ((x - 1 / 2) * d, diamond_y ((x - 1 / 2) * d))) \ False\ using assms by (smt (verit,ccfv_threshold)) (*qnt_rm_unused example*) lemma assumes \\z y x. P z y\ \P z y \ False\ shows False using assms by (smt (verit)) lemma "max (x::int) y \ y" supply [[smt_trace]] by (smt (verit))+ context begin abbreviation finite' :: "'a set \ bool" where "finite' A \ finite A \ A \ {}" lemma fixes f :: "'b \ 'c :: linorder" assumes \\(S::'b::type set) f::'b::type \ 'c::linorder. finite' S \ arg_min_on f S \ S\ \\(S::'a::type set) f::'a::type \ 'c::linorder. finite' S \ arg_min_on f S \ S\ \\(S::'b::type set) (y::'b::type) f::'b::type \ 'c::linorder. finite S \ S \ {} \ y \ S \ f (arg_min_on f S) \ f y\ \\(S::'a::type set) (y::'a::type) f::'a::type \ 'c::linorder. finite S \ S \ {} \ y \ S \ f (arg_min_on f S) \ f y\ \\(f::'b::type \ 'c::linorder) (g::'a::type \ 'b::type) x::'a::type. (f \ g) x = f (g x)\ \\(F::'b::type set) h::'b::type \ 'a::type. finite F \ finite (h ` F)\ \\(F::'b::type set) h::'b::type \ 'b::type. finite F \ finite (h ` F)\ \\(F::'a::type set) h::'a::type \ 'b::type. finite F \ finite (h ` F)\ \\(F::'a::type set) h::'a::type \ 'a::type. finite F \ finite (h ` F)\ \\(b::'a::type) (f::'b::type \ 'a::type) A::'b::type set. b \ f ` A \ (\x::'b::type. b = f x \ x \ A \ False) \ False\ \\(b::'b::type) (f::'b::type \ 'b::type) A::'b::type set. b \ f ` A \ (\x::'b::type. b = f x \ x \ A \ False) \ False\ \\(b::'b::type) (f::'a::type \ 'b::type) A::'a::type set. b \ f ` A \ (\x::'a::type. b = f x \ x \ A \ False) \ False\ \\(b::'a::type) (f::'a::type \ 'a::type) A::'a::type set. b \ f ` A \ (\x::'a::type. b = f x \ x \ A \ False) \ False\ \\(b::'a::type) (f::'b::type \ 'a::type) (x::'b::type) A::'b::type set. b = f x \ x \ A \ b \ f ` A \ \\(b::'b::type) (f::'b::type \ 'b::type) (x::'b::type) A::'b::type set. b = f x \ x \ A \ b \ f ` A \ \\(b::'b::type) (f::'a::type \ 'b::type) (x::'a::type) A::'a::type set. b = f x \ x \ A \ b \ f ` A \ \\(b::'a::type) (f::'a::type \ 'a::type) (x::'a::type) A::'a::type set. b = f x \ x \ A \ b \ f ` A \ \\(f::'b::type \ 'a::type) A::'b::type set. (f ` A = {}) = (A = {}) \ \\(f::'b::type \ 'b::type) A::'b::type set. (f ` A = {}) = (A = {}) \ \\(f::'a::type \ 'b::type) A::'a::type set. (f ` A = {}) = (A = {}) \ \\(f::'a::type \ 'a::type) A::'a::type set. (f ` A = {}) = (A = {}) \ \\(f::'b::type \ 'c::linorder) (A::'b::type set) (x::'b::type) y::'b::type. inj_on f A \ f x = f y \ x \ A \ y \ A \ x = y\ \\(x::'c::linorder) y::'c::linorder. (x < y) = (x \ y \ x \ y)\ \inj_on (f::'b::type \ 'c::linorder) ((g::'a::type \ 'b::type) ` (B::'a::type set))\ \finite (B::'a::type set)\ \(B::'a::type set) \ {}\ \arg_min_on ((f::'b::type \ 'c::linorder) \ (g::'a::type \ 'b::type)) (B::'a::type set) \ B\ \\x::'a::type. x \ (B::'a::type set) \ ((f::'b::type \ 'c::linorder) \ (g::'a::type \ 'b::type)) x < (f \ g) (arg_min_on (f \ g) B)\ \\(f::'b::type \ 'c::linorder) (P::'b::type \ bool) a::'b::type. inj_on f (Collect P) \ P a \ (\y::'b::type. P y \ f a \ f y) \ arg_min f P = a\ \\(S::'b::type set) f::'b::type \ 'c::linorder. finite' S \ arg_min_on f S \ S\ \\(S::'a::type set) f::'a::type \ 'c::linorder. finite' S \ arg_min_on f S \ S\ \\(S::'b::type set) (y::'b::type) f::'b::type \ 'c::linorder. finite S \ S \ {} \ y \ S \ f (arg_min_on f S) \ f y\ \\(S::'a::type set) (y::'a::type) f::'a::type \ 'c::linorder. finite S \ S \ {} \ y \ S \ f (arg_min_on f S) \ f y\ \\(f::'b::type \ 'c::linorder) (g::'a::type \ 'b::type) x::'a::type. (f \ g) x = f (g x)\ \\(F::'b::type set) h::'b::type \ 'a::type. finite F \ finite (h ` F)\ \\(F::'b::type set) h::'b::type \ 'b::type. finite F \ finite (h ` F)\ \\(F::'a::type set) h::'a::type \ 'b::type. finite F \ finite (h ` F)\ \\(F::'a::type set) h::'a::type \ 'a::type. finite F \ finite (h ` F)\ \\(b::'a::type) (f::'b::type \ 'a::type) A::'b::type set. b \ f ` A \ (\x::'b::type. b = f x \ x \ A \ False) \ False\ \\(b::'b::type) (f::'b::type \ 'b::type) A::'b::type set. b \ f ` A \ (\x::'b::type. b = f x \ x \ A \ False) \ False\ \\(b::'b::type) (f::'a::type \ 'b::type) A::'a::type set. b \ f ` A \ (\x::'a::type. b = f x \ x \ A \ False) \ False\ \\(b::'a::type) (f::'a::type \ 'a::type) A::'a::type set. b \ f ` A \ (\x::'a::type. b = f x \ x \ A \ False) \ False\ \\(b::'a::type) (f::'b::type \ 'a::type) (x::'b::type) A::'b::type set. b = f x \ x \ A \ b \ f ` A \ \\(b::'b::type) (f::'b::type \ 'b::type) (x::'b::type) A::'b::type set. b = f x \ x \ A \ b \ f ` A \ \\(b::'b::type) (f::'a::type \ 'b::type) (x::'a::type) A::'a::type set. b = f x \ x \ A \ b \ f ` A \ \\(b::'a::type) (f::'a::type \ 'a::type) (x::'a::type) A::'a::type set. b = f x \ x \ A \ b \ f ` A \ \\(f::'b::type \ 'a::type) A::'b::type set. (f ` A = {}) = (A = {}) \ \\(f::'b::type \ 'b::type) A::'b::type set. (f ` A = {}) = (A = {}) \ \\(f::'a::type \ 'b::type) A::'a::type set. (f ` A = {}) = (A = {}) \ \\(f::'a::type \ 'a::type) A::'a::type set. (f ` A = {}) = (A = {})\ \\(f::'b::type \ 'c::linorder) (A::'b::type set) (x::'b::type) y::'b::type. inj_on f A \ f x = f y \ x \ A \ y \ A \ x = y\ \\(x::'c::linorder) y::'c::linorder. (x < y) = (x \ y \ x \ y)\ \arg_min_on (f::'b::type \ 'c::linorder) ((g::'a::type \ 'b::type) ` (B::'a::type set)) \ g (arg_min_on (f \ g) B) \ shows False using assms by (smt (verit)) end experiment begin private datatype abort = Rtype_error | Rtimeout_error private datatype ('a) error_result = Rraise " 'a " \ \\ Should only be a value of type exn \\ | Rabort " abort " private datatype( 'a, 'b) result = Rval " 'a " | Rerr " ('b) error_result " lemma fixes clock :: \'astate \ nat\ and fun_evaluate_match :: \'astate \ 'vsemv_env \ _ \ ('pat \ 'exp0) list \ _ \ 'astate*((('v)list),('v))result\ assumes "fix_clock (st::'astate) (fun_evaluate st (env::'vsemv_env) [e::'exp0]) = (st'::'astate, r::('v list, 'v) result)" "clock (fst (fun_evaluate (st::'astate) (env::'vsemv_env) [e::'exp0])) \ clock st" "\(b::nat) (a::nat) c::nat. b \ a \ c \ b \ c \ a" "\(a::'astate) p::'astate \ ('v list, 'v) result. (a = fst p) = (\b::('v list, 'v) result. p = (a, b))" "\y::'v error_result. (\x1::'v. y = Rraise x1 \ False) \ (\x2::abort. y = Rabort x2 \ False) \ False" "\(f1::'v \ 'astate \ ('v list, 'v) result) (f2::abort \ 'astate \ ('v list, 'v) result) x1::'v. (case Rraise x1 of Rraise (x::'v) \ f1 x | Rabort (x::abort) \ f2 x) = f1 x1" "\(f1::'v \ 'astate \ ('v list, 'v) result) (f2::abort \ 'astate \ ('v list, 'v) result) x2::abort. (case Rabort x2 of Rraise (x::'v) \ f1 x | Rabort (x::abort) \ f2 x) = f2 x2" "\(s1::'astate) (s2::'astate) (x::('v list, 'v) result) s::'astate. fix_clock s1 (s2, x) = (s, x) \ clock s \ clock s2" "\(s::'astate) (s'::'astate) res::('v list, 'v) result. fix_clock s (s', res) = (update_clock (\_::nat. if clock s' \ clock s then clock s' else clock s) s', res)" "\(x2::'v error_result) x1::'v. (r::('v list, 'v) result) = Rerr x2 \ x2 = Rraise x1 \ clock (fst (fun_evaluate_match (st'::'astate) (env::'vsemv_env) x1 (pes::('pat \ 'exp0) list) x1)) \ clock st'" shows "((r::('v list, 'v) result) = Rerr (x2::'v error_result) \ clock (fst (case x2 of Rraise (v2::'v) \ fun_evaluate_match (st'::'astate) (env::'vsemv_env) v2 (pes::('pat \ 'exp0) list) v2 | Rabort (abort::abort) \ (st', Rerr (Rabort abort)))) \ clock (st::'astate))" using assms by (smt (verit)) end context fixes piecewise_C1 :: "('real :: {one,zero,ord} \ 'a :: {one,zero,ord}) \ 'real set \ bool" and joinpaths :: "('real \ 'a) \ ('real \ 'a) \ 'real \ 'a" begin notation piecewise_C1 (infixr "piecewise'_C1'_differentiable'_on" 50) notation joinpaths (infixr "+++" 75) lemma \(\v1. \v0. (rec_join v0 = v1 \ (v0 = [] \ (\uu. 0) = v1 \ False) \ (\v2. v0 = [v2] \ v1 = coeff_cube_to_path v2 \ False) \ (\v2 v3 v4. v0 = v2 # v3 # v4 \ v1 = coeff_cube_to_path v2 +++ rec_join (v3 # v4) \ False) \ False) = (rec_join v0 = rec_join v0 \ (v0 = [] \ (\uu. 0) = rec_join v0 \ False) \ (\v2. v0 = [v2] \ rec_join v0 = coeff_cube_to_path v2 \ False) \ (\v2 v3 v4. v0 = v2 # v3 # v4 \ rec_join v0 = coeff_cube_to_path v2 +++ rec_join (v3 # v4) \ False) \ False)) \ (\v0 v1. rec_join v0 = v1 \ (v0 = [] \ (\uu. 0) = v1 \ False) \ (\v2. v0 = [v2] \ v1 = coeff_cube_to_path v2 \ False) \ (\v2 v3 v4. v0 = v2 # v3 # v4 \ v1 = coeff_cube_to_path v2 +++ rec_join (v3 # v4) \ False) \ False) = (\v0. rec_join v0 = rec_join v0 \ (v0 = [] \ (\uu. 0) = rec_join v0 \ False) \ (\v2. v0 = [v2] \ rec_join v0 = coeff_cube_to_path v2 \ False) \ (\v2 v3 v4. v0 = v2 # v3 # v4 \ rec_join v0 = coeff_cube_to_path v2 +++ rec_join (v3 # v4) \ False) \ False)\ by (smt (verit)) end section \Monomorphization examples\ definition Pred :: "'a \ bool" where "Pred x = True" lemma poly_Pred: "Pred x \ (Pred [x] \ \ Pred [x])" by (simp add: Pred_def) lemma "Pred (1::int)" by (smt (verit) poly_Pred) axiomatization g :: "'a \ nat" axiomatization where g1: "g (Some x) = g [x]" and g2: "g None = g []" and g3: "g xs = length xs" lemma "g (Some (3::int)) = g (Some True)" by (smt (verit) g1 g2 g3 list.size) +experiment +begin + +lemma duplicate_goal: \A \ A \ A\ + by auto + +datatype 'a M_nres = is_fail: FAIL | SPEC "'a \ bool" + +definition "is_res m x \ case m of FAIL \ True | SPEC P \ P x" + +datatype ('a,'s) M_state = M_STATE (run: "'s \ ('a\'s) M_nres") + +(*Courtesy of Peter Lammich +https://isabelle.zulipchat.com/#narrow/stream/247541-Mirror.3A-Isabelle-Users-Mailing-List/topic/.5Bisabelle.5D.20smt.20.28verit.29.3A.20exception.20THM.200.20raised.20.28line.20312.20.2E.2E.2E/near/290088165 +*) +lemma "\\x y. (\xa s. is_fail (run (x xa) s) \ + is_fail (run (y xa) s) = is_fail (run (x xa) s) \ + (\a b. is_res (run (y xa) s) (a, b) = is_res (run (x xa) s) (a, b))) +\ + (\s. is_fail (run (B x) s) \ + is_fail (run (B y) s) = is_fail (run (B x) s) \ + (\a b. is_res (run (B y) s) (a, b) = is_res (run (B x) s) (a, b))); + \y. \x ya. (\xa s. is_fail (run (x xa) s) \ + is_fail (run (ya xa) s) = is_fail (run (x xa) s) \ + (\a b. is_res (run (ya xa) s) (a, b) = is_res (run (x xa) s) (a, b))) +\ + (\s. is_fail (run (C y x) s) \ + is_fail (run (C y ya) s) = is_fail (run (C y x) s) \ + (\a b. is_res (run (C y ya) s) (a, b) = is_res (run (C y x) s) (a, +b)))\ + \ \x y. (\xa s. + is_fail (run (x xa) s) \ + is_fail (run (y xa) s) = is_fail (run (x xa) s) \ + (\a b. is_res (run (y xa) s) (a, b) = is_res (run (x xa) s) (a, b))) +\ + (\s. is_fail (run (B x) s) \ + (\a b. is_res (run (B x) s) (a, b) \ is_fail (run (C a x) b)) \ + (is_fail (run (B y) s) \ (\a b. is_res (run (B y) s) (a, b) \ +is_fail (run (C a y) b))) = + (is_fail (run (B x) s) \ (\a b. is_res (run (B x) s) (a, b) \ +is_fail (run (C a x) b))) \ + (\a b. (is_fail (run (B y) s) \ + (\aa ba. is_res (run (B y) s) (aa, ba) \ is_res (run (C aa y) +ba) (a, b))) = + (is_fail (run (B x) s) \ + (\aa ba. is_res (run (B x) s) (aa, ba) \ is_res (run (C aa x) +ba) (a, b)))))" + apply (rule duplicate_goal) + subgoal + supply [[verit_compress_proofs=true]] + by (smt (verit)) + subgoal + supply [[verit_compress_proofs=false]] + by (smt (verit)) + done + +(*Example of Reordering in skolemization*) +lemma + fixes Abs_ExpList :: "'freeExp_list \ 'exp_list" and + Abs_Exp:: "'freeExp_set \ 'exp" and + exprel:: "('freeExp \ 'freeExp) set" and + map2 :: "('freeExp \ 'exp) \ 'freeExp_list \ 'exp_list" + assumes "\Xs. Abs_ExpList Xs \ map2 (\U. Abs_Exp (myImage exprel {U})) Xs" + "\P z. (\U. z = Abs_Exp (myImage exprel {U}) \ P) \ P" + "\(ys::'exp_list) (f::'freeExp \ _). (\xs. ys = map2 f xs) = (\y\myset ys. \x. y = f x)" + shows "\Us. z = Abs_ExpList Us" + apply (rule duplicate_goal) + subgoal + supply [[verit_compress_proofs=true]] + using assms + by (smt (verit,del_insts)) + subgoal + using assms + supply [[verit_compress_proofs=false]] + by (smt (verit,del_insts)) + done + +end end \ No newline at end of file diff --git a/src/HOL/Tools/SMT/lethe_proof.ML b/src/HOL/Tools/SMT/lethe_proof.ML --- a/src/HOL/Tools/SMT/lethe_proof.ML +++ b/src/HOL/Tools/SMT/lethe_proof.ML @@ -1,787 +1,788 @@ (* Title: HOL/Tools/SMT/Lethe_Proof.ML Author: Mathias Fleury, ENS Rennes Author: Sascha Boehme, TU Muenchen Lethe proofs: parsing and abstract syntax tree. *) signature LETHE_PROOF = sig (*proofs*) datatype lethe_step = Lethe_Step of { id: string, rule: string, prems: string list, proof_ctxt: term list, concl: term, fixes: string list} datatype lethe_replay_node = Lethe_Replay_Node of { id: string, rule: string, args: term list, prems: string list, proof_ctxt: term list, concl: term, bounds: (string * typ) list, declarations: (string * term) list, insts: term Symtab.table, subproof: (string * typ) list * term list * term list * lethe_replay_node list} val mk_replay_node: string -> string -> term list -> string list -> term list -> term -> (string * typ) list -> term Symtab.table -> (string * term) list -> (string * typ) list * term list * term list * lethe_replay_node list -> lethe_replay_node (*proof parser*) val parse: typ Symtab.table -> term Symtab.table -> string list -> Proof.context -> lethe_step list * Proof.context val parse_replay: typ Symtab.table -> term Symtab.table -> string list -> Proof.context -> lethe_replay_node list * Proof.context val step_prefix : string val input_rule: string val keep_app_symbols: string -> bool val keep_raw_lifting: string -> bool val normalized_input_rule: string val la_generic_rule : string val rewrite_rule : string val simp_arith_rule : string val lethe_deep_skolemize_rule : string val lethe_def : string + val is_lethe_def : string -> bool val subproof_rule : string val local_input_rule : string val not_not_rule : string val contract_rule : string val ite_intro_rule : string val eq_congruent_rule : string val eq_congruent_pred_rule : string val skolemization_steps : string list val theory_resolution2_rule: string val equiv_pos2_rule: string val and_pos_rule: string val th_resolution_rule: string val is_skolemization: string -> bool val is_skolemization_step: lethe_replay_node -> bool val number_of_steps: lethe_replay_node list -> int end; structure Lethe_Proof: LETHE_PROOF = struct open SMTLIB_Proof datatype raw_lethe_node = Raw_Lethe_Node of { id: string, rule: string, args: SMTLIB.tree, prems: string list, concl: SMTLIB.tree, declarations: (string * SMTLIB.tree) list, subproof: raw_lethe_node list} fun mk_raw_node id rule args prems declarations concl subproof = Raw_Lethe_Node {id = id, rule = rule, args = args, prems = prems, declarations = declarations, concl = concl, subproof = subproof} datatype lethe_node = Lethe_Node of { id: string, rule: string, prems: string list, proof_ctxt: term list, concl: term} fun mk_node id rule prems proof_ctxt concl = Lethe_Node {id = id, rule = rule, prems = prems, proof_ctxt = proof_ctxt, concl = concl} datatype lethe_replay_node = Lethe_Replay_Node of { id: string, rule: string, args: term list, prems: string list, proof_ctxt: term list, concl: term, bounds: (string * typ) list, insts: term Symtab.table, declarations: (string * term) list, subproof: (string * typ) list * term list * term list * lethe_replay_node list} fun mk_replay_node id rule args prems proof_ctxt concl bounds insts declarations subproof = Lethe_Replay_Node {id = id, rule = rule, args = args, prems = prems, proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, insts = insts, declarations = declarations, subproof = subproof} datatype lethe_step = Lethe_Step of { id: string, rule: string, prems: string list, proof_ctxt: term list, concl: term, fixes: string list} fun mk_step id rule prems proof_ctxt concl fixes = Lethe_Step {id = id, rule = rule, prems = prems, proof_ctxt = proof_ctxt, concl = concl, fixes = fixes} val step_prefix = ".c" val input_rule = "input" val la_generic_rule = "la_generic" val normalized_input_rule = "__normalized_input" (*arbitrary*) val rewrite_rule = "__rewrite" (*arbitrary*) val subproof_rule = "subproof" val local_input_rule = "__local_input" (*arbitrary*) val simp_arith_rule = "simp_arith" val lethe_def = "__skolem_definition" (*arbitrary*) val not_not_rule = "not_not" val contract_rule = "contraction" val eq_congruent_pred_rule = "eq_congruent_pred" val eq_congruent_rule = "eq_congruent" val ite_intro_rule = "ite_intro" val default_skolem_rule = "sko_forall" (*arbitrary, but must be one of the skolems*) val theory_resolution2_rule = "__theory_resolution2" (*arbitrary*) val equiv_pos2_rule = "equiv_pos2" val th_resolution_rule = "th_resolution" val and_pos_rule = "and_pos" +val is_lethe_def = String.isSuffix lethe_def val skolemization_steps = ["sko_forall", "sko_ex"] val is_skolemization = member (op =) skolemization_steps val keep_app_symbols = member (op =) [eq_congruent_pred_rule, eq_congruent_rule, ite_intro_rule, and_pos_rule] val keep_raw_lifting = member (op =) [eq_congruent_pred_rule, eq_congruent_rule, ite_intro_rule, and_pos_rule] val is_SH_trivial = member (op =) [not_not_rule, contract_rule] fun is_skolemization_step (Lethe_Replay_Node {id, ...}) = is_skolemization id (* Even the lethe developers do not know if the following rule can still appear in proofs: *) val lethe_deep_skolemize_rule = "deep_skolemize" fun number_of_steps [] = 0 | number_of_steps ((Lethe_Replay_Node {subproof = (_, _, _, subproof), ...}) :: pf) = 1 + number_of_steps subproof + number_of_steps pf (* proof parser *) fun node_of p cx = ([], cx) ||>> `(with_fresh_names (term_of p)) |>> snd fun synctactic_var_subst old_name new_name (u $ v) = (synctactic_var_subst old_name new_name u $ synctactic_var_subst old_name new_name v) | synctactic_var_subst old_name new_name (Abs (v, T, u)) = Abs (if String.isPrefix old_name v then new_name else v, T, synctactic_var_subst old_name new_name u) | synctactic_var_subst old_name new_name (Free (v, T)) = if String.isPrefix old_name v then Free (new_name, T) else Free (v, T) | synctactic_var_subst _ _ t = t fun synctatic_rew_in_lhs_subst old_name new_name (Const(\<^const_name>\HOL.eq\, T) $ t1 $ t2) = Const(\<^const_name>\HOL.eq\, T) $ synctactic_var_subst old_name new_name t1 $ t2 | synctatic_rew_in_lhs_subst old_name new_name (Const(\<^const_name>\Trueprop\, T) $ t1) = Const(\<^const_name>\Trueprop\, T) $ (synctatic_rew_in_lhs_subst old_name new_name t1) | synctatic_rew_in_lhs_subst _ _ t = t fun add_bound_variables_to_ctxt cx = fold (update_binding o (fn (s, SOME typ) => (s, Term (Free (s, type_of cx typ))))) local fun extract_symbols bds = bds |> map (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, SMTLIB.Sym y], typ) => [([x, y], typ)] | t => raise (Fail ("match error " ^ @{make_string} t))) |> flat (* onepoint can bind a variable to another variable or to a constant *) fun extract_qnt_symbols cx bds = bds |> map (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, SMTLIB.Sym y], typ) => (case node_of (SMTLIB.Sym y) cx of ((_, []), _) => [([x], typ)] | _ => [([x, y], typ)]) | (SMTLIB.S (SMTLIB.Sym "=" :: SMTLIB.Sym x :: _), typ) => [([x], typ)] | t => raise (Fail ("match error " ^ @{make_string} t))) |> flat fun extract_symbols_map bds = bds |> map (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, _], typ) => [([x], typ)]) |> flat in fun declared_csts _ "__skolem_definition" [(SMTLIB.S [SMTLIB.Sym x, typ, _], _)] = [(x, typ)] | declared_csts _ "__skolem_definition" t = raise (Fail ("unrecognized skolem_definition " ^ @{make_string} t)) | declared_csts _ _ _ = [] fun skolems_introduced_by_rule (SMTLIB.S bds) = fold (fn (SMTLIB.S [SMTLIB.Sym "=", _, SMTLIB.Sym y]) => curry (op ::) y) bds [] (*FIXME there is probably a way to use the information given by onepoint*) fun bound_vars_by_rule _ "bind" (bds) = extract_symbols bds | bound_vars_by_rule cx "onepoint" bds = extract_qnt_symbols cx bds | bound_vars_by_rule _ "sko_forall" bds = extract_symbols_map bds | bound_vars_by_rule _ "sko_ex" bds = extract_symbols_map bds | bound_vars_by_rule _ "__skolem_definition" [(SMTLIB.S [SMTLIB.Sym x, typ, _], _)] = [([x], SOME typ)] | bound_vars_by_rule _ "__skolem_definition" [(SMTLIB.S [_, SMTLIB.Sym x, _], _)] = [([x], NONE)] | bound_vars_by_rule _ _ _ = [] (* Lethe adds "?" before some variables. *) fun remove_all_qm (SMTLIB.Sym v :: l) = SMTLIB.Sym (perhaps (try (unprefix "?")) v) :: remove_all_qm l | remove_all_qm (SMTLIB.S l :: l') = SMTLIB.S (remove_all_qm l) :: remove_all_qm l' | remove_all_qm (SMTLIB.Key v :: l) = SMTLIB.Key v :: remove_all_qm l | remove_all_qm (v :: l) = v :: remove_all_qm l | remove_all_qm [] = [] fun remove_all_qm2 (SMTLIB.Sym v) = SMTLIB.Sym (perhaps (try (unprefix "?")) v) | remove_all_qm2 (SMTLIB.S l) = SMTLIB.S (remove_all_qm l) | remove_all_qm2 (SMTLIB.Key v) = SMTLIB.Key v | remove_all_qm2 v = v end datatype step_kind = ASSUME | ANCHOR | NO_STEP | NORMAL_STEP | SKOLEM fun parse_raw_proof_steps (limit : string option) (ls : SMTLIB.tree list) (cx : name_bindings) : (raw_lethe_node list * SMTLIB.tree list * name_bindings) = let fun rotate_pair (a, (b, c)) = ((a, b), c) fun step_kind [] = (NO_STEP, SMTLIB.S [], []) | step_kind ((p as SMTLIB.S (SMTLIB.Sym "anchor" :: _)) :: l) = (ANCHOR, p, l) | step_kind ((p as SMTLIB.S (SMTLIB.Sym "assume" :: _)) :: l) = (ASSUME, p, l) | step_kind ((p as SMTLIB.S (SMTLIB.Sym "step" :: _)) :: l) = (NORMAL_STEP, p, l) | step_kind ((p as SMTLIB.S (SMTLIB.Sym "define-fun" :: _)) :: l) = (SKOLEM, p, l) | step_kind p = raise (Fail ("step_kind unrec: " ^ @{make_string} p)) fun parse_skolem (SMTLIB.S [SMTLIB.Sym "define-fun", SMTLIB.Sym id, _, typ, SMTLIB.S (SMTLIB.Sym "!" :: t :: [SMTLIB.Key _, SMTLIB.Sym name])]) cx = (*replace the name binding by the constant instead of the full term in order to reduce the size of the generated terms and therefore the reconstruction time*) let val (l, cx) = (fst oo SMTLIB_Proof.extract_and_update_name_bindings) t cx |> apsnd (SMTLIB_Proof.update_name_binding (name, SMTLIB.Sym id)) in (mk_raw_node (id ^ lethe_def) lethe_def (SMTLIB.S [SMTLIB.Sym id, typ, l]) [] [] (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym id, l]) [], cx) end | parse_skolem (SMTLIB.S [SMTLIB.Sym "define-fun", SMTLIB.Sym id, _, typ, SMTLIB.S l]) cx = let val (l, cx) = (fst oo SMTLIB_Proof.extract_and_update_name_bindings) (SMTLIB.S l ) cx in (mk_raw_node (id ^ lethe_def) lethe_def (SMTLIB.S [SMTLIB.Sym id, typ, l]) [] [] (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym id, l]) [], cx) end | parse_skolem t _ = raise Fail ("unrecognized Lethe proof " ^ \<^make_string> t) fun get_id_cx (SMTLIB.S ((SMTLIB.Sym _) :: (SMTLIB.Sym id) :: l), cx) = (id, (l, cx)) | get_id_cx t = raise Fail ("unrecognized Lethe proof " ^ \<^make_string> t) fun get_id (SMTLIB.S ((SMTLIB.Sym _) :: (SMTLIB.Sym id) :: l)) = (id, l) | get_id t = raise Fail ("unrecognized Lethe proof " ^ \<^make_string> t) fun parse_source (SMTLIB.Key "premises" :: SMTLIB.S source ::l, cx) = (SOME (map (fn (SMTLIB.Sym id) => id) source), (l, cx)) | parse_source (l, cx) = (NONE, (l, cx)) fun parse_rule (SMTLIB.Key "rule" :: SMTLIB.Sym r :: l, cx) = (r, (l, cx)) | parse_rule t = raise Fail ("unrecognized Lethe proof " ^ \<^make_string> t) fun parse_anchor_step (SMTLIB.S (SMTLIB.Sym "anchor" :: SMTLIB.Key "step" :: SMTLIB.Sym r :: l), cx) = (r, (l, cx)) | parse_anchor_step t = raise Fail ("unrecognized Lethe proof " ^ \<^make_string> t) fun parse_args (SMTLIB.Key "args" :: args :: l, cx) = let val ((args, cx), _) = SMTLIB_Proof.extract_and_update_name_bindings args cx in (args, (l, cx)) end | parse_args (l, cx) = (SMTLIB.S [], (l, cx)) fun parse_and_clausify_conclusion (SMTLIB.S (SMTLIB.Sym "cl" :: []) :: l, cx) = (SMTLIB.Sym "false", (l, cx)) | parse_and_clausify_conclusion (SMTLIB.S (SMTLIB.Sym "cl" :: concl) :: l, cx) = let val (concl, cx) = fold_map (fst oo SMTLIB_Proof.extract_and_update_name_bindings) concl cx in (SMTLIB.S (SMTLIB.Sym "or" :: concl), (l, cx)) end | parse_and_clausify_conclusion t = raise Fail ("unrecognized Lethe proof " ^ \<^make_string> t) val parse_normal_step = get_id_cx ##> parse_and_clausify_conclusion #> rotate_pair ##> parse_rule #> rotate_pair ##> parse_source #> rotate_pair ##> parse_args #> rotate_pair fun to_raw_node subproof ((((id, concl), rule), prems), args) = mk_raw_node id rule args (the_default [] prems) [] concl subproof fun at_discharge NONE _ = false | at_discharge (SOME id) p = p |> get_id |> fst |> (fn id2 => id = id2) in case step_kind ls of (NO_STEP, _, _) => ([],[], cx) | (NORMAL_STEP, p, l) => if at_discharge limit p then ([], ls, cx) else let (*ignores content of "discharge": Isabelle is keeping track of it via the context*) val (s, (_, cx)) = (p, cx) |> parse_normal_step |>> (to_raw_node []) val (rp, rl, cx) = parse_raw_proof_steps limit l cx in (s :: rp, rl, cx) end | (ASSUME, p, l) => let val (id, t :: []) = p |> get_id val ((t, cx), _) = SMTLIB_Proof.extract_and_update_name_bindings t cx val s = mk_raw_node id input_rule (SMTLIB.S []) [] [] t [] val (rp, rl, cx) = parse_raw_proof_steps limit l cx in (s :: rp, rl, cx) end | (ANCHOR, p, l) => let val (anchor_id, (anchor_args, (_, cx))) = (p, cx) |> (parse_anchor_step ##> parse_args) val (subproof, discharge_step :: remaining_proof, cx) = parse_raw_proof_steps (SOME anchor_id) l cx val (curss, (_, cx)) = parse_normal_step (discharge_step, cx) val s = to_raw_node subproof (fst curss, anchor_args) val (rp, rl, cx) = parse_raw_proof_steps limit remaining_proof cx in (s :: rp, rl, cx) end | (SKOLEM, p, l) => let val (s, cx) = parse_skolem p cx val (rp, rl, cx) = parse_raw_proof_steps limit l cx in (s :: rp, rl, cx) end end fun proof_ctxt_of_rule "bind" t = t | proof_ctxt_of_rule "sko_forall" t = t | proof_ctxt_of_rule "sko_ex" t = t | proof_ctxt_of_rule "let" t = t | proof_ctxt_of_rule "onepoint" t = t | proof_ctxt_of_rule _ _ = [] fun args_of_rule "bind" t = t | args_of_rule "la_generic" t = t | args_of_rule _ _ = [] fun insts_of_forall_inst "forall_inst" t = map (fn SMTLIB.S [_, SMTLIB.Sym x, a] => (x, a)) t | insts_of_forall_inst _ _ = [] fun id_of_last_step prems = if null prems then [] else let val Lethe_Replay_Node {id, ...} = List.last prems in [id] end fun extract_assumptions_from_subproof subproof = let fun extract_assumptions_from_subproof (Lethe_Replay_Node {rule, concl, ...}) assms = if rule = local_input_rule then concl :: assms else assms in fold extract_assumptions_from_subproof subproof [] end fun normalized_rule_name id rule = (case (rule = input_rule, can SMTLIB_Interface.role_and_index_of_assert_name id) of (true, true) => normalized_input_rule | (true, _) => local_input_rule | _ => rule) fun is_assm_repetition id rule = rule = input_rule andalso can SMTLIB_Interface.role_and_index_of_assert_name id fun extract_skolem ([SMTLIB.Sym var, typ, choice]) = (var, typ, choice) | extract_skolem t = raise Fail ("fail to parse type" ^ @{make_string} t) (* The preprocessing takes care of: 1. unfolding the shared terms 2. extract the declarations of skolems to make sure that there are not unfolded *) fun preprocess compress step = let fun expand_assms cs = map (fn t => case AList.lookup (op =) cs t of NONE => t | SOME a => a) fun expand_lonely_arguments (args as SMTLIB.S [SMTLIB.Sym "=", _, _]) = [args] | expand_lonely_arguments (x as SMTLIB.S [SMTLIB.Sym var, _]) = [SMTLIB.S [SMTLIB.Sym "=", x, SMTLIB.Sym var]] fun preprocess (Raw_Lethe_Node {id, rule, args, prems, concl, subproof, ...}) (cx, remap_assms) = let val (skolem_names, stripped_args) = args |> (fn SMTLIB.S args => args) |> map (fn SMTLIB.S [SMTLIB.Key "=", x, y] => SMTLIB.S [SMTLIB.Sym "=", x, y] | x => x) |> (rule = "bind" orelse rule = "onepoint") ? flat o (map expand_lonely_arguments) |> `(if rule = lethe_def then single o extract_skolem else K []) ||> SMTLIB.S val (subproof, (cx, _)) = fold_map preprocess subproof (cx, remap_assms) |> apfst flat val remap_assms = (if rule = "or" then (id, hd prems) :: remap_assms else remap_assms) (* declare variables in the context *) val declarations = if rule = lethe_def then skolem_names |> map (fn (name, _, choice) => (name, choice)) else [] in if compress andalso rule = "or" then ([], (cx, remap_assms)) else ([Raw_Lethe_Node {id = id, rule = rule, args = stripped_args, prems = expand_assms remap_assms prems, declarations = declarations, concl = concl, subproof = subproof}], (cx, remap_assms)) end in preprocess step end fun filter_split _ [] = ([], []) | filter_split f (a :: xs) = (if f a then apfst (curry op :: a) else apsnd (curry op :: a)) (filter_split f xs) fun extract_types_of_args (SMTLIB.S [var, typ, t as SMTLIB.S [SMTLIB.Sym "choice", _, _]]) = (SMTLIB.S [var, typ, t], SOME typ) |> single | extract_types_of_args (SMTLIB.S t) = let fun extract_types_of_arg (SMTLIB.S [eq, SMTLIB.S [var, typ], t]) = (SMTLIB.S [eq, var, t], SOME typ) | extract_types_of_arg t = (t, NONE) in t |> map extract_types_of_arg end fun collect_skolem_defs (Raw_Lethe_Node {rule, subproof = subproof, args, ...}) = (if is_skolemization rule then map (fn id => id ^ lethe_def) (skolems_introduced_by_rule args) else []) @ flat (map collect_skolem_defs subproof) (*The postprocessing does: 1. translate the terms to Isabelle syntax, taking care of free variables 2. remove the ambiguity in the proof terms: x \ y |- x = x means y = x. To remove ambiguity, we use the fact that y is a free variable and replace the term by: xy \ y |- xy = x. This is now does not have an ambiguity and we can safely move the "xy \ y" to the proof assumptions. *) fun postprocess_proof compress ctxt step cx = let fun postprocess (Raw_Lethe_Node {id, rule, args, prems, declarations, concl, subproof}) (cx, rew) = let val _ = (SMT_Config.verit_msg ctxt) (fn () => @{print} ("id =", id, "concl =", concl)) val args = extract_types_of_args args val globally_bound_vars = declared_csts cx rule args val cx = fold (update_binding o (fn (s, typ) => (s, Term (Free (s, type_of cx typ))))) globally_bound_vars cx (*find rebound variables specific to the LHS of the equivalence symbol*) val bound_vars = bound_vars_by_rule cx rule args val bound_vars_no_typ = map fst bound_vars val rhs_vars = fold (fn [t', t] => t <> t' ? (curry (op ::) t) | _ => fn x => x) bound_vars_no_typ [] fun not_already_bound cx t = SMTLIB_Proof.lookup_binding cx t = None andalso not (member (op =) rhs_vars t) val (shadowing_vars, rebound_lhs_vars) = bound_vars |> filter_split (fn ([t, _], typ) => not_already_bound cx t | _ => true) |>> map (apfst (hd)) |>> (fn vars => vars @ flat (map (fn ([_, t], typ) => [(t, typ)] | _ => []) bound_vars)) val subproof_rew = fold (fn [t, t'] => curry (op ::) (t, t ^ t')) (map fst rebound_lhs_vars) rew val subproof_rewriter = fold (fn (t, t') => synctatic_rew_in_lhs_subst t t') subproof_rew val ((concl, bounds), cx') = node_of concl cx val extra_lhs_vars = map (fn ([a,b], typ) => (a, a^b, typ)) rebound_lhs_vars val old_lhs_vars = map (fn (a, _, typ) => (a, typ)) extra_lhs_vars val new_lhs_vars = map (fn (_, newvar, typ) => (newvar, typ)) extra_lhs_vars (* postprocess conclusion *) val concl = SMTLIB_Isar.unskolemize_names ctxt (subproof_rewriter concl) val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> ("id =", id, "concl =", concl)) val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> ("id =", id, "cx' =", cx', "bound_vars =", bound_vars)) val bound_tvars = map (fn (s, SOME typ) => (s, type_of cx typ)) (shadowing_vars @ new_lhs_vars) val subproof_cx = add_bound_variables_to_ctxt cx (shadowing_vars @ new_lhs_vars) cx fun could_unify (Bound i, Bound j) = i = j | could_unify (Var v, Var v') = v = v' | could_unify (Free v, Free v') = v = v' | could_unify (Const (v, ty), Const (v', ty')) = v = v' andalso ty = ty' | could_unify (Abs (_, ty, bdy), Abs (_, ty', bdy')) = ty = ty' andalso could_unify (bdy, bdy') | could_unify (u $ v, u' $ v') = could_unify (u, u') andalso could_unify (v, v') | could_unify _ = false fun is_alpha_renaming t = t |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> could_unify handle TERM _ => false val alpha_conversion = rule = "bind" andalso is_alpha_renaming concl val can_remove_subproof = compress andalso (is_skolemization rule orelse alpha_conversion) val (fixed_subproof : lethe_replay_node list, _) = fold_map postprocess (if can_remove_subproof then [] else subproof) (subproof_cx, subproof_rew) val unsk_and_rewrite = SMTLIB_Isar.unskolemize_names ctxt o subproof_rewriter (* postprocess assms *) val stripped_args = map fst args val sanitized_args = proof_ctxt_of_rule rule stripped_args val arg_cx = add_bound_variables_to_ctxt cx (shadowing_vars @ old_lhs_vars) subproof_cx val (termified_args, _) = fold_map node_of sanitized_args arg_cx |> apfst (map fst) val normalized_args = map unsk_and_rewrite termified_args val subproof_assms = proof_ctxt_of_rule rule normalized_args (* postprocess arguments *) val rule_args = args_of_rule rule stripped_args val (termified_args, _) = fold_map term_of rule_args subproof_cx val normalized_args = map unsk_and_rewrite termified_args val rule_args = map subproof_rewriter normalized_args val raw_insts = insts_of_forall_inst rule stripped_args fun termify_term (x, t) cx = let val (t, cx) = term_of t cx in ((x, t), cx) end val (termified_args, _) = fold_map termify_term raw_insts subproof_cx val insts = Symtab.empty |> fold (fn (x, t) => fn insts => Symtab.update_new (x, t) insts) termified_args |> Symtab.map (K unsk_and_rewrite) (* declarations *) val (declarations, _) = fold_map termify_term declarations cx |> apfst (map (apsnd unsk_and_rewrite)) (* fix step *) val _ = if bounds <> [] then raise (Fail "found dangling variable in concl") else () - val skolem_defs = (if is_skolemization rule then map (fn id => id ^ lethe_def) (skolems_introduced_by_rule (SMTLIB.S (map fst args))) else []) - val skolems_of_subproof = (if is_skolemization rule + val skolems_of_subproof = (if compress andalso is_skolemization rule then flat (map collect_skolem_defs subproof) else []) val fixed_prems = prems @ (if is_assm_repetition id rule then [id] else []) @ skolem_defs @ skolems_of_subproof @ (id_of_last_step fixed_subproof) (* fix subproof *) val normalized_rule = normalized_rule_name id rule |> (if compress andalso alpha_conversion then K "refl" else I) val extra_assms2 = (if rule = subproof_rule then extract_assumptions_from_subproof fixed_subproof else []) val step = mk_replay_node id normalized_rule rule_args fixed_prems subproof_assms concl [] insts declarations (bound_tvars, subproof_assms, extra_assms2, fixed_subproof) in (step, (cx', rew)) end in postprocess step (cx, []) |> (fn (step, (cx, _)) => (step, cx)) end fun combine_proof_steps ((step1 : lethe_replay_node) :: step2 :: steps) = let val (Lethe_Replay_Node {id = id1, rule = rule1, args = args1, prems = prems1, proof_ctxt = proof_ctxt1, concl = concl1, bounds = bounds1, insts = insts1, declarations = declarations1, subproof = (bound_sub1, assms_sub1, assms_extra1, subproof1)}) = step1 val (Lethe_Replay_Node {id = id2, rule = rule2, args = args2, prems = prems2, proof_ctxt = proof_ctxt2, concl = concl2, bounds = bounds2, insts = insts2, declarations = declarations2, subproof = (bound_sub2, assms_sub2, assms_extra2, subproof2)}) = step2 val goals1 = (case concl1 of _ $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.Not\, _) $a) $ b)) => [a,b] | _ => []) val goal2 = (case concl2 of _ $ a => a) in if rule1 = equiv_pos2_rule andalso rule2 = th_resolution_rule andalso member (op =) prems2 id1 andalso member (op =) goals1 goal2 then mk_replay_node id2 theory_resolution2_rule args2 (filter_out (curry (op =) id1) prems2) proof_ctxt2 concl2 bounds2 insts2 declarations2 (bound_sub2, assms_sub2, assms_extra2, combine_proof_steps subproof2) :: combine_proof_steps steps else mk_replay_node id1 rule1 args1 prems1 proof_ctxt1 concl1 bounds1 insts1 declarations1 (bound_sub1, assms_sub1, assms_extra1, combine_proof_steps subproof1) :: combine_proof_steps (step2 :: steps) end | combine_proof_steps steps = steps val linearize_proof = let fun map_node_concl f (Lethe_Node {id, rule, prems, proof_ctxt, concl}) = mk_node id rule prems proof_ctxt (f concl) fun linearize (Lethe_Replay_Node {id = id, rule = rule, args = _, prems = prems, proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, insts = _, declarations = _, subproof = (bounds', assms, inputs, subproof)}) = let val bounds = distinct (op =) bounds val bounds' = distinct (op =) bounds' fun mk_prop_of_term concl = concl |> fastype_of concl = \<^typ>\bool\ ? curry (op $) \<^term>\Trueprop\ fun remove_assumption_id assumption_id prems = filter_out (curry (op =) assumption_id) prems fun add_assumption assumption concl = \<^Const>\Pure.imp for \mk_prop_of_term assumption\ \mk_prop_of_term concl\\ fun inline_assumption assumption assumption_id (Lethe_Node {id, rule, prems, proof_ctxt, concl}) = mk_node id rule (remove_assumption_id assumption_id prems) proof_ctxt (add_assumption assumption concl) fun find_input_steps_and_inline [] = [] | find_input_steps_and_inline (Lethe_Node {id = id', rule, prems, concl, ...} :: steps) = if rule = input_rule then find_input_steps_and_inline (map (inline_assumption concl id') steps) else mk_node (id') rule prems [] concl :: find_input_steps_and_inline steps fun free_bounds bounds (concl) = fold (fn (var, typ) => fn t => Logic.all (Free (var, typ)) t) bounds concl val subproof = subproof |> flat o map linearize |> map (map_node_concl (fold add_assumption (assms @ inputs))) |> map (map_node_concl (free_bounds (bounds @ bounds'))) |> find_input_steps_and_inline val concl = free_bounds bounds concl in subproof @ [mk_node id rule prems proof_ctxt concl] end in linearize end fun rule_of (Lethe_Replay_Node {rule,...}) = rule fun subproof_of (Lethe_Replay_Node {subproof = (_, _, _, subproof),...}) = subproof (* Massage Skolems for Sledgehammer. We have to make sure that there is an "arrow" in the graph for skolemization steps. A. The normal easy case This function detects the steps of the form P \ Q :skolemization Q :resolution with P and replace them by Q :skolemization Throwing away the step "P \ Q" completely. This throws away a lot of information, but it does not matter too much for Sledgehammer. B. Skolems in subproofs Supporting this is more or less hopeless as long as the Isar reconstruction of Sledgehammer does not support more features like definitions. lethe is able to generate proofs with skolemization happening in subproofs inside the formula. (assume "A \ P" ... P \ Q :skolemization in the subproof ...) hence A \ P \ A \ Q :lemma ... R :something with some rule and replace them by R :skolemization with some rule Without any subproof *) fun remove_skolem_definitions_proof steps = let fun replace_equivalent_by_imp (judgement $ ((Const(\<^const_name>\HOL.eq\, typ) $ arg1) $ arg2)) = judgement $ ((Const(\<^const_name>\HOL.implies\, typ) $ arg1) $ arg2) | replace_equivalent_by_imp a = a (*This case is probably wrong*) fun remove_skolem_definitions (Lethe_Replay_Node {id = id, rule = rule, args = args, prems = prems, proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, insts = insts, declarations = declarations, subproof = (vars, assms', extra_assms', subproof)}) (prems_to_remove, skolems) = let val prems = prems |> filter_out (member (op =) prems_to_remove) val trivial_step = is_SH_trivial rule fun has_skolem_substep st NONE = if is_skolemization (rule_of st) then SOME (rule_of st) else fold has_skolem_substep (subproof_of st) NONE | has_skolem_substep _ a = a val promote_to_skolem = exists (fn t => member (op =) skolems t) prems val promote_from_assms = fold has_skolem_substep subproof NONE <> NONE val promote_step = promote_to_skolem orelse promote_from_assms val skolem_step_to_skip = is_skolemization rule orelse (promote_from_assms andalso length prems > 1) val is_skolem = is_skolemization rule orelse promote_step val prems = prems |> filter_out (fn t => member (op =) skolems t) |> is_skolem ? filter_out (String.isPrefix id) val rule = (if promote_step then default_skolem_rule else rule) val subproof = subproof |> (is_skolem ? K []) (*subproofs of skolemization steps are useless for SH*) |> map (fst o (fn st => remove_skolem_definitions st (prems_to_remove, skolems))) (*no new definitions in subproofs*) |> flat val concl = concl |> is_skolem ? replace_equivalent_by_imp val step = (if skolem_step_to_skip orelse rule = lethe_def orelse trivial_step then [] else mk_replay_node id rule args prems proof_ctxt concl bounds insts declarations (vars, assms', extra_assms', subproof) |> single) val defs = (if rule = lethe_def orelse trivial_step then id :: prems_to_remove else prems_to_remove) val skolems = (if skolem_step_to_skip then id :: skolems else skolems) in (step, (defs, skolems)) end in fold_map remove_skolem_definitions steps ([], []) |> fst |> flat end local (*TODO useful?*) fun remove_pattern (SMTLIB.S (SMTLIB.Sym "!" :: t :: [SMTLIB.Key _, SMTLIB.S _])) = t | remove_pattern (SMTLIB.S xs) = SMTLIB.S (map remove_pattern xs) | remove_pattern p = p fun import_proof_and_post_process typs funs lines ctxt = let val compress = SMT_Config.compress_verit_proofs ctxt val smtlib_lines_without_qm = lines |> map single |> map SMTLIB.parse |> map remove_all_qm2 |> map remove_pattern val (raw_steps, _, _) = parse_raw_proof_steps NONE smtlib_lines_without_qm SMTLIB_Proof.empty_name_binding fun process step (cx, cx') = let fun postprocess step (cx, cx') = let val (step, cx) = postprocess_proof compress ctxt step cx in (step, (cx, cx')) end in uncurry (fold_map postprocess) (preprocess compress step (cx, cx')) end val step = (empty_context ctxt typs funs, []) |> fold_map process raw_steps |> (fn (steps, (cx, _)) => (flat steps, cx)) |> compress? apfst combine_proof_steps in step end in fun parse typs funs lines ctxt = let val (u, env) = import_proof_and_post_process typs funs lines ctxt val t = u |> remove_skolem_definitions_proof |> flat o (map linearize_proof) fun node_to_step (Lethe_Node {id, rule, prems, concl, ...}) = mk_step id rule prems [] concl [] in (map node_to_step t, ctxt_of env) end fun parse_replay typs funs lines ctxt = let val (u, env) = import_proof_and_post_process typs funs lines ctxt val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> u) in (u, ctxt_of env) end end end; diff --git a/src/HOL/Tools/SMT/lethe_replay_methods.ML b/src/HOL/Tools/SMT/lethe_replay_methods.ML --- a/src/HOL/Tools/SMT/lethe_replay_methods.ML +++ b/src/HOL/Tools/SMT/lethe_replay_methods.ML @@ -1,1191 +1,1217 @@ (* Title: HOL/Tools/SMT/verit_replay_methods.ML Author: Mathias Fleury, MPII, JKU, University Freiburg Proof method for replaying veriT proofs. *) signature LETHE_REPLAY_METHODS = sig datatype verit_rule = False | True | (*input: a repeated (normalized) assumption of assumption of in a subproof*) Normalized_Input | Local_Input | (*Subproof:*) Subproof | (*Conjunction:*) And | Not_And | And_Pos | And_Neg | (*Disjunction""*) Or | Or_Pos | Not_Or | Or_Neg | (*Disjunction:*) Implies | Implies_Neg1 | Implies_Neg2 | Implies_Pos | Not_Implies1 | Not_Implies2 | (*Equivalence:*) Equiv_neg1 | Equiv_pos1 | Equiv_pos2 | Equiv_neg2 | Not_Equiv1 | Not_Equiv2 | Equiv1 | Equiv2 | (*If-then-else:*) ITE_Pos1 | ITE_Pos2 | ITE_Neg1 | ITE_Neg2 | Not_ITE1 | Not_ITE2 | ITE_Intro | ITE1 | ITE2 | (*Equality:*) Eq_Congruent | Eq_Reflexive | Eq_Transitive | Eq_Congruent_Pred | Trans | Refl | Cong | (*Arithmetics:*) LA_Disequality | LA_Generic | LA_Tautology | LIA_Generic | LA_Totality | LA_RW_Eq | NLA_Generic | (*Quantifiers:*) Forall_Inst | Qnt_Rm_Unused | Qnt_Join | Onepoint | Bind | Skolem_Forall | Skolem_Ex | (*Resolution:*) Theory_Resolution | Resolution | (*Temporary rules, that the verit developers want to remove:*) AC_Simp | Bfun_Elim | Qnt_CNF | (*Used to introduce skolem constants*) Definition | (*Former cong rules*) Bool_Simplify | Or_Simplify | Not_Simplify | And_Simplify | Equiv_Simplify | Implies_Simplify | Connective_Def | Minus_Simplify | Comp_Simplify | Eq_Simplify | ITE_Simplify | Sum_Simplify | Prod_Simplify | All_Simplify | Qnt_Simplify | Not_Not | Tautological_Clause | Duplicate_Literals | (*Unsupported rule*) Unsupported | (*For compression*) Theory_Resolution2 | (*Extended rules*) Symm | Not_Symm | Reordering | Tmp_Quantifier_Simplify val requires_subproof_assms : string list -> string -> bool val requires_local_input: string list -> string -> bool val string_of_verit_rule: verit_rule -> string type lethe_tac_args = Proof.context -> thm list -> term list -> term -> thm type lethe_tac = Proof.context -> thm list -> term -> thm type lethe_tac2 = Proof.context -> thm list -> term list -> term Symtab.table -> term -> thm val bind: lethe_tac_args val and_rule: lethe_tac val and_neg_rule: lethe_tac val and_pos: lethe_tac val rewrite_and_simplify: lethe_tac val rewrite_bool_simplify: lethe_tac val rewrite_comp_simplify: lethe_tac val rewrite_minus_simplify: lethe_tac val rewrite_not_simplify: lethe_tac val rewrite_eq_simplify: lethe_tac val rewrite_equiv_simplify: lethe_tac val rewrite_implies_simplify: lethe_tac val rewrite_or_simplify: lethe_tac val cong: lethe_tac val rewrite_connective_def: lethe_tac val duplicate_literals: lethe_tac val eq_congruent: lethe_tac val eq_congruent_pred: lethe_tac val eq_reflexive: lethe_tac val eq_transitive: lethe_tac val equiv1: lethe_tac val equiv2: lethe_tac val equiv_neg1: lethe_tac val equiv_neg2: lethe_tac val equiv_pos1: lethe_tac val equiv_pos2: lethe_tac val false_rule: lethe_tac val forall_inst: lethe_tac2 val implies_rules: lethe_tac val implies_neg1: lethe_tac val implies_neg2: lethe_tac val implies_pos: lethe_tac val ite1: lethe_tac val ite2: lethe_tac val ite_intro: lethe_tac val ite_neg1: lethe_tac val ite_neg2: lethe_tac val ite_pos1: lethe_tac val ite_pos2: lethe_tac val rewrite_ite_simplify: lethe_tac val la_disequality: lethe_tac val la_generic: lethe_tac_args val la_rw_eq: lethe_tac val lia_generic: lethe_tac val refl: lethe_tac val normalized_input: lethe_tac val not_and_rule: lethe_tac val not_equiv1: lethe_tac val not_equiv2: lethe_tac val not_implies1: lethe_tac val not_implies2: lethe_tac val not_ite1: lethe_tac val not_ite2: lethe_tac val not_not: lethe_tac val not_or_rule: lethe_tac val or: lethe_tac val or_neg_rule: lethe_tac val or_pos_rule: lethe_tac val theory_resolution2: lethe_tac val prod_simplify: lethe_tac val qnt_join: lethe_tac val qnt_rm_unused: lethe_tac val onepoint: lethe_tac val qnt_simplify: lethe_tac val all_simplify: lethe_tac val unit_res: lethe_tac val skolem_ex: lethe_tac val skolem_forall: lethe_tac val subproof: lethe_tac val sum_simplify: lethe_tac val tautological_clause: lethe_tac val tmp_AC_rule: lethe_tac val bfun_elim: lethe_tac val qnt_cnf: lethe_tac val trans: lethe_tac val symm: lethe_tac val not_symm: lethe_tac val reordering: lethe_tac (* val : lethe_tac *) val REPEAT_CHANGED: ('a -> tactic) -> 'a -> tactic val TRY': ('a -> tactic) -> 'a -> tactic end; structure Lethe_Replay_Methods: LETHE_REPLAY_METHODS = struct type lethe_tac_args = Proof.context -> thm list -> term list -> term -> thm type lethe_tac = Proof.context -> thm list -> term -> thm type lethe_tac2 = Proof.context -> thm list -> term list -> term Symtab.table -> term -> thm (*Some general comments on the proof format: 1. Double negations are not always removed. This means for example that the equivalence rules cannot assume that the double negations have already been removed. Therefore, we match the term, instantiate the theorem, then use simp (to remove double negations), and finally use assumption. 2. The reconstruction for rule forall_inst is buggy and tends to be very fragile, because the rule is doing much more that is supposed to do. Moreover types can make trivial goals (for the boolean structure) impossible to prove. 3. Duplicate literals are sometimes removed, mostly by the SAT solver. Rules unsupported on purpose: * Distinct_Elim, XOR, let (we don't need them). * deep_skolemize (because it is not clear if verit still generates using it). *) datatype verit_rule = False | True | (*input: a repeated (normalized) assumption of assumption of in a subproof*) Normalized_Input | Local_Input | (*Subproof:*) Subproof | (*Conjunction:*) And | Not_And | And_Pos | And_Neg | (*Disjunction""*) Or | Or_Pos | Not_Or | Or_Neg | (*Disjunction:*) Implies | Implies_Neg1 | Implies_Neg2 | Implies_Pos | Not_Implies1 | Not_Implies2 | (*Equivalence:*) Equiv_neg1 | Equiv_pos1 | Equiv_pos2 | Equiv_neg2 | Not_Equiv1 | Not_Equiv2 | Equiv1 | Equiv2 | (*If-then-else:*) ITE_Pos1 | ITE_Pos2 | ITE_Neg1 | ITE_Neg2 | Not_ITE1 | Not_ITE2 | ITE_Intro | ITE1 | ITE2 | (*Equality:*) Eq_Congruent | Eq_Reflexive | Eq_Transitive | Eq_Congruent_Pred | Trans | Refl | Cong | (*Arithmetics:*) LA_Disequality | LA_Generic | LA_Tautology | LIA_Generic | LA_Totality | LA_RW_Eq | NLA_Generic | (*Quantifiers:*) Forall_Inst | Qnt_Rm_Unused | Qnt_Join | Onepoint | Bind | Skolem_Forall | Skolem_Ex | (*Resolution:*) Theory_Resolution | Resolution | (*Temporary rules, that the verit developpers want to remove:*) AC_Simp | Bfun_Elim | Qnt_CNF | (*Used to introduce skolem constants*) Definition | (*Former cong rules*) Bool_Simplify | Or_Simplify | Not_Simplify | And_Simplify | Equiv_Simplify | Implies_Simplify | Connective_Def | Minus_Simplify | Comp_Simplify | Eq_Simplify | ITE_Simplify | Sum_Simplify | Prod_Simplify | All_Simplify | Qnt_Simplify | Not_Not | Tautological_Clause | Duplicate_Literals | (*Unsupported rule*) Unsupported | (*For compression*) Theory_Resolution2 | (*Extended rules*) Symm | Not_Symm | Reordering | Tmp_Quantifier_Simplify fun string_of_verit_rule Bind = "Bind" | string_of_verit_rule Cong = "Cong" | string_of_verit_rule Refl = "Refl" | string_of_verit_rule Equiv1 = "Equiv1" | string_of_verit_rule Equiv2 = "Equiv2" | string_of_verit_rule Equiv_pos1 = "Equiv_pos1" | string_of_verit_rule Equiv_pos2 = "Equiv_pos2" | string_of_verit_rule Equiv_neg1 = "Equiv_neg1" | string_of_verit_rule Equiv_neg2 = "Equiv_neg2" | string_of_verit_rule Skolem_Forall = "Skolem_Forall" | string_of_verit_rule Skolem_Ex = "Skolem_Ex" | string_of_verit_rule Eq_Reflexive = "Eq_Reflexive" | string_of_verit_rule Theory_Resolution = "Theory_Resolution" | string_of_verit_rule Theory_Resolution2 = "Theory_Resolution2" | string_of_verit_rule Forall_Inst = "forall_inst" | string_of_verit_rule Or = "Or" | string_of_verit_rule Not_Or = "Not_Or" | string_of_verit_rule Resolution = "Resolution" | string_of_verit_rule Eq_Congruent = "eq_congruent" | string_of_verit_rule Trans = "trans" | string_of_verit_rule False = "false" | string_of_verit_rule And = "and" | string_of_verit_rule And_Pos = "and_pos" | string_of_verit_rule Not_And = "not_and" | string_of_verit_rule And_Neg = "and_neg" | string_of_verit_rule Or_Pos = "or_pos" | string_of_verit_rule Or_Neg = "or_neg" | string_of_verit_rule AC_Simp = "ac_simp" | string_of_verit_rule Not_Equiv1 = "not_equiv1" | string_of_verit_rule Not_Equiv2 = "not_equiv2" | string_of_verit_rule Not_Implies1 = "not_implies1" | string_of_verit_rule Not_Implies2 = "not_implies2" | string_of_verit_rule Implies_Neg1 = "implies_neg1" | string_of_verit_rule Implies_Neg2 = "implies_neg2" | string_of_verit_rule Implies = "implies" | string_of_verit_rule Bfun_Elim = "bfun_elim" | string_of_verit_rule ITE1 = "ite1" | string_of_verit_rule ITE2 = "ite2" | string_of_verit_rule Not_ITE1 = "not_ite1" | string_of_verit_rule Not_ITE2 = "not_ite2" | string_of_verit_rule ITE_Pos1 = "ite_pos1" | string_of_verit_rule ITE_Pos2 = "ite_pos2" | string_of_verit_rule ITE_Neg1 = "ite_neg1" | string_of_verit_rule ITE_Neg2 = "ite_neg2" | string_of_verit_rule ITE_Intro = "ite_intro" | string_of_verit_rule LA_Disequality = "la_disequality" | string_of_verit_rule LA_Generic = "la_generic" | string_of_verit_rule LIA_Generic = "lia_generic" | string_of_verit_rule LA_Tautology = "la_tautology" | string_of_verit_rule LA_RW_Eq = "la_rw_eq" | string_of_verit_rule LA_Totality = "LA_Totality" | string_of_verit_rule NLA_Generic = "nla_generic" | string_of_verit_rule Eq_Transitive = "eq_transitive" | string_of_verit_rule Qnt_Rm_Unused = "qnt_remove_unused" | string_of_verit_rule Onepoint = "onepoint" | string_of_verit_rule Qnt_Join = "qnt_join" | string_of_verit_rule Eq_Congruent_Pred = "eq_congruent_pred" | string_of_verit_rule Normalized_Input = Lethe_Proof.normalized_input_rule | string_of_verit_rule Local_Input = Lethe_Proof.local_input_rule | string_of_verit_rule Subproof = "subproof" | string_of_verit_rule Bool_Simplify = "bool_simplify" | string_of_verit_rule Equiv_Simplify = "equiv_simplify" | string_of_verit_rule Eq_Simplify = "eq_simplify" | string_of_verit_rule Not_Simplify = "not_simplify" | string_of_verit_rule And_Simplify = "and_simplify" | string_of_verit_rule Or_Simplify = "or_simplify" | string_of_verit_rule ITE_Simplify = "ite_simplify" | string_of_verit_rule Implies_Simplify = "implies_simplify" | string_of_verit_rule Connective_Def = "connective_def" | string_of_verit_rule Minus_Simplify = "minus_simplify" | string_of_verit_rule Sum_Simplify = "sum_simplify" | string_of_verit_rule Prod_Simplify = "prod_simplify" | string_of_verit_rule All_Simplify = "all_simplify" | string_of_verit_rule Comp_Simplify = "comp_simplify" | string_of_verit_rule Qnt_Simplify = "qnt_simplify" | string_of_verit_rule Symm = "symm" | string_of_verit_rule Not_Symm = "not_symm" | string_of_verit_rule Reordering = "reordering" | string_of_verit_rule Not_Not = Lethe_Proof.not_not_rule | string_of_verit_rule Tautological_Clause = "tautology" | string_of_verit_rule Duplicate_Literals = Lethe_Proof.contract_rule | string_of_verit_rule Qnt_CNF = "qnt_cnf" | string_of_verit_rule r = "Unknown rule: " ^ \<^make_string> r fun replay_error ctxt msg rule thms t = SMT_Replay_Methods.replay_error ctxt msg (string_of_verit_rule rule) thms t (* utility function *) fun eqsubst_all ctxt thms = K ((REPEAT o HEADGOAL) (EqSubst.eqsubst_asm_tac ctxt [0] thms)) THEN' K ((REPEAT o HEADGOAL) (EqSubst.eqsubst_tac ctxt [0] thms)) fun simplify_tac ctxt thms = ctxt |> empty_simpset |> put_simpset HOL_basic_ss |> (fn ctxt => ctxt addsimps @{thms simp_thms} addsimps thms) |> Simplifier.asm_full_simp_tac (* sko_forall requires the assumptions to be able to prove the equivalence in case of double skolemization. See comment below. *) fun requires_subproof_assms _ t = member (op =) ["refl", "sko_forall", "sko_ex", "cong"] t fun requires_local_input _ t = member (op =) [Lethe_Proof.local_input_rule] t (*This is a weaker simplification form. It is weaker, but is also less likely to loop*) fun partial_simplify_tac ctxt thms = ctxt |> empty_simpset |> put_simpset HOL_basic_ss |> (fn ctxt => ctxt addsimps @{thms simp_thms} addsimps thms) |> Simplifier.full_simp_tac val try_provers = SMT_Replay_Methods.try_provers "verit" fun TRY' tac = fn i => TRY (tac i) fun REPEAT' tac = fn i => REPEAT (tac i) fun REPEAT_CHANGED tac = fn i => REPEAT (CHANGED (tac i)) (* Bind *) (*The bind rule is non-obvious due to the handling of quantifiers: "\x y a. x = y ==> (\b. P a b x) \ (\b. P' a b y)" ------------------------------------------------------ \a. (\b x. P a b x) \ (\b y. P' a b y) is a valid application.*) val bind_thms = [@{lemma \(\x x'. x = x' \ P x = Q x') \ (\x. P x) = (\y. Q y)\ by blast}, @{lemma \(\x x'. x = x' \ P x = Q x') \ (\x. P x) = (\y. Q y)\ by blast}, @{lemma \(\x x'. x = x' \ P x = Q x') \ (\x. P x = Q x)\ by blast}, @{lemma \(\x x'. x = x' \ P x = Q x') \ (\x. P x = Q x)\ by blast}] val bind_thms_same_name = [@{lemma \(\x. P x = Q x) \ (\x. P x) = (\y. Q y)\ by blast}, @{lemma \(\x. P x = Q x) \ (\x. P x) = (\y. Q y)\ by blast}, @{lemma \(\x. P x = Q x) \ (\x. P x = Q x)\ by blast}, @{lemma \(\x. P x = Q x) \ (\x. P x = Q x)\ by blast}] fun extract_quantified_names_q (_ $ Abs (name, _, t)) = apfst (curry (op ::) name) (extract_quantified_names_q t) | extract_quantified_names_q t = ([], t) fun extract_forall_quantified_names_q (Const(\<^const_name>\HOL.All\, _) $ Abs (name, ty, t)) = (name, ty) :: (extract_forall_quantified_names_q t) | extract_forall_quantified_names_q _ = [] fun extract_all_forall_quantified_names_q (Const(\<^const_name>\HOL.All\, _) $ Abs (name, _, t)) = name :: (extract_all_forall_quantified_names_q t) | extract_all_forall_quantified_names_q (t $ u) = extract_all_forall_quantified_names_q t @ extract_all_forall_quantified_names_q u | extract_all_forall_quantified_names_q _ = [] val extract_quantified_names_ba = SMT_Replay_Methods.dest_prop #> extract_quantified_names_q ##> HOLogic.dest_eq ##> fst ##> extract_quantified_names_q ##> fst val extract_quantified_names = extract_quantified_names_ba #> (op @) val extract_all_forall_quantified_names = SMT_Replay_Methods.dest_prop #> HOLogic.dest_eq #> fst #> extract_all_forall_quantified_names_q fun extract_all_exists_quantified_names_q (Const(\<^const_name>\HOL.Ex\, _) $ Abs (name, _, t)) = name :: (extract_all_exists_quantified_names_q t) | extract_all_exists_quantified_names_q (t $ u) = extract_all_exists_quantified_names_q t @ extract_all_exists_quantified_names_q u | extract_all_exists_quantified_names_q _ = [] val extract_all_exists_quantified_names = SMT_Replay_Methods.dest_prop #> HOLogic.dest_eq #> fst #> extract_all_exists_quantified_names_q +fun extract_all_forall_exists_quantified_names_q (Const(\<^const_name>\HOL.Ex\, _) $ Abs (name, _, t)) = + name :: (extract_all_forall_exists_quantified_names_q t) + | extract_all_forall_exists_quantified_names_q (Const(\<^const_name>\HOL.All\, _) $ Abs (name, _, t)) = + name :: (extract_all_forall_exists_quantified_names_q t) + | extract_all_forall_exists_quantified_names_q (t $ u) = + extract_all_forall_exists_quantified_names_q t @ extract_all_forall_exists_quantified_names_q u + | extract_all_forall_exists_quantified_names_q _ = [] + val extract_bind_names = HOLogic.dest_eq #> apply2 (fn (Free (name, _)) => name) fun combine_quant ctxt ((n1, n2) :: bounds) (n1' :: formula) = TRY' (if n1 = n1' then if n1 <> n2 then resolve_tac ctxt bind_thms THEN' TRY'(resolve_tac ctxt [@{thm refl}]) THEN' combine_quant ctxt bounds formula else resolve_tac ctxt bind_thms_same_name THEN' combine_quant ctxt bounds formula else resolve_tac ctxt @{thms allI} THEN' combine_quant ctxt ((n1, n2) :: bounds) formula) | combine_quant _ _ _ = K all_tac fun bind_quants ctxt args t = combine_quant ctxt (map extract_bind_names args) (extract_quantified_names t) fun generalize_prems_q [] prems = prems | generalize_prems_q (_ :: quants) prems = generalize_prems_q quants (@{thm spec} OF [prems]) fun generalize_prems t = generalize_prems_q (fst (extract_quantified_names_ba t)) fun bind ctxt [prems] args t = SMT_Replay_Methods.prove ctxt t (fn _ => bind_quants ctxt args t THEN' TRY' (SOLVED' (resolve_tac ctxt [generalize_prems t prems] THEN_ALL_NEW TRY' (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms refl})))) | bind ctxt thms _ t = replay_error ctxt "invalid bind application" Bind thms t (* Congruence/Refl *) fun cong ctxt thms = try_provers ctxt (string_of_verit_rule Cong) [ ("basic", SMT_Replay_Methods.cong_basic ctxt thms), ("unfolding then reflexivity", SMT_Replay_Methods.cong_unfolding_trivial ctxt thms), ("unfolding then auto", SMT_Replay_Methods.cong_unfolding_first ctxt thms), ("full", SMT_Replay_Methods.cong_full ctxt thms)] thms fun refl ctxt thm t = (case find_first (fn thm => t = Thm.full_prop_of thm) thm of SOME thm => thm | NONE => (case try (fn t => SMT_Replay_Methods.match_instantiate ctxt t @{thm refl}) t of NONE => cong ctxt thm t | SOME thm => thm)) (* Instantiation *) local fun dropWhile _ [] = [] | dropWhile f (x :: xs) = if f x then dropWhile f xs else x :: xs in fun forall_inst ctxt _ _ insts t = let fun instantiate_and_solve i ({context = ctxt, prems = [prem], ...}: Subgoal.focus) = let val t = Thm.prop_of prem val quants = t |> SMT_Replay_Methods.dest_prop |> extract_forall_quantified_names_q val insts = map (Symtab.lookup insts o fst) (rev quants) |> dropWhile (curry (op =) NONE) |> rev fun option_map _ NONE = NONE | option_map f (SOME a) = SOME (f a) fun instantiate_with inst prem = Drule.infer_instantiate' ctxt [NONE, inst] @{thm spec} OF [prem] val inst_thm = fold instantiate_with (map (option_map (Thm.cterm_of ctxt)) insts) prem in (Method.insert_tac ctxt [inst_thm] THEN' TRY' (fn i => assume_tac ctxt i) - THEN' TRY' (partial_simplify_tac ctxt @{thms eq_commute})) i + THEN' TRY' (partial_simplify_tac ctxt @{thms eq_commute ac_simps}) + THEN' TRY' (blast_tac ctxt)) i end | instantiate_and_solve _ ({context = ctxt, prems = thms, ...}: Subgoal.focus) = replay_error ctxt "invalid application" Forall_Inst thms t in SMT_Replay_Methods.prove ctxt t (fn _ => match_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF [@{thm impI}]] THEN' (fn i => Subgoal.FOCUS (instantiate_and_solve i) ctxt i)) end end (* Or *) fun or _ (thm :: _) _ = thm | or ctxt thms t = replay_error ctxt "invalid bind application" Or thms t (* Implication *) val implies_pos_thm = [@{lemma \\(A \ B) \ \A \ B\ by blast}, @{lemma \\(\A \ B) \ A \ B\ by blast}] fun implies_pos ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt implies_pos_thm) (* Skolemization *) - +local + fun split _ [] = ([], []) + | split f (a :: xs) = + split f xs + |> (if f a then apfst (curry (op ::) a) else apsnd (curry (op ::) a)) +in fun extract_rewrite_rule_assumption _ thms = let fun is_rewrite_rule thm = (case Thm.prop_of thm of \<^term>\Trueprop\ $ (Const(\<^const_name>\HOL.eq\, _) $ _ $ Free(_, _)) => true | _ => false) fun is_context_rule thm = (case Thm.prop_of thm of \<^term>\Trueprop\ $ (Const(\<^const_name>\HOL.eq\, _) $ Free(_, _) $ Free(_, _)) => true | _ => false) - val ctxt_eq = + val (ctxt_eq, other) = thms - |> filter is_context_rule - val rew = - thms - |> filter_out is_context_rule - |> filter is_rewrite_rule + |> split is_context_rule + val (rew, other) = + other + |> split is_rewrite_rule in - (ctxt_eq, rew) + (ctxt_eq, rew, other) end +end +(* +Without compression, we have to rewrite skolems only once. However, it can happen than the same +skolem constant is used multiple times with a different name under the forall. +For strictness, we use the multiple rewriting only when compressing is activated. +*) local - fun rewrite_all_skolems thm_indirect ctxt (SOME thm :: thms) = - EqSubst.eqsubst_tac ctxt [0] [thm_indirect OF [thm]] - THEN' (partial_simplify_tac ctxt (@{thms eq_commute})) - THEN' rewrite_all_skolems thm_indirect ctxt thms - | rewrite_all_skolems thm_indirect ctxt (NONE :: thms) = rewrite_all_skolems thm_indirect ctxt thms + fun rewrite_all_skolems thm_indirect ctxt ((v,SOME thm) :: thms) = + let + val rewrite_sk_thms = + List.mapPartial (fn tm => SOME (tm OF [thm]) handle THM _ => NONE) thm_indirect + val multiple_rew = if SMT_Config.compress_verit_proofs ctxt then REPEAT_CHANGED else fn x => x + in + multiple_rew (EqSubst.eqsubst_tac ctxt [0] rewrite_sk_thms + THEN' SOLVED' (K (HEADGOAL (partial_simplify_tac ctxt (@{thms eq_commute}))))) + THEN' rewrite_all_skolems thm_indirect ctxt thms + end + | rewrite_all_skolems thm_indirect ctxt ((_,NONE) :: thms) = rewrite_all_skolems thm_indirect ctxt thms | rewrite_all_skolems _ _ [] = K (all_tac) fun extract_var_name (thm :: thms) = let val name = Thm.concl_of thm |> SMT_Replay_Methods.dest_prop |> HOLogic.dest_eq |> fst |> (fn Const(_,_) $ Abs(name, _, _) => [(name,@{thm sym} OF [thm])] | _ => []) in name :: extract_var_name thms end | extract_var_name [] = [] fun skolem_tac extractor thm1 thm2 ctxt thms t = let - val (ctxt_eq, ts) = extract_rewrite_rule_assumption ctxt thms + val (ctxt_eq, ts, other) = extract_rewrite_rule_assumption ctxt thms + fun ordered_definitions () = let val var_order = extractor t val thm_names_with_var = extract_var_name ts |> flat - in map (AList.lookup (op =) thm_names_with_var) var_order end - + in map (fn v => (v, AList.lookup (op =) thm_names_with_var v)) var_order end in SMT_Replay_Methods.prove ctxt t (fn _ => K (unfold_tac ctxt ctxt_eq) - THEN' ((SOLVED' (K (unfold_tac ctxt (map (fn thm => thm1 OF [@{thm sym} OF [thm]]) ts)))) - ORELSE' - (rewrite_all_skolems thm2 ctxt (ordered_definitions ()) - THEN' partial_simplify_tac ctxt @{thms eq_commute}))) + THEN' rewrite_all_skolems thm2 ctxt (ordered_definitions ()) + THEN' (eqsubst_all ctxt (map (fn thm => thm RS sym) other)) + THEN_ALL_NEW TRY' (resolve_tac ctxt @{thms refl}) + THEN' K (unfold_tac ctxt ctxt_eq) + THEN' TRY' (partial_simplify_tac ctxt (@{thms eq_commute}))) end in val skolem_forall = - skolem_tac extract_all_forall_quantified_names @{thm verit_sko_forall_indirect} - @{thm verit_sko_forall_indirect2} + skolem_tac extract_all_forall_exists_quantified_names_q @{thm verit_sko_forall_indirect} + @{thms verit_sko_forall_indirect2 verit_sko_ex_indirect2} val skolem_ex = - skolem_tac extract_all_exists_quantified_names @{thm verit_sko_ex_indirect} - @{thm verit_sko_ex_indirect2} + skolem_tac extract_all_forall_exists_quantified_names_q @{thm verit_sko_ex_indirect} + @{thms verit_sko_ex_indirect2 verit_sko_forall_indirect2} end fun eq_reflexive ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t @{thm refl} local fun not_not_prove ctxt _ = K (unfold_tac ctxt @{thms not_not}) THEN' match_tac ctxt @{thms verit_or_simplify_1} fun duplicate_literals_prove ctxt prems = Method.insert_tac ctxt prems THEN' match_tac ctxt @{thms ccontr} THEN' K (unfold_tac ctxt @{thms de_Morgan_disj not_not}) THEN' TRY o CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms conjE}) THEN' CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms disjE}) THEN' REPEAT' (ematch_tac ctxt @{thms notE} THEN' TRY' (assume_tac ctxt)) fun tautological_clause_prove ctxt _ = match_tac ctxt @{thms verit_or_neg} THEN' K (unfold_tac ctxt @{thms not_not disj_assoc[symmetric]}) THEN' TRY' (match_tac ctxt @{thms notE} THEN_ALL_NEW assume_tac ctxt) val theory_resolution2_lemma = @{lemma \a \ a = b \ b\ by blast} in fun prove_abstract abstracter tac ctxt thms t = let val thms = map (Conv.fconv_rule Thm.eta_long_conversion) thms val t' = Thm.eta_long_conversion (Object_Logic.dest_judgment ctxt (Thm.cterm_of ctxt t)) val (_, t2) = Logic.dest_equals (Thm.prop_of t') val thm = SMT_Replay_Methods.prove_abstract ctxt thms t2 tac ( fold_map (abstracter o SMT_Replay_Methods.dest_thm) thms ##>> abstracter (SMT_Replay_Methods.dest_prop t2)) in @{thm verit_Pure_trans} OF [t', thm] end val not_not = prove_abstract SMT_Replay_Methods.abstract_bool_shallow not_not_prove val tautological_clause = prove_abstract SMT_Replay_Methods.abstract_bool_shallow tautological_clause_prove val duplicate_literals = prove_abstract SMT_Replay_Methods.abstract_bool_shallow duplicate_literals_prove val unit_res = prove_abstract SMT_Replay_Methods.abstract_bool_shallow SMT_Replay_Methods.prop_tac (*match_instantiate does not work*) fun theory_resolution2 ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => match_tac ctxt [theory_resolution2_lemma OF prems]) end fun normalized_input ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt prems THEN' TRY' (K (unfold_tac ctxt @{thms SMT.trigger_def})) THEN' TRY' (partial_simplify_tac ctxt @{thms eq_commute})) val false_rule_thm = @{lemma \\False\ by blast} fun false_rule ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t false_rule_thm (* Transitivity *) val trans_bool_thm = @{lemma \P = Q \ Q \ P\ by blast} fun trans ctxt thms t = let val prop_of = HOLogic.dest_Trueprop o Thm.full_prop_of fun combine_thms [thm1, thm2] = (case (prop_of thm1, prop_of thm2) of ((Const(\<^const_name>\HOL.eq\, _) $ t1 $ t2), (Const(\<^const_name>\HOL.eq\, _) $ t3 $ t4)) => if t2 aconv t3 then thm1 RSN (1, thm2 RSN (2, @{thm trans})) else if t2 aconv t4 then thm1 RSN ((1, ((thm2 RS sym)) RSN (2, @{thm trans}))) else if t1 aconv t4 then thm2 RSN (1, thm1 RSN (2, @{thm trans})) else raise Fail "invalid trans theorem" | _ => trans_bool_thm OF [thm1, thm2]) | combine_thms (thm1 :: thm2 :: thms) = combine_thms (combine_thms [thm1, thm2] :: thms) | combine_thms thms = replay_error ctxt "invalid bind application" Trans thms t val t' = Thm.eta_long_conversion (Object_Logic.dest_judgment ctxt (Thm.cterm_of ctxt t)) val (_, t2) = Logic.dest_equals (Thm.prop_of t') val thms = map (Conv.fconv_rule Thm.eta_long_conversion) thms val trans_thm = combine_thms thms in (case (prop_of trans_thm, t2) of ((Const(\<^const_name>\HOL.eq\, _) $ t1 $ _), (Const(\<^const_name>\HOL.eq\, _) $ t3 $ _)) => if t1 aconv t3 then trans_thm else trans_thm RS sym | _ => trans_thm (*to be on the safe side*)) end fun tmp_AC_rule ctxt thms t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt thms THEN' REPEAT_ALL_NEW (partial_simplify_tac ctxt @{thms eq_commute} THEN' TRY' (simplify_tac ctxt @{thms ac_simps conj_ac}) THEN' TRY' (Classical.fast_tac ctxt))) (* And/Or *) local fun not_and_rule_prove ctxt prems = Method.insert_tac ctxt prems THEN' K (unfold_tac ctxt @{thms de_Morgan_conj}) THEN_ALL_NEW TRY' (assume_tac ctxt) fun or_pos_prove ctxt _ = K (unfold_tac ctxt @{thms de_Morgan_disj not_not}) THEN' match_tac ctxt @{thms verit_and_pos} THEN' K (unfold_tac ctxt @{thms de_Morgan_conj de_Morgan_disj not_not}) THEN' TRY' (assume_tac ctxt) fun not_or_rule_prove ctxt prems = Method.insert_tac ctxt prems THEN' K (unfold_tac ctxt @{thms de_Morgan_disj not_not}) THEN' TRY' (REPEAT' (match_tac ctxt @{thms conjI})) THEN_ALL_NEW ((REPEAT' (ematch_tac ctxt @{thms conjE})) THEN' TRY' (assume_tac ctxt)) fun and_rule_prove ctxt prems = Method.insert_tac ctxt prems THEN' (fn i => REPEAT (dresolve_tac ctxt @{thms conjE} i THEN assume_tac ctxt (i+1))) THEN' TRY' (assume_tac ctxt) fun and_neg_rule_prove ctxt _ = match_tac ctxt @{thms verit_and_neg} THEN' K (unfold_tac ctxt @{thms de_Morgan_conj not_not}) THEN' TRY' (assume_tac ctxt) fun prover tac = prove_abstract SMT_Replay_Methods.abstract_prop tac in val and_rule = prover and_rule_prove val not_and_rule = prover not_and_rule_prove val not_or_rule = prover not_or_rule_prove val or_pos_rule = prover or_pos_prove val and_neg_rule = prover and_neg_rule_prove val or_neg_rule = prove_abstract SMT_Replay_Methods.abstract_bool (fn ctxt => fn _ => resolve_tac ctxt @{thms verit_or_neg} THEN' K (unfold_tac ctxt @{thms not_not}) THEN_ALL_NEW (REPEAT' (SOLVED' (match_tac ctxt @{thms disjI1} THEN_ALL_NEW assume_tac ctxt) ORELSE' (match_tac ctxt @{thms disjI2} THEN' TRY' (assume_tac ctxt))))) fun and_pos ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => REPEAT_CHANGED (resolve_tac ctxt @{thms verit_and_pos}) THEN' TRY' (assume_tac ctxt)) end (* Equivalence Transformation *) local fun prove_equiv equiv_thm ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt [equiv_thm OF [thm]] THEN' TRY' (assume_tac ctxt)) | prove_equiv _ ctxt thms t = replay_error ctxt "invalid application in some equiv rule" Unsupported thms t in val not_equiv1 = prove_equiv @{lemma \\(A \ B) \ A \ B\ by blast} val not_equiv2 = prove_equiv @{lemma \\(A \ B) \ \A \ \B\ by blast} val equiv1 = prove_equiv @{lemma \(A \ B) \ \A \ B\ by blast} val equiv2 = prove_equiv @{lemma \(A \ B) \ A \ \B\ by blast} val not_implies1 = prove_equiv @{lemma \\(A \ B) \ A\ by blast} val not_implies2 = prove_equiv @{lemma \\(A \B) \ \B\ by blast} end (* Equivalence Lemma *) (*equiv_pos2 is very important for performance. We have tried various tactics, including a specialisation of SMT_Replay_Methods.match_instantiate, but there never was a measurable and consistent gain.*) local fun prove_equiv_pos_neg2 thm ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t thm in val equiv_pos1_thm = @{lemma \\(a \ b) \ a \ \b\ by blast+} val equiv_pos1 = prove_equiv_pos_neg2 equiv_pos1_thm val equiv_pos2_thm = @{lemma \\a b. (a \ b) \ \a \ b\ by blast+} val equiv_pos2 = prove_equiv_pos_neg2 equiv_pos2_thm val equiv_neg1_thm = @{lemma \(a \ b) \ \a \ \b\ by blast+} val equiv_neg1 = prove_equiv_pos_neg2 equiv_neg1_thm val equiv_neg2_thm = @{lemma \(a \ b) \ a \ b\ by blast} val equiv_neg2 = prove_equiv_pos_neg2 equiv_neg2_thm end local fun implies_pos_neg_term ctxt thm (\<^term>\Trueprop\ $ (\<^term>\HOL.disj\ $ (\<^term>\HOL.implies\ $ a $ b) $ _)) = Drule.infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) [a, b]) thm | implies_pos_neg_term ctxt thm t = replay_error ctxt "invalid application in Implies" Unsupported [thm] t fun prove_implies_pos_neg thm ctxt _ t = let val thm = implies_pos_neg_term ctxt thm t in SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt [thm] THEN' TRY' (assume_tac ctxt)) end in val implies_neg1_thm = @{lemma \(a \ b) \ a\ by blast} val implies_neg1 = prove_implies_pos_neg implies_neg1_thm val implies_neg2_thm = @{lemma \(a \ b) \ \b\ by blast} val implies_neg2 = prove_implies_pos_neg implies_neg2_thm val implies_thm = @{lemma \(a \ b) \ \a \ b\ by blast} fun implies_rules ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt [implies_thm OF prems] THEN' TRY' (assume_tac ctxt)) end (* BFun *) local val bfun_theorems = @{thms verit_bfun_elim} in fun bfun_elim ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt prems THEN' TRY' (eqsubst_all ctxt bfun_theorems) THEN' TRY' (simplify_tac ctxt @{thms eq_commute all_conj_distrib ex_disj_distrib})) end (* If-Then-Else *) local fun prove_ite ite_thm thm ctxt = resolve_tac ctxt [ite_thm OF [thm]] THEN' TRY' (assume_tac ctxt) in val ite_pos1_thm = @{lemma \\(if x then P else Q) \ x \ Q\ by auto} fun ite_pos1 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt [ite_pos1_thm]) val ite_pos2_thms = @{lemma \\(if x then P else Q) \ \x \ P\ \\(if \x then P else Q) \ x \ P\ by auto} fun ite_pos2 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt ite_pos2_thms) val ite_neg1_thms = @{lemma \(if x then P else Q) \ x \ \Q\ \(if x then P else \Q) \ x \ Q\ by auto} fun ite_neg1 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt ite_neg1_thms) val ite_neg2_thms = @{lemma \(if x then P else Q) \ \x \ \P\ \(if \x then P else Q) \ x \ \P\ \(if x then \P else Q) \ \x \ P\ \(if \x then \P else Q) \ x \ P\ by auto} fun ite_neg2 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt ite_neg2_thms) val ite1_thm = @{lemma \(if x then P else Q) \ x \ Q\ by (auto split: if_splits) } fun ite1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite ite1_thm thm ctxt) val ite2_thm = @{lemma \(if x then P else Q) \ \x \ P\ by (auto split: if_splits) } fun ite2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite ite2_thm thm ctxt) val not_ite1_thm = @{lemma \\(if x then P else Q) \ x \ \Q\ by (auto split: if_splits) } fun not_ite1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite not_ite1_thm thm ctxt) val not_ite2_thm = @{lemma \\(if x then P else Q) \ \x \ \P\ by (auto split: if_splits) } fun not_ite2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite not_ite2_thm thm ctxt) (*ite_intro can introduce new terms that are in the proof exactly the sames as the old one, but are not internally, hence the possible reordering.*) fun ite_intro ctxt _ t = let fun simplify_ite ctxt thms = ctxt |> empty_simpset |> put_simpset HOL_basic_ss |> (fn ctxt => ctxt addsimps thms @ @{thms if_True if_False refl simp_thms if_cancel}) |> Raw_Simplifier.add_eqcong @{thm verit_ite_if_cong} |> Simplifier.full_simp_tac in SMT_Replay_Methods.prove ctxt t (fn _ => simplify_ite ctxt [] THEN' TRY' (simplify_ite ctxt @{thms eq_commute})) end end (* Quantifiers *) local val rm_unused = @{lemma \(\x. P) = P\ \(\x. P) = P\ by blast+} fun qnt_cnf_tac ctxt = simplify_tac ctxt @{thms de_Morgan_conj de_Morgan_disj imp_conv_disj iff_conv_conj_imp if_bool_eq_disj ex_simps all_simps ex_disj_distrib verit_connective_def(3) all_conj_distrib} THEN' TRY' (Blast.blast_tac ctxt) in fun qnt_rm_unused ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => K (unfold_tac ctxt rm_unused)) fun onepoint ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt prems THEN' Simplifier.full_simp_tac (put_simpset HOL_basic_ss (empty_simpset ctxt) addsimps @{thms HOL.simp_thms HOL.all_simps} addsimprocs [@{simproc HOL.defined_All}, @{simproc HOL.defined_Ex}]) THEN' TRY' (Blast.blast_tac ctxt) THEN' TRY' (Metis_Tactic.metis_tac [] ATP_Problem_Generate.combsN ctxt [])) fun qnt_join ctxt _ t = SMT_Replay_Methods.prove ctxt t Classical.fast_tac fun qnt_cnf ctxt _ t = SMT_Replay_Methods.prove ctxt t qnt_cnf_tac fun qnt_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => partial_simplify_tac ctxt []) end (* Equality *) fun eq_transitive ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => REPEAT_CHANGED (resolve_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF @{thms impI}]) THEN' REPEAT' (resolve_tac ctxt @{thms impI}) THEN' REPEAT' (eresolve_tac ctxt @{thms conjI}) THEN' REPEAT' (fn i => dresolve_tac ctxt @{thms verit_eq_transitive} i THEN assume_tac ctxt (i+1)) THEN' resolve_tac ctxt @{thms refl}) local fun find_rew rews t t' = (case AList.lookup (op =) rews (t, t') of SOME thm => SOME (thm COMP @{thm symmetric}) | NONE => (case AList.lookup (op =) rews (t', t) of SOME thm => SOME thm | NONE => NONE)) fun eq_pred_conv rews t ctxt ctrm = (case find_rew rews t (Thm.term_of ctrm) of SOME thm => Conv.rewr_conv thm ctrm | NONE => (case t of f $ arg => (Conv.fun_conv (eq_pred_conv rews f ctxt) then_conv Conv.arg_conv (eq_pred_conv rews arg ctxt)) ctrm | Abs (_, _, f) => Conv.abs_conv (eq_pred_conv rews f o snd) ctxt ctrm | _ => Conv.all_conv ctrm)) fun eq_pred_rewrite_tac ({context = ctxt, prems, concl, ...}: Subgoal.focus) = let val rews = prems |> map_filter (try (apfst (HOLogic.dest_eq o Thm.term_of o Object_Logic.dest_judgment ctxt o Thm.cconcl_of) o `(fn x => x))) |> map (apsnd (fn x => @{thm eq_reflection} OF [x])) fun conv_left conv = Conv.arg_conv (Conv.arg_conv conv) fun conv_left_neg conv = Conv.arg_conv (Conv.arg_conv (Conv.arg_conv conv)) val (t1, conv_term) = (case Thm.term_of (Object_Logic.dest_judgment ctxt concl) of Const(_, _) $ (Const(\<^const_name>\HOL.Not\, _) $ t1) $ _ => (t1, conv_left) | Const(_, _) $ t1 $ (Const(\<^const_name>\HOL.Not\, _) $ _) => (t1, conv_left_neg) | Const(_, _) $ t1 $ _ => (t1, conv_left) | t1 => (t1, conv_left)) fun throwing_CONVERSION cv i st = Seq.single (Conv.gconv_rule cv i st) in throwing_CONVERSION (conv_term (eq_pred_conv rews t1 ctxt)) end in fun eq_congruent_pred ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => REPEAT' (resolve_tac ctxt [@{thm disj_not1[of \_ = _\]} RSN (1, @{thm iffD2}) OF @{thms impI}]) THEN' (fn i => Subgoal.FOCUS (fn focus => eq_pred_rewrite_tac focus i) ctxt i) THEN' (resolve_tac ctxt @{thms refl excluded_middle excluded_middle[of \\_\, unfolded not_not]} ORELSE' assume_tac ctxt)) val eq_congruent = eq_congruent_pred end (* Subproof *) fun subproof ctxt [prem] t = SMT_Replay_Methods.prove ctxt t (fn _ => (REPEAT' (resolve_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF [@{thm impI}], @{thm disj_not1[of \\_\, unfolded not_not]} RSN (1, @{thm iffD2}) OF [@{thm impI}]]) THEN' resolve_tac ctxt [prem] THEN_ALL_NEW assume_tac ctxt THEN' TRY' (assume_tac ctxt)) ORELSE' TRY' (Method.insert_tac ctxt [prem] THEN' Blast.blast_tac ctxt)) | subproof ctxt prems t = replay_error ctxt "invalid application, only one assumption expected" Subproof prems t (* Simplifying Steps *) (* The reconstruction is a bit tricky: At first we only rewrite only based on the rules that are passed as argument. Then we simply use simp_thms to reconstruct all the proofs (and these theorems cover all the simplification below). *) local fun rewrite_only_thms ctxt thms = ctxt |> empty_simpset |> put_simpset HOL_basic_ss |> (fn ctxt => ctxt addsimps thms) |> Simplifier.full_simp_tac fun rewrite_only_thms_tmp ctxt thms = rewrite_only_thms ctxt thms THEN' TRY' (K (Clasimp.auto_tac ctxt)) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) fun rewrite_thms ctxt thms = ctxt |> empty_simpset |> put_simpset HOL_basic_ss |> Raw_Simplifier.add_eqcong @{thm eq_reflection[OF imp_cong]} |> (fn ctxt => ctxt addsimps thms addsimps @{thms simp_thms}) |> Simplifier.full_simp_tac fun chain_rewrite_thms ctxt thms = TRY' (rewrite_only_thms ctxt thms) THEN' TRY' (rewrite_thms ctxt thms) THEN' TRY' (K (Clasimp.auto_tac ctxt)) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) fun chain_rewrite_two_step_with_ac_simps ctxt thms1 thms2 = TRY' (rewrite_only_thms ctxt thms1) THEN' TRY' (rewrite_thms ctxt thms2) THEN' TRY' (K (Clasimp.auto_tac ctxt)) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) fun chain_rewrite_thms_two_step ctxt thms1 thms2 thms3 thms4 = chain_rewrite_two_step_with_ac_simps ctxt thms1 thms2 THEN' TRY' (chain_rewrite_two_step_with_ac_simps ctxt thms3 thms4) val imp_refl = @{lemma \(P \ P) = True\ by blast} in fun rewrite_bool_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (chain_rewrite_thms ctxt @{thms verit_bool_simplify})) fun rewrite_and_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (chain_rewrite_two_step_with_ac_simps ctxt @{thms verit_and_simplify verit_and_simplify1} @{thms verit_and_simplify})) fun rewrite_or_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (chain_rewrite_two_step_with_ac_simps ctxt @{thms verit_or_simplify verit_or_simplify_1} @{thms verit_or_simplify}) THEN' TRY' (REPEAT' (match_tac ctxt @{thms verit_and_pos}) THEN' assume_tac ctxt)) fun rewrite_not_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (rewrite_only_thms_tmp ctxt @{thms verit_not_simplify})) fun rewrite_equiv_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (rewrite_only_thms_tmp ctxt @{thms verit_equiv_simplify})) fun rewrite_eq_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (chain_rewrite_two_step_with_ac_simps ctxt @{thms verit_eq_simplify} (Named_Theorems.get ctxt @{named_theorems ac_simps}))) fun rewrite_implies_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (rewrite_only_thms_tmp ctxt @{thms verit_implies_simplify})) (* It is more efficient to first fold the implication symbols. That is however not enough when symbols appears within expressions, hence we also unfold implications in the other direction. *) fun rewrite_connective_def ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => chain_rewrite_thms_two_step ctxt (imp_refl :: @{thms not_not verit_connective_def[symmetric]}) (@{thms verit_connective_def[symmetric]}) (imp_refl :: @{thms not_not verit_connective_def}) (@{thms verit_connective_def})) fun rewrite_minus_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => chain_rewrite_two_step_with_ac_simps ctxt @{thms arith_simps verit_minus_simplify} (Named_Theorems.get ctxt @{named_theorems ac_simps} @ @{thms numerals arith_simps arith_special numeral_class.numeral.numeral_One})) fun rewrite_comp_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => chain_rewrite_thms ctxt @{thms verit_comp_simplify}) fun rewrite_ite_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => (rewrite_only_thms_tmp ctxt @{thms verit_ite_simplify})) end (* Simplifications *) local fun simplify_arith ctxt thms = partial_simplify_tac ctxt (thms @ Named_Theorems.get ctxt @{named_theorems ac_simps} @ @{thms arith_simps}) in fun sum_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => simplify_arith ctxt @{thms verit_sum_simplify arith_special} THEN' TRY' (K (Clasimp.auto_tac ctxt))) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) fun prod_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => simplify_arith ctxt @{thms verit_prod_simplify} THEN' TRY' (K (Clasimp.auto_tac ctxt))) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) end fun all_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => TRY' (K (Clasimp.auto_tac ctxt))) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) (* Arithmetics *) local val la_rw_eq_thm = @{lemma \(a :: nat) = b \ (a \ b) \ (a \ b)\ by auto} in fun la_rw_eq ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t la_rw_eq_thm fun la_generic_prove args ctxt _ = SMT_Replay_Arith.la_farkas args ctxt fun la_generic ctxt _ args = prove_abstract (SMT_Replay_Methods.abstract_arith_shallow ctxt) (la_generic_prove args) ctxt [] fun lia_generic ctxt _ t = SMT_Replay_Methods.prove_arith_rewrite SMT_Replay_Methods.abstract_neq ctxt t fun la_disequality ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t @{thm verit_la_disequality} end (* Symm and Not_Symm*) local fun prove_symm symm_thm ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt [symm_thm OF [thm]] THEN' TRY' (assume_tac ctxt)) | prove_symm _ ctxt thms t = replay_error ctxt "invalid application in some symm rule" Unsupported thms t in val symm_thm = @{lemma \(B = A) \ A = B\ by blast } val symm = prove_symm symm_thm val not_symm_thm = @{lemma \\(B = A) \ \(A = B)\ by blast } val not_symm = prove_symm not_symm_thm end (* Reordering *) local fun prove_reordering ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ =>( Method.insert_tac ctxt [thm] THEN' match_tac ctxt @{thms ccontr} THEN' K (unfold_tac ctxt @{thms de_Morgan_disj}) THEN' TRY o CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms conjE}) THEN' CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms disjE}) THEN' REPEAT'(ematch_tac ctxt @{thms notE} THEN' TRY' (assume_tac ctxt)) )) | prove_reordering ctxt thms t = replay_error ctxt "invalid application in some reordering rule" Unsupported thms t in val reordering = prove_reordering end end; diff --git a/src/HOL/Tools/SMT/smt_systems.ML b/src/HOL/Tools/SMT/smt_systems.ML --- a/src/HOL/Tools/SMT/smt_systems.ML +++ b/src/HOL/Tools/SMT/smt_systems.ML @@ -1,263 +1,263 @@ (* Title: HOL/Tools/SMT/smt_systems.ML Author: Sascha Boehme, TU Muenchen Setup SMT solvers. *) signature SMT_SYSTEMS = sig val cvc_extensions: bool Config.T val z3_extensions: bool Config.T end; structure SMT_Systems: SMT_SYSTEMS = struct val mashN = "mash" val mepoN = "mepo" val meshN = "mesh" (* helper functions *) fun check_tool var () = (case getenv var of "" => NONE | s => if File.is_file (Path.variable var |> Path.expand |> Path.platform_exe) then SOME [s] else NONE); fun make_avail name () = getenv (name ^ "_SOLVER") <> "" fun make_command name () = [getenv (name ^ "_SOLVER")] fun outcome_of unsat sat unknown timeout solver_name line = if String.isPrefix unsat line then SMT_Solver.Unsat else if String.isPrefix sat line then SMT_Solver.Sat else if String.isPrefix unknown line then SMT_Solver.Unknown else if String.isPrefix timeout line then SMT_Solver.Time_Out else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^ quote solver_name ^ " failed -- enable tracing using the " ^ quote (Config.name_of SMT_Config.trace) ^ " option for details")) (* When used with bitvectors, CVC4 can produce error messages like: $ISABELLE_TMP_PREFIX/... No set-logic command was given before this point. These message should be ignored. *) fun is_blank_or_error_line "" = true | is_blank_or_error_line s = String.isPrefix "(error " s orelse String.isPrefix (getenv "ISABELLE_TMP_PREFIX") s fun on_first_line test_outcome solver_name lines = let val split_first = (fn [] => ("", []) | l :: ls => (l, ls)) val (l, ls) = split_first (drop_prefix is_blank_or_error_line lines) in (test_outcome solver_name l, ls) end fun on_first_non_unsupported_line test_outcome solver_name lines = on_first_line test_outcome solver_name (filter (curry (op <>) "unsupported") lines) (* CVC4 and cvc5 *) val cvc_extensions = Attrib.setup_config_bool \<^binding>\cvc_extensions\ (K false) local fun cvc4_options ctxt = ["--no-stats", "--random-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), "--lang=smt2"] @ (case SMT_Config.get_timeout ctxt of NONE => [] | SOME t => ["--tlimit", string_of_int (Time.toMilliseconds t)]) fun cvc5_options ctxt = ["--no-stats", "--sat-random-seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), "--lang=smt2"] @ (case SMT_Config.get_timeout ctxt of NONE => [] | SOME t => ["--tlimit", string_of_int (Time.toMilliseconds t)]) fun select_class ctxt = if Config.get ctxt cvc_extensions then if Config.get ctxt SMT_Config.higher_order then CVC_Interface.hosmtlib_cvcC else CVC_Interface.smtlib_cvcC else if Config.get ctxt SMT_Config.higher_order then SMTLIB_Interface.hosmtlibC else if Config.get ctxt SMT_Config.native_bv then SMTLIB_Interface.bvsmlibC else SMTLIB_Interface.smtlibC in val cvc4: SMT_Solver.solver_config = { name = "cvc4", class = select_class, avail = make_avail "CVC4", command = make_command "CVC4", options = cvc4_options, smt_options = [(":produce-unsat-cores", "true")], good_slices = (* FUDGE *) [((2, 512, meshN), ["--full-saturate-quant", "--inst-when=full-last-call", "--inst-no-entail", "--term-db-mode=relevant", "--multi-trigger-linear"]), ((2, 64, meshN), ["--decision=internal", "--simplification=none", "--full-saturate-quant"]), ((2, 1024, meshN), ["--relevant-triggers", "--full-saturate-quant"]), ((2, 256, mepoN), ["--trigger-sel=max", "--full-saturate-quant"]), ((2, 32, meshN), ["--term-db-mode=relevant", "--full-saturate-quant"]), ((2, 128, meshN), ["--no-e-matching", "--full-saturate-quant"]), ((2, 256, meshN), ["--finite-model-find", "--fmf-inst-engine"])], outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), parse_proof = SOME (K CVC_Proof_Parse.parse_proof), replay = NONE } val cvc5: SMT_Solver.solver_config = { name = "cvc5", class = select_class, avail = make_avail "CVC5", command = make_command "CVC5", options = cvc5_options, smt_options = [(":produce-unsat-cores", "true")], good_slices = (* FUDGE *) [((2, 512, meshN), ["--full-saturate-quant", "--inst-when=full-last-call", "--inst-no-entail", "--term-db-mode=relevant", "--multi-trigger-linear"]), ((2, 64, meshN), ["--decision=internal", "--simplification=none", "--full-saturate-quant"]), ((2, 1024, meshN), ["--relevant-triggers", "--full-saturate-quant"]), ((2, 256, mepoN), ["--trigger-sel=max", "--full-saturate-quant"]), ((2, 32, meshN), ["--term-db-mode=relevant", "--full-saturate-quant"]), ((2, 128, meshN), ["--no-e-matching", "--full-saturate-quant"]), ((2, 256, meshN), ["--finite-model-find", "--fmf-inst-engine"])], outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), parse_proof = SOME (K CVC_Proof_Parse.parse_proof), replay = NONE } end (* veriT *) local fun select_class ctxt = if Config.get ctxt SMT_Config.higher_order then SMTLIB_Interface.hosmtlibC else SMTLIB_Interface.smtlibC fun veriT_options ctxt = ["--proof-with-sharing", "--proof-define-skolems", "--proof-prune", "--proof-merge", "--disable-print-success", "--disable-banner"] @ - Verit_Proof.veriT_current_strategy (Context.Proof ctxt) @ + Verit_Strategies.veriT_current_strategy (Context.Proof ctxt) @ (case SMT_Config.get_timeout ctxt of NONE => [] | SOME t => ["--max-time=" ^ string_of_int (Time.toMilliseconds t)]) in val veriT: SMT_Solver.solver_config = { name = "verit", class = select_class, avail = is_some o check_tool "ISABELLE_VERIT", command = the o check_tool "ISABELLE_VERIT", options = veriT_options, smt_options = [(":produce-proofs", "true")], good_slices = (* FUDGE *) [((2, 1024, meshN), []), ((2, 512, mashN), []), ((2, 64, meshN), []), ((2, 128, meshN), []), ((2, 256, mepoN), []), ((2, 32, meshN), [])], outcome = on_first_non_unsupported_line (outcome_of "unsat" "sat" "unknown" "Time limit exceeded"), parse_proof = SOME (K Lethe_Proof_Parse.parse_proof), replay = SOME Verit_Replay.replay } end (* Z3 *) val z3_extensions = Attrib.setup_config_bool \<^binding>\z3_extensions\ (K false) local fun z3_options ctxt = ["smt.random_seed=" ^ string_of_int (Config.get ctxt SMT_Config.random_seed), "smt.refine_inj_axioms=false"] @ (case SMT_Config.get_timeout ctxt of NONE => [] | SOME t => ["-T:" ^ string_of_int (Real.ceil (Time.toReal t))]) @ ["-smt2"] fun select_class ctxt = if Config.get ctxt z3_extensions then Z3_Interface.smtlib_z3C else if Config.get ctxt SMT_Config.native_bv then SMTLIB_Interface.bvsmlibC else SMTLIB_Interface.smtlibC in val z3: SMT_Solver.solver_config = { name = "z3", class = select_class, avail = make_avail "Z3", command = make_command "Z3", options = z3_options, smt_options = [(":produce-proofs", "true")], good_slices = (* FUDGE *) [((2, 1024, meshN), []), ((2, 512, mepoN), []), ((2, 64, meshN), []), ((2, 256, meshN), []), ((2, 128, mashN), []), ((2, 32, meshN), [])], outcome = on_first_line (outcome_of "unsat" "sat" "unknown" "timeout"), parse_proof = SOME Z3_Replay.parse_proof, replay = SOME Z3_Replay.replay } end (* smt tactic *) val parse_smt_options = Scan.optional (Args.parens (Args.name -- Scan.option (\<^keyword>\,\ |-- Args.name)) >> apfst SOME) (NONE, NONE) fun smt_method ((solver, stgy), thms) ctxt facts = let val default_solver = SMT_Config.solver_of ctxt val solver = the_default default_solver solver val _ = if solver = "z3" andalso stgy <> NONE then warning ("No strategy is available for z3. Ignoring " ^ quote (the stgy)) else () val ctxt = ctxt - |> (if stgy <> NONE then Context.proof_map (Verit_Proof.select_veriT_stgy (the stgy)) else I) + |> (if stgy <> NONE then Context.proof_map (Verit_Strategies.select_veriT_stgy (the stgy)) else I) |> Context.Proof |> SMT_Config.select_solver solver |> Context.proof_of in HEADGOAL (SOLVED' (SMT_Solver.smt_tac ctxt (thms @ facts))) end val _ = Theory.setup (Method.setup \<^binding>\smt\ (Scan.lift parse_smt_options -- Attrib.thms >> (METHOD oo smt_method)) "Call to the SMT solvers veriT or z3") (* overall setup *) val _ = Theory.setup ( SMT_Solver.add_solver cvc4 #> SMT_Solver.add_solver cvc5 #> SMT_Solver.add_solver veriT #> SMT_Solver.add_solver z3) end; diff --git a/src/HOL/Tools/SMT/verit_replay.ML b/src/HOL/Tools/SMT/verit_replay.ML --- a/src/HOL/Tools/SMT/verit_replay.ML +++ b/src/HOL/Tools/SMT/verit_replay.ML @@ -1,322 +1,314 @@ (* Title: HOL/Tools/SMT/verit_replay.ML Author: Mathias Fleury, MPII VeriT proof parsing and replay. *) signature VERIT_REPLAY = sig val replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm end; structure Verit_Replay: VERIT_REPLAY = struct fun subst_only_free pairs = let fun substf u = (case Termtab.lookup pairs u of SOME u' => u' | NONE => (case u of (Abs(a,T,t)) => Abs(a, T, substf t) | (t$u') => substf t $ substf u' | u => u)) in substf end; fun under_fixes f unchanged_prems (prems, nthms) names args insts decls (concl, ctxt) = let val thms1 = unchanged_prems @ map (SMT_Replay.varify ctxt) prems val _ = SMT_Config.verit_msg ctxt (fn () => \<^print> ("names =", names)) val thms2 = map snd nthms val _ = SMT_Config.verit_msg ctxt (fn () => \<^print> ("prems=", prems)) val _ = SMT_Config.verit_msg ctxt (fn () => \<^print> ("nthms=", nthms)) val _ = SMT_Config.verit_msg ctxt (fn () => \<^print> ("thms1=", thms1)) val _ = SMT_Config.verit_msg ctxt (fn () => \<^print> ("thms2=", thms2)) in (f ctxt (thms1 @ thms2) args insts decls concl) end (** Replaying **) fun replay_thm method_for rewrite_rules ll_defs ctxt assumed unchanged_prems prems nthms concl_transformation global_transformation args insts - (Verit_Proof.VeriT_Replay_Node {id, rule, concl, bounds, declarations = decls, ...}) = + (Lethe_Proof.Lethe_Replay_Node {id, rule, concl, bounds, declarations = decls, ...}) = let val _ = SMT_Config.verit_msg ctxt (fn () => \<^print> id) val rewrite = let val thy = Proof_Context.theory_of (empty_simpset ctxt) in Raw_Simplifier.rewrite_term thy rewrite_rules [] - #> not (null ll_defs andalso Verit_Proof.keep_raw_lifting rule) ? SMTLIB_Isar.unlift_term ll_defs + #> not (null ll_defs andalso Lethe_Proof.keep_raw_lifting rule) ? SMTLIB_Isar.unlift_term ll_defs end - val rewrite_concl = if Verit_Proof.keep_app_symbols rule then + val rewrite_concl = if Lethe_Proof.keep_app_symbols rule then filter (curry Term.could_unify (Thm.concl_of @{thm SMT.fun_app_def}) o Thm.concl_of) rewrite_rules else rewrite_rules val post = let val thy = Proof_Context.theory_of (empty_simpset ctxt) in Raw_Simplifier.rewrite_term thy rewrite_concl [] #> Object_Logic.atomize_term ctxt #> not (null ll_defs) ? SMTLIB_Isar.unlift_term ll_defs #> SMTLIB_Isar.unskolemize_names ctxt #> HOLogic.mk_Trueprop end val concl = concl |> Term.subst_free concl_transformation |> subst_only_free global_transformation |> post in - if rule = Verit_Proof.input_rule then + if rule = Lethe_Proof.input_rule then (case Symtab.lookup assumed id of SOME (_, thm) => thm | _ => raise Fail ("assumption " ^ @{make_string} id ^ " not found")) else under_fixes (method_for rule) unchanged_prems (prems, nthms) (map fst bounds) (map rewrite args) (Symtab.map (K rewrite) insts) decls (concl, ctxt) |> Simplifier.simplify (empty_simpset ctxt addsimps rewrite_rules) end -fun add_used_asserts_in_step (Verit_Proof.VeriT_Replay_Node {prems, +fun add_used_asserts_in_step (Lethe_Proof.Lethe_Replay_Node {prems, subproof = (_, _, _, subproof), ...}) = union (op =) (map_filter (try (snd o SMTLIB_Interface.role_and_index_of_assert_name)) prems @ flat (map (fn x => add_used_asserts_in_step x []) subproof)) fun remove_rewrite_rules_from_rules n = - (fn (step as Verit_Proof.VeriT_Replay_Node {id, ...}) => + (fn (step as Lethe_Proof.Lethe_Replay_Node {id, ...}) => (case try (snd o SMTLIB_Interface.role_and_index_of_assert_name) id of NONE => SOME step | SOME a => if a < n then NONE else SOME step)) fun replay_theorem_step rewrite_rules ll_defs assumed inputs proof_prems - (step as Verit_Proof.VeriT_Replay_Node {id, rule, prems, bounds, args, insts, + (step as Lethe_Proof.Lethe_Replay_Node {id, rule, prems, bounds, args, insts, subproof = (fixes, assms, input, subproof), concl, ...}) state = let val (proofs, stats, ctxt, concl_tranformation, global_transformation) = state val (_, ctxt) = Variable.variant_fixes (map fst bounds) ctxt |> (fn (names, ctxt) => (names, fold Variable.declare_term [SMTLIB_Isar.unskolemize_names ctxt concl] ctxt)) val (names, sub_ctxt) = Variable.variant_fixes (map fst fixes) ctxt ||> fold Variable.declare_term (map Free fixes) val export_vars = concl_tranformation @ (ListPair.zip (map Free fixes, map Free (ListPair.zip (names, map snd fixes)))) val post = let val thy = Proof_Context.theory_of (empty_simpset ctxt) in - Raw_Simplifier.rewrite_term thy ((if Verit_Proof.keep_raw_lifting rule then tl rewrite_rules else rewrite_rules)) [] + Raw_Simplifier.rewrite_term thy ((if Lethe_Proof.keep_raw_lifting rule then tl rewrite_rules else rewrite_rules)) [] #> Object_Logic.atomize_term ctxt - #> not (null ll_defs andalso Verit_Proof.keep_raw_lifting rule) ? SMTLIB_Isar.unlift_term ll_defs + #> not (null ll_defs andalso Lethe_Proof.keep_raw_lifting rule) ? SMTLIB_Isar.unlift_term ll_defs #> SMTLIB_Isar.unskolemize_names ctxt #> HOLogic.mk_Trueprop end val assms = map (subst_only_free global_transformation o Term.subst_free (export_vars) o post) assms val input = map (subst_only_free global_transformation o Term.subst_free (export_vars) o post) input val (all_proof_prems', sub_ctxt2) = Assumption.add_assumes (map (Thm.cterm_of sub_ctxt) (assms @ input)) sub_ctxt fun is_refl thm = Thm.concl_of thm |> (fn (_ $ t) => t) |> HOLogic.dest_eq |> (op =) handle TERM _=> false val all_proof_prems' = all_proof_prems' |> filter_out is_refl val proof_prems' = take (length assms) all_proof_prems' val input = drop (length assms) all_proof_prems' val all_proof_prems = proof_prems @ proof_prems' val replay = replay_theorem_step rewrite_rules ll_defs assumed (input @ inputs) all_proof_prems val (proofs', stats, _, _, sub_global_rew) = fold replay subproof (proofs, stats, sub_ctxt2, export_vars, global_transformation) val export_thm = singleton (Proof_Context.export sub_ctxt2 ctxt) (*for sko_ex and sko_forall, assumptions are in proofs', but the definition of the skolem function is in proofs *) val nthms = prems + |> filter_out Lethe_Proof.is_lethe_def |> map (apsnd export_thm) o map_filter (Symtab.lookup (if (null subproof) then proofs else proofs')) - val nthms' = (if Verit_Proof.is_skolemization rule - then prems else []) + val nthms' = (if Lethe_Proof.is_skolemization rule then prems else []) + |> filter Lethe_Proof.is_lethe_def |> map_filter (Symtab.lookup proofs) val args = map (Term.subst_free concl_tranformation o subst_only_free global_transformation) args val insts = Symtab.map (K (Term.subst_free concl_tranformation o subst_only_free global_transformation)) insts val proof_prems = if Lethe_Replay_Methods.requires_subproof_assms prems rule then proof_prems else [] val local_inputs = if Lethe_Replay_Methods.requires_local_input prems rule then input @ inputs else [] val replay = Timing.timing (replay_thm Verit_Replay_Methods.method_for rewrite_rules ll_defs ctxt assumed [] (proof_prems @ local_inputs) (nthms @ nthms') concl_tranformation global_transformation args insts) val ({elapsed, ...}, thm) = SMT_Config.with_time_limit ctxt SMT_Config.reconstruction_step_timeout replay step handle Timeout.TIMEOUT _ => raise SMT_Failure.SMT SMT_Failure.Time_Out val stats' = Symtab.cons_list (rule, Time.toNanoseconds elapsed) stats -(* val _ = ((Time.toMilliseconds elapsed > 10 andalso (rule = "cong")) ? @{print}) - ("WARNING slow " ^ id ^ @{make_string} rule ^ ": " ^ string_of_int (Time.toMilliseconds elapsed) ^ " " - ^ @{make_string} (proof_prems @ local_inputs)) - val _ = ((Time.toMilliseconds elapsed > 10 andalso (rule = "cong")) ? @{print}) - ( (proof_prems @ local_inputs)) - val _ = ((Time.toMilliseconds elapsed > 10 andalso (rule = "cong")) ? @{print}) - thm - val _ = ((Time.toMilliseconds elapsed > 40) ? @{print}) - ("WARNING slow " ^ id ^ @{make_string} rule ^ ": " ^ string_of_int (Time.toMilliseconds elapsed)) *) val proofs = Symtab.update (id, (map fst bounds, thm)) proofs in (proofs, stats', ctxt, concl_tranformation, sub_global_rew) end fun replay_definition_step rewrite_rules ll_defs _ _ _ - (Verit_Proof.VeriT_Replay_Node {id, declarations = raw_declarations, subproof = (_, _, _, subproof), ...}) state = + (Lethe_Proof.Lethe_Replay_Node {id, declarations = raw_declarations, subproof = (_, _, _, subproof), ...}) state = let val _ = if null subproof then () else raise (Fail ("unrecognized veriT proof, definition has a subproof")) val (proofs, stats, ctxt, concl_tranformation, global_transformation) = state val global_transformer = subst_only_free global_transformation val rewrite = let val thy = Proof_Context.theory_of ctxt in Raw_Simplifier.rewrite_term thy (rewrite_rules) [] #> not (null ll_defs) ? SMTLIB_Isar.unlift_term ll_defs end val start0 = Timing.start () val declaration = map (apsnd (rewrite o global_transformer)) raw_declarations val (names, ctxt) = Variable.variant_fixes (map fst declaration) ctxt ||> fold Variable.declare_term (map Free (map (apsnd fastype_of) declaration)) val old_names = map Free (map (fn (a, b) => (a, fastype_of b)) declaration) val new_names = map Free (ListPair.zip (names, map (fastype_of o snd) declaration)) fun update_mapping (a, b) tab = if a <> b andalso Termtab.lookup tab a = NONE then Termtab.update_new (a, b) tab else tab val global_transformation = global_transformation |> fold update_mapping (ListPair.zip (old_names, new_names)) val global_transformer = subst_only_free global_transformation val generate_definition = (fn (name, term) => (HOLogic.mk_Trueprop (Const(\<^const_name>\HOL.eq\, fastype_of term --> fastype_of term --> @{typ bool}) $ Free (name, fastype_of term) $ term))) #> global_transformer #> Thm.cterm_of ctxt val decls = map generate_definition declaration val (defs, ctxt) = Assumption.add_assumes decls ctxt val thms_with_old_name = ListPair.zip (map fst declaration, defs) val proofs = fold (fn (name, thm) => Symtab.update (id, ([name], @{thm sym} OF [thm]))) thms_with_old_name proofs val total = Time.toNanoseconds (#elapsed (Timing.result start0)) val stats = Symtab.cons_list ("choice", total) stats in (proofs, stats, ctxt, concl_tranformation, global_transformation) end fun replay_assumed assms ll_defs rewrite_rules stats ctxt term = let val rewrite = let val thy = Proof_Context.theory_of (empty_simpset ctxt) in Raw_Simplifier.rewrite_term thy rewrite_rules [] #> not (null ll_defs) ? SMTLIB_Isar.unlift_term ll_defs end val replay = Timing.timing (SMT_Replay_Methods.prove ctxt (rewrite term)) val ({elapsed, ...}, thm) = SMT_Config.with_time_limit ctxt SMT_Config.reconstruction_step_timeout replay (fn _ => Method.insert_tac ctxt (map snd assms) THEN' Classical.fast_tac ctxt) handle Timeout.TIMEOUT _ => raise SMT_Failure.SMT SMT_Failure.Time_Out - val stats' = Symtab.cons_list (Verit_Proof.input_rule, Time.toNanoseconds elapsed) stats + val stats' = Symtab.cons_list (Lethe_Proof.input_rule, Time.toNanoseconds elapsed) stats in (thm, stats') end fun replay_step rewrite_rules ll_defs assumed inputs proof_prems - (step as Verit_Proof.VeriT_Replay_Node {rule, ...}) state = - if rule = Verit_Proof.veriT_def + (step as Lethe_Proof.Lethe_Replay_Node {rule, ...}) state = + if rule = Lethe_Proof.lethe_def then replay_definition_step rewrite_rules ll_defs assumed inputs proof_prems step state else replay_theorem_step rewrite_rules ll_defs assumed inputs proof_prems step state fun replay outer_ctxt ({context = ctxt, typs, terms, rewrite_rules, assms, ll_defs, ...} : SMT_Translate.replay_data) output = let val rewrite_rules = filter_out (fn thm => Term.could_unify (Thm.prop_of @{thm verit_eq_true_simplify}, Thm.prop_of thm)) rewrite_rules val num_ll_defs = length ll_defs val index_of_id = Integer.add (~ num_ll_defs) val id_of_index = Integer.add num_ll_defs val start0 = Timing.start () val (actual_steps, ctxt2) = - Verit_Proof.parse_replay typs terms output ctxt + Lethe_Proof.parse_replay typs terms output ctxt val parsing_time = Time.toNanoseconds (#elapsed (Timing.result start0)) fun step_of_assume (j, (_, th)) = - Verit_Proof.VeriT_Replay_Node { + Lethe_Proof.Lethe_Replay_Node { id = SMTLIB_Interface.assert_name_of_role_and_index SMT_Util.Axiom (id_of_index j), - rule = Verit_Proof.input_rule, + rule = Lethe_Proof.input_rule, args = [], prems = [], proof_ctxt = [], concl = Thm.prop_of th |> Raw_Simplifier.rewrite_term (Proof_Context.theory_of (empty_simpset ctxt addsimps rewrite_rules)) rewrite_rules [], bounds = [], insts = Symtab.empty, declarations = [], subproof = ([], [], [], [])} val used_assert_ids = fold add_used_asserts_in_step actual_steps [] fun normalize_tac ctxt = let val thy = Proof_Context.theory_of (empty_simpset ctxt) in Raw_Simplifier.rewrite_term thy rewrite_rules [] end val used_assm_js = map_filter (fn id => let val i = index_of_id id in if i >= 0 then SOME (i, nth assms i) else NONE end) used_assert_ids val assm_steps = map step_of_assume used_assm_js - fun extract (Verit_Proof.VeriT_Replay_Node {id, rule, concl, bounds, ...}) = + fun extract (Lethe_Proof.Lethe_Replay_Node {id, rule, concl, bounds, ...}) = (id, rule, concl, map fst bounds) - fun cond rule = rule = Verit_Proof.input_rule + fun cond rule = rule = Lethe_Proof.input_rule val add_asssert = SMT_Replay.add_asserted Symtab.update Symtab.empty extract cond val ((_, _), (ctxt3, assumed)) = add_asssert outer_ctxt rewrite_rules (map (apfst fst) assms) (map_filter (remove_rewrite_rules_from_rules num_ll_defs) assm_steps) ctxt2 val used_rew_js = map_filter (fn id => let val i = index_of_id id in if i < 0 then SOME (id, normalize_tac ctxt (nth ll_defs id)) else NONE end) used_assert_ids val (assumed, stats) = fold (fn ((id, thm)) => fn (assumed, stats) => let val (thm, stats) = replay_assumed assms ll_defs rewrite_rules stats ctxt thm val name = SMTLIB_Interface.assert_name_of_role_and_index SMT_Util.Axiom id in (Symtab.update (name, ([], thm)) assumed, stats) end) used_rew_js (assumed, Symtab.cons_list ("parsing", parsing_time) Symtab.empty) val ctxt4 = ctxt3 |> put_simpset (SMT_Replay.make_simpset ctxt3 []) |> Config.put SAT.solver (Config.get ctxt3 SMT_Config.sat_solver) - val len = Verit_Proof.number_of_steps actual_steps + val len = Lethe_Proof.number_of_steps actual_steps fun steps_with_depth _ [] = [] - | steps_with_depth i (p :: ps) = (i + Verit_Proof.number_of_steps [p], p) :: - steps_with_depth (i + Verit_Proof.number_of_steps [p]) ps + | steps_with_depth i (p :: ps) = (i + Lethe_Proof.number_of_steps [p], p) :: + steps_with_depth (i + Lethe_Proof.number_of_steps [p]) ps val actual_steps = steps_with_depth 0 actual_steps val start = Timing.start () val print_runtime_statistics = SMT_Replay.intermediate_statistics ctxt4 start len fun blockwise f (i, x) (next, y) = (if i > next then print_runtime_statistics i else (); (if i > next then i + 10 else next, f x y)) val global_transformation : term Termtab.table = Termtab.empty val (_, (proofs, stats, ctxt5, _, _)) = fold (blockwise (replay_step rewrite_rules ll_defs assumed [] [])) actual_steps (1, (assumed, stats, ctxt4, [], global_transformation)) val total = Time.toMilliseconds (#elapsed (Timing.result start)) - val (_, (_, Verit_Proof.VeriT_Replay_Node {id, ...})) = split_last actual_steps + val (_, (_, Lethe_Proof.Lethe_Replay_Node {id, ...})) = split_last actual_steps val _ = print_runtime_statistics len val thm_with_defs = Symtab.lookup proofs id |> the |> snd |> singleton (Proof_Context.export ctxt5 outer_ctxt) val _ = SMT_Config.statistics_msg ctxt5 (Pretty.string_of o SMT_Replay.pretty_statistics "veriT" total) stats val _ = SMT_Replay.spying (SMT_Config.spy_verit ctxt) ctxt (fn () => SMT_Replay.print_stats (Symtab.dest stats)) "spy_verit" in Verit_Replay_Methods.discharge ctxt [thm_with_defs] @{term False} end end diff --git a/src/HOL/Tools/SMT/verit_proof.ML b/src/HOL/Tools/SMT/verit_strategies.ML rename from src/HOL/Tools/SMT/verit_proof.ML rename to src/HOL/Tools/SMT/verit_strategies.ML --- a/src/HOL/Tools/SMT/verit_proof.ML +++ b/src/HOL/Tools/SMT/verit_strategies.ML @@ -1,895 +1,126 @@ (* Title: HOL/Tools/SMT/Verit_Proof.ML - Author: Mathias Fleury, ENS Rennes - Author: Sascha Boehme, TU Muenchen + Author: Mathias Fleury, ENS Rennes, MPI, JKU, Freiburg University VeriT proofs: parsing and abstract syntax tree. *) -signature VERIT_PROOF = +signature VERIT_STRATEGIES = sig - (*proofs*) - datatype veriT_step = VeriT_Step of { - id: string, - rule: string, - prems: string list, - proof_ctxt: term list, - concl: term, - fixes: string list} - - datatype veriT_replay_node = VeriT_Replay_Node of { - id: string, - rule: string, - args: term list, - prems: string list, - proof_ctxt: term list, - concl: term, - bounds: (string * typ) list, - declarations: (string * term) list, - insts: term Symtab.table, - subproof: (string * typ) list * term list * term list * veriT_replay_node list} - - (*proof parser*) - val parse: typ Symtab.table -> term Symtab.table -> string list -> - Proof.context -> veriT_step list * Proof.context - val parse_replay: typ Symtab.table -> term Symtab.table -> string list -> - Proof.context -> veriT_replay_node list * Proof.context - - val step_prefix : string - val input_rule: string - val keep_app_symbols: string -> bool - val keep_raw_lifting: string -> bool - val normalized_input_rule: string - val la_generic_rule : string - val rewrite_rule : string - val simp_arith_rule : string - val veriT_deep_skolemize_rule : string - val veriT_def : string - val subproof_rule : string - val local_input_rule : string - val not_not_rule : string - val contract_rule : string - val ite_intro_rule : string - val eq_congruent_rule : string - val eq_congruent_pred_rule : string - val skolemization_steps : string list - val theory_resolution2_rule: string - val equiv_pos2_rule: string - val th_resolution_rule: string - val and_pos_rule: string - - val is_skolemization: string -> bool - val is_skolemization_step: veriT_replay_node -> bool - - val number_of_steps: veriT_replay_node list -> int - (*Strategy related*) val veriT_strategy : string Config.T val veriT_current_strategy : Context.generic -> string list val all_veriT_stgies: Context.generic -> string list; val select_veriT_stgy: string -> Context.generic -> Context.generic; val valid_veriT_stgy: string -> Context.generic -> bool; val verit_add_stgy: string * string list -> Context.generic -> Context.generic val verit_rm_stgy: string -> Context.generic -> Context.generic (*Global tactic*) val verit_tac: Proof.context -> thm list -> int -> tactic val verit_tac_stgy: string -> Proof.context -> thm list -> int -> tactic end; -structure Verit_Proof: VERIT_PROOF = +structure Verit_Strategies: VERIT_STRATEGIES = struct open SMTLIB_Proof val veriT_strategy_default_name = "default"; (*FUDGE*) val veriT_strategy_del_insts_name = "del_insts"; (*FUDGE*) val veriT_strategy_rm_insts_name = "ccfv_SIG"; (*FUDGE*) val veriT_strategy_ccfv_insts_name = "ccfv_threshold"; (*FUDGE*) val veriT_strategy_best_name = "best"; (*FUDGE*) val veriT_strategy_best = ["--index-sorts", "--index-fresh-sorts", "--triggers-new", "--triggers-sel-rm-specific"]; val veriT_strategy_del_insts = ["--index-sorts", "--index-fresh-sorts", "--ccfv-breadth", "--inst-deletion", "--index-SAT-triggers", "--inst-deletion-loops", "--inst-deletion-track-vars", "--inst-deletion", "--index-SAT-triggers"]; val veriT_strategy_rm_insts = ["--index-SIG", "--triggers-new", "--triggers-sel-rm-specific"]; val veriT_strategy_ccfv_insts = ["--index-sorts", "--index-fresh-sorts", "--triggers-new", "--triggers-sel-rm-specific", "--triggers-restrict-combine", "--inst-deletion", "--index-SAT-triggers", "--inst-deletion-loops", "--inst-deletion-track-vars", "--inst-deletion", "--index-SAT-triggers", "--inst-sorts-threshold=100000", "--ematch-exp=10000000", "--ccfv-index=100000", "--ccfv-index-full=1000"] val veriT_strategy_default = []; type verit_strategy = {default_strategy: string, strategies: (string * string list) list} fun mk_verit_strategy default_strategy strategies : verit_strategy = {default_strategy=default_strategy,strategies=strategies} val empty_data = mk_verit_strategy veriT_strategy_best_name [(veriT_strategy_default_name, veriT_strategy_default), (veriT_strategy_del_insts_name, veriT_strategy_del_insts), (veriT_strategy_rm_insts_name, veriT_strategy_rm_insts), (veriT_strategy_ccfv_insts_name, veriT_strategy_ccfv_insts), (veriT_strategy_best_name, veriT_strategy_best)] fun merge_data ({strategies=strategies1,...}:verit_strategy, {default_strategy,strategies=strategies2}:verit_strategy) : verit_strategy = mk_verit_strategy default_strategy (AList.merge (op =) (op =) (strategies1, strategies2)) structure Data = Generic_Data ( type T = verit_strategy val empty = empty_data val merge = merge_data ) fun veriT_current_strategy ctxt = let val {default_strategy,strategies} = (Data.get ctxt) in AList.lookup (op=) strategies default_strategy |> the end val veriT_strategy = Attrib.setup_config_string \<^binding>\smt_verit_strategy\ (K veriT_strategy_best_name); fun valid_veriT_stgy stgy context = let val {strategies,...} = Data.get context in AList.defined (op =) strategies stgy end fun select_veriT_stgy stgy context = let val {strategies,...} = Data.get context val upd = Data.map (K (mk_verit_strategy stgy strategies)) in if not (AList.defined (op =) strategies stgy) then error ("Trying to select unknown veriT strategy: " ^ quote stgy) else upd context end fun verit_add_stgy stgy context = let val {default_strategy,strategies} = Data.get context in Data.map (K (mk_verit_strategy default_strategy (AList.update (op =) stgy strategies))) context end fun verit_rm_stgy stgy context = let val {default_strategy,strategies} = Data.get context in Data.map (K (mk_verit_strategy default_strategy (AList.delete (op =) stgy strategies))) context end fun all_veriT_stgies context = let val {strategies,...} = Data.get context in map fst strategies end val select_verit = SMT_Config.select_solver "verit" fun verit_tac ctxt = SMT_Solver.smt_tac (Config.put SMT_Config.native_bv false ((Context.proof_map select_verit ctxt))) fun verit_tac_stgy stgy ctxt = verit_tac (Context.proof_of (select_veriT_stgy stgy (Context.Proof ctxt))) -datatype raw_veriT_node = Raw_VeriT_Node of { - id: string, - rule: string, - args: SMTLIB.tree, - prems: string list, - concl: SMTLIB.tree, - declarations: (string * SMTLIB.tree) list, - subproof: raw_veriT_node list} - -fun mk_raw_node id rule args prems declarations concl subproof = - Raw_VeriT_Node {id = id, rule = rule, args = args, prems = prems, declarations = declarations, - concl = concl, subproof = subproof} - -datatype veriT_node = VeriT_Node of { - id: string, - rule: string, - prems: string list, - proof_ctxt: term list, - concl: term} - -fun mk_node id rule prems proof_ctxt concl = - VeriT_Node {id = id, rule = rule, prems = prems, proof_ctxt = proof_ctxt, concl = concl} - -datatype veriT_replay_node = VeriT_Replay_Node of { - id: string, - rule: string, - args: term list, - prems: string list, - proof_ctxt: term list, - concl: term, - bounds: (string * typ) list, - insts: term Symtab.table, - declarations: (string * term) list, - subproof: (string * typ) list * term list * term list * veriT_replay_node list} - -fun mk_replay_node id rule args prems proof_ctxt concl bounds insts declarations subproof = - VeriT_Replay_Node {id = id, rule = rule, args = args, prems = prems, proof_ctxt = proof_ctxt, - concl = concl, bounds = bounds, insts = insts, declarations = declarations, - subproof = subproof} - -datatype veriT_step = VeriT_Step of { - id: string, - rule: string, - prems: string list, - proof_ctxt: term list, - concl: term, - fixes: string list} - -fun mk_step id rule prems proof_ctxt concl fixes = - VeriT_Step {id = id, rule = rule, prems = prems, proof_ctxt = proof_ctxt, concl = concl, - fixes = fixes} - -val step_prefix = ".c" -val input_rule = "input" -val la_generic_rule = "la_generic" -val normalized_input_rule = "__normalized_input" (*arbitrary*) -val rewrite_rule = "__rewrite" (*arbitrary*) -val subproof_rule = "subproof" -val local_input_rule = "__local_input" (*arbitrary*) -val simp_arith_rule = "simp_arith" -val veriT_def = "__skolem_definition" (*arbitrary*) -val not_not_rule = "not_not" -val contract_rule = "contraction" -val eq_congruent_pred_rule = "eq_congruent_pred" -val eq_congruent_rule = "eq_congruent" -val ite_intro_rule = "ite_intro" -val default_skolem_rule = "sko_forall" (*arbitrary, but must be one of the skolems*) -val theory_resolution2_rule = "__theory_resolution2" (*arbitrary*) -val equiv_pos2_rule = "equiv_pos2" -val th_resolution_rule = "th_resolution" -val and_pos_rule = "and_pos" - -val skolemization_steps = ["sko_forall", "sko_ex"] -val is_skolemization = member (op =) skolemization_steps -val keep_app_symbols = member (op =) [eq_congruent_pred_rule, eq_congruent_rule, ite_intro_rule, and_pos_rule] -val keep_raw_lifting = member (op =) [eq_congruent_pred_rule, eq_congruent_rule, ite_intro_rule, and_pos_rule] -val is_SH_trivial = member (op =) [not_not_rule, contract_rule] - -fun is_skolemization_step (VeriT_Replay_Node {id, ...}) = is_skolemization id - -(* Even the veriT developers do not know if the following rule can still appear in proofs: *) -val veriT_deep_skolemize_rule = "deep_skolemize" - -fun number_of_steps [] = 0 - | number_of_steps ((VeriT_Replay_Node {subproof = (_, _, _, subproof), ...}) :: pf) = - 1 + number_of_steps subproof + number_of_steps pf - -(* proof parser *) - -fun node_of p cx = - ([], cx) - ||>> `(with_fresh_names (term_of p)) - |>> snd - -fun find_type_in_formula (Abs (v, T, u)) var_name = - if String.isPrefix var_name v then SOME T else find_type_in_formula u var_name - | find_type_in_formula (u $ v) var_name = - (case find_type_in_formula u var_name of - NONE => find_type_in_formula v var_name - | some_T => some_T) - | find_type_in_formula (Free(v, T)) var_name = - if String.isPrefix var_name v then SOME T else NONE - | find_type_in_formula _ _ = NONE - -fun synctactic_var_subst old_name new_name (u $ v) = - (synctactic_var_subst old_name new_name u $ synctactic_var_subst old_name new_name v) - | synctactic_var_subst old_name new_name (Abs (v, T, u)) = - Abs (if String.isPrefix old_name v then new_name else v, T, - synctactic_var_subst old_name new_name u) - | synctactic_var_subst old_name new_name (Free (v, T)) = - if String.isPrefix old_name v then Free (new_name, T) else Free (v, T) - | synctactic_var_subst _ _ t = t - -fun synctatic_rew_in_lhs_subst old_name new_name (Const(\<^const_name>\HOL.eq\, T) $ t1 $ t2) = - Const(\<^const_name>\HOL.eq\, T) $ synctactic_var_subst old_name new_name t1 $ t2 - | synctatic_rew_in_lhs_subst old_name new_name (Const(\<^const_name>\Trueprop\, T) $ t1) = - Const(\<^const_name>\Trueprop\, T) $ (synctatic_rew_in_lhs_subst old_name new_name t1) - | synctatic_rew_in_lhs_subst _ _ t = t - -fun add_bound_variables_to_ctxt cx = - fold (update_binding o - (fn (s, SOME typ) => (s, Term (Free (s, type_of cx typ))))) - -local - - fun extract_symbols bds = - bds - |> map (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, SMTLIB.Sym y], typ) => [([x, y], typ)] - | t => raise (Fail ("match error " ^ @{make_string} t))) - |> flat - - (* onepoint can bind a variable to another variable or to a constant *) - fun extract_qnt_symbols cx bds = - bds - |> map (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, SMTLIB.Sym y], typ) => - (case node_of (SMTLIB.Sym y) cx of - ((_, []), _) => [([x], typ)] - | _ => [([x, y], typ)]) - | (SMTLIB.S (SMTLIB.Sym "=" :: SMTLIB.Sym x :: _), typ) => [([x], typ)] - | t => raise (Fail ("match error " ^ @{make_string} t))) - |> flat - - fun extract_symbols_map bds = - bds - |> map (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, _], typ) => [([x], typ)]) - |> flat -in - -fun declared_csts _ "__skolem_definition" [(SMTLIB.S [SMTLIB.Sym x, typ, _], _)] = [(x, typ)] - | declared_csts _ "__skolem_definition" t = raise (Fail ("unrecognized skolem_definition " ^ @{make_string} t)) - | declared_csts _ _ _ = [] - -fun skolems_introduced_by_rule (SMTLIB.S bds) = - fold (fn (SMTLIB.S [SMTLIB.Sym "=", _, SMTLIB.Sym y]) => curry (op ::) y) bds [] - -(*FIXME there is probably a way to use the information given by onepoint*) -fun bound_vars_by_rule _ "bind" (bds) = extract_symbols bds - | bound_vars_by_rule cx "onepoint" bds = extract_qnt_symbols cx bds - | bound_vars_by_rule _ "sko_forall" bds = extract_symbols_map bds - | bound_vars_by_rule _ "sko_ex" bds = extract_symbols_map bds - | bound_vars_by_rule _ "__skolem_definition" [(SMTLIB.S [SMTLIB.Sym x, typ, _], _)] = [([x], SOME typ)] - | bound_vars_by_rule _ "__skolem_definition" [(SMTLIB.S [_, SMTLIB.Sym x, _], _)] = [([x], NONE)] - | bound_vars_by_rule _ _ _ = [] - -(* VeriT adds "?" before some variables. *) -fun remove_all_qm (SMTLIB.Sym v :: l) = - SMTLIB.Sym (perhaps (try (unprefix "?")) v) :: remove_all_qm l - | remove_all_qm (SMTLIB.S l :: l') = SMTLIB.S (remove_all_qm l) :: remove_all_qm l' - | remove_all_qm (SMTLIB.Key v :: l) = SMTLIB.Key v :: remove_all_qm l - | remove_all_qm (v :: l) = v :: remove_all_qm l - | remove_all_qm [] = [] - -fun remove_all_qm2 (SMTLIB.Sym v) = SMTLIB.Sym (perhaps (try (unprefix "?")) v) - | remove_all_qm2 (SMTLIB.S l) = SMTLIB.S (remove_all_qm l) - | remove_all_qm2 (SMTLIB.Key v) = SMTLIB.Key v - | remove_all_qm2 v = v - -end - -datatype step_kind = ASSUME | ANCHOR | NO_STEP | NORMAL_STEP | SKOLEM - -fun parse_raw_proof_steps (limit : string option) (ls : SMTLIB.tree list) (cx : name_bindings) : - (raw_veriT_node list * SMTLIB.tree list * name_bindings) = - let - fun rotate_pair (a, (b, c)) = ((a, b), c) - fun step_kind [] = (NO_STEP, SMTLIB.S [], []) - | step_kind ((p as SMTLIB.S (SMTLIB.Sym "anchor" :: _)) :: l) = (ANCHOR, p, l) - | step_kind ((p as SMTLIB.S (SMTLIB.Sym "assume" :: _)) :: l) = (ASSUME, p, l) - | step_kind ((p as SMTLIB.S (SMTLIB.Sym "step" :: _)) :: l) = (NORMAL_STEP, p, l) - | step_kind ((p as SMTLIB.S (SMTLIB.Sym "define-fun" :: _)) :: l) = (SKOLEM, p, l) - fun parse_skolem (SMTLIB.S [SMTLIB.Sym "define-fun", SMTLIB.Sym id, _, typ, - SMTLIB.S (SMTLIB.Sym "!" :: t :: [SMTLIB.Key _, SMTLIB.Sym name])]) cx = - (*replace the name binding by the constant instead of the full term in order to reduce - the size of the generated terms and therefore the reconstruction time*) - let val (l, cx) = (fst oo SMTLIB_Proof.extract_and_update_name_bindings) t cx - |> apsnd (SMTLIB_Proof.update_name_binding (name, SMTLIB.Sym id)) - in - (mk_raw_node (id ^ veriT_def) veriT_def (SMTLIB.S [SMTLIB.Sym id, typ, l]) [] [] - (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym id, l]) [], cx) - end - | parse_skolem (SMTLIB.S [SMTLIB.Sym "define-fun", SMTLIB.Sym id, _, typ, SMTLIB.S l]) cx = - let val (l, cx) = (fst oo SMTLIB_Proof.extract_and_update_name_bindings) (SMTLIB.S l ) cx - in - (mk_raw_node (id ^ veriT_def) veriT_def (SMTLIB.S [SMTLIB.Sym id, typ, l]) [] [] - (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym id, l]) [], cx) - end - | parse_skolem t _ = raise Fail ("unrecognized VeriT proof " ^ \<^make_string> t) - fun get_id_cx (SMTLIB.S ((SMTLIB.Sym _) :: (SMTLIB.Sym id) :: l), cx) = (id, (l, cx)) - | get_id_cx t = raise Fail ("unrecognized VeriT proof " ^ \<^make_string> t) - fun get_id (SMTLIB.S ((SMTLIB.Sym _) :: (SMTLIB.Sym id) :: l)) = (id, l) - | get_id t = raise Fail ("unrecognized VeriT proof " ^ \<^make_string> t) - fun parse_source (SMTLIB.Key "premises" :: SMTLIB.S source ::l, cx) = - (SOME (map (fn (SMTLIB.Sym id) => id) source), (l, cx)) - | parse_source (l, cx) = (NONE, (l, cx)) - fun parse_rule (SMTLIB.Key "rule" :: SMTLIB.Sym r :: l, cx) = (r, (l, cx)) - | parse_rule t = raise Fail ("unrecognized VeriT proof " ^ \<^make_string> t) - fun parse_anchor_step (SMTLIB.S (SMTLIB.Sym "anchor" :: SMTLIB.Key "step" :: SMTLIB.Sym r :: l), cx) = (r, (l, cx)) - | parse_anchor_step t = raise Fail ("unrecognized VeriT proof " ^ \<^make_string> t) - fun parse_args (SMTLIB.Key "args" :: args :: l, cx) = - let val ((args, cx), _) = SMTLIB_Proof.extract_and_update_name_bindings args cx - in (args, (l, cx)) end - | parse_args (l, cx) = (SMTLIB.S [], (l, cx)) - fun parse_and_clausify_conclusion (SMTLIB.S (SMTLIB.Sym "cl" :: []) :: l, cx) = - (SMTLIB.Sym "false", (l, cx)) - | parse_and_clausify_conclusion (SMTLIB.S (SMTLIB.Sym "cl" :: concl) :: l, cx) = - let val (concl, cx) = fold_map (fst oo SMTLIB_Proof.extract_and_update_name_bindings) concl cx - in (SMTLIB.S (SMTLIB.Sym "or" :: concl), (l, cx)) end - | parse_and_clausify_conclusion t = raise Fail ("unrecognized VeriT proof " ^ \<^make_string> t) - val parse_normal_step = - get_id_cx - ##> parse_and_clausify_conclusion - #> rotate_pair - ##> parse_rule - #> rotate_pair - ##> parse_source - #> rotate_pair - ##> parse_args - #> rotate_pair - - fun to_raw_node subproof ((((id, concl), rule), prems), args) = - mk_raw_node id rule args (the_default [] prems) [] concl subproof - fun at_discharge NONE _ = false - | at_discharge (SOME id) p = p |> get_id |> fst |> (fn id2 => id = id2) - in - case step_kind ls of - (NO_STEP, _, _) => ([],[], cx) - | (NORMAL_STEP, p, l) => - if at_discharge limit p then ([], ls, cx) else - let - val (s, (_, cx)) = (p, cx) - |> parse_normal_step - ||> (fn i => i) - |>> (to_raw_node []) - val (rp, rl, cx) = parse_raw_proof_steps limit l cx - in (s :: rp, rl, cx) end - | (ASSUME, p, l) => - let - val (id, t :: []) = p - |> get_id - val ((t, cx), _) = SMTLIB_Proof.extract_and_update_name_bindings t cx - val s = mk_raw_node id input_rule (SMTLIB.S []) [] [] t [] - val (rp, rl, cx) = parse_raw_proof_steps limit l cx - in (s :: rp, rl, cx) end - | (ANCHOR, p, l) => - let - val (anchor_id, (anchor_args, (_, cx))) = (p, cx) |> (parse_anchor_step ##> parse_args) - val (subproof, discharge_step :: remaining_proof, cx) = parse_raw_proof_steps (SOME anchor_id) l cx - val (curss, (_, cx)) = parse_normal_step (discharge_step, cx) - val s = to_raw_node subproof (fst curss, anchor_args) - val (rp, rl, cx) = parse_raw_proof_steps limit remaining_proof cx - in (s :: rp, rl, cx) end - | (SKOLEM, p, l) => - let - val (s, cx) = parse_skolem p cx - val (rp, rl, cx) = parse_raw_proof_steps limit l cx - in (s :: rp, rl, cx) end - end - -fun proof_ctxt_of_rule "bind" t = t - | proof_ctxt_of_rule "sko_forall" t = t - | proof_ctxt_of_rule "sko_ex" t = t - | proof_ctxt_of_rule "let" t = t - | proof_ctxt_of_rule "onepoint" t = t - | proof_ctxt_of_rule _ _ = [] - -fun args_of_rule "bind" t = t - | args_of_rule "la_generic" t = t - | args_of_rule "lia_generic" t = t - | args_of_rule _ _ = [] - -fun insts_of_forall_inst "forall_inst" t = map (fn SMTLIB.S [_, SMTLIB.Sym x, a] => (x, a)) t - | insts_of_forall_inst _ _ = [] - -fun id_of_last_step prems = - if null prems then [] - else - let val VeriT_Replay_Node {id, ...} = List.last prems in [id] end - -fun extract_assumptions_from_subproof subproof = - let fun extract_assumptions_from_subproof (VeriT_Replay_Node {rule, concl, ...}) assms = - if rule = local_input_rule then concl :: assms else assms - in - fold extract_assumptions_from_subproof subproof [] - end - -fun normalized_rule_name id rule = - (case (rule = input_rule, can SMTLIB_Interface.role_and_index_of_assert_name id) of - (true, true) => normalized_input_rule - | (true, _) => local_input_rule - | _ => rule) - -fun is_assm_repetition id rule = - rule = input_rule andalso can SMTLIB_Interface.role_and_index_of_assert_name id - -fun extract_skolem ([SMTLIB.Sym var, typ, choice]) = (var, typ, choice) - | extract_skolem t = raise Fail ("fail to parse type" ^ @{make_string} t) - -(* The preprocessing takes care of: - 1. unfolding the shared terms - 2. extract the declarations of skolems to make sure that there are not unfolded -*) -fun preprocess compress step = - let - fun expand_assms cs = - map (fn t => case AList.lookup (op =) cs t of NONE => t | SOME a => a) - fun expand_lonely_arguments (args as SMTLIB.S [SMTLIB.Sym "=", _, _]) = [args] - | expand_lonely_arguments (x as SMTLIB.S [SMTLIB.Sym var, _]) = [SMTLIB.S [SMTLIB.Sym "=", x, SMTLIB.Sym var]] - - fun preprocess (Raw_VeriT_Node {id, rule, args, prems, concl, subproof, ...}) (cx, remap_assms) = - let - val (skolem_names, stripped_args) = args - |> (fn SMTLIB.S args => args) - |> map - (fn SMTLIB.S [SMTLIB.Key "=", x, y] => SMTLIB.S [SMTLIB.Sym "=", x, y] - | x => x) - |> (rule = "bind" orelse rule = "onepoint") ? flat o (map expand_lonely_arguments) - |> `(if rule = veriT_def then single o extract_skolem else K []) - ||> SMTLIB.S - val (subproof, (cx, _)) = fold_map preprocess subproof (cx, remap_assms) |> apfst flat - val remap_assms = (if rule = "or" then (id, hd prems) :: remap_assms else remap_assms) - (* declare variables in the context *) - val declarations = - if rule = veriT_def - then skolem_names |> map (fn (name, _, choice) => (name, choice)) - else [] - in - if compress andalso rule = "or" - then ([], (cx, remap_assms)) - else ([Raw_VeriT_Node {id = id, rule = rule, args = stripped_args, - prems = expand_assms remap_assms prems, declarations = declarations, concl = concl, subproof = subproof}], - (cx, remap_assms)) - end - in preprocess step end - -fun filter_split _ [] = ([], []) - | filter_split f (a :: xs) = - (if f a then apfst (curry op :: a) else apsnd (curry op :: a)) (filter_split f xs) - -fun collect_skolem_defs (Raw_VeriT_Node {rule, subproof = subproof, args, ...}) = - (if is_skolemization rule then map (fn id => id ^ veriT_def) (skolems_introduced_by_rule args) else []) @ - flat (map collect_skolem_defs subproof) - -fun extract_types_of_args (SMTLIB.S [var, typ, t as SMTLIB.S [SMTLIB.Sym "choice", _, _]]) = - (SMTLIB.S [var, typ, t], SOME typ) - |> single - | extract_types_of_args (SMTLIB.S t) = - let - fun extract_types_of_arg (SMTLIB.S [eq, SMTLIB.S [var, typ], t]) = (SMTLIB.S [eq, var, t], SOME typ) - | extract_types_of_arg t = (t, NONE) - in - t - |> map extract_types_of_arg - end - -(*The postprocessing does: - 1. translate the terms to Isabelle syntax, taking care of free variables - 2. remove the ambiguity in the proof terms: - x \ y |- x = x - means y = x. To remove ambiguity, we use the fact that y is a free variable and replace the term - by: - xy \ y |- xy = x. - This is now does not have an ambiguity and we can safely move the "xy \ y" to the proof - assumptions. -*) -fun postprocess_proof compress ctxt step cx = - let - fun postprocess (Raw_VeriT_Node {id, rule, args, prems, declarations, concl, subproof}) (cx, rew) = - let - val _ = (SMT_Config.verit_msg ctxt) (fn () => @{print} ("id =", id, "concl =", concl)) - - val (args) = extract_types_of_args args - val globally_bound_vars = declared_csts cx rule args - val cx = fold (update_binding o (fn (s, typ) => (s, Term (Free (s, type_of cx typ))))) - globally_bound_vars cx - - (*find rebound variables specific to the LHS of the equivalence symbol*) - val bound_vars = bound_vars_by_rule cx rule args - val bound_vars_no_typ = map fst bound_vars - val rhs_vars = - fold (fn [t', t] => t <> t' ? (curry (op ::) t) | _ => fn x => x) bound_vars_no_typ [] - fun not_already_bound cx t = SMTLIB_Proof.lookup_binding cx t = None andalso - not (member (op =) rhs_vars t) - val (shadowing_vars, rebound_lhs_vars) = bound_vars - |> filter_split (fn ([t, _], typ) => not_already_bound cx t | _ => true) - |>> map (apfst (hd)) - |>> (fn vars => vars @ flat (map (fn ([_, t], typ) => [(t, typ)] | _ => []) bound_vars)) - val subproof_rew = fold (fn [t, t'] => curry (op ::) (t, t ^ t')) - (map fst rebound_lhs_vars) rew - val subproof_rewriter = fold (fn (t, t') => synctatic_rew_in_lhs_subst t t') - subproof_rew - - val ((concl, bounds), cx') = node_of concl cx - - val extra_lhs_vars = map (fn ([a,b], typ) => (a, a^b, typ)) rebound_lhs_vars - val old_lhs_vars = map (fn (a, _, typ) => (a, typ)) extra_lhs_vars - val new_lhs_vars = map (fn (_, newvar, typ) => (newvar, typ)) extra_lhs_vars - - (* postprocess conclusion *) - val concl = SMTLIB_Isar.unskolemize_names ctxt (subproof_rewriter concl) - - val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> ("id =", id, "concl =", concl)) - val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> ("id =", id, "cx' =", cx', - "bound_vars =", bound_vars)) - - val bound_tvars = - map (fn (s, SOME typ) => (s, type_of cx typ)) - (shadowing_vars @ new_lhs_vars) - val subproof_cx = - add_bound_variables_to_ctxt cx (shadowing_vars @ new_lhs_vars) cx - - fun could_unify (Bound i, Bound j) = i = j - | could_unify (Var v, Var v') = v = v' - | could_unify (Free v, Free v') = v = v' - | could_unify (Const (v, ty), Const (v', ty')) = v = v' andalso ty = ty' - | could_unify (Abs (_, ty, bdy), Abs (_, ty', bdy')) = ty = ty' andalso could_unify (bdy, bdy') - | could_unify (u $ v, u' $ v') = could_unify (u, u') andalso could_unify (v, v') - | could_unify _ = false - fun is_alpha_renaming t = - t - |> HOLogic.dest_Trueprop - |> HOLogic.dest_eq - |> could_unify - handle TERM _ => false - val alpha_conversion = rule = "bind" andalso is_alpha_renaming concl - - val can_remove_subproof = - compress andalso (is_skolemization rule orelse alpha_conversion) - val (fixed_subproof : veriT_replay_node list, _) = - fold_map postprocess (if can_remove_subproof then [] else subproof) - (subproof_cx, subproof_rew) - - val unsk_and_rewrite = SMTLIB_Isar.unskolemize_names ctxt o subproof_rewriter - - (* postprocess assms *) - val stripped_args = map fst args - val sanitized_args = proof_ctxt_of_rule rule stripped_args - - val arg_cx = add_bound_variables_to_ctxt cx (shadowing_vars @ old_lhs_vars) subproof_cx - val (termified_args, _) = fold_map node_of sanitized_args arg_cx |> apfst (map fst) - val normalized_args = map unsk_and_rewrite termified_args - - val subproof_assms = proof_ctxt_of_rule rule normalized_args - - (* postprocess arguments *) - val rule_args = args_of_rule rule stripped_args - val (termified_args, _) = fold_map term_of rule_args subproof_cx - val normalized_args = map unsk_and_rewrite termified_args - val rule_args = map subproof_rewriter normalized_args - - val raw_insts = insts_of_forall_inst rule stripped_args - fun termify_term (x, t) cx = let val (t, cx) = term_of t cx in ((x, t), cx) end - val (termified_args, _) = fold_map termify_term raw_insts subproof_cx - val insts = Symtab.empty - |> fold (fn (x, t) => fn insts => Symtab.update_new (x, t) insts) termified_args - |> Symtab.map (K unsk_and_rewrite) - - (* declarations *) - val (declarations, _) = fold_map termify_term declarations cx - |> apfst (map (apsnd unsk_and_rewrite)) - - (* fix step *) - val _ = if bounds <> [] then raise (Fail "found dangling variable in concl") else () - val skolem_defs = (if is_skolemization rule - then map (fn id => id ^ veriT_def) (skolems_introduced_by_rule (SMTLIB.S (map fst args))) else []) - val skolems_of_subproof = (if is_skolemization rule - then flat (map collect_skolem_defs subproof) else []) - val fixed_prems = - prems @ (if is_assm_repetition id rule then [id] else []) @ - skolem_defs @ skolems_of_subproof @ (id_of_last_step fixed_subproof) - - (* fix subproof *) - val normalized_rule = normalized_rule_name id rule - |> (if compress andalso alpha_conversion then K "refl" else I) - - val extra_assms2 = - (if rule = subproof_rule then extract_assumptions_from_subproof fixed_subproof else []) - - val step = mk_replay_node id normalized_rule rule_args fixed_prems subproof_assms concl - [] insts declarations (bound_tvars, subproof_assms, extra_assms2, fixed_subproof) - - in - (step, (cx', rew)) - end - in - postprocess step (cx, []) - |> (fn (step, (cx, _)) => (step, cx)) - end - -fun combine_proof_steps ((step1 : veriT_replay_node) :: step2 :: steps) = - let - val (VeriT_Replay_Node {id = id1, rule = rule1, args = args1, prems = prems1, - proof_ctxt = proof_ctxt1, concl = concl1, bounds = bounds1, insts = insts1, - declarations = declarations1, - subproof = (bound_sub1, assms_sub1, assms_extra1, subproof1)}) = step1 - val (VeriT_Replay_Node {id = id2, rule = rule2, args = args2, prems = prems2, - proof_ctxt = proof_ctxt2, concl = concl2, bounds = bounds2, insts = insts2, - declarations = declarations2, - subproof = (bound_sub2, assms_sub2, assms_extra2, subproof2)}) = step2 - val goals1 = - (case concl1 of - _ $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ - (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.Not\, _) $a) $ b)) => [a,b] - | _ => []) - val goal2 = (case concl2 of _ $ a => a) - in - if rule1 = equiv_pos2_rule andalso rule2 = th_resolution_rule andalso member (op =) prems2 id1 - andalso member (op =) goals1 goal2 - then - mk_replay_node id2 theory_resolution2_rule args2 (filter_out (curry (op =) id1) prems2) - proof_ctxt2 concl2 bounds2 insts2 declarations2 - (bound_sub2, assms_sub2, assms_extra2, combine_proof_steps subproof2) :: - combine_proof_steps steps - else - mk_replay_node id1 rule1 args1 prems1 - proof_ctxt1 concl1 bounds1 insts1 declarations1 - (bound_sub1, assms_sub1, assms_extra1, combine_proof_steps subproof1) :: - combine_proof_steps (step2 :: steps) - end - | combine_proof_steps steps = steps - - -val linearize_proof = - let - fun map_node_concl f (VeriT_Node {id, rule, prems, proof_ctxt, concl}) = - mk_node id rule prems proof_ctxt (f concl) - fun linearize (VeriT_Replay_Node {id = id, rule = rule, args = _, prems = prems, - proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, insts = _, declarations = _, - subproof = (bounds', assms, inputs, subproof)}) = - let - val bounds = distinct (op =) bounds - val bounds' = distinct (op =) bounds' - fun mk_prop_of_term concl = - concl |> fastype_of concl = \<^typ>\bool\ ? curry (op $) \<^term>\Trueprop\ - fun remove_assumption_id assumption_id prems = - filter_out (curry (op =) assumption_id) prems - fun add_assumption assumption concl = - \<^Const>\Pure.imp for \mk_prop_of_term assumption\ \mk_prop_of_term concl\\ - fun inline_assumption assumption assumption_id - (VeriT_Node {id, rule, prems, proof_ctxt, concl}) = - mk_node id rule (remove_assumption_id assumption_id prems) proof_ctxt - (add_assumption assumption concl) - fun find_input_steps_and_inline [] = [] - | find_input_steps_and_inline - (VeriT_Node {id = id', rule, prems, concl, ...} :: steps) = - if rule = input_rule then - find_input_steps_and_inline (map (inline_assumption concl id') steps) - else - mk_node (id') rule prems [] concl :: find_input_steps_and_inline steps - - fun free_bounds bounds (concl) = - fold (fn (var, typ) => fn t => Logic.all (Free (var, typ)) t) bounds concl - val subproof = subproof - |> flat o map linearize - |> map (map_node_concl (fold add_assumption (assms @ inputs))) - |> map (map_node_concl (free_bounds (bounds @ bounds'))) - |> find_input_steps_and_inline - val concl = free_bounds bounds concl - in - subproof @ [mk_node id rule prems proof_ctxt concl] - end - in linearize end - -fun rule_of (VeriT_Replay_Node {rule,...}) = rule -fun subproof_of (VeriT_Replay_Node {subproof = (_, _, _, subproof),...}) = subproof - - -(* Massage Skolems for Sledgehammer. - -We have to make sure that there is an "arrow" in the graph for skolemization steps. - - -A. The normal easy case - -This function detects the steps of the form - P \ Q :skolemization - Q :resolution with P -and replace them by - Q :skolemization -Throwing away the step "P \ Q" completely. This throws away a lot of information, but it does not -matter too much for Sledgehammer. - - -B. Skolems in subproofs -Supporting this is more or less hopeless as long as the Isar reconstruction of Sledgehammer -does not support more features like definitions. veriT is able to generate proofs with skolemization -happening in subproofs inside the formula. - (assume "A \ P" - ... - P \ Q :skolemization in the subproof - ...) - hence A \ P \ A \ Q :lemma - ... - R :something with some rule -and replace them by - R :skolemization with some rule -Without any subproof -*) -fun remove_skolem_definitions_proof steps = - let - fun replace_equivalent_by_imp (judgement $ ((Const(\<^const_name>\HOL.eq\, typ) $ arg1) $ arg2)) = - judgement $ ((Const(\<^const_name>\HOL.implies\, typ) $ arg1) $ arg2) - | replace_equivalent_by_imp a = a (*This case is probably wrong*) - fun remove_skolem_definitions (VeriT_Replay_Node {id = id, rule = rule, args = args, - prems = prems, - proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, insts = insts, - declarations = declarations, - subproof = (vars, assms', extra_assms', subproof)}) (prems_to_remove, skolems) = - let - val prems = prems - |> filter_out (member (op =) prems_to_remove) - val trivial_step = is_SH_trivial rule - fun has_skolem_substep st NONE = if is_skolemization (rule_of st) then SOME (rule_of st) - else fold has_skolem_substep (subproof_of st) NONE - | has_skolem_substep _ a = a - val promote_to_skolem = exists (fn t => member (op =) skolems t) prems - val promote_from_assms = fold has_skolem_substep subproof NONE <> NONE - val promote_step = promote_to_skolem orelse promote_from_assms - val skolem_step_to_skip = is_skolemization rule orelse - (promote_from_assms andalso length prems > 1) - val is_skolem = is_skolemization rule orelse promote_step - val prems = prems - |> filter_out (fn t => member (op =) skolems t) - |> is_skolem ? filter_out (String.isPrefix id) - val rule = (if promote_step then default_skolem_rule else rule) - val subproof = subproof - |> (is_skolem ? K []) (*subproofs of skolemization steps are useless for SH*) - |> map (fst o (fn st => remove_skolem_definitions st (prems_to_remove, skolems))) - (*no new definitions in subproofs*) - |> flat - val concl = concl - |> is_skolem ? replace_equivalent_by_imp - val step = (if skolem_step_to_skip orelse rule = veriT_def orelse trivial_step then [] - else mk_replay_node id rule args prems proof_ctxt concl bounds insts declarations - (vars, assms', extra_assms', subproof) - |> single) - val defs = (if rule = veriT_def orelse trivial_step then id :: prems_to_remove - else prems_to_remove) - val skolems = (if skolem_step_to_skip then id :: skolems else skolems) - in - (step, (defs, skolems)) - end - in - fold_map remove_skolem_definitions steps ([], []) - |> fst - |> flat - end - -local - fun import_proof_and_post_process typs funs lines ctxt = - let - val compress = SMT_Config.compress_verit_proofs ctxt - val smtlib_lines_without_qm = - lines - |> map single - |> map SMTLIB.parse - |> map remove_all_qm2 - val (raw_steps, _, _) = - parse_raw_proof_steps NONE smtlib_lines_without_qm SMTLIB_Proof.empty_name_binding - - fun process step (cx, cx') = - let fun postprocess step (cx, cx') = - let val (step, cx) = postprocess_proof compress ctxt step cx - in (step, (cx, cx')) end - in uncurry (fold_map postprocess) (preprocess compress step (cx, cx')) end - val step = - (empty_context ctxt typs funs, []) - |> fold_map process raw_steps - |> (fn (steps, (cx, _)) => (flat steps, cx)) - |> compress? apfst combine_proof_steps - in step end -in - -fun parse typs funs lines ctxt = - let - val (u, env) = import_proof_and_post_process typs funs lines ctxt - val t = u - |> remove_skolem_definitions_proof - |> flat o (map linearize_proof) - fun node_to_step (VeriT_Node {id, rule, prems, concl, ...}) = - mk_step id rule prems [] concl [] - in - (map node_to_step t, ctxt_of env) - end - -fun parse_replay typs funs lines ctxt = - let - val (u, env) = import_proof_and_post_process typs funs lines ctxt - val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> u) - in - (u, ctxt_of env) - end -end - end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML @@ -1,531 +1,531 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_isar.ML Author: Jasmin Blanchette, TU Muenchen Author: Steffen Juilf Smolka, TU Muenchen Isar proof reconstruction from ATP proofs. *) signature SLEDGEHAMMER_ISAR = sig type atp_step_name = ATP_Proof.atp_step_name type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step type 'a atp_proof = 'a ATP_Proof.atp_proof type stature = ATP_Problem_Generate.stature type one_line_params = Sledgehammer_Proof_Methods.one_line_params val trace : bool Config.T type isar_params = bool * (string option * string option) * Time.time * real option * bool * bool * (term, string) atp_step list * thm val proof_text : Proof.context -> bool -> bool option -> bool -> (unit -> isar_params) -> int -> one_line_params -> string end; structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR = struct open ATP_Util open ATP_Problem open ATP_Problem_Generate open ATP_Proof open ATP_Proof_Reconstruct open Sledgehammer_Util open Sledgehammer_Proof_Methods open Sledgehammer_Isar_Proof open Sledgehammer_Isar_Preplay open Sledgehammer_Isar_Compress open Sledgehammer_Isar_Minimize structure String_Redirect = ATP_Proof_Redirect( type key = atp_step_name val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') val string_of = fst) open String_Redirect val trace = Attrib.setup_config_bool \<^binding>\sledgehammer_isar_trace\ (K false) val e_definition_rule = "definition" val e_skolemize_rule = "skolemize" val leo2_extcnf_forall_neg_rule = "extcnf_forall_neg" val satallax_skolemize_rule = "tab_ex" val vampire_skolemisation_rule = "skolemisation" val veriT_la_generic_rule = "la_generic" val veriT_simp_arith_rule = "simp_arith" -val veriT_skolemize_rules = Verit_Proof.skolemization_steps +val veriT_skolemize_rules = Lethe_Proof.skolemization_steps val z3_skolemize_rule = Z3_Proof.string_of_rule Z3_Proof.Skolemize val z3_th_lemma_rule_prefix = Z3_Proof.string_of_rule (Z3_Proof.Th_Lemma "") val zipperposition_cnf_rule = "cnf" val symbol_introduction_rules = [e_definition_rule, e_skolemize_rule, leo2_extcnf_forall_neg_rule, satallax_skolemize_rule, spass_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule, zipperposition_cnf_rule, zipperposition_define_rule] @ veriT_skolemize_rules fun is_ext_rule rule = (rule = leo2_extcnf_equal_neg_rule) val is_maybe_ext_rule = is_ext_rule orf String.isPrefix satallax_tab_rule_prefix val is_symbol_introduction_rule = member (op =) symbol_introduction_rules fun is_arith_rule rule = String.isPrefix z3_th_lemma_rule_prefix rule orelse rule = veriT_simp_arith_rule orelse rule = veriT_la_generic_rule fun raw_label_of_num num = (num, 0) fun label_of_clause [(num, _)] = raw_label_of_num num | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0) fun add_global_fact ss = apsnd (union (op =) ss) fun add_fact_of_dependency [(_, ss as _ :: _)] = add_global_fact ss | add_fact_of_dependency names = apfst (insert (op =) (label_of_clause names)) fun add_line_pass1 (line as (name, role, t, rule, [])) lines = (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or definitions. *) if role = Conjecture orelse role = Negated_Conjecture then line :: lines else if t aconv \<^prop>\True\ then map (replace_dependencies_in_line (name, [])) lines else if role = Definition orelse role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then line :: lines else if role = Axiom then lines (* axioms (facts) need no proof lines *) else map (replace_dependencies_in_line (name, [])) lines | add_line_pass1 line lines = line :: lines fun add_lines_pass2 res [] = rev res | add_lines_pass2 res ((line as (name, role, t, rule, deps)) :: lines) = let fun normalize role = role = Conjecture ? (HOLogic.dest_Trueprop #> s_not #> HOLogic.mk_Trueprop) val norm_t = normalize role t val is_duplicate = exists (fn (prev_name, prev_role, prev_t, _, _) => (prev_role = Hypothesis andalso prev_t aconv t) orelse (member (op =) deps prev_name andalso Term.aconv_untyped (normalize prev_role prev_t, norm_t))) res fun looks_boring () = t aconv \<^prop>\False\ orelse length deps < 2 fun is_symbol_introduction_line (_, _, _, rule', deps') = is_symbol_introduction_rule rule' andalso member (op =) deps' name fun is_before_symbol_introduction_rule () = exists is_symbol_introduction_line lines in if is_duplicate orelse (role = Plain andalso not (is_symbol_introduction_rule rule) andalso not (is_ext_rule rule) andalso not (is_arith_rule rule) andalso not (null lines) andalso looks_boring () andalso not (is_before_symbol_introduction_rule ())) then add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines) else add_lines_pass2 (line :: res) lines end type isar_params = bool * (string option * string option) * Time.time * real option * bool * bool * (term, string) atp_step list * thm val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method, SATx_Method] val basic_simp_based_methods = [Auto_Method, Simp_Method, Fastforce_Method, Force_Method] val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method] val arith_methods = basic_arith_methods @ basic_simp_based_methods @ basic_systematic_methods val systematic_methods = basic_systematic_methods @ basic_arith_methods @ basic_simp_based_methods @ [Metis_Method (SOME full_typesN, NONE), Metis_Method (SOME no_typesN, NONE)] val rewrite_methods = basic_simp_based_methods @ basic_systematic_methods @ basic_arith_methods val skolem_methods = Moura_Method :: systematic_methods fun isar_proof_text ctxt debug num_chained isar_proofs smt_proofs isar_params (one_line_params as ((used_facts, (_, one_line_play)), banner, subgoal, subgoal_count)) = let val _ = if debug then writeln "Constructing Isar proof..." else () fun generate_proof_text () = let val (verbose, alt_metis_args, preplay_timeout, compress, try0, minimize, atp_proof0, goal) = isar_params () in if null atp_proof0 then one_line_proof_text ctxt 0 one_line_params else let val systematic_methods' = insert (op =) (Metis_Method alt_metis_args) systematic_methods fun massage_methods (meths as meth :: _) = if not try0 then [meth] else if smt_proofs then insert (op =) (SMT_Method SMT_Z3) meths else meths val (params, _, concl_t) = strip_subgoal goal subgoal ctxt val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params val ctxt = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd val do_preplay = preplay_timeout <> Time.zeroTime val compress = (case compress of NONE => if isar_proofs = NONE andalso do_preplay then 1000.0 else 10.0 | SOME n => n) fun is_fixed ctxt = Variable.is_declared ctxt orf Name.is_skolem fun introduced_symbols_of ctxt t = Term.add_frees t [] |> filter_out (is_fixed ctxt o fst) |> rev fun get_role keep_role ((num, _), role, t, rule, _) = if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE val trace = Config.get ctxt trace val string_of_atp_steps = let val to_string = ATP_Proof.string_of_atp_step (Syntax.string_of_term ctxt) I in enclose "[\n" "\n]" o cat_lines o map (enclose " " "," o to_string) end val atp_proof = atp_proof0 |> trace ? tap (tracing o prefix "atp_proof0 = " o string_of_atp_steps) |> distinct (op =) (* Zipperposition generates duplicate lines *) |> (fn lines => fold_rev add_line_pass1 lines []) |> add_lines_pass2 [] |> trace ? tap (tracing o prefix "atp_proof = " o string_of_atp_steps) val conjs = map_filter (fn (name, role, _, _, _) => if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE) atp_proof val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof fun add_lemma ((label, goal), rule) ctxt = let val (obtains, proof_methods) = (if is_symbol_introduction_rule rule then (introduced_symbols_of ctxt goal, skolem_methods) else if is_arith_rule rule then ([], arith_methods) else ([], rewrite_methods)) ||> massage_methods val prove = Prove { qualifiers = [], obtains = obtains, label = label, goal = goal, subproofs = [], facts = ([], []), proof_methods = proof_methods, comment = ""} in (prove, ctxt |> not (null obtains) ? (Variable.add_fixes (map fst obtains) #> snd)) end val (lems, _) = fold_map add_lemma (map_filter (get_role (member (op =) [Definition, Lemma])) atp_proof) ctxt val bot = #1 (List.last atp_proof) val refute_graph = atp_proof |> map (fn (name, _, _, _, from) => (from, name)) |> make_refute_graph bot |> fold (Atom_Graph.default_node o rpair ()) conjs val axioms = axioms_of_refute_graph refute_graph conjs val tainted = tainted_atoms_of_refute_graph refute_graph conjs val is_clause_tainted = exists (member (op =) tainted) val steps = Symtab.empty |> fold (fn (name as (s, _), role, t, rule, _) => Symtab.update_new (s, (rule, t |> (if is_clause_tainted [name] then HOLogic.dest_Trueprop #> role <> Conjecture ? s_not #> fold exists_of (map Var (Term.add_vars t [])) #> HOLogic.mk_Trueprop else I)))) atp_proof fun is_referenced_in_step _ (Let _) = false | is_referenced_in_step l (Prove {subproofs, facts = (ls, _), ...}) = member (op =) ls l orelse exists (is_referenced_in_proof l) subproofs and is_referenced_in_proof l (Proof {steps, ...}) = exists (is_referenced_in_step l) steps fun insert_lemma_in_step lem (step as Prove {qualifiers, obtains, label, goal, subproofs, facts = (ls, gs), proof_methods, comment}) = let val l' = the (label_of_isar_step lem) in if member (op =) ls l' then [lem, step] else let val refs = map (is_referenced_in_proof l') subproofs in if length (filter I refs) = 1 then [Prove { qualifiers = qualifiers, obtains = obtains, label = label, goal = goal, subproofs = map2 (fn false => I | true => insert_lemma_in_proof lem) refs subproofs, facts = (ls, gs), proof_methods = proof_methods, comment = comment}] else [lem, step] end end and insert_lemma_in_steps lem [] = [lem] | insert_lemma_in_steps lem (step :: steps) = if is_referenced_in_step (the (label_of_isar_step lem)) step then insert_lemma_in_step lem step @ steps else step :: insert_lemma_in_steps lem steps and insert_lemma_in_proof lem (proof as Proof {steps, ...}) = isar_proof_with_steps proof (insert_lemma_in_steps lem steps) val rule_of_clause_id = fst o the o Symtab.lookup steps o fst val finish_off = close_form #> rename_bound_vars fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> finish_off | prop_of_clause names = let val lits = map (HOLogic.dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names) in (case List.partition (can HOLogic.dest_not) lits of (negs as _ :: _, pos as _ :: _) => s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) | _ => fold (curry s_disj) lits \<^term>\False\) end |> HOLogic.mk_Trueprop |> finish_off fun maybe_show outer c = if outer andalso eq_set (op =) (c, conjs) then [Show] else [] fun isar_steps outer predecessor accum [] = accum |> (if tainted = [] then (* e.g., trivial, empty proof by Z3 *) cons (Prove { qualifiers = if outer then [Show] else [], obtains = [], label = no_label, goal = concl_t, subproofs = [], facts = sort_facts (the_list predecessor, []), proof_methods = massage_methods systematic_methods', comment = ""}) else I) |> rev | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) = let val l = label_of_clause c val t = prop_of_clause c val rule = rule_of_clause_id id val introduces_symbols = is_symbol_introduction_rule rule val deps = ([], []) |> fold add_fact_of_dependency gamma |> is_maybe_ext_rule rule ? add_global_fact [short_thm_name ctxt ext] |> sort_facts val meths = (if introduces_symbols then skolem_methods else if is_arith_rule rule then arith_methods else systematic_methods') |> massage_methods fun prove subproofs facts = Prove { qualifiers = maybe_show outer c, obtains = [], label = l, goal = t, subproofs = subproofs, facts = facts, proof_methods = meths, comment = ""} fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs in if is_clause_tainted c then (case gamma of [g] => if introduces_symbols andalso is_clause_tainted g andalso not (null accum) then let val fixes = introduced_symbols_of ctxt (prop_of_clause g) val subproof = Proof {fixes = fixes, assumptions = [], steps = rev accum} in isar_steps outer (SOME l) [prove [subproof] ([], [])] infs end else steps_of_rest (prove [] deps) | _ => steps_of_rest (prove [] deps)) else steps_of_rest (if introduces_symbols then (case introduced_symbols_of ctxt t of [] => prove [] deps | skos => Prove { qualifiers = [], obtains = skos, label = l, goal = t, subproofs = [], facts = deps, proof_methods = meths, comment = ""}) else prove [] deps) end | isar_steps outer predecessor accum (Cases cases :: infs) = let fun isar_case (c, subinfs) = isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs val c = succedent_of_cases cases val l = label_of_clause c val t = prop_of_clause c val step = Prove { qualifiers = maybe_show outer c, obtains = [], label = l, goal = t, subproofs = map isar_case (filter_out (null o snd) cases), facts = sort_facts (the_list predecessor, []), proof_methods = massage_methods systematic_methods', comment = ""} in isar_steps outer (SOME l) (step :: accum) infs end and isar_proof outer fixes assumptions lems infs = let val steps = fold_rev insert_lemma_in_steps lems (isar_steps outer NONE [] infs) in Proof {fixes = fixes, assumptions = assumptions, steps = steps} end val canonical_isar_proof = refute_graph |> trace ? tap (tracing o prefix "Refute graph:\n" o string_of_refute_graph) |> redirect_graph axioms tainted bot |> trace ? tap (tracing o prefix "Direct proof:\n" o string_of_direct_proof) |> isar_proof true params assms lems |> postprocess_isar_proof_remove_show_stuttering |> postprocess_isar_proof_remove_unreferenced_steps I |> relabel_isar_proof_canonically val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty val _ = fold_isar_steps (fn meth => K (set_preplay_outcomes_of_isar_step ctxt preplay_timeout preplay_data meth [])) (steps_of_isar_proof canonical_isar_proof) () fun str_of_preplay_outcome outcome = if Lazy.is_finished outcome then string_of_play_outcome (Lazy.force outcome) else "?" fun str_of_meth l meth = string_of_proof_method [] meth ^ " " ^ str_of_preplay_outcome (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth) fun comment_of l = map (str_of_meth l) #> commas fun trace_isar_proof label proof = if trace then tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^ string_of_isar_proof ctxt subgoal subgoal_count (comment_isar_proof comment_of proof) ^ "\n") else () fun comment_of l (meth :: _) = (case (verbose, Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)) of (false, Played _) => "" | (_, outcome) => string_of_play_outcome outcome) val (play_outcome, isar_proof) = canonical_isar_proof |> tap (trace_isar_proof "Original") |> compress_isar_proof ctxt compress preplay_timeout preplay_data |> tap (trace_isar_proof "Compressed") |> postprocess_isar_proof_remove_unreferenced_steps (do_preplay ? keep_fastest_method_of_isar_step (!preplay_data) #> minimize ? minimize_isar_step_dependencies ctxt preplay_data) |> tap (trace_isar_proof "Minimized") |> `(if do_preplay then preplay_outcome_of_isar_proof (!preplay_data) else K (Play_Timed_Out Time.zeroTime)) ||> (comment_isar_proof comment_of #> chain_isar_proof #> kill_useless_labels_in_isar_proof #> relabel_isar_proof_nicely #> rationalize_obtains_in_isar_proofs ctxt) in (case (num_chained, add_isar_steps (steps_of_isar_proof isar_proof) 0) of (0, 1) => one_line_proof_text ctxt 0 (if is_less (play_outcome_ord (play_outcome, one_line_play)) then (case isar_proof of Proof {steps = [Prove {facts = (_, gfs), proof_methods = meth :: _, ...}], ...} => let val used_facts' = map_filter (fn s => if exists (fn (s', (sc, _)) => s' = s andalso sc = Chained) used_facts then NONE else SOME (s, (Global, General))) gfs in ((used_facts', (meth, play_outcome)), banner, subgoal, subgoal_count) end) else one_line_params) ^ (if isar_proofs = SOME true then "\n(No Isar proof available)" else "") | (_, num_steps) => let val msg = (if verbose then [string_of_int num_steps ^ " step" ^ plural_s num_steps] else []) @ (if do_preplay then [string_of_play_outcome play_outcome] else []) in one_line_proof_text ctxt 0 one_line_params ^ (if isar_proofs <> NONE orelse (case play_outcome of Played _ => true | _ => false) then "\n\nIsar proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^ Active.sendback_markup_command (string_of_isar_proof ctxt subgoal subgoal_count isar_proof) else "") end) end end in if debug then generate_proof_text () else (case try generate_proof_text () of SOME s => s | NONE => one_line_proof_text ctxt 0 one_line_params ^ (if isar_proofs = SOME true then "\nWarning: Isar proof construction failed" else "")) end fun isar_proof_would_be_a_good_idea (_, play) = (case play of Played _ => false | Play_Timed_Out time => time > Time.zeroTime | Play_Failed => true) fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained (one_line_params as ((_, preplay), _, _, _)) = (if isar_proofs = SOME true orelse (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then isar_proof_text ctxt debug num_chained isar_proofs smt_proofs isar_params else one_line_proof_text ctxt num_chained) one_line_params end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML @@ -1,198 +1,198 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML Author: Jasmin Blanchette, TU Muenchen Author: Steffen Juilf Smolka, TU Muenchen Reconstructors. *) signature SLEDGEHAMMER_PROOF_METHODS = sig type stature = ATP_Problem_Generate.stature datatype SMT_backend = SMT_Z3 | SMT_Verit of string datatype proof_method = Metis_Method of string option * string option | Meson_Method | SMT_Method of SMT_backend | SATx_Method | Blast_Method | Simp_Method | Auto_Method | Fastforce_Method | Force_Method | Moura_Method | Linarith_Method | Presburger_Method | Algebra_Method datatype play_outcome = Played of Time.time | Play_Timed_Out of Time.time | Play_Failed type one_line_params = ((string * stature) list * (proof_method * play_outcome)) * string * int * int val is_proof_method_direct : proof_method -> bool val string_of_proof_method : string list -> proof_method -> string val tac_of_proof_method : Proof.context -> thm list * thm list -> proof_method -> int -> tactic val string_of_play_outcome : play_outcome -> string val play_outcome_ord : play_outcome ord val one_line_proof_text : Proof.context -> int -> one_line_params -> string end; structure Sledgehammer_Proof_Methods : SLEDGEHAMMER_PROOF_METHODS = struct open ATP_Util open ATP_Problem_Generate open ATP_Proof_Reconstruct datatype SMT_backend = SMT_Z3 | SMT_Verit of string datatype proof_method = Metis_Method of string option * string option | Meson_Method | SMT_Method of SMT_backend | SATx_Method | Blast_Method | Simp_Method | Auto_Method | Fastforce_Method | Force_Method | Moura_Method | Linarith_Method | Presburger_Method | Algebra_Method datatype play_outcome = Played of Time.time | Play_Timed_Out of Time.time | Play_Failed type one_line_params = ((string * stature) list * (proof_method * play_outcome)) * string * int * int fun is_proof_method_direct (Metis_Method _) = true | is_proof_method_direct Meson_Method = true | is_proof_method_direct (SMT_Method _) = true | is_proof_method_direct Simp_Method = true | is_proof_method_direct _ = false fun is_proof_method_multi_goal Auto_Method = true | is_proof_method_multi_goal _ = false fun maybe_paren s = s |> not (Symbol_Pos.is_identifier s) ? enclose "(" ")" fun string_of_proof_method ss meth = let val meth_s = (case meth of Metis_Method (NONE, NONE) => "metis" | Metis_Method (type_enc_opt, lam_trans_opt) => "metis (" ^ commas (map_filter I [type_enc_opt, lam_trans_opt]) ^ ")" | Meson_Method => "meson" | SMT_Method SMT_Z3 => "smt (z3)" | SMT_Method (SMT_Verit strategy) => "smt (" ^ commas ("verit" :: (if strategy = "default" then [] else [strategy])) ^ ")" | SATx_Method => "satx" | Blast_Method => "blast" | Simp_Method => if null ss then "simp" else "simp add:" | Auto_Method => "auto" | Fastforce_Method => "fastforce" | Force_Method => "force" | Moura_Method => "moura" | Linarith_Method => "linarith" | Presburger_Method => "presburger" | Algebra_Method => "algebra") in maybe_paren (space_implode " " (meth_s :: ss)) end fun tac_of_proof_method ctxt (local_facts, global_facts) meth = let fun tac_of_metis (type_enc_opt, lam_trans_opt) = let val ctxt = ctxt |> Config.put Metis_Tactic.verbose false |> Config.put Metis_Tactic.trace false in SELECT_GOAL (Metis_Tactic.metis_method ((Option.map single type_enc_opt, lam_trans_opt), global_facts) ctxt local_facts) end fun tac_of_smt SMT_Z3 = SMT_Solver.smt_tac - | tac_of_smt (SMT_Verit strategy) = Verit_Proof.verit_tac_stgy strategy + | tac_of_smt (SMT_Verit strategy) = Verit_Strategies.verit_tac_stgy strategy in (case meth of Metis_Method options => tac_of_metis options | SMT_Method backend => tac_of_smt backend ctxt (local_facts @ global_facts) | _ => Method.insert_tac ctxt local_facts THEN' (case meth of Meson_Method => Meson_Tactic.meson_general_tac ctxt global_facts | Simp_Method => Simplifier.asm_full_simp_tac (ctxt addsimps global_facts) | _ => Method.insert_tac ctxt global_facts THEN' (case meth of SATx_Method => SAT.satx_tac ctxt | Blast_Method => blast_tac ctxt | Auto_Method => SELECT_GOAL (Clasimp.auto_tac ctxt) | Fastforce_Method => Clasimp.fast_force_tac ctxt | Force_Method => Clasimp.force_tac ctxt | Moura_Method => moura_tac ctxt | Linarith_Method => Lin_Arith.tac ctxt | Presburger_Method => Cooper.tac true [] [] ctxt | Algebra_Method => Groebner.algebra_tac [] [] ctxt))) end fun string_of_play_outcome (Played time) = string_of_ext_time (false, time) | string_of_play_outcome (Play_Timed_Out time) = if time = Time.zeroTime then "" else string_of_ext_time (true, time) ^ ", timed out" | string_of_play_outcome Play_Failed = "failed" fun play_outcome_ord (Played time1, Played time2) = int_ord (apply2 Time.toMilliseconds (time1, time2)) | play_outcome_ord (Played _, _) = LESS | play_outcome_ord (_, Played _) = GREATER | play_outcome_ord (Play_Timed_Out time1, Play_Timed_Out time2) = int_ord (apply2 Time.toMilliseconds (time1, time2)) | play_outcome_ord (Play_Timed_Out _, _) = LESS | play_outcome_ord (_, Play_Timed_Out _) = GREATER | play_outcome_ord (Play_Failed, Play_Failed) = EQUAL fun apply_on_subgoal _ 1 = "by " | apply_on_subgoal 1 _ = "apply " | apply_on_subgoal i n = "prefer " ^ string_of_int i ^ " " ^ apply_on_subgoal 1 n (* FIXME *) fun proof_method_command meth i n used_chaineds _(*num_chained*) extras = let val (indirect_ss, direct_ss) = if is_proof_method_direct meth then ([], extras) else (extras, []) in (if null indirect_ss then "" else "using " ^ space_implode " " indirect_ss ^ " ") ^ apply_on_subgoal i n ^ string_of_proof_method direct_ss meth ^ (if is_proof_method_multi_goal meth andalso n <> 1 then "[1]" else "") end fun try_command_line banner play command = let val s = string_of_play_outcome play in banner ^ Active.sendback_markup_command command ^ (s |> s <> "" ? enclose " (" ")") end fun one_line_proof_text _ num_chained ((used_facts, (meth, play)), banner, subgoal, subgoal_count) = let val (chained, extra) = List.partition (fn (_, (sc, _)) => sc = Chained) used_facts in map fst extra |> proof_method_command meth subgoal subgoal_count (map fst chained) num_chained |> (if play = Play_Failed then prefix "One-line proof reconstruction failed: " else try_command_line banner play) end end; diff --git a/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML b/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML @@ -1,263 +1,263 @@ (* Title: HOL/Tools/Sledgehammer/sledgehammer_prover.ML Author: Fabian Immler, TU Muenchen Author: Makarius Author: Jasmin Blanchette, TU Muenchen Generic prover abstraction for Sledgehammer. *) signature SLEDGEHAMMER_PROVER = sig type atp_failure = ATP_Proof.atp_failure type stature = ATP_Problem_Generate.stature type type_enc = ATP_Problem_Generate.type_enc type fact = Sledgehammer_Fact.fact type proof_method = Sledgehammer_Proof_Methods.proof_method type play_outcome = Sledgehammer_Proof_Methods.play_outcome type base_slice = Sledgehammer_ATP_Systems.base_slice type atp_slice = Sledgehammer_ATP_Systems.atp_slice datatype mode = Auto_Try | Try | Normal | Minimize | MaSh datatype induction_rules = Include | Exclude | Instantiate val induction_rules_of_string : string -> induction_rules option val maybe_filter_out_induction_rules : induction_rules -> fact list -> fact list type params = {debug : bool, verbose : bool, overlord : bool, spy : bool, provers : string list, type_enc : string option, strict : bool, lam_trans : string option, uncurried_aliases : bool option, learn : bool, fact_filter : string option, induction_rules : induction_rules option, max_facts : int option, fact_thresholds : real * real, max_mono_iters : int option, max_new_mono_instances : int option, max_proofs : int, isar_proofs : bool option, compress : real option, try0 : bool, smt_proofs : bool, minimize : bool, slices : int, timeout : Time.time, preplay_timeout : Time.time, expect : string} val string_of_params : params -> string val slice_timeout : int -> int -> Time.time -> Time.time type prover_problem = {comment : string, state : Proof.state, goal : thm, subgoal : int, subgoal_count : int, factss : (string * fact list) list, found_proof : string -> unit} datatype prover_slice_extra = ATP_Slice of atp_slice | SMT_Slice of string list type prover_slice = base_slice * prover_slice_extra type prover_result = {outcome : atp_failure option, used_facts : (string * stature) list, used_from : fact list, preferred_methss : proof_method * proof_method list list, run_time : Time.time, message : (unit -> (string * stature) list * (proof_method * play_outcome)) -> string} type prover = params -> prover_problem -> prover_slice -> prover_result val SledgehammerN : string val str_of_mode : mode -> string val overlord_file_location_of_prover : string -> string * string val proof_banner : mode -> string -> string val is_atp : string -> bool val bunches_of_proof_methods : Proof.context -> bool -> bool -> string -> proof_method list list val facts_of_filter : string -> (string * fact list) list -> fact list val facts_of_basic_slice : base_slice -> (string * fact list) list -> fact list val is_fact_chained : (('a * stature) * 'b) -> bool val filter_used_facts : bool -> (''a * stature) list -> ((''a * stature) * 'b) list -> ((''a * stature) * 'b) list val repair_monomorph_context : int option -> int -> int option -> int -> Proof.context -> Proof.context val supported_provers : Proof.context -> unit end; structure Sledgehammer_Prover : SLEDGEHAMMER_PROVER = struct open ATP_Proof open ATP_Util open ATP_Problem open ATP_Problem_Generate open ATP_Proof_Reconstruct open Metis_Tactic open Sledgehammer_Fact open Sledgehammer_Proof_Methods open Sledgehammer_ATP_Systems (* Identifier that distinguishes Sledgehammer from other tools that could use "Async_Manager". *) val SledgehammerN = "Sledgehammer" datatype mode = Auto_Try | Try | Normal | Minimize | MaSh fun str_of_mode Auto_Try = "Auto Try" | str_of_mode Try = "Try" | str_of_mode Normal = "Normal" | str_of_mode Minimize = "Minimize" | str_of_mode MaSh = "MaSh" datatype induction_rules = Include | Exclude | Instantiate fun induction_rules_of_string "include" = SOME Include | induction_rules_of_string "exclude" = SOME Exclude | induction_rules_of_string "instantiate" = SOME Instantiate | induction_rules_of_string _ = NONE val is_atp = member (op =) all_atps type params = {debug : bool, verbose : bool, overlord : bool, spy : bool, provers : string list, type_enc : string option, strict : bool, lam_trans : string option, uncurried_aliases : bool option, learn : bool, fact_filter : string option, induction_rules : induction_rules option, max_facts : int option, fact_thresholds : real * real, max_mono_iters : int option, max_new_mono_instances : int option, max_proofs : int, isar_proofs : bool option, compress : real option, try0 : bool, smt_proofs : bool, minimize : bool, slices : int, timeout : Time.time, preplay_timeout : Time.time, expect : string} fun string_of_params (params : params) = YXML.content_of (ML_Pretty.string_of_polyml (ML_system_pretty (params, 100))) fun maybe_filter_out_induction_rules induction_rules : fact list -> fact list = induction_rules = Exclude ? filter_out (fn ((_, (_, ATP_Problem_Generate.Induction)), _) => true | _ => false) fun slice_timeout slice_size slices timeout = let val max_threads = Multithreading.max_threads () val batches = (slices + max_threads - 1) div max_threads in seconds (Real.fromInt slice_size * Time.toReal timeout / Real.fromInt batches) end type prover_problem = {comment : string, state : Proof.state, goal : thm, subgoal : int, subgoal_count : int, factss : (string * fact list) list, found_proof : string -> unit} datatype prover_slice_extra = ATP_Slice of atp_slice | SMT_Slice of string list type prover_slice = base_slice * prover_slice_extra type prover_result = {outcome : atp_failure option, used_facts : (string * stature) list, used_from : fact list, preferred_methss : proof_method * proof_method list list, run_time : Time.time, message : (unit -> (string * stature) list * (proof_method * play_outcome)) -> string} type prover = params -> prover_problem -> prover_slice -> prover_result fun overlord_file_location_of_prover prover = (getenv "ISABELLE_HOME_USER", "prob_" ^ prover) fun proof_banner mode prover_name = (case mode of Auto_Try => "Auto Sledgehammer (" ^ prover_name ^ ") found a proof: " | Try => "Sledgehammer (" ^ prover_name ^ ") found a proof: " | _ => "Try this: ") fun bunches_of_proof_methods ctxt smt_proofs needs_full_types desperate_lam_trans = let val misc_methodss = [[Simp_Method, Auto_Method, Blast_Method, Linarith_Method, Meson_Method, Metis_Method (NONE, NONE), Fastforce_Method, Force_Method, Presburger_Method]] val metis_methodss = [Metis_Method (SOME full_typesN, NONE) :: Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans) :: (if needs_full_types then [Metis_Method (SOME really_full_type_enc, NONE), Metis_Method (SOME full_typesN, SOME desperate_lam_trans)] else [Metis_Method (SOME no_typesN, SOME desperate_lam_trans)])] val smt_methodss = if smt_proofs then - [map (SMT_Method o SMT_Verit) (Verit_Proof.all_veriT_stgies (Context.Proof ctxt)), + [map (SMT_Method o SMT_Verit) (Verit_Strategies.all_veriT_stgies (Context.Proof ctxt)), [SMT_Method SMT_Z3]] else [] in misc_methodss @ metis_methodss @ smt_methodss end fun facts_of_filter fact_filter factss = (case AList.lookup (op =) factss fact_filter of SOME facts => facts | NONE => snd (hd factss)) fun facts_of_basic_slice (_, num_facts, fact_filter) factss = facts_of_filter fact_filter factss |> take num_facts fun is_fact_chained ((_, (sc, _)), _) = sc = Chained fun filter_used_facts keep_chained used = filter ((member (eq_fst (op =)) used o fst) orf (if keep_chained then is_fact_chained else K false)) val max_fact_instances = 10 (* FUDGE *) fun repair_monomorph_context max_iters best_max_iters max_new_instances best_max_new_instances = Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters) #> Config.put Monomorph.max_new_instances (max_new_instances |> the_default best_max_new_instances) #> Config.put Monomorph.max_thm_instances max_fact_instances fun supported_provers ctxt = let val local_provers = sort_strings (local_atps @ SMT_Config.available_solvers_of ctxt) val remote_provers = sort_strings remote_atps in writeln ("Supported provers: " ^ commas (local_provers @ remote_provers)) end end;