# HG changeset patch # User Mathias Fleury # Date 1602521984 -7200 # Node ID b44e894796d5160642338720642d94768908a274 # Parent 2c7f0ef8323aef728a64c06191b5efd1d97d4223 add reconstruction for the SMT solver veriT * * * Improved veriT reconstruction diff -r 2c7f0ef8323a -r b44e894796d5 NEWS --- a/NEWS Mon Oct 12 17:42:15 2020 +0200 +++ b/NEWS Mon Oct 12 18:59:44 2020 +0200 @@ -158,6 +158,8 @@ are in working order again, as opposed to outputting "GaveUp" on nearly all problems. +* SMT reconstruction: It is now possible to reconstruct proofs from the +SMT solver veriT via the tactic veriT_smt. *** FOL *** diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/ROOT --- a/src/HOL/ROOT Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/ROOT Mon Oct 12 18:59:44 2020 +0200 @@ -932,6 +932,9 @@ SMT_Examples SMT_Word_Examples SMT_Tests + theories[condition=VERIT_SOLVER] + SMT_Tests_Verit + SMT_Examples_Verit session "HOL-SPARK" in "SPARK" = "HOL-Word" + theories diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Real.thy --- a/src/HOL/Real.thy Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Real.thy Mon Oct 12 18:59:44 2020 +0200 @@ -1738,6 +1738,44 @@ for x y :: real by auto +lemma [smt_arith_multiplication]: + fixes A B :: real and p n :: real + assumes "A \ B" "0 < n" "p > 0" + shows "(A / n) * p \ (B / n) * p" + using assms by (auto simp: field_simps) + +lemma [smt_arith_multiplication]: + fixes A B :: real and p n :: real + assumes "A < B" "0 < n" "p > 0" + shows "(A / n) * p < (B / n) * p" + using assms by (auto simp: field_simps) + +lemma [smt_arith_multiplication]: + fixes A B :: real and p n :: int + assumes "A \ B" "0 < n" "p > 0" + shows "(A / n) * p \ (B / n) * p" + using assms by (auto simp: field_simps) + +lemma [smt_arith_multiplication]: + fixes A B :: real and p n :: int + assumes "A < B" "0 < n" "p > 0" + shows "(A / n) * p < (B / n) * p" + using assms by (auto simp: field_simps) + +lemmas [smt_arith_multiplication] = + verit_le_mono_div[THEN mult_left_mono, unfolded int_distrib, of _ _ \nat (floor (_ :: real))\ \nat (floor (_ :: real))\] + div_le_mono[THEN mult_left_mono, unfolded int_distrib, of _ _ \nat (floor (_ :: real))\ \nat (floor (_ :: real))\] + verit_le_mono_div_int[THEN mult_left_mono, unfolded int_distrib, of _ _ \floor (_ :: real)\ \floor (_ :: real)\] + zdiv_mono1[THEN mult_left_mono, unfolded int_distrib, of _ _ \floor (_ :: real)\ \floor (_ :: real)\] + arg_cong[of _ _ \\a :: real. a / real (n::nat) * real (p::nat)\ for n p :: nat, THEN sym] + arg_cong[of _ _ \\a :: real. a / real_of_int n * real_of_int p\ for n p :: int, THEN sym] + arg_cong[of _ _ \\a :: real. a / n * p\ for n p :: real, THEN sym] + +lemmas [smt_arith_simplify] = + floor_one floor_numeral div_by_1 times_divide_eq_right + nonzero_mult_div_cancel_left division_ring_divide_zero div_0 + divide_minus_left zero_less_divide_iff + subsection \Setup for Argo\ diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/SMT.thy --- a/src/HOL/SMT.thy Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/SMT.thy Mon Oct 12 18:59:44 2020 +0200 @@ -6,7 +6,7 @@ section \Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2\ theory SMT - imports Divides + imports Divides Numeral_Simprocs keywords "smt_status" :: diag begin @@ -199,10 +199,18 @@ 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 @@ -214,6 +222,10 @@ 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 @@ -229,13 +241,6 @@ \b \ c \ (if b then x else y) \ (if c then x else y)\ by auto -lemma verit_ite_intro_simp: - \(if c then (a :: 'a) = (if c then P else Q') else Q) = (if c then a = P else Q)\ - \(if c then R else b = (if c then R' else Q')) = - (if c then R else b = Q')\ - \(if c then a' = a' else b' = b')\ - by (auto split: if_splits) - lemma verit_or_neg: \(A \ B) \ B \ \A\ \(\A \ B) \ B \ A\ @@ -250,13 +255,21 @@ \(\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_tmp_bfun_elim: +lemma verit_bfun_elim: \(if b then P True else P False) = P b\ - by (cases b) auto + \(\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\ @@ -283,6 +296,313 @@ \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_eq 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 divmod_step_eq 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\ @@ -308,6 +628,7 @@ 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\ @@ -316,12 +637,6 @@ ML_file \Tools/SMT/verit_replay.ML\ ML_file \Tools/SMT/smt_systems.ML\ -method_setup smt = \ - Scan.optional Attrib.thms [] >> - (fn thms => fn ctxt => - METHOD (fn facts => HEADGOAL (SMT_Solver.smt_tac ctxt (thms @ facts)))) -\ "apply an SMT solver to the current goal" - subsection \Configuration\ @@ -372,7 +687,7 @@ declare [[cvc3_options = ""]] declare [[cvc4_options = "--full-saturate-quant --inst-when=full-last-call --inst-no-entail --term-db-mode=relevant --multi-trigger-linear"]] -declare [[verit_options = "--index-fresh-sorts"]] +declare [[verit_options = ""]] declare [[z3_options = ""]] text \ @@ -541,7 +856,7 @@ "0 * x = 0" "1 * x = x" "x + y = y + x" - by (auto simp add: mult_2) + by auto lemma [z3_rule]: (* for def-axiom *) "P = Q \ P \ Q" diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/SMT_Examples/SMT_Examples_Verit.certs --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT_Examples/SMT_Examples_Verit.certs Mon Oct 12 18:59:44 2020 +0200 @@ -0,0 +1,8792 @@ +c055bd110ab5330ffea7ea071d4ae520b8d8969f 12 0 +unsat +(assume a0 (! (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 (a0 t5 t8)) +(step t10 (cl (not false)) :rule false) +(step t11 (cl) :rule resolution :premises (t9 t10)) +22855940eecd8c13f695c6c889897cf3ad4802c6 15 0 +unsat +(assume a0 (! (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 (a0 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)) +5151c0ef6f1906845c0fb5bdaf2e74bf1b74e293 17 0 +unsat +(assume a0 (! (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 (a0 t10 t13)) +(step t15 (cl (not false)) :rule false) +(step t16 (cl) :rule resolution :premises (t14 t15)) +203f8541d7252e23e071560483b681b0d332320d 29 0 +unsat +(assume a0 (! (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 (a0 t22 t25)) +(step t27 (cl (not false)) :rule false) +(step t28 (cl) :rule resolution :premises (t26 t27)) +f08b7e9d00093c5579cef71fa922f4b1ea326f1a 12 0 +unsat +(assume a0 (! (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 (a0 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)) +0a3e14bc8f00b5bf07f38962ea996d3c2d20ffc8 12 0 +unsat +(assume a0 (! (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 (a0 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)) +9f6c7e127b3fe4f4c06297dd72a18082c9f01f8a 9 0 +unsat +(assume a0 (! (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 (a0 t2 t5)) +(step t7 (cl (not false)) :rule false) +(step t8 (cl) :rule resolution :premises (t6 t7)) +b2829fdf215ad44dbeaf94aa26094427083ac2a9 59 0 +unsat +(assume a0 (! (or a$ (or b$ (or c$ d$))) :named @p_1)) +(assume a2 (! (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 a3 (! (or (! (not (! (and b$ (! (or x$ (not x$)) :named @p_13)) :named @p_14)) :named @p_17) c$) :named @p_18)) +(assume a4 (! (or (! (not (! (or d$ false) :named @p_22)) :named @p_24) c$) :named @p_25)) +(assume a5 (! (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 (a0 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 (a2 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 (a3 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 (a4 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 (a5 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)) +e3b0f8bc56ad2ce76e39c6e618b58c09ec3f7474 38 0 +unsat +(assume a0 (! (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 a1 (! (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 veriT_vr0) (:= ?v1 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 (a0 t3 t4)) +(anchor :step t6 :args ((:= veriT_vr0 veriT_vr2) (:= veriT_vr1 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 (a1 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)) +2b7ca9ab80d78db50690054e655456fb60950de8 333 0 +unsat +(assume a0 (not x0$)) +(assume a1 (not x30$)) +(assume a2 (not x29$)) +(assume a3 (not x59$)) +(assume a4 (! (or x1$ (or x31$ x0$)) :named @p_57)) +(assume a6 (! (or x3$ (or x33$ x2$)) :named @p_60)) +(assume a7 (! (or x4$ (or x34$ x3$)) :named @p_63)) +(assume a8 (or x35$ x4$)) +(assume a9 (! (or x5$ (or x36$ x30$)) :named @p_66)) +(assume a11 (! (or x7$ (or x38$ (or x6$ x32$))) :named @p_69)) +(assume a13 (! (or x9$ (or x40$ (or x8$ x34$))) :named @p_72)) +(assume a16 (! (or x11$ (or x43$ (or x10$ x37$))) :named @p_75)) +(assume a18 (! (or x13$ (or x45$ (or x12$ x39$))) :named @p_78)) +(assume a20 (! (or x47$ (or x14$ x41$)) :named @p_81)) +(assume a21 (! (or x15$ (or x48$ x42$)) :named @p_84)) +(assume a23 (! (or x17$ (or x50$ (or x16$ x44$))) :named @p_87)) +(assume a25 (! (or x19$ (or x52$ (or x18$ x46$))) :named @p_90)) +(assume a28 (! (or x21$ (or x55$ (or x20$ x49$))) :named @p_93)) +(assume a30 (! (or x23$ (or x57$ (or x22$ x51$))) :named @p_96)) +(assume a32 (! (or x59$ (or x24$ x53$)) :named @p_99)) +(assume a33 (or x25$ x54$)) +(assume a35 (! (or x27$ (or x26$ x56$)) :named @p_102)) +(assume a37 (! (or x29$ (or x28$ x58$)) :named @p_105)) +(assume a41 (or (! (not x2$) :named @p_1) (! (not x32$) :named @p_2))) +(assume a42 (or @p_1 (! (not x1$) :named @p_3))) +(assume a43 (or @p_2 @p_3)) +(assume a47 (or (! (not x4$) :named @p_4) (! (not x34$) :named @p_5))) +(assume a48 (or @p_4 (! (not x3$) :named @p_6))) +(assume a49 (or @p_5 @p_6)) +(assume a54 (or (! (not x6$) :named @p_7) (! (not x37$) :named @p_8))) +(assume a55 (or @p_7 (! (not x5$) :named @p_9))) +(assume a56 (or @p_7 (! (not x31$) :named @p_10))) +(assume a57 (or @p_8 @p_9)) +(assume a58 (or @p_8 @p_10)) +(assume a59 (or @p_9 @p_10)) +(assume a63 (or (! (not x38$) :named @p_11) @p_7)) +(assume a64 (or @p_11 @p_2)) +(assume a66 (or (! (not x8$) :named @p_12) (! (not x39$) :named @p_13))) +(assume a67 (or @p_12 (! (not x7$) :named @p_14))) +(assume a68 (or @p_12 (! (not x33$) :named @p_15))) +(assume a69 (or @p_13 @p_14)) +(assume a70 (or @p_13 @p_15)) +(assume a71 (or @p_14 @p_15)) +(assume a78 (or (! (not x41$) :named @p_16) (! (not x9$) :named @p_17))) +(assume a79 (or @p_16 (! (not x35$) :named @p_18))) +(assume a80 (or @p_17 @p_18)) +(assume a81 (or (! (not x10$) :named @p_19) (! (not x42$) :named @p_20))) +(assume a82 (or @p_19 (! (not x36$) :named @p_21))) +(assume a83 (or @p_20 @p_21)) +(assume a90 (or (! (not x12$) :named @p_22) (! (not x44$) :named @p_23))) +(assume a91 (or @p_22 (! (not x11$) :named @p_24))) +(assume a92 (or @p_22 @p_11)) +(assume a93 (or @p_23 @p_24)) +(assume a94 (or @p_23 @p_11)) +(assume a95 (or @p_24 @p_11)) +(assume a99 (or (! (not x45$) :named @p_25) @p_22)) +(assume a100 (or @p_25 @p_13)) +(assume a102 (or (! (not x14$) :named @p_26) (! (not x46$) :named @p_27))) +(assume a103 (or @p_26 (! (not x13$) :named @p_28))) +(assume a104 (or @p_26 (! (not x40$) :named @p_29))) +(assume a105 (or @p_27 @p_28)) +(assume a106 (or @p_27 @p_29)) +(assume a107 (or @p_28 @p_29)) +(assume a113 (or (! (not x48$) :named @p_41) @p_20)) +(assume a114 (or (! (not x16$) :named @p_30) (! (not x49$) :named @p_31))) +(assume a115 (or @p_30 (! (not x15$) :named @p_32))) +(assume a116 (or @p_30 (! (not x43$) :named @p_33))) +(assume a117 (or @p_31 @p_32)) +(assume a118 (or @p_31 @p_33)) +(assume a119 (or @p_32 @p_33)) +(assume a126 (or (! (not x18$) :named @p_34) (! (not x51$) :named @p_35))) +(assume a127 (or @p_34 (! (not x17$) :named @p_36))) +(assume a128 (or @p_34 @p_25)) +(assume a129 (or @p_35 @p_36)) +(assume a130 (or @p_35 @p_25)) +(assume a131 (or @p_36 @p_25)) +(assume a134 (or (! (not x19$) :named @p_37) @p_27)) +(assume a138 (or (! (not x53$) :named @p_38) @p_37)) +(assume a139 (or @p_38 (! (not x47$) :named @p_39))) +(assume a140 (or @p_37 @p_39)) +(assume a141 (or (! (not x20$) :named @p_40) (! (not x54$) :named @p_42))) +(assume a142 (or @p_40 @p_41)) +(assume a143 (or @p_42 @p_41)) +(assume a150 (or (! (not x22$) :named @p_43) (! (not x56$) :named @p_44))) +(assume a151 (or @p_43 (! (not x21$) :named @p_45))) +(assume a152 (or @p_43 (! (not x50$) :named @p_46))) +(assume a153 (or @p_44 @p_45)) +(assume a154 (or @p_44 @p_46)) +(assume a155 (or @p_45 @p_46)) +(assume a162 (or (! (not x24$) :named @p_47) (! (not x58$) :named @p_48))) +(assume a163 (or @p_47 (! (not x23$) :named @p_49))) +(assume a164 (or @p_47 (! (not x52$) :named @p_50))) +(assume a165 (or @p_48 @p_49)) +(assume a166 (or @p_48 @p_50)) +(assume a167 (or @p_49 @p_50)) +(assume a172 (or (! (not x26$) :named @p_51) (! (not x25$) :named @p_52))) +(assume a173 (or @p_51 (! (not x55$) :named @p_53))) +(assume a174 (or @p_52 @p_53)) +(assume a178 (or (! (not x28$) :named @p_54) (! (not x27$) :named @p_55))) +(assume a179 (or @p_54 (! (not x57$) :named @p_56))) +(assume a180 (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 (a4 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 (a6 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 (a7 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 (a9 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 (a11 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 (a13 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 (a16 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 (a18 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 (a20 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 (a21 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 (a23 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 (a25 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 (a28 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 (a30 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 (a32 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 (a35 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 (a37 t150 t151)) +(step t153 (cl x1$ x31$ x0$) :rule or :premises (t104)) +(step t154 (cl x1$ x31$) :rule resolution :premises (t153 a0)) +(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 (a8)) +(step t158 (cl x5$ x36$ x30$) :rule or :premises (t113)) +(step t159 (cl x5$ x36$) :rule resolution :premises (t158 a1)) +(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 a3)) +(step t172 (cl x25$ x54$) :rule or :premises (a33)) +(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 a2)) +(step t176 (cl @p_1 @p_2) :rule or :premises (a41)) +(step t177 (cl @p_1 @p_3) :rule or :premises (a42)) +(step t178 (cl @p_2 @p_3) :rule or :premises (a43)) +(step t179 (cl @p_4 @p_5) :rule or :premises (a47)) +(step t180 (cl @p_4 @p_6) :rule or :premises (a48)) +(step t181 (cl @p_5 @p_6) :rule or :premises (a49)) +(step t182 (cl @p_7 @p_8) :rule or :premises (a54)) +(step t183 (cl @p_7 @p_9) :rule or :premises (a55)) +(step t184 (cl @p_7 @p_10) :rule or :premises (a56)) +(step t185 (cl @p_8 @p_9) :rule or :premises (a57)) +(step t186 (cl @p_8 @p_10) :rule or :premises (a58)) +(step t187 (cl @p_9 @p_10) :rule or :premises (a59)) +(step t188 (cl @p_11 @p_7) :rule or :premises (a63)) +(step t189 (cl @p_11 @p_2) :rule or :premises (a64)) +(step t190 (cl @p_12 @p_13) :rule or :premises (a66)) +(step t191 (cl @p_12 @p_14) :rule or :premises (a67)) +(step t192 (cl @p_12 @p_15) :rule or :premises (a68)) +(step t193 (cl @p_13 @p_14) :rule or :premises (a69)) +(step t194 (cl @p_13 @p_15) :rule or :premises (a70)) +(step t195 (cl @p_14 @p_15) :rule or :premises (a71)) +(step t196 (cl @p_16 @p_17) :rule or :premises (a78)) +(step t197 (cl @p_16 @p_18) :rule or :premises (a79)) +(step t198 (cl @p_17 @p_18) :rule or :premises (a80)) +(step t199 (cl @p_19 @p_20) :rule or :premises (a81)) +(step t200 (cl @p_19 @p_21) :rule or :premises (a82)) +(step t201 (cl @p_20 @p_21) :rule or :premises (a83)) +(step t202 (cl @p_22 @p_23) :rule or :premises (a90)) +(step t203 (cl @p_22 @p_24) :rule or :premises (a91)) +(step t204 (cl @p_22 @p_11) :rule or :premises (a92)) +(step t205 (cl @p_23 @p_24) :rule or :premises (a93)) +(step t206 (cl @p_23 @p_11) :rule or :premises (a94)) +(step t207 (cl @p_24 @p_11) :rule or :premises (a95)) +(step t208 (cl @p_25 @p_22) :rule or :premises (a99)) +(step t209 (cl @p_25 @p_13) :rule or :premises (a100)) +(step t210 (cl @p_26 @p_27) :rule or :premises (a102)) +(step t211 (cl @p_26 @p_28) :rule or :premises (a103)) +(step t212 (cl @p_26 @p_29) :rule or :premises (a104)) +(step t213 (cl @p_27 @p_28) :rule or :premises (a105)) +(step t214 (cl @p_27 @p_29) :rule or :premises (a106)) +(step t215 (cl @p_28 @p_29) :rule or :premises (a107)) +(step t216 (cl @p_41 @p_20) :rule or :premises (a113)) +(step t217 (cl @p_30 @p_31) :rule or :premises (a114)) +(step t218 (cl @p_30 @p_32) :rule or :premises (a115)) +(step t219 (cl @p_30 @p_33) :rule or :premises (a116)) +(step t220 (cl @p_31 @p_32) :rule or :premises (a117)) +(step t221 (cl @p_31 @p_33) :rule or :premises (a118)) +(step t222 (cl @p_32 @p_33) :rule or :premises (a119)) +(step t223 (cl @p_34 @p_35) :rule or :premises (a126)) +(step t224 (cl @p_34 @p_36) :rule or :premises (a127)) +(step t225 (cl @p_34 @p_25) :rule or :premises (a128)) +(step t226 (cl @p_35 @p_36) :rule or :premises (a129)) +(step t227 (cl @p_35 @p_25) :rule or :premises (a130)) +(step t228 (cl @p_36 @p_25) :rule or :premises (a131)) +(step t229 (cl @p_37 @p_27) :rule or :premises (a134)) +(step t230 (cl @p_38 @p_37) :rule or :premises (a138)) +(step t231 (cl @p_38 @p_39) :rule or :premises (a139)) +(step t232 (cl @p_37 @p_39) :rule or :premises (a140)) +(step t233 (cl @p_40 @p_42) :rule or :premises (a141)) +(step t234 (cl @p_40 @p_41) :rule or :premises (a142)) +(step t235 (cl @p_42 @p_41) :rule or :premises (a143)) +(step t236 (cl @p_43 @p_44) :rule or :premises (a150)) +(step t237 (cl @p_43 @p_45) :rule or :premises (a151)) +(step t238 (cl @p_43 @p_46) :rule or :premises (a152)) +(step t239 (cl @p_44 @p_45) :rule or :premises (a153)) +(step t240 (cl @p_44 @p_46) :rule or :premises (a154)) +(step t241 (cl @p_45 @p_46) :rule or :premises (a155)) +(step t242 (cl @p_47 @p_48) :rule or :premises (a162)) +(step t243 (cl @p_47 @p_49) :rule or :premises (a163)) +(step t244 (cl @p_47 @p_50) :rule or :premises (a164)) +(step t245 (cl @p_48 @p_49) :rule or :premises (a165)) +(step t246 (cl @p_48 @p_50) :rule or :premises (a166)) +(step t247 (cl @p_49 @p_50) :rule or :premises (a167)) +(step t248 (cl @p_51 @p_52) :rule or :premises (a172)) +(step t249 (cl @p_51 @p_53) :rule or :premises (a173)) +(step t250 (cl @p_52 @p_53) :rule or :premises (a174)) +(step t251 (cl @p_54 @p_55) :rule or :premises (a178)) +(step t252 (cl @p_54 @p_56) :rule or :premises (a179)) +(step t253 (cl @p_55 @p_56) :rule or :premises (a180)) +(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)) +3f50b6bc75bba3bc427e4cc7f64b27bbceb4d4e3 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 a0 (! (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 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 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 (a0 t3 t6)) +(anchor :step t8 :args ((:= veriT_vr0 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 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 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 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)) +90ffc0572f49cac01988fa0e245f0987edde82e3 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 a0 (forall ((?v0 Bool) (?v1 A$)) (= (p$ ?v0 ?v1) ?v0))) +(assume a1 (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 (a0)) +(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 (a1)) +(anchor :step t5 :args ((:= ?v1 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)) +(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 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 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 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 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 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)) +2745032ad717e20a36605714726c17c0f3260db8 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 a0 (forall ((?v0 Bool) (?v1 A$)) (= (p$ ?v0 ?v1) ?v0))) +(assume a2 (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 (a0)) +(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 (a2)) +(anchor :step t5 :args ((:= ?v1 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)) +(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 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 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 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 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 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)) +9c015c55d0136140309760a09ce754c95d1d3faa 12 0 +unsat +(assume a0 (! (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 (a0 t5 t8)) +(step t10 (cl (not false)) :rule false) +(step t11 (cl) :rule resolution :premises (t9 t10)) +08eb77f45478246ccd84e33e8e66f49351a834c5 12 0 +unsat +(assume a0 (! (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 (a0 t5 t8)) +(step t10 (cl (not false)) :rule false) +(step t11 (cl) :rule resolution :premises (t9 t10)) +d2da4c1f8287a8ebd2829df6fb08f0a8e938fa66 57 0 +unsat +(assume a0 (! (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 a1 (! (not (! (=> @p_2 (! (p$ y$) :named @p_35)) :named @p_39)) :named @p_34)) +(anchor :step t3 :args ((:= ?v0 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 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 (a0 t6 t7)) +(anchor :step t9 :args ((:= veriT_vr0 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 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 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 (a1 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)) +413fc207ba9e12ba020fc37967fff3270d5abb5e 15 0 +unsat +(assume a0 (! (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 (a0 t8 t11)) +(step t13 (cl (not false)) :rule false) +(step t14 (cl) :rule resolution :premises (t12 t13)) +845b74b5125cd3ba120ea66952b9f97a5c6ab1c3 9 0 +unsat +(assume a0 (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 a0)) +(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)) +d8dbfe74fd522039391998d9bbfb9e867f02f283 18 0 +unsat +(assume a0 (! (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 (a0 t11 t14)) +(step t16 (cl (not false)) :rule false) +(step t17 (cl) :rule resolution :premises (t15 t16)) +aa08521b4449c3ee379b738c697dd0a12da6c519 52 0 +unsat +(assume a0 (! (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 (a0 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)) +1a8db76be6f6a2bddea4084033566e745253b622 19 0 +unsat +(assume a0 (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 (a0)) +(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)) +229e9450bd215085107290b86bc3fa4754765364 14 0 +unsat +(assume a0 (! (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 (a0 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)) +b1115b3662d14722175185e9377414ecb2b1046b 9 0 +unsat +(assume a1 (! (= y$ (! (+ x$ 4) :named @p_1)) :named @p_2)) +(assume a2 (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 (a1 t4 t5)) +(step t7 (cl @p_6 (not @p_5)) :rule la_generic :args (1 1)) +(step t8 (cl) :rule resolution :premises (t7 t6 a2)) +dfc587dec547748e78f3739a9819c7922a81f17f 6 0 +unsat +(assume a0 (! (< (+ (* 3 x$) (* 7 a$)) 4) :named @p_3)) +(assume a1 (! (< 3 (* 2 x$)) :named @p_1)) +(assume a2 (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 a0 a1 a2)) +dedf92567622c8873796e3213079c20d4c0be727 20 0 +unsat +(assume a0 (! (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 a0)) +(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)) +282e1f4ae1fe9c95cfac6cd19305613aa9bb9fdc 29 0 +unsat +(assume a0 (! (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 (a0 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)) +db4fd717bbfc4454488919af499429897e517aaf 62 0 +unsat +(assume a0 (! (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 (a0 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)) +930494d7814037632f0f2cf405656cd7e35a6de4 21 0 +unsat +(assume a0 (! (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 (a0 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)) +af075a1bc11466b86b2e32272fabf79492e0766f 16 0 +unsat +(assume a0 (! (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 a0)) +(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)) +feb8f580c0ed1105dae26acee8e38d20f9cad1e9 16 0 +unsat +(assume a0 (! (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 a0)) +(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)) +91440d65260ff516e2ae884704f5e854d05caa5d 44 0 +unsat +(assume a0 (! (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 0) (:= ?v1 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 (a0 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)) +10c5b784e1c17a8b9f972e3ab359ad726230a233 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 a0 (! (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 veriT_vr0) (:= ?v1 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 (a0 t3 t6)) +(anchor :step t8 :args (veriT_vr0 veriT_vr1)) +(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 veriT_vr2) (:= veriT_vr1 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 veriT_sk0) (:= veriT_vr3 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)) +37c31d8c4e1a148b5be59fa2fa31b2125c42764a 371 0 +unsat +(assume a0 (! (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 (a0 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)) +2bf1326b4ee67b65a19f29be926aaada4159282d 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 a0 (! (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 veriT_vr0) (:= ?v1 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 (a0 t3 t6)) +(anchor :step t8 :args (veriT_vr0 veriT_vr1)) +(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 veriT_vr2) (:= veriT_vr1 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 veriT_sk0) (:= veriT_vr3 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)) +99600e9e235b6ce7e9b8a3d2c559c08161c6c86d 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 a0 (! (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 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 (a0 t3 t6)) +(anchor :step t8 :args (veriT_vr0)) +(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 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 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)) +78fb89c5958debfdc5cd4bf36e64314b72b330d7 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 a0 (! (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 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 (a0 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 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 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 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)) +3d6f96b90383a5a25927e677a0f1a6c0d59e14ee 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 a0 (! (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 (a0 t3 t6)) +(anchor :step t8 :args ((:= ?v1 veriT_vr0) (:= ?v2 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 veriT_vr2) (:= veriT_vr1 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 veriT_sk0) (:= veriT_vr3 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)) +55b72bfb066df27d9fdbc3ae3e0d5919e30bcc28 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 a0 (! (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 (a0 t3 t6)) +(anchor :step t8 :args ((:= ?v1 veriT_vr0) (:= ?v2 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 veriT_vr1)) +(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 veriT_vr2) (:= veriT_vr1 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 veriT_sk0) (:= veriT_vr3 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)) +6a29c7940c898207fc3ec2e9201980e49631f59c 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 a0 (! (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 (a0 t3 t6)) +(anchor :step t8 :args ((:= ?v1 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 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 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)) +0dca9cb26a065d9037b253edeea7dfd938c4d74a 7 0 +unsat +(assume a0 (! (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 a0)) +(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)) +00f0e6a8fcdecc425e1f539e94cddf8d2f165175 11 0 +unsat +(assume a0 (! (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 (a0 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)) +c564c917673f8ddcc3c8e83b6c17c836ab96aabf 21 0 +unsat +(assume a0 (! (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 (a0 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)) +faa000b2df6428276051b48e21683f2504617d26 33 0 +unsat +(assume a0 (! (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 a1 (! (<= 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 (a0 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 a1)) +0e67342ae09b659ec99568f03952d769c3c5f476 76 0 +unsat +(assume a4 (! (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 a1 (! (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 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 (a4 t3 t4)) +(anchor :step t6 :args ((:= veriT_vr0 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 (a1 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)) +473e52f76301e5561016da909d7bb09254857479 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 a0 (! (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 a1 (! (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 a2 (! (forall ((?v0 Nat$)) (! (<= 0 @p_1) :named @p_127)) :named @p_125)) +(assume a4 (! (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 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 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 (a0 t5 t6)) +(anchor :step t8 :args (veriT_vr0)) +(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 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 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 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 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 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 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 veriT_vr5))) +(anchor :step t17.t1 :args ((:= veriT_vr6 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 (a1 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 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 (a2 t31 t32)) +(anchor :step t34 :args ((:= veriT_vr8 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 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 (a4 t37 t38)) +(anchor :step t40 :args ((:= veriT_vr12 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 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 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 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 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 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)) +7b70726cf72a0c74489606bae1ad61426f0b0ab9 7 0 +unsat +(assume a0 (! (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 a0)) +(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)) +bf831287a0298d285882760b346e495880eb16d1 35 0 +unsat +(assume a0 (! (forall ((?v0 A$) (?v1 B$)) (! (= ?v0 (! (fst$ (! (pair$ ?v0 ?v1) :named @p_3)) :named @p_5)) :named @p_7)) :named @p_1)) +(assume a1 (! (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 veriT_vr0) (:= ?v1 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 (a0 t3 t4)) +(anchor :step t6 :args ((:= veriT_vr0 veriT_vr2) (:= veriT_vr1 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 (a1 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)) +998e9357ccbe8a5d4cb8c349e083e5fd9d181b25 67 0 +unsat +(assume a1 (! (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 a3 (! (forall ((?v0 B$) (?v1 A$)) (! (= ?v1 (! (snd$a (! (pair$ ?v0 ?v1) :named @p_19)) :named @p_21)) :named @p_23)) :named @p_17)) +(assume a4 (! (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 veriT_vr4) (:= ?v1 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 (a1 t4 t5)) +(anchor :step t7 :args ((:= veriT_vr4 veriT_vr6) (:= veriT_vr5 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 veriT_vr12) (:= ?v1 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 (a3 t10 t11)) +(anchor :step t13 :args ((:= veriT_vr12 veriT_vr14) (:= veriT_vr13 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 (a4 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)) +fc4647afe32c480bbb8b8a4c6dcf695e2b726323 97 0 +unsat +(assume a1 (! (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 a2 (! (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 veriT_vr6) (:= ?v1 veriT_vr7) (:= ?v2 veriT_vr8) (:= ?v3 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 (a1 t3 t4)) +(anchor :step t6 :args ((:= veriT_vr6 veriT_vr10) (:= veriT_vr7 veriT_vr11) (:= veriT_vr8 veriT_vr12) (:= veriT_vr9 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 (a2 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)) +e0b4b05e5f6d1b962b2a9456c742aa35d01413a7 25 0 +unsat +(assume a0 (! (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 (a0 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)) +12292726bb00f25b0aea5ab0221cdea71601b8fd 38 0 +unsat +(assume a0 (! (forall ((?v0 A$)) (! (= ?v0 (! (id$ ?v0) :named @p_4)) :named @p_6)) :named @p_2)) +(assume a1 (forall ((?v0 Bool)) (= (id$a ?v0) ?v0))) +(assume a2 (! (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 (a1)) +(anchor :step t5 :args ((:= ?v0 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 (a0 t5 t6)) +(anchor :step t8 :args ((:= veriT_vr0 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 (a2 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)) +fab16bc820af74a6055b3e7aedc3c727e6f5f6af 27 0 +unsat +(assume a0 (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 (a0)) +(anchor :step t3 :args ((:= ?v0 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)) +8f7e19d83778dca8772549964febf837f10fc07f 60 0 +unsat +(assume a0 (! (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 a1 (! (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 veriT_vr0) (:= ?v1 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 (a0 t3 t4)) +(anchor :step t6 :args ((:= veriT_vr0 veriT_vr2) (:= veriT_vr1 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 (a1 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)) +4e293cd43d036c71e0f6573a40238cbbbdfc63d9 27 0 +unsat +(assume a0 (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 (a0)) +(anchor :step t3 :args ((:= ?v0 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)) +0466d7eca133f7c37fadd223e6ede987a03df89f 125 0 +unsat +(assume a0 (! (forall ((?v0 Int)) (! (= (! (fun_app$ uu$ ?v0) :named @p_2) (! (+ ?v0 1) :named @p_5)) :named @p_7)) :named @p_1)) +(assume a1 (! (forall ((?v0 Int_int_fun$)) (! (= nil$ (! (map$ ?v0 nil$) :named @p_22)) :named @p_24)) :named @p_21)) +(assume a2 (! (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 a3 (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 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 (a0 t5 t6)) +(anchor :step t8 :args (veriT_vr0)) +(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 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 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 (a1 t14 t15)) +(anchor :step t17 :args ((:= veriT_vr2 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 veriT_vr4) (:= ?v1 veriT_vr5) (:= ?v2 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 (a2 t20 t21)) +(anchor :step t23 :args ((:= veriT_vr4 veriT_vr7) (:= veriT_vr5 veriT_vr8) (:= veriT_vr6 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 a3 t28 t39 t41 t51 t53)) +fea4fb556b09fdb802b089e39cc40016c69868f5 23 0 +unsat +(assume a0 (! (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 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 (a0 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)) +b6f9f8108480d266f9742c4ac46381a03c300fd2 107 0 +unsat +(assume a0 (! (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 a1 (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 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 (a0 t3 t4)) +(anchor :step t6 :args ((:= veriT_vr0 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 a1 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)) +a9a0192e0312157bc10296025518980a110117e0 48 0 +unsat +(assume a2 (! (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 a3 (! (less_eq$ (! (sup$ (collect$ uu$)) :named @p_2) (! (sup$ (collect$ uua$)) :named @p_1)) :named @p_31)) +(assume a4 (! (less_eq$ @p_1 @p_2) :named @p_32)) +(assume a5 (not (! (less_eq$ @p_2 @p_2) :named @p_33))) +(anchor :step t5 :args ((:= ?v0 veriT_vr4) (:= ?v1 veriT_vr5) (:= ?v2 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 (a2 t5 t6)) +(anchor :step t8 :args ((:= veriT_vr4 veriT_vr7) (:= veriT_vr5 veriT_vr8) (:= veriT_vr6 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 a3 a4 a5)) +(step t23 (cl) :rule resolution :premises (t21 t10 t22)) +33354ae90736acf767a389a3432693e110da91cb 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_sk5 () A$ (! (choice ((veriT_vr100 A$)) (not (! (=> (! (member$ veriT_vr100 @p_380) :named @p_828) (! (not (! (member$ veriT_vr100 @p_381) :named @p_830)) :named @p_831)) :named @p_827))) :named @p_802)) +(define-fun veriT_sk6 () A$ (! (choice ((veriT_vr103 A$)) (not (! (=> (! (member$ veriT_vr103 (! (set$ xs$) :named @p_715)) :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 a0 (! (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 a1 (! (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 a2 (! (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 a3 (! (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 a4 (! (forall ((?v0 A_list$)) (! (finite$ @p_2) :named @p_120)) :named @p_118)) +(assume a5 (! (forall ((?v0 A$)) (! (member$ ?v0 top$) :named @p_129)) :named @p_128)) +(assume a6 (not (! (= top$ bot$) :named @p_717))) +(assume a9 (! (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 a10 (! (= (! (fun_app$b card$a @p_715) :named @p_1077) (! (size$ @p_9) :named @p_8)) :named @p_1076)) +(assume a11 (! (= card$ card$a) :named @p_1079)) +(assume a12 (! (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 a13 (! (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 a14 (! (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 a15 (! (forall ((?v0 A_list$)) (! (= @p_2 (! (uminus$ (! (coset$ ?v0) :named @p_5)) :named @p_253)) :named @p_255)) :named @p_249)) +(assume a16 (! (forall ((?v0 A_list$)) (! (= @p_5 (! (uminus$ @p_2) :named @p_270)) :named @p_272)) :named @p_266)) +(assume a17 (! (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 a18 (! (= uu$ eq_set$) :named @p_1100)) +(assume a19 (! (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 a20 (! (= 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 a21 (! (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 a22 (! (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 a23 (! (forall ((?v0 A_list$)) (! (= @p_2 (! (set$ @p_14) :named @p_552)) :named @p_554)) :named @p_548)) +(assume a24 (! (= subset$ less_eq$) :named @p_1090)) +(assume a25 (! (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 a26 (! (forall ((?v0 A_set$) (?v1 A_set$)) (! (=> @p_18 @p_13) :named @p_602)) :named @p_593)) +(assume a27 (! (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 a28 (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 veriT_vr0) (:= ?v1 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 (a0 t28 t29)) +(anchor :step t31 :args ((:= veriT_vr0 veriT_vr2) (:= veriT_vr1 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 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 (a1 t34 t35)) +(anchor :step t37 :args ((:= veriT_vr4 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 veriT_vr6) (:= ?v1 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 (a2 t40 t41)) +(anchor :step t43 :args ((:= veriT_vr6 veriT_vr8) (:= veriT_vr7 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 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 (a3 t46 t47)) +(anchor :step t49 :args ((:= veriT_vr10 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 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 (a4 t52 t53)) +(anchor :step t55 :args ((:= veriT_vr12 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 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 (a5 t58 t59)) +(anchor :step t61 :args ((:= veriT_vr14 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 veriT_vr26) (:= ?v1 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 (a9 t64 t65)) +(anchor :step t67 :args ((:= veriT_vr26 veriT_vr28) (:= veriT_vr27 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 veriT_vr30) (:= ?v1 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 (a12 t70 t71)) +(anchor :step t73 :args ((:= veriT_vr30 veriT_vr32) (:= veriT_vr31 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 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 (a13 t76 t77)) +(anchor :step t79 :args ((:= veriT_vr34 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 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 (a14 t82 t83)) +(anchor :step t85 :args ((:= veriT_vr36 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 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 (a15 t88 t89)) +(anchor :step t91 :args ((:= veriT_vr38 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 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 (a16 t94 t95)) +(anchor :step t97 :args ((:= veriT_vr40 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 veriT_vr42) (:= ?v1 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 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 (a17 t100 t101)) +(anchor :step t103 :args (veriT_vr42 veriT_vr43)) +(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 veriT_vr45) (:= veriT_vr43 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 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 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 veriT_vr49) (:= veriT_vr46 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 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 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 veriT_vr49) (:= veriT_vr50 veriT_vr50))) +(anchor :step t112.t1 :args ((:= veriT_vr51 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 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 (a19 t115 t116)) +(anchor :step t118 :args ((:= veriT_vr53 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 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 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 (a20 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 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 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 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 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 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 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 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 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 veriT_vr61) (:= ?v1 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 (a21 t185 t186)) +(anchor :step t188 :args ((:= veriT_vr61 veriT_vr63) (:= veriT_vr62 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 veriT_vr65) (:= ?v1 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 (a22 t191 t192)) +(anchor :step t194 :args ((:= veriT_vr65 veriT_vr67) (:= veriT_vr66 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 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 (a23 t197 t198)) +(anchor :step t200 :args ((:= veriT_vr69 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 veriT_vr71) (:= ?v1 veriT_vr72))) +(anchor :step t203.t1 :args ((:= ?v2 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 (a25 t203 t204)) +(anchor :step t206 :args ((:= veriT_vr71 veriT_vr74) (:= veriT_vr72 veriT_vr75))) +(anchor :step t206.t1 :args ((:= veriT_vr73 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 veriT_vr77) (:= ?v1 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 (a26 t209 t210)) +(anchor :step t212 :args ((:= veriT_vr77 veriT_vr79) (:= veriT_vr78 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 veriT_vr81) (:= ?v1 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 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 (a27 t215 t216)) +(anchor :step t218 :args (veriT_vr81 veriT_vr82)) +(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 veriT_vr84) (:= veriT_vr82 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 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 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 veriT_vr88) (:= veriT_vr85 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 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 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 veriT_vr88) (:= veriT_vr89 veriT_vr89))) +(anchor :step t227.t1 :args ((:= veriT_vr90 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 (a28)) +(step t249 (cl (! (not @p_711) :named @p_1099) @p_712) :rule not_equiv2 :premises (a28)) +(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 a6)) +(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_799) @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_1005) (! (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 (not (! (not (! (not @p_703) :named @p_798)) :named @p_854)) @p_703) :rule not_not) +(step t353 (cl (or @p_798 (! (and (! (=> (! (fun_app$ @p_799 @p_792) :named @p_800) (! (forall ((veriT_vr90 A$)) (! (=> (! (member$ veriT_vr90 @p_380) :named @p_804) (! (not (! (member$ veriT_vr90 @p_381) :named @p_807)) :named @p_809)) :named @p_811)) :named @p_803)) :named @p_813) (! (=> (! (forall ((veriT_vr91 A$)) (! (=> (! (member$ veriT_vr91 @p_380) :named @p_816) (! (not (! (member$ veriT_vr91 @p_381) :named @p_818)) :named @p_819)) :named @p_820)) :named @p_815) @p_800) :named @p_821)) :named @p_801))) :rule forall_inst :args ((:= veriT_vr88 @p_381) (:= veriT_vr89 @p_9))) +(anchor :step t354) +(assume t354.h1 @p_801) +(anchor :step t354.t2 :args ((:= veriT_vr90 veriT_vr99))) +(step t354.t2.t1 (cl (! (= veriT_vr90 veriT_vr99) :named @p_806)) :rule refl) +(step t354.t2.t2 (cl (= @p_804 (! (member$ veriT_vr99 @p_380) :named @p_805))) :rule cong :premises (t354.t2.t1)) +(step t354.t2.t3 (cl @p_806) :rule refl) +(step t354.t2.t4 (cl (= @p_807 (! (member$ veriT_vr99 @p_381) :named @p_808))) :rule cong :premises (t354.t2.t3)) +(step t354.t2.t5 (cl (= @p_809 (! (not @p_808) :named @p_810))) :rule cong :premises (t354.t2.t4)) +(step t354.t2.t6 (cl (= @p_811 (! (=> @p_805 @p_810) :named @p_812))) :rule cong :premises (t354.t2.t2 t354.t2.t5)) +(step t354.t2 (cl (= @p_803 (! (forall ((veriT_vr99 A$)) @p_812) :named @p_814))) :rule bind) +(step t354.t3 (cl (= @p_813 (! (=> @p_800 @p_814) :named @p_822))) :rule cong :premises (t354.t2)) +(anchor :step t354.t4 :args ((:= veriT_vr91 veriT_vr99))) +(step t354.t4.t1 (cl (! (= veriT_vr91 veriT_vr99) :named @p_817)) :rule refl) +(step t354.t4.t2 (cl (= @p_816 @p_805)) :rule cong :premises (t354.t4.t1)) +(step t354.t4.t3 (cl @p_817) :rule refl) +(step t354.t4.t4 (cl (= @p_818 @p_808)) :rule cong :premises (t354.t4.t3)) +(step t354.t4.t5 (cl (= @p_819 @p_810)) :rule cong :premises (t354.t4.t4)) +(step t354.t4.t6 (cl (= @p_820 @p_812)) :rule cong :premises (t354.t4.t2 t354.t4.t5)) +(step t354.t4 (cl (= @p_815 @p_814)) :rule bind) +(step t354.t5 (cl (= @p_821 (! (=> @p_814 @p_800) :named @p_823))) :rule cong :premises (t354.t4)) +(step t354.t6 (cl (! (= @p_801 (! (and @p_822 @p_823) :named @p_826)) :named @p_824)) :rule cong :premises (t354.t3 t354.t5)) +(step t354.t7 (cl (not @p_824) (! (not @p_801) :named @p_825) @p_826) :rule equiv_pos2) +(step t354.t8 (cl @p_826) :rule th_resolution :premises (t354.h1 t354.t6 t354.t7)) +(anchor :step t354.t9 :args ((:= veriT_vr99 veriT_vr100))) +(step t354.t9.t1 (cl (! (= veriT_vr99 veriT_vr100) :named @p_829)) :rule refl) +(step t354.t9.t2 (cl (= @p_805 @p_828)) :rule cong :premises (t354.t9.t1)) +(step t354.t9.t3 (cl @p_829) :rule refl) +(step t354.t9.t4 (cl (= @p_808 @p_830)) :rule cong :premises (t354.t9.t3)) +(step t354.t9.t5 (cl (= @p_810 @p_831)) :rule cong :premises (t354.t9.t4)) +(step t354.t9.t6 (cl (= @p_812 @p_827)) :rule cong :premises (t354.t9.t2 t354.t9.t5)) +(step t354.t9 (cl (= @p_814 (! (forall ((veriT_vr100 A$)) @p_827) :named @p_832))) :rule bind) +(step t354.t10 (cl (= @p_822 (! (=> @p_800 @p_832) :named @p_833))) :rule cong :premises (t354.t9)) +(step t354.t11 (cl (= @p_823 (! (=> @p_832 @p_800) :named @p_834))) :rule cong :premises (t354.t9)) +(step t354.t12 (cl (! (= @p_826 (! (and @p_833 @p_834) :named @p_836)) :named @p_835)) :rule cong :premises (t354.t10 t354.t11)) +(step t354.t13 (cl (not @p_835) (not @p_826) @p_836) :rule equiv_pos2) +(step t354.t14 (cl @p_836) :rule th_resolution :premises (t354.t8 t354.t12 t354.t13)) +(anchor :step t354.t15 :args ((:= veriT_vr100 veriT_sk5))) +(step t354.t15.t1 (cl (! (= veriT_vr100 veriT_sk5) :named @p_839)) :rule refl) +(step t354.t15.t2 (cl (= @p_828 (! (member$ veriT_sk5 @p_380) :named @p_838))) :rule cong :premises (t354.t15.t1)) +(step t354.t15.t3 (cl @p_839) :rule refl) +(step t354.t15.t4 (cl (= @p_830 (! (member$ veriT_sk5 @p_381) :named @p_840))) :rule cong :premises (t354.t15.t3)) +(step t354.t15.t5 (cl (= @p_831 (! (not @p_840) :named @p_841))) :rule cong :premises (t354.t15.t4)) +(step t354.t15.t6 (cl (= @p_827 (! (=> @p_838 @p_841) :named @p_837))) :rule cong :premises (t354.t15.t2 t354.t15.t5)) +(step t354.t15 (cl (= @p_832 @p_837)) :rule sko_forall) +(step t354.t16 (cl (= @p_834 (! (=> @p_837 @p_800) :named @p_842))) :rule cong :premises (t354.t15)) +(step t354.t17 (cl (! (= @p_836 (! (and @p_833 @p_842) :named @p_844)) :named @p_843)) :rule cong :premises (t354.t16)) +(step t354.t18 (cl (not @p_843) (not @p_836) @p_844) :rule equiv_pos2) +(step t354.t19 (cl @p_844) :rule th_resolution :premises (t354.t14 t354.t17 t354.t18)) +(anchor :step t354.t20 :args ((:= veriT_vr100 veriT_vr101))) +(step t354.t20.t1 (cl (! (= veriT_vr100 veriT_vr101) :named @p_847)) :rule refl) +(step t354.t20.t2 (cl (= @p_828 (! (member$ veriT_vr101 @p_380) :named @p_846))) :rule cong :premises (t354.t20.t1)) +(step t354.t20.t3 (cl @p_847) :rule refl) +(step t354.t20.t4 (cl (= @p_830 (! (member$ veriT_vr101 @p_381) :named @p_848))) :rule cong :premises (t354.t20.t3)) +(step t354.t20.t5 (cl (= @p_831 (! (not @p_848) :named @p_849))) :rule cong :premises (t354.t20.t4)) +(step t354.t20.t6 (cl (= @p_827 (! (=> @p_846 @p_849) :named @p_850))) :rule cong :premises (t354.t20.t2 t354.t20.t5)) +(step t354.t20 (cl (= @p_832 (! (forall ((veriT_vr101 A$)) @p_850) :named @p_845))) :rule bind) +(step t354.t21 (cl (= @p_833 (! (=> @p_800 @p_845) :named @p_851))) :rule cong :premises (t354.t20)) +(step t354.t22 (cl (! (= @p_844 (! (and @p_851 @p_842) :named @p_852)) :named @p_853)) :rule cong :premises (t354.t21)) +(step t354.t23 (cl (not @p_853) (not @p_844) @p_852) :rule equiv_pos2) +(step t354.t24 (cl @p_852) :rule th_resolution :premises (t354.t19 t354.t22 t354.t23)) +(step t354 (cl @p_825 @p_852) :rule subproof :discharge (h1)) +(step t355 (cl @p_798 @p_801) :rule or :premises (t353)) +(step t356 (cl (! (or @p_798 @p_852) :named @p_855) @p_854) :rule or_neg) +(step t357 (cl @p_855 @p_703) :rule th_resolution :premises (t352 t356)) +(step t358 (cl @p_855 (! (not @p_852) :named @p_1004)) :rule or_neg) +(step t359 (cl @p_855) :rule th_resolution :premises (t355 t354 t357 t358)) +(step t360 (cl (or @p_798 (! (and (! (=> @p_762 (! (forall ((veriT_vr90 A$)) (! (=> (! (member$ veriT_vr90 @p_715) :named @p_859) @p_809) :named @p_864)) :named @p_858)) :named @p_866) (! (=> (! (forall ((veriT_vr91 A$)) (! (=> (! (member$ veriT_vr91 @p_715) :named @p_869) @p_819) :named @p_871)) :named @p_868) @p_762) :named @p_872)) :named @p_856))) :rule forall_inst :args ((:= veriT_vr88 @p_381) (:= veriT_vr89 xs$))) +(anchor :step t361) +(assume t361.h1 @p_856) +(anchor :step t361.t2 :args ((:= veriT_vr90 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_715) :named @p_860))) :rule cong :premises (t361.t2.t1)) +(step t361.t2.t3 (cl @p_861) :rule refl) +(step t361.t2.t4 (cl (= @p_807 (! (member$ veriT_vr102 @p_381) :named @p_862))) :rule cong :premises (t361.t2.t3)) +(step t361.t2.t5 (cl (= @p_809 (! (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_762 @p_867) :named @p_873))) :rule cong :premises (t361.t2)) +(anchor :step t361.t4 :args ((:= veriT_vr91 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_818 @p_862)) :rule cong :premises (t361.t4.t3)) +(step t361.t4.t5 (cl (= @p_819 @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_762) :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 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_762 @p_883) :named @p_884))) :rule cong :premises (t361.t9)) +(step t361.t11 (cl (= @p_874 (! (=> @p_883 @p_762) :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 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_715) :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_762) :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 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_715) :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_762 @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_798 @p_856) :rule or :premises (t360)) +(step t363 (cl (! (or @p_798 @p_903) :named @p_905) @p_854) :rule or_neg) +(step t364 (cl @p_905 @p_703) :rule th_resolution :premises (t352 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 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 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 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 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 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 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 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 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 @p_837 @p_838) :rule implies_neg1) +(step t387 (cl @p_837 (! (not @p_841) :named @p_1003)) :rule implies_neg2) +(step t388 (cl (not @p_1003) @p_840) :rule not_not) +(step t389 (cl @p_837 @p_840) :rule th_resolution :premises (t388 t387)) +(step t390 (cl (not @p_842) (not @p_837) @p_800) :rule implies_pos) +(step t391 (cl @p_1004 @p_842) :rule and_pos) +(step t392 (cl @p_798 @p_852) :rule or :premises (t359)) +(step t393 (cl @p_852) :rule resolution :premises (t392 t229)) +(step t394 (cl @p_842) :rule resolution :premises (t391 t393)) +(step t395 (cl (not @p_902) @p_1005 @p_896) :rule implies_pos) +(step t396 (cl @p_1006 @p_902) :rule and_pos) +(step t397 (cl @p_798 @p_903) :rule or :premises (t366)) +(step t398 (cl @p_903) :rule resolution :premises (t397 t229)) +(step t399 (cl @p_902) :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_896) :named @p_1022) (! (forall ((veriT_vr104 A$)) (or (not @p_897) @p_900)) :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_vr104 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_896) :rule not_not) +(step t434 (cl @p_1027 @p_896) :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_838) :named @p_1037) @p_841) :named @p_1033))) :rule forall_inst :args ((:= veriT_vr59 veriT_sk5))) +(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_841) :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_837) :rule resolution :premises (t448 t389 t386)) +(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 a10 t319 t344 t405 t293 a11 t266)) +(step t498 (cl (not (! (= @p_43 @p_799) :named @p_1091)) (! (not (! (= @p_21 @p_939) :named @p_1094)) :named @p_1093) @p_1019 @p_1005) :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_1005 @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_1005 @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_1005 @p_1092 @p_1040 @p_1043 @p_1067 @p_1068) :rule contraction :premises (t503)) +(step t505 (cl @p_1019 @p_1005 @p_1040) :rule resolution :premises (t504 a24 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 a18)) +(step t519 (cl rhs$) :rule resolution :premises (t471 t294 t243 t240 t246 t462 t497 t419 t416 t415 t410 t452 t422 t425 t505 t440 t282 t395 t284 t320 t328 t321 t509 t510 t301 t518 t248 t247 t296 t421 t418 t414 t427 t285 t399 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_837) :rule resolution :premises (t454 t526)) +(step t529 (cl @p_800) :rule resolution :premises (t390 t528 t394)) +(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 a18 t521)) +(step t534 (cl (! (not (! (= @p_799 @p_799) :named @p_1110)) :named @p_1108) (not @p_796) @p_762 (! (not @p_800) :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_896) :rule resolution :premises (t395 t538 t399)) +(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 a11 t523 t266 t293 t327 t344 a10 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 a11 t523 t266 t270 a10 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 a24 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 a24 t588 t586)) +042a030a3ede99a1c051b4a3202d4371c64c6c35 654 0 +unsat +(assume a0 (! (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 a1 (! (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 a2 (! (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 a3 (! (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 a4 (! (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 a5 (! (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 a6 (! (member$ (! (pair$ k$ g$) :named @p_403) one_chain_typeI$) :named @p_402)) +(assume a7 (! (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 a8 (! (finite$ bot$) :named @p_414)) +(assume a9 (! (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 a10 (! (= i$ (! (fun_app$b (pair$a 1.0) 0.0) :named @p_417)) :named @p_499)) +(assume a11 (! (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 a12 (! (= j$ (! (fun_app$b (pair$a 0.0) 1.0) :named @p_419)) :named @p_500)) +(assume a13 (! (forall ((?v0 Real_real_prod_set$)) (! (= bot$ (! (inf$ ?v0 bot$) :named @p_196)) :named @p_198)) :named @p_195)) +(assume a14 (! (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 a15 (! (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 a16 (! (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 a17 (! (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 a18 (! (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 a19 (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 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 (a0 t21 t22)) +(anchor :step t24 :args ((:= veriT_vr0 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 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 (a1 t27 t28)) +(anchor :step t30 :args ((:= veriT_vr2 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 veriT_vr4) (:= ?v1 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 (a2 t33 t34)) +(anchor :step t36 :args ((:= veriT_vr4 veriT_vr6) (:= veriT_vr5 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 veriT_vr8) (:= ?v1 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 (a3 t39 t40)) +(anchor :step t42 :args ((:= veriT_vr8 veriT_vr10) (:= veriT_vr9 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 veriT_vr12) (:= ?v1 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 (a4 t45 t46)) +(anchor :step t48 :args ((:= veriT_vr12 veriT_vr14) (:= veriT_vr13 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 veriT_vr16) (:= ?v1 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 (a5 t51 t52)) +(anchor :step t54 :args ((:= veriT_vr16 veriT_vr18) (:= veriT_vr17 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 veriT_vr20) (:= ?v1 veriT_vr21) (:= ?v2 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 (a7 t57 t58)) +(anchor :step t60 :args ((:= veriT_vr20 veriT_vr23) (:= veriT_vr21 veriT_vr24) (:= veriT_vr22 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 veriT_vr26) (:= ?v1 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 (a9 t63 t64)) +(anchor :step t66 :args ((:= veriT_vr26 veriT_vr28) (:= veriT_vr27 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 veriT_vr30) (:= ?v1 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 (a11 t69 t70)) +(anchor :step t72 :args ((:= veriT_vr30 veriT_vr32) (:= veriT_vr31 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 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 (a13 t75 t76)) +(anchor :step t78 :args ((:= veriT_vr34 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 veriT_vr36) (:= ?v1 veriT_vr37) (:= ?v2 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 (a14 t81 t82)) +(anchor :step t84 :args ((:= veriT_vr36 veriT_vr39) (:= veriT_vr37 veriT_vr40) (:= veriT_vr38 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 veriT_vr42) (:= ?v1 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 (a15 t87 t88)) +(anchor :step t90 :args ((:= veriT_vr42 veriT_vr44) (:= veriT_vr43 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 veriT_vr46) (:= ?v1 veriT_vr47) (:= ?v2 veriT_vr48) (:= ?v3 veriT_vr49) (:= ?v4 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 (a16 t93 t94)) +(anchor :step t96 :args (veriT_vr46 veriT_vr47 veriT_vr48 veriT_vr49 veriT_vr50)) +(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 veriT_vr51) (:= veriT_vr47 veriT_vr52) (:= veriT_vr48 veriT_vr53) (:= veriT_vr49 veriT_vr54) (:= veriT_vr50 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 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 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 (a17 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 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 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 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 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 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 (a18 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 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 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 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 a6)) +(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 a6)) +(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 a19)) +(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 a8)) +(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 a8)) +(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 a19)) +(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 a10 a12 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)) +1b1ef677d2a459ea9bcbad2db18e951f812782ca 6 0 +unsat +(assume a0 (! (< 0.0 (+ x$ (! (* 2.0 y$) :named @p_1))) :named @p_2)) +(assume a1 (! (< 0.0 (- x$ @p_1)) :named @p_3)) +(assume a2 (! (< 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 a0 a1 a2)) +8983b40ecc1de094c5b355c5a66dc3ff3a0c8509 735 0 +unsat +(assume a0 (! (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 a3 (! (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 a4 (! (< 0.0 d$) :named @p_269)) +(assume a5 (! (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 a7 (! (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 a8 (! (forall ((?v0 Real) (?v1 Real)) (! (= (! (divide$ @p_4 ?v1) :named @p_142) (! (- @p_5) :named @p_147)) :named @p_149)) :named @p_140)) +(assume a9 (! (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 a10 (! (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 a11 (! (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 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 (a0 t10 t11)) +(anchor :step t13 :args ((:= veriT_vr0 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 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 (a3 t16 t17)) +(anchor :step t19 :args ((:= veriT_vr6 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 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 (a5 t22 t23)) +(anchor :step t25 :args ((:= veriT_vr8 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 veriT_vr10) (:= ?v1 veriT_vr11) (:= ?v2 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 (a7 t28 t29)) +(anchor :step t31 :args (veriT_vr10 veriT_vr11 veriT_vr12)) +(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 veriT_vr13) (:= veriT_vr11 veriT_vr14) (:= veriT_vr12 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 veriT_vr16) (:= ?v1 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 (a8 t37 t38)) +(anchor :step t40 :args ((:= veriT_vr16 veriT_vr18) (:= veriT_vr17 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 veriT_vr20) (:= ?v1 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 (a9 t43 t44)) +(anchor :step t46 :args ((:= veriT_vr20 veriT_vr22) (:= veriT_vr21 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 veriT_vr24) (:= ?v1 veriT_vr25) (:= ?v2 veriT_vr26) (:= ?v3 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 (a10 t49 t50)) +(anchor :step t52 :args ((:= veriT_vr24 veriT_vr28) (:= veriT_vr25 veriT_vr29) (:= veriT_vr26 veriT_vr30) (:= veriT_vr27 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 (a11 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_281)) :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) :named @p_234) @p_211)) :named @p_499))) :rule qnt_cnf) +(step t71 (cl (or @p_233 (! (forall ((veriT_vr28 Real) (veriT_vr29 Real) (veriT_vr30 Real) (veriT_vr31 Real)) (or @p_234 @p_214)) :named @p_488))) :rule qnt_cnf) +(step t72 (cl (not (! (not @p_233) :named @p_493)) @p_218) :rule not_not) +(step t73 (cl (or (! (not @p_184) :named @p_236) (! (= (! (times$ (! (- @p_36) :named @p_237) (! (diamond_y$ @p_3) :named @p_235)) :named @p_446) (! (- (! (times$ @p_36 @p_235) :named @p_445)) :named @p_426)) :named @p_389))) :rule forall_inst :args ((:= veriT_vr22 @p_36) (:= veriT_vr23 @p_235))) +(step t74 (cl (or @p_236 (! (= (! (times$ @p_237 d$) :named @p_477) (! (- (times$ @p_36 d$)) :named @p_471)) :named @p_390))) :rule forall_inst :args ((:= veriT_vr22 @p_36) (:= veriT_vr23 d$))) +(step t75 (cl (not (! (not @p_236) :named @p_246)) @p_184) :rule not_not) +(step t76 (cl (or @p_236 (! (= (! (times$ (! (- 0.0) :named @p_241) d$) :named @p_242) (! (- (! (times$ 0.0 d$) :named @p_239)) :named @p_240)) :named @p_238))) :rule forall_inst :args ((:= veriT_vr22 0.0) (:= veriT_vr23 d$))) +(anchor :step t77) +(assume t77.h1 @p_238) +(step t77.t2 (cl (! (= 0.0 @p_241) :named @p_261)) :rule minus_simplify) +(step t77.t3 (cl (= @p_242 @p_239)) :rule cong :premises (t77.t2)) +(step t77.t4 (cl (! (= @p_238 (! (= @p_239 @p_240) :named @p_243)) :named @p_244)) :rule cong :premises (t77.t3)) +(step t77.t5 (cl (not @p_244) (! (not @p_238) :named @p_245) @p_243) :rule equiv_pos2) +(step t77.t6 (cl @p_243) :rule th_resolution :premises (t77.h1 t77.t4 t77.t5)) +(step t77 (cl @p_245 @p_243) :rule subproof :discharge (h1)) +(step t78 (cl @p_236 @p_238) :rule or :premises (t76)) +(step t79 (cl (! (or @p_236 @p_243) :named @p_247) @p_246) :rule or_neg) +(step t80 (cl @p_247 @p_184) :rule th_resolution :premises (t75 t79)) +(step t81 (cl @p_247 (! (not @p_243) :named @p_476)) :rule or_neg) +(step t82 (cl @p_247) :rule th_resolution :premises (t78 t77 t80 t81)) +(step t83 (cl (or @p_236 (! (= (! (times$ (! (- 1.0) :named @p_250) @p_36) :named @p_251) (! (- (! (times$ 1.0 @p_36) :named @p_348)) :named @p_249)) :named @p_248))) :rule forall_inst :args ((:= veriT_vr22 1.0) (:= veriT_vr23 @p_36))) +(anchor :step t84) +(assume t84.h1 @p_248) +(step t84.t2 (cl (= @p_250 (- 1.0))) :rule minus_simplify) +(step t84.t3 (cl (= @p_251 (! (times$ (- 1.0) @p_36) :named @p_252))) :rule cong :premises (t84.t2)) +(step t84.t4 (cl (! (= @p_248 (! (= @p_249 @p_252) :named @p_253)) :named @p_254)) :rule cong :premises (t84.t3)) +(step t84.t5 (cl (not @p_254) (! (not @p_248) :named @p_255) @p_253) :rule equiv_pos2) +(step t84.t6 (cl @p_253) :rule th_resolution :premises (t84.h1 t84.t4 t84.t5)) +(step t84 (cl @p_255 @p_253) :rule subproof :discharge (h1)) +(step t85 (cl @p_236 @p_248) :rule or :premises (t83)) +(step t86 (cl (! (or @p_236 @p_253) :named @p_256) @p_246) :rule or_neg) +(step t87 (cl @p_256 @p_184) :rule th_resolution :premises (t75 t86)) +(step t88 (cl @p_256 (! (not @p_253) :named @p_554)) :rule or_neg) +(step t89 (cl @p_256) :rule th_resolution :premises (t85 t84 t87 t88)) +(step t90 (cl (not (! (not (! (not @p_161) :named @p_257)) :named @p_266)) @p_161) :rule not_not) +(step t91 (cl (or @p_257 (! (= (! (divide$ @p_241 2.0) :named @p_262) (! (- (! (divide$ 0.0 2.0) :named @p_259)) :named @p_260)) :named @p_258))) :rule forall_inst :args ((:= veriT_vr18 0.0) (:= veriT_vr19 2.0))) +(anchor :step t92) +(assume t92.h1 @p_258) +(step t92.t2 (cl @p_261) :rule minus_simplify) +(step t92.t3 (cl (= @p_262 @p_259)) :rule cong :premises (t92.t2)) +(step t92.t4 (cl (! (= @p_258 (! (= @p_259 @p_260) :named @p_263)) :named @p_264)) :rule cong :premises (t92.t3)) +(step t92.t5 (cl (not @p_264) (! (not @p_258) :named @p_265) @p_263) :rule equiv_pos2) +(step t92.t6 (cl @p_263) :rule th_resolution :premises (t92.h1 t92.t4 t92.t5)) +(step t92 (cl @p_265 @p_263) :rule subproof :discharge (h1)) +(step t93 (cl @p_257 @p_258) :rule or :premises (t91)) +(step t94 (cl (! (or @p_257 @p_263) :named @p_267) @p_266) :rule or_neg) +(step t95 (cl @p_267 @p_161) :rule th_resolution :premises (t90 t94)) +(step t96 (cl @p_267 (! (not @p_263) :named @p_407)) :rule or_neg) +(step t97 (cl @p_267) :rule th_resolution :premises (t93 t92 t95 t96)) +(step t98 (cl (not (! (not (! (not @p_139) :named @p_268)) :named @p_288)) @p_139) :rule not_not) +(step t99 (cl (or @p_268 (! (= (! (< @p_259 @p_3) :named @p_271) (! (and (! (=> (! (< 0.0 2.0) :named @p_272) @p_269) :named @p_273) (! (=> (! (< 2.0 0.0) :named @p_275) (! (< d$ 0.0) :named @p_277)) :named @p_276) (! (not (! (= 2.0 0.0) :named @p_279)) :named @p_280)) :named @p_282)) :named @p_270))) :rule forall_inst :args ((:= veriT_vr13 0.0) (:= veriT_vr14 2.0) (:= veriT_vr15 d$))) +(anchor :step t100) +(assume t100.h1 @p_270) +(step t100.t2 (cl (= @p_272 true)) :rule comp_simplify) +(step t100.t3 (cl (= @p_273 (! (=> true @p_269) :named @p_274))) :rule cong :premises (t100.t2)) +(step t100.t4 (cl (= @p_274 @p_269)) :rule implies_simplify) +(step t100.t5 (cl (= @p_273 @p_269)) :rule trans :premises (t100.t3 t100.t4)) +(step t100.t6 (cl (= @p_275 false)) :rule comp_simplify) +(step t100.t7 (cl (= @p_276 (! (=> false @p_277) :named @p_278))) :rule cong :premises (t100.t6)) +(step t100.t8 (cl (= @p_278 true)) :rule implies_simplify) +(step t100.t9 (cl (= @p_276 true)) :rule trans :premises (t100.t7 t100.t8)) +(step t100.t10 (cl (= @p_279 false)) :rule eq_simplify) +(step t100.t11 (cl (= @p_280 @p_228)) :rule cong :premises (t100.t10)) +(step t100.t12 (cl @p_281) :rule not_simplify) +(step t100.t13 (cl (= @p_280 true)) :rule trans :premises (t100.t11 t100.t12)) +(step t100.t14 (cl (= @p_282 (! (and @p_269 true true) :named @p_283))) :rule cong :premises (t100.t5 t100.t9 t100.t13)) +(step t100.t15 (cl (= @p_283 (! (and @p_269) :named @p_284))) :rule and_simplify) +(step t100.t16 (cl (= @p_284 @p_269)) :rule and_simplify) +(step t100.t17 (cl (= @p_282 @p_269)) :rule trans :premises (t100.t14 t100.t15 t100.t16)) +(step t100.t18 (cl (! (= @p_270 (! (= @p_271 @p_269) :named @p_285)) :named @p_286)) :rule cong :premises (t100.t17)) +(step t100.t19 (cl (not @p_286) (! (not @p_270) :named @p_287) @p_285) :rule equiv_pos2) +(step t100.t20 (cl @p_285) :rule th_resolution :premises (t100.h1 t100.t18 t100.t19)) +(step t100 (cl @p_287 @p_285) :rule subproof :discharge (h1)) +(step t101 (cl @p_268 @p_270) :rule or :premises (t99)) +(step t102 (cl (! (or @p_268 @p_285) :named @p_289) @p_288) :rule or_neg) +(step t103 (cl @p_289 @p_139) :rule th_resolution :premises (t98 t102)) +(step t104 (cl @p_289 (! (not @p_285) :named @p_391)) :rule or_neg) +(step t105 (cl @p_289) :rule th_resolution :premises (t101 t100 t103 t104)) +(step t106 (cl (not (! (not (! (not @p_84) :named @p_290)) :named @p_303)) @p_84) :rule not_not) +(step t107 (cl (or @p_290 (! (= @p_36 (! (- @p_3 (! (ite (! (< @p_2 0.0) :named @p_292) (! (- @p_2) :named @p_294) @p_2) :named @p_295)) :named @p_296)) :named @p_291))) :rule forall_inst :args ((:= veriT_vr9 @p_2))) +(anchor :step t108) +(assume t108.h1 @p_291) +(step t108.t2 (cl (= @p_3 @p_294)) :rule minus_simplify) +(step t108.t3 (cl (= @p_295 (! (ite @p_292 @p_3 @p_2) :named @p_293))) :rule cong :premises (t108.t2)) +(step t108.t4 (cl (= @p_296 (! (- @p_3 @p_293) :named @p_297))) :rule cong :premises (t108.t3)) +(step t108.t5 (cl (! (= @p_291 (! (= @p_36 @p_297) :named @p_300)) :named @p_298)) :rule cong :premises (t108.t4)) +(step t108.t6 (cl (not @p_298) (! (not @p_291) :named @p_299) @p_300) :rule equiv_pos2) +(step t108.t7 (cl @p_300) :rule th_resolution :premises (t108.h1 t108.t5 t108.t6)) +(step t108.t8 (cl (! (= @p_300 (! (and (! (= @p_36 (! (- @p_3 @p_293) :named @p_417)) :named @p_393) (! (ite @p_292 (! (= @p_3 @p_293) :named @p_395) (! (= @p_2 @p_293) :named @p_413)) :named @p_394)) :named @p_301)) :named @p_302)) :rule ite_intro) +(step t108.t9 (cl (not @p_302) (not @p_300) @p_301) :rule equiv_pos2) +(step t108.t10 (cl @p_301) :rule th_resolution :premises (t108.t7 t108.t8 t108.t9)) +(step t108 (cl @p_299 @p_301) :rule subproof :discharge (h1)) +(step t109 (cl @p_290 @p_291) :rule or :premises (t107)) +(step t110 (cl (! (or @p_290 @p_301) :named @p_304) @p_303) :rule or_neg) +(step t111 (cl @p_304 @p_84) :rule th_resolution :premises (t106 t110)) +(step t112 (cl @p_304 (! (not @p_301) :named @p_392)) :rule or_neg) +(step t113 (cl @p_304) :rule th_resolution :premises (t109 t108 t111 t112)) +(step t114 (cl (or @p_290 (! (= (! (diamond_y$ @p_235) :named @p_306) (- @p_3 (! (ite (! (< @p_235 0.0) :named @p_308) (! (- @p_235) :named @p_309) @p_235) :named @p_307))) :named @p_305))) :rule forall_inst :args ((:= veriT_vr9 @p_235))) +(anchor :step t115) +(assume t115.h1 @p_305) +(step t115.t2 (cl (! (= @p_305 (! (and (= @p_306 (- @p_3 @p_307)) (! (ite @p_308 (! (= @p_309 @p_307) :named @p_450) (! (= @p_235 @p_307) :named @p_397)) :named @p_396)) :named @p_310)) :named @p_311)) :rule ite_intro) +(step t115.t3 (cl (not @p_311) (! (not @p_305) :named @p_312) @p_310) :rule equiv_pos2) +(step t115.t4 (cl @p_310) :rule th_resolution :premises (t115.h1 t115.t2 t115.t3)) +(step t115 (cl @p_312 @p_310) :rule subproof :discharge (h1)) +(step t116 (cl @p_290 @p_305) :rule or :premises (t114)) +(step t117 (cl (! (or @p_290 @p_310) :named @p_313) @p_303) :rule or_neg) +(step t118 (cl @p_313 @p_84) :rule th_resolution :premises (t106 t117)) +(step t119 (cl @p_313 (! (not @p_310) :named @p_398)) :rule or_neg) +(step t120 (cl @p_313) :rule th_resolution :premises (t116 t115 t118 t119)) +(step t121 (cl (or @p_290 (! (= @p_235 (- @p_3 (! (ite (! (< @p_3 0.0) :named @p_316) @p_2 @p_3) :named @p_315))) :named @p_314))) :rule forall_inst :args ((:= veriT_vr9 @p_3))) +(anchor :step t122) +(assume t122.h1 @p_314) +(step t122.t2 (cl (! (= @p_314 (! (and (! (= @p_235 (! (- @p_3 @p_315) :named @p_418)) :named @p_400) (! (ite @p_316 (= @p_2 @p_315) (! (= @p_3 @p_315) :named @p_402)) :named @p_401)) :named @p_317)) :named @p_318)) :rule ite_intro) +(step t122.t3 (cl (not @p_318) (! (not @p_314) :named @p_319) @p_317) :rule equiv_pos2) +(step t122.t4 (cl @p_317) :rule th_resolution :premises (t122.h1 t122.t2 t122.t3)) +(step t122 (cl @p_319 @p_317) :rule subproof :discharge (h1)) +(step t123 (cl @p_290 @p_314) :rule or :premises (t121)) +(step t124 (cl (! (or @p_290 @p_317) :named @p_320) @p_303) :rule or_neg) +(step t125 (cl @p_320 @p_84) :rule th_resolution :premises (t106 t124)) +(step t126 (cl @p_320 (! (not @p_317) :named @p_399)) :rule or_neg) +(step t127 (cl @p_320) :rule th_resolution :premises (t123 t122 t125 t126)) +(step t128 (cl (or @p_290 (! (= (! (diamond_y$ 0.0) :named @p_322) (! (- @p_3 (! (ite (! (< 0.0 0.0) :named @p_323) @p_241 0.0) :named @p_324)) :named @p_326)) :named @p_321))) :rule forall_inst :args ((:= veriT_vr9 0.0))) +(anchor :step t129) +(assume t129.h1 @p_321) +(step t129.t2 (cl (= @p_323 false)) :rule comp_simplify) +(step t129.t3 (cl @p_261) :rule minus_simplify) +(step t129.t4 (cl (= @p_324 (! (ite false 0.0 0.0) :named @p_325))) :rule cong :premises (t129.t2 t129.t3)) +(step t129.t5 (cl (= 0.0 @p_325)) :rule ite_simplify) +(step t129.t6 (cl (= 0.0 @p_324)) :rule trans :premises (t129.t4 t129.t5)) +(step t129.t7 (cl (= @p_326 (! (- @p_3 0.0) :named @p_327))) :rule cong :premises (t129.t6)) +(step t129.t8 (cl (= @p_3 @p_327)) :rule minus_simplify) +(step t129.t9 (cl (= @p_3 @p_326)) :rule trans :premises (t129.t7 t129.t8)) +(step t129.t10 (cl (! (= @p_321 (! (= @p_3 @p_322) :named @p_328)) :named @p_329)) :rule cong :premises (t129.t9)) +(step t129.t11 (cl (not @p_329) (! (not @p_321) :named @p_330) @p_328) :rule equiv_pos2) +(step t129.t12 (cl @p_328) :rule th_resolution :premises (t129.h1 t129.t10 t129.t11)) +(step t129 (cl @p_330 @p_328) :rule subproof :discharge (h1)) +(step t130 (cl @p_290 @p_321) :rule or :premises (t128)) +(step t131 (cl (! (or @p_290 @p_328) :named @p_331) @p_303) :rule or_neg) +(step t132 (cl @p_331 @p_84) :rule th_resolution :premises (t106 t131)) +(step t133 (cl @p_331 (! (not @p_328) :named @p_574)) :rule or_neg) +(step t134 (cl @p_331) :rule th_resolution :premises (t130 t129 t132 t133)) +(step t135 (cl (or (! (not @p_60) :named @p_332) (! (= (! (fun_app$ uub$ @p_36) :named @p_525) (! (pair$ @p_2 (! (times$ (- (* 2.0 @p_36) 1.0) @p_36) :named @p_509)) :named @p_508)) :named @p_403))) :rule forall_inst :args ((:= veriT_vr7 @p_36))) +(step t136 (cl (or @p_332 (! (= (! (fun_app$ uub$ @p_7) :named @p_533) (! (pair$ @p_2 (! (times$ (- (* 2.0 @p_7) 1.0) @p_36) :named @p_505)) :named @p_504)) :named @p_404))) :rule forall_inst :args ((:= veriT_vr7 @p_7))) +(step t137 (cl (not (! (not @p_332) :named @p_344)) @p_60) :rule not_not) +(step t138 (cl (or @p_332 (! (= (! (fun_app$ uub$ 0.0) :named @p_334) (! (pair$ @p_2 (! (times$ (! (- (! (* 2.0 0.0) :named @p_335) 1.0) :named @p_336) @p_36) :named @p_338)) :named @p_339)) :named @p_333))) :rule forall_inst :args ((:= veriT_vr7 0.0))) +(anchor :step t139) +(assume t139.h1 @p_333) +(step t139.t2 (cl (= 0.0 @p_335)) :rule prod_simplify) +(step t139.t3 (cl (= @p_336 (! (- 0.0 1.0) :named @p_337))) :rule cong :premises (t139.t2)) +(step t139.t4 (cl (= (- 1.0) @p_337)) :rule minus_simplify) +(step t139.t5 (cl (= (- 1.0) @p_336)) :rule trans :premises (t139.t3 t139.t4)) +(step t139.t6 (cl (= @p_252 @p_338)) :rule cong :premises (t139.t5)) +(step t139.t7 (cl (= @p_339 (! (pair$ @p_2 @p_252) :named @p_340))) :rule cong :premises (t139.t6)) +(step t139.t8 (cl (! (= @p_333 (! (= @p_334 @p_340) :named @p_341)) :named @p_342)) :rule cong :premises (t139.t7)) +(step t139.t9 (cl (not @p_342) (! (not @p_333) :named @p_343) @p_341) :rule equiv_pos2) +(step t139.t10 (cl @p_341) :rule th_resolution :premises (t139.h1 t139.t8 t139.t9)) +(step t139 (cl @p_343 @p_341) :rule subproof :discharge (h1)) +(step t140 (cl @p_332 @p_333) :rule or :premises (t138)) +(step t141 (cl (! (or @p_332 @p_341) :named @p_345) @p_344) :rule or_neg) +(step t142 (cl @p_345 @p_60) :rule th_resolution :premises (t137 t141)) +(step t143 (cl @p_345 (! (not @p_341) :named @p_544)) :rule or_neg) +(step t144 (cl @p_345) :rule th_resolution :premises (t140 t139 t142 t143)) +(step t145 (cl (or @p_332 (! (= (! (fun_app$ uub$ 1.0) :named @p_347) (! (pair$ @p_2 (! (times$ (! (- (! (* 2.0 1.0) :named @p_349) 1.0) :named @p_350) @p_36) :named @p_352)) :named @p_353)) :named @p_346))) :rule forall_inst :args ((:= veriT_vr7 1.0))) +(anchor :step t146) +(assume t146.h1 @p_346) +(step t146.t2 (cl (= 2.0 @p_349)) :rule prod_simplify) +(step t146.t3 (cl (= @p_350 (! (- 2.0 1.0) :named @p_351))) :rule cong :premises (t146.t2)) +(step t146.t4 (cl (= 1.0 @p_351)) :rule minus_simplify) +(step t146.t5 (cl (= 1.0 @p_350)) :rule trans :premises (t146.t3 t146.t4)) +(step t146.t6 (cl (= @p_348 @p_352)) :rule cong :premises (t146.t5)) +(step t146.t7 (cl (= @p_353 (! (pair$ @p_2 @p_348) :named @p_354))) :rule cong :premises (t146.t6)) +(step t146.t8 (cl (! (= @p_346 (! (= @p_347 @p_354) :named @p_355)) :named @p_356)) :rule cong :premises (t146.t7)) +(step t146.t9 (cl (not @p_356) (! (not @p_346) :named @p_357) @p_355) :rule equiv_pos2) +(step t146.t10 (cl @p_355) :rule th_resolution :premises (t146.h1 t146.t8 t146.t9)) +(step t146 (cl @p_357 @p_355) :rule subproof :discharge (h1)) +(step t147 (cl @p_332 @p_346) :rule or :premises (t145)) +(step t148 (cl (! (or @p_332 @p_355) :named @p_358) @p_344) :rule or_neg) +(step t149 (cl @p_358 @p_60) :rule th_resolution :premises (t137 t148)) +(step t150 (cl @p_358 (! (not @p_355) :named @p_538)) :rule or_neg) +(step t151 (cl @p_358) :rule th_resolution :premises (t147 t146 t149 t150)) +(step t152 (cl (or (! (not @p_34) :named @p_360) (! (= (! (fun_app$ uuc$ @p_36) :named @p_526) (! (pair$ (! (times$ (! (- @p_36 @p_7) :named @p_424) d$) :named @p_359) (! (diamond_y$ @p_359) :named @p_496)) :named @p_495)) :named @p_405))) :rule forall_inst :args ((:= veriT_vr1 @p_36))) +(step t153 (cl (or @p_360 (! (= (! (fun_app$ uuc$ @p_7) :named @p_363) (! (pair$ (! (times$ (! (- @p_7 @p_7) :named @p_364) d$) :named @p_361) (! (diamond_y$ @p_361) :named @p_365)) :named @p_367)) :named @p_362))) :rule forall_inst :args ((:= veriT_vr1 @p_7))) +(anchor :step t154) +(assume t154.h1 @p_362) +(step t154.t2 (cl (= 0.0 @p_364)) :rule minus_simplify) +(step t154.t3 (cl (= @p_239 @p_361)) :rule cong :premises (t154.t2)) +(step t154.t4 (cl (= @p_365 (! (diamond_y$ @p_239) :named @p_366))) :rule cong :premises (t154.t3)) +(step t154.t5 (cl (= @p_367 (! (pair$ @p_239 @p_366) :named @p_368))) :rule cong :premises (t154.t3 t154.t4)) +(step t154.t6 (cl (! (= @p_362 (! (= @p_363 @p_368) :named @p_369)) :named @p_370)) :rule cong :premises (t154.t5)) +(step t154.t7 (cl (not @p_370) (! (not @p_362) :named @p_371) @p_369) :rule equiv_pos2) +(step t154.t8 (cl @p_369) :rule th_resolution :premises (t154.h1 t154.t6 t154.t7)) +(step t154 (cl @p_371 @p_369) :rule subproof :discharge (h1)) +(step t155 (cl @p_360 @p_362) :rule or :premises (t153)) +(step t156 (cl (! (or @p_360 @p_369) :named @p_373) (! (not @p_360) :named @p_372)) :rule or_neg) +(step t157 (cl (not @p_372) @p_34) :rule not_not) +(step t158 (cl @p_373 @p_34) :rule th_resolution :premises (t157 t156)) +(step t159 (cl @p_373 (! (not @p_369) :named @p_534)) :rule or_neg) +(step t160 (cl @p_373) :rule th_resolution :premises (t155 t154 t158 t159)) +(step t161 (cl (or @p_360 (! (= (! (fun_app$ uuc$ 0.0) :named @p_376) (! (pair$ (! (times$ (! (- 0.0 @p_7) :named @p_379) d$) :named @p_374) (! (diamond_y$ @p_374) :named @p_380)) :named @p_382)) :named @p_375))) :rule forall_inst :args ((:= veriT_vr1 0.0))) +(anchor :step t162) +(assume t162.h1 @p_375) +(step t162.t2 (cl (= (! (- @p_7) :named @p_378) @p_379)) :rule minus_simplify) +(step t162.t3 (cl (= (! (times$ @p_378 d$) :named @p_377) @p_374)) :rule cong :premises (t162.t2)) +(step t162.t4 (cl (= @p_380 (! (diamond_y$ @p_377) :named @p_381))) :rule cong :premises (t162.t3)) +(step t162.t5 (cl (= @p_382 (! (pair$ @p_377 @p_381) :named @p_383))) :rule cong :premises (t162.t3 t162.t4)) +(step t162.t6 (cl (! (= @p_375 (! (= @p_376 @p_383) :named @p_384)) :named @p_385)) :rule cong :premises (t162.t5)) +(step t162.t7 (cl (not @p_385) (! (not @p_375) :named @p_386) @p_384) :rule equiv_pos2) +(step t162.t8 (cl @p_384) :rule th_resolution :premises (t162.h1 t162.t6 t162.t7)) +(step t162 (cl @p_386 @p_384) :rule subproof :discharge (h1)) +(step t163 (cl @p_360 @p_375) :rule or :premises (t161)) +(step t164 (cl (! (or @p_360 @p_384) :named @p_387) @p_372) :rule or_neg) +(step t165 (cl @p_387 @p_34) :rule th_resolution :premises (t157 t164)) +(step t166 (cl @p_387 (! (not @p_384) :named @p_545)) :rule or_neg) +(step t167 (cl @p_387) :rule th_resolution :premises (t163 t162 t165 t166)) +(step t168 (cl (or @p_360 (! (= (! (fun_app$ uuc$ 1.0) :named @p_539) (! (pair$ (! (times$ (- 1.0 @p_7) d$) :named @p_388) (! (diamond_y$ @p_388) :named @p_490)) :named @p_489)) :named @p_406))) :rule forall_inst :args ((:= veriT_vr1 1.0))) +(step t169 (cl @p_236 @p_389) :rule or :premises (t73)) +(step t170 (cl @p_389) :rule resolution :premises (t169 t48)) +(step t171 (cl @p_236 @p_390) :rule or :premises (t74)) +(step t172 (cl @p_390) :rule resolution :premises (t171 t48)) +(step t173 (cl @p_236 @p_243) :rule or :premises (t82)) +(step t174 (cl @p_243) :rule resolution :premises (t173 t48)) +(step t175 (cl @p_236 @p_253) :rule or :premises (t89)) +(step t176 (cl @p_253) :rule resolution :premises (t175 t48)) +(step t177 (cl @p_257 @p_263) :rule or :premises (t97)) +(step t178 (cl @p_263) :rule resolution :premises (t177 t42)) +(step t179 (cl @p_391 @p_271 (not @p_269)) :rule equiv_pos1) +(step t180 (cl @p_268 @p_285) :rule or :premises (t105)) +(step t181 (cl @p_391 @p_271) :rule resolution :premises (t179 a4)) +(step t182 (cl @p_285) :rule resolution :premises (t180 t36)) +(step t183 (cl @p_271) :rule resolution :premises (t181 t182)) +(step t184 (cl @p_392 @p_393) :rule and_pos) +(step t185 (cl (not @p_394) (! (not @p_292) :named @p_410) @p_395) :rule ite_pos2) +(step t186 (cl @p_392 @p_394) :rule and_pos) +(step t187 (cl @p_290 @p_301) :rule or :premises (t113)) +(step t188 (cl @p_301) :rule resolution :premises (t187 t27)) +(step t189 (cl @p_393) :rule resolution :premises (t184 t188)) +(step t190 (cl @p_394) :rule resolution :premises (t186 t188)) +(step t191 (cl (not @p_396) @p_308 @p_397) :rule ite_pos1) +(step t192 (cl @p_398 @p_396) :rule and_pos) +(step t193 (cl @p_290 @p_310) :rule or :premises (t120)) +(step t194 (cl @p_310) :rule resolution :premises (t193 t27)) +(step t195 (cl @p_396) :rule resolution :premises (t192 t194)) +(step t196 (cl @p_399 @p_400) :rule and_pos) +(step t197 (cl (not @p_401) @p_316 @p_402) :rule ite_pos1) +(step t198 (cl @p_399 @p_401) :rule and_pos) +(step t199 (cl @p_290 @p_317) :rule or :premises (t127)) +(step t200 (cl @p_317) :rule resolution :premises (t199 t27)) +(step t201 (cl @p_400) :rule resolution :premises (t196 t200)) +(step t202 (cl @p_401) :rule resolution :premises (t198 t200)) +(step t203 (cl @p_290 @p_328) :rule or :premises (t134)) +(step t204 (cl @p_328) :rule resolution :premises (t203 t27)) +(step t205 (cl @p_332 @p_403) :rule or :premises (t135)) +(step t206 (cl @p_403) :rule resolution :premises (t205 t21)) +(step t207 (cl @p_332 @p_404) :rule or :premises (t136)) +(step t208 (cl @p_404) :rule resolution :premises (t207 t21)) +(step t209 (cl @p_332 @p_341) :rule or :premises (t144)) +(step t210 (cl @p_341) :rule resolution :premises (t209 t21)) +(step t211 (cl @p_332 @p_355) :rule or :premises (t151)) +(step t212 (cl @p_355) :rule resolution :premises (t211 t21)) +(step t213 (cl @p_360 @p_405) :rule or :premises (t152)) +(step t214 (cl @p_405) :rule resolution :premises (t213 t15)) +(step t215 (cl @p_360 @p_369) :rule or :premises (t160)) +(step t216 (cl @p_369) :rule resolution :premises (t215 t15)) +(step t217 (cl @p_360 @p_384) :rule or :premises (t167)) +(step t218 (cl @p_384) :rule resolution :premises (t217 t15)) +(step t219 (cl @p_360 @p_406) :rule or :premises (t168)) +(step t220 (cl @p_406) :rule resolution :premises (t219 t15)) +(step t221 (cl (! (not @p_271) :named @p_409) (! (not @p_316) :named @p_408) @p_407) :rule la_generic :args (2.0 2.0 1)) +(step t222 (cl @p_408) :rule resolution :premises (t221 t178 t183)) +(step t223 (cl @p_402) :rule resolution :premises (t197 t222 t202)) +(step t224 (cl @p_409 @p_292 @p_407) :rule la_generic :args (2.0 2.0 1)) +(step t225 (cl @p_292) :rule resolution :premises (t224 t178 t183)) +(step t226 (cl @p_395) :rule resolution :premises (t185 t225 t190)) +(step t227 (cl (! (not (! (= @p_3 @p_2) :named @p_414)) :named @p_412) (! (not (! (= 0.0 0.0) :named @p_411)) :named @p_461) @p_410 @p_316) :rule eq_congruent_pred) +(step t228 (cl @p_411) :rule eq_reflexive) +(step t229 (cl @p_412 @p_410 @p_316) :rule th_resolution :premises (t227 t228)) +(step t230 (cl (! (not @p_395) :named @p_415) (! (not @p_413) :named @p_416) @p_414) :rule eq_transitive) +(step t231 (cl @p_410 @p_316 @p_415 @p_416) :rule th_resolution :premises (t229 t230)) +(step t232 (cl @p_416) :rule resolution :premises (t231 t225 t226 t222)) +(step t233 (cl (! (not (! (= @p_3 @p_3) :named @p_423)) :named @p_420) (not (! (= @p_293 @p_315) :named @p_419)) (! (= @p_417 @p_418) :named @p_421)) :rule eq_congruent) +(step t234 (cl @p_415 (! (not @p_402) :named @p_422) @p_419) :rule eq_transitive) +(step t235 (cl @p_420 @p_421 @p_415 @p_422) :rule th_resolution :premises (t233 t234)) +(step t236 (cl @p_423) :rule eq_reflexive) +(step t237 (cl @p_421 @p_415 @p_422) :rule th_resolution :premises (t235 t236)) +(step t238 (cl (not (! (= @p_378 @p_424) :named @p_437)) (! (not (! (= d$ d$) :named @p_425)) :named @p_479) (! (= @p_377 @p_359) :named @p_547)) :rule eq_congruent) +(step t239 (cl @p_425) :rule eq_reflexive) +(step t240 (cl (or (! (= @p_36 @p_426) :named @p_427) (! (not (! (<= @p_36 @p_426) :named @p_464)) :named @p_428) (! (not (! (<= @p_426 @p_36) :named @p_444)) :named @p_429))) :rule la_disequality) +(step t241 (cl @p_427 @p_428 @p_429) :rule or :premises (t240)) +(step t242 (cl (or (! (= @p_378 (! (- @p_235 @p_7) :named @p_430)) :named @p_431) (! (not (! (<= @p_378 @p_430) :named @p_435)) :named @p_432) (! (not (! (<= @p_430 @p_378) :named @p_434)) :named @p_433))) :rule la_disequality) +(step t243 (cl @p_431 @p_432 @p_433) :rule or :premises (t242)) +(step t244 (cl @p_434 (! (not @p_400) :named @p_436) @p_422) :rule la_generic :args (1.0 (- 1) (- 1))) +(step t245 (cl @p_434) :rule resolution :premises (t244 t201 t223)) +(step t246 (cl @p_435 @p_436 @p_422) :rule la_generic :args (1.0 1 1)) +(step t247 (cl @p_435) :rule resolution :premises (t246 t201 t223)) +(step t248 (cl @p_431) :rule resolution :premises (t243 t247 t245)) +(step t249 (cl (! (not @p_431) :named @p_443) (not (! (= @p_424 @p_430) :named @p_438)) @p_437) :rule eq_transitive) +(step t250 (cl (! (not (! (= @p_235 @p_36) :named @p_439)) :named @p_455) (! (not (! (= @p_7 @p_7) :named @p_442)) :named @p_441) @p_438) :rule eq_congruent) +(step t251 (cl @p_436 (! (not @p_421) :named @p_449) (! (not @p_393) :named @p_440) @p_439) :rule eq_transitive) +(step t252 (cl @p_436 @p_440 @p_439 @p_415 @p_422) :rule th_resolution :premises (t251 t237)) +(step t253 (cl @p_441 @p_438 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t250 t252)) +(step t254 (cl @p_442) :rule eq_reflexive) +(step t255 (cl @p_438 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t253 t254)) +(step t256 (cl @p_443 @p_437 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t249 t255)) +(step t257 (cl @p_444 @p_440 @p_415 (! (not (! (= @p_445 @p_426) :named @p_447)) :named @p_465)) :rule la_generic :args (2.0 2 2 1)) +(step t258 (cl (not (! (= @p_446 @p_445) :named @p_448)) (! (not @p_389) :named @p_459) @p_447) :rule eq_transitive) +(step t259 (cl (not (! (= @p_36 @p_237) :named @p_451)) (! (not (! (= @p_235 @p_235) :named @p_458)) :named @p_457) @p_448) :rule eq_congruent) +(step t260 (cl @p_440 @p_449 @p_436 (! (not @p_397) :named @p_452) (! (not @p_450) :named @p_453) (! (not (! (= @p_237 @p_309) :named @p_456)) :named @p_454) @p_451) :rule eq_transitive) +(step t261 (cl @p_440 @p_436 @p_452 @p_453 @p_454 @p_451 @p_415 @p_422) :rule th_resolution :premises (t260 t237)) +(step t262 (cl @p_455 @p_456) :rule eq_congruent) +(step t263 (cl @p_456 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t262 t252)) +(step t264 (cl @p_440 @p_436 @p_452 @p_453 @p_451 @p_415 @p_422 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t261 t263)) +(step t265 (cl @p_440 @p_436 @p_452 @p_453 @p_451 @p_415 @p_422) :rule contraction :premises (t264)) +(step t266 (cl @p_457 @p_448 @p_440 @p_436 @p_452 @p_453 @p_415 @p_422) :rule th_resolution :premises (t259 t265)) +(step t267 (cl @p_458) :rule eq_reflexive) +(step t268 (cl @p_448 @p_440 @p_436 @p_452 @p_453 @p_415 @p_422) :rule th_resolution :premises (t266 t267)) +(step t269 (cl @p_459 @p_447 @p_440 @p_436 @p_452 @p_453 @p_415 @p_422) :rule th_resolution :premises (t258 t268)) +(step t270 (cl @p_444 @p_440 @p_415 @p_459 @p_436 @p_452 @p_453 @p_422) :rule th_resolution :premises (t257 t269)) +(step t271 (cl @p_444 @p_452 @p_453) :rule resolution :premises (t270 t170 t189 t226 t201 t223)) +(step t272 (cl (! (not (! (< @p_36 0.0) :named @p_462)) :named @p_460) @p_440 @p_415) :rule la_generic :args (1.0 1 1)) +(step t273 (cl @p_460) :rule resolution :premises (t272 t189 t226)) +(step t274 (cl @p_455 @p_461 (! (not @p_308) :named @p_463) @p_462) :rule eq_congruent_pred) +(step t275 (cl @p_455 @p_463 @p_462) :rule th_resolution :premises (t274 t228)) +(step t276 (cl @p_463 @p_462 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t275 t252)) +(step t277 (cl @p_463) :rule resolution :premises (t276 t273 t189 t226 t201 t223)) +(step t278 (cl @p_397) :rule resolution :premises (t191 t277 t195)) +(step t279 (cl @p_464 @p_440 @p_415 @p_465) :rule la_generic :args (2.0 (- 2) (- 2) (- 1))) +(step t280 (cl @p_464 @p_440 @p_415 @p_459 @p_436 @p_452 @p_453 @p_422) :rule th_resolution :premises (t279 t269)) +(step t281 (cl @p_464 @p_453) :rule resolution :premises (t280 t170 t189 t226 t278 t201 t223)) +(step t282 (cl (or (! (= 0.0 @p_426) :named @p_466) (! (not (! (<= 0.0 @p_426) :named @p_470)) :named @p_467) (! (not (! (<= @p_426 0.0) :named @p_469)) :named @p_468))) :rule la_disequality) +(step t283 (cl @p_466 @p_467 @p_468) :rule or :premises (t282)) +(step t284 (cl @p_429 @p_469 @p_440 @p_415) :rule la_generic :args (1.0 1.0 (- 1) (- 1))) +(step t285 (cl @p_429 @p_469) :rule resolution :premises (t284 t189 t226)) +(step t286 (cl @p_428 @p_470 @p_440 @p_415) :rule la_generic :args (1.0 1.0 1 1)) +(step t287 (cl @p_428 @p_470) :rule resolution :premises (t286 t189 t226)) +(step t288 (cl (or (! (= @p_426 @p_471) :named @p_472) (! (not (! (<= @p_426 @p_471) :named @p_486)) :named @p_473) (! (not (! (<= @p_471 @p_426) :named @p_475)) :named @p_474))) :rule la_disequality) +(step t289 (cl @p_472 @p_473 @p_474) :rule or :premises (t288)) +(step t290 (cl @p_428 @p_475 @p_440 @p_415 (! (not (! (= @p_471 @p_239) :named @p_478)) :named @p_487) @p_476) :rule la_generic :args (2.0 2.0 2 2 (- 2) (- 1))) +(step t291 (cl (! (not @p_390) :named @p_485) (! (not (! (= @p_477 @p_239) :named @p_480)) :named @p_567) @p_478) :rule eq_transitive) +(step t292 (cl (! (not (! (= 0.0 @p_237) :named @p_482)) :named @p_481) @p_479 @p_480) :rule eq_congruent) +(step t293 (cl @p_481 @p_480) :rule th_resolution :premises (t292 t239)) +(step t294 (cl (! (not @p_466) :named @p_483) (! (not @p_427) :named @p_484) @p_440 @p_449 @p_436 @p_452 @p_453 @p_454 @p_482) :rule eq_transitive) +(step t295 (cl @p_483 @p_484 @p_440 @p_436 @p_452 @p_453 @p_454 @p_482 @p_415 @p_422) :rule th_resolution :premises (t294 t237)) +(step t296 (cl @p_483 @p_484 @p_440 @p_436 @p_452 @p_453 @p_482 @p_415 @p_422 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t295 t263)) +(step t297 (cl @p_483 @p_484 @p_440 @p_436 @p_452 @p_453 @p_482 @p_415 @p_422) :rule contraction :premises (t296)) +(step t298 (cl @p_480 @p_483 @p_484 @p_440 @p_436 @p_452 @p_453 @p_415 @p_422) :rule th_resolution :premises (t293 t297)) +(step t299 (cl @p_485 @p_478 @p_483 @p_484 @p_440 @p_436 @p_452 @p_453 @p_415 @p_422) :rule th_resolution :premises (t291 t298)) +(step t300 (cl @p_428 @p_475 @p_440 @p_415 @p_476 @p_485 @p_483 @p_484 @p_436 @p_452 @p_453 @p_422) :rule th_resolution :premises (t290 t299)) +(step t301 (cl @p_428 @p_475 @p_483 @p_484 @p_453) :rule resolution :premises (t300 t172 t174 t189 t226 t278 t201 t223)) +(step t302 (cl @p_429 @p_486 @p_440 @p_415 @p_487 @p_476) :rule la_generic :args (2.0 2.0 (- 2) (- 2) 2 1)) +(step t303 (cl @p_429 @p_486 @p_440 @p_415 @p_476 @p_485 @p_483 @p_484 @p_436 @p_452 @p_453 @p_422) :rule th_resolution :premises (t302 t299)) +(step t304 (cl @p_429 @p_486 @p_483 @p_484 @p_453) :rule resolution :premises (t303 t172 t174 t189 t226 t278 t201 t223)) +(step t305 (cl @p_233 @p_488) :rule or :premises (t71)) +(step t306 (cl (or (! (not @p_488) :named @p_491) (! (or (! (not (! (= @p_354 @p_489) :named @p_540)) :named @p_500) (! (= @p_348 @p_490) :named @p_513)) :named @p_492))) :rule forall_inst :args ((:= veriT_vr28 @p_388) (:= veriT_vr29 @p_490) (:= veriT_vr30 @p_2) (:= veriT_vr31 @p_348))) +(step t307 (cl @p_491 @p_492) :rule or :premises (t306)) +(step t308 (cl (! (or @p_233 @p_492) :named @p_494) @p_493) :rule or_neg) +(step t309 (cl @p_494 @p_218) :rule th_resolution :premises (t72 t308)) +(step t310 (cl @p_494 (! (not @p_492) :named @p_512)) :rule or_neg) +(step t311 (cl @p_494) :rule th_resolution :premises (t305 t307 t309 t310)) +(step t312 (cl (or @p_491 (! (or (! (not (! (= @p_340 @p_495) :named @p_546)) :named @p_515) (! (= @p_252 @p_496) :named @p_516)) :named @p_497))) :rule forall_inst :args ((:= veriT_vr28 @p_2) (:= veriT_vr29 @p_252) (:= veriT_vr30 @p_359) (:= veriT_vr31 @p_496))) +(step t313 (cl @p_491 @p_497) :rule or :premises (t312)) +(step t314 (cl (! (or @p_233 @p_497) :named @p_498) @p_493) :rule or_neg) +(step t315 (cl @p_498 @p_218) :rule th_resolution :premises (t72 t314)) +(step t316 (cl @p_498 (! (not @p_497) :named @p_514)) :rule or_neg) +(step t317 (cl @p_498) :rule th_resolution :premises (t305 t313 t315 t316)) +(step t318 (cl @p_233 @p_499) :rule or :premises (t70)) +(step t319 (cl (or (! (not @p_499) :named @p_501) (! (or @p_500 (! (= @p_2 @p_388) :named @p_518)) :named @p_502))) :rule forall_inst :args ((:= veriT_vr28 @p_388) (:= veriT_vr29 @p_490) (:= veriT_vr30 @p_2) (:= veriT_vr31 @p_348))) +(step t320 (cl @p_501 @p_502) :rule or :premises (t319)) +(step t321 (cl (! (or @p_233 @p_502) :named @p_503) @p_493) :rule or_neg) +(step t322 (cl @p_503 @p_218) :rule th_resolution :premises (t72 t321)) +(step t323 (cl @p_503 (! (not @p_502) :named @p_517)) :rule or_neg) +(step t324 (cl @p_503) :rule th_resolution :premises (t318 t320 t322 t323)) +(step t325 (cl (or @p_501 (! (or (! (not (! (= @p_504 @p_368) :named @p_535)) :named @p_520) (! (= @p_2 @p_239) :named @p_521)) :named @p_506))) :rule forall_inst :args ((:= veriT_vr28 @p_2) (:= veriT_vr29 @p_505) (:= veriT_vr30 @p_239) (:= veriT_vr31 @p_366))) +(step t326 (cl @p_501 @p_506) :rule or :premises (t325)) +(step t327 (cl (! (or @p_233 @p_506) :named @p_507) @p_493) :rule or_neg) +(step t328 (cl @p_507 @p_218) :rule th_resolution :premises (t72 t327)) +(step t329 (cl @p_507 (! (not @p_506) :named @p_519)) :rule or_neg) +(step t330 (cl @p_507) :rule th_resolution :premises (t318 t326 t328 t329)) +(step t331 (cl (or @p_501 (! (or (! (not (! (= @p_508 @p_495) :named @p_527)) :named @p_523) (! (= @p_2 @p_359) :named @p_524)) :named @p_510))) :rule forall_inst :args ((:= veriT_vr28 @p_2) (:= veriT_vr29 @p_509) (:= veriT_vr30 @p_359) (:= veriT_vr31 @p_496))) +(step t332 (cl @p_501 @p_510) :rule or :premises (t331)) +(step t333 (cl (! (or @p_233 @p_510) :named @p_511) @p_493) :rule or_neg) +(step t334 (cl @p_511 @p_218) :rule th_resolution :premises (t72 t333)) +(step t335 (cl @p_511 (! (not @p_510) :named @p_522)) :rule or_neg) +(step t336 (cl @p_511) :rule th_resolution :premises (t318 t332 t334 t335)) +(step t337 (cl @p_512 @p_500 @p_513) :rule or_pos) +(step t338 (cl @p_233 @p_492) :rule or :premises (t311)) +(step t339 (cl @p_492) :rule resolution :premises (t338 t54)) +(step t340 (cl @p_514 @p_515 @p_516) :rule or_pos) +(step t341 (cl @p_233 @p_497) :rule or :premises (t317)) +(step t342 (cl @p_497) :rule resolution :premises (t341 t54)) +(step t343 (cl @p_517 @p_500 @p_518) :rule or_pos) +(step t344 (cl @p_233 @p_502) :rule or :premises (t324)) +(step t345 (cl @p_502) :rule resolution :premises (t344 t54)) +(step t346 (cl @p_519 @p_520 @p_521) :rule or_pos) +(step t347 (cl @p_233 @p_506) :rule or :premises (t330)) +(step t348 (cl @p_506) :rule resolution :premises (t347 t54)) +(step t349 (cl @p_522 @p_523 @p_524) :rule or_pos) +(step t350 (cl @p_233 @p_510) :rule or :premises (t336)) +(step t351 (cl @p_510) :rule resolution :premises (t350 t54)) +(step t352 (cl (! (not @p_403) :named @p_531) (not (! (= @p_525 @p_526) :named @p_528)) (! (not @p_405) :named @p_532) @p_527) :rule eq_transitive) +(step t353 (cl (! (not @p_227) :named @p_530) (not (! (= @p_36 @p_36) :named @p_529)) @p_528) :rule eq_congruent) +(step t354 (cl @p_529) :rule eq_reflexive) +(step t355 (cl @p_530 @p_528) :rule th_resolution :premises (t353 t354)) +(step t356 (cl @p_531 @p_532 @p_527 @p_530) :rule th_resolution :premises (t352 t355)) +(step t357 (cl @p_527) :rule resolution :premises (t356 t69 t206 t214)) +(step t358 (cl (! (not @p_404) :named @p_537) (not (! (= @p_533 @p_363) :named @p_536)) @p_534 @p_535) :rule eq_transitive) +(step t359 (cl @p_530 @p_441 @p_536) :rule eq_congruent) +(step t360 (cl @p_530 @p_536) :rule th_resolution :premises (t359 t254)) +(step t361 (cl @p_537 @p_534 @p_535 @p_530) :rule th_resolution :premises (t358 t360)) +(step t362 (cl @p_535) :rule resolution :premises (t361 t69 t208 t216)) +(step t363 (cl @p_538 (not (! (= @p_347 @p_539) :named @p_541)) (! (not @p_406) :named @p_543) @p_540) :rule eq_transitive) +(step t364 (cl @p_530 (not (! (= 1.0 1.0) :named @p_542)) @p_541) :rule eq_congruent) +(step t365 (cl @p_542) :rule eq_reflexive) +(step t366 (cl @p_530 @p_541) :rule th_resolution :premises (t364 t365)) +(step t367 (cl @p_538 @p_543 @p_540 @p_530) :rule th_resolution :premises (t363 t366)) +(step t368 (cl @p_540) :rule resolution :premises (t367 t69 t212 t220)) +(step t369 (cl @p_544 (! (not (! (= @p_334 @p_376) :named @p_553)) :named @p_552) @p_545 (not (! (= @p_495 @p_383) :named @p_548)) @p_546) :rule eq_transitive) +(step t370 (cl (! (not @p_547) :named @p_550) (! (not (! (= @p_496 @p_381) :named @p_551)) :named @p_549) @p_548) :rule eq_congruent) +(step t371 (cl @p_479 @p_547 @p_443 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t238 t256)) +(step t372 (cl @p_547 @p_443 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t371 t239)) +(step t373 (cl @p_549 @p_548 @p_443 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t370 t372)) +(step t374 (cl @p_550 @p_551) :rule eq_congruent) +(step t375 (cl @p_551 @p_443 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t374 t372)) +(step t376 (cl @p_548 @p_443 @p_436 @p_440 @p_415 @p_422 @p_443 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t373 t375)) +(step t377 (cl @p_548 @p_443 @p_436 @p_440 @p_415 @p_422) :rule contraction :premises (t376)) +(step t378 (cl @p_544 @p_552 @p_545 @p_546 @p_443 @p_436 @p_440 @p_415 @p_422) :rule th_resolution :premises (t369 t377)) +(step t379 (cl @p_530 @p_461 @p_553) :rule eq_congruent) +(step t380 (cl @p_530 @p_553) :rule th_resolution :premises (t379 t228)) +(step t381 (cl @p_544 @p_545 @p_546 @p_443 @p_436 @p_440 @p_415 @p_422 @p_530) :rule th_resolution :premises (t378 t380)) +(step t382 (cl @p_546) :rule resolution :premises (t381 t189 t226 t201 t69 t210 t218 t248 t223)) +(step t383 (cl @p_524) :rule resolution :premises (t349 t357 t351)) +(step t384 (cl @p_521) :rule resolution :premises (t346 t362 t348)) +(step t385 (cl @p_513) :rule resolution :premises (t337 t368 t339)) +(step t386 (cl @p_518) :rule resolution :premises (t343 t368 t345)) +(step t387 (cl @p_516) :rule resolution :premises (t340 t382 t342)) +(step t388 (cl (! (not (! (= @p_309 @p_249) :named @p_559)) :named @p_556) @p_554 (! (not @p_516) :named @p_557) (! (not (! (= @p_36 @p_496) :named @p_555)) :named @p_569) @p_440 @p_449 @p_436 @p_452 @p_450) :rule eq_transitive) +(step t389 (cl (! (not @p_524) :named @p_558) @p_555) :rule eq_congruent) +(step t390 (cl @p_556 @p_554 @p_557 @p_440 @p_449 @p_436 @p_452 @p_450 @p_558) :rule th_resolution :premises (t388 t389)) +(step t391 (cl @p_556 @p_554 @p_557 @p_440 @p_436 @p_452 @p_450 @p_558 @p_415 @p_422) :rule th_resolution :premises (t390 t237)) +(step t392 (cl (not (! (= @p_235 @p_348) :named @p_560)) @p_559) :rule eq_congruent) +(step t393 (cl @p_436 @p_449 @p_440 (! (not (! (= @p_36 @p_490) :named @p_561)) :named @p_564) (! (not @p_513) :named @p_562) @p_560) :rule eq_transitive) +(step t394 (cl (! (not @p_518) :named @p_563) @p_561) :rule eq_congruent) +(step t395 (cl @p_436 @p_449 @p_440 @p_562 @p_560 @p_563) :rule th_resolution :premises (t393 t394)) +(step t396 (cl @p_436 @p_440 @p_562 @p_560 @p_563 @p_415 @p_422) :rule th_resolution :premises (t395 t237)) +(step t397 (cl @p_559 @p_436 @p_440 @p_562 @p_563 @p_415 @p_422) :rule th_resolution :premises (t392 t396)) +(step t398 (cl @p_554 @p_557 @p_440 @p_436 @p_452 @p_450 @p_558 @p_415 @p_422 @p_436 @p_440 @p_562 @p_563 @p_415 @p_422) :rule th_resolution :premises (t391 t397)) +(step t399 (cl @p_554 @p_557 @p_440 @p_436 @p_452 @p_450 @p_558 @p_415 @p_422 @p_562 @p_563) :rule contraction :premises (t398)) +(step t400 (cl @p_450) :rule resolution :premises (t399 t189 t226 t278 t201 t223 t385 t387 t386 t176 t383)) +(step t401 (cl (not (! (= @p_36 @p_348) :named @p_565)) (! (= @p_237 @p_249) :named @p_566)) :rule eq_congruent) +(step t402 (cl @p_564 @p_562 @p_565) :rule eq_transitive) +(step t403 (cl @p_562 @p_565 @p_563) :rule th_resolution :premises (t402 t394)) +(step t404 (cl @p_566 @p_562 @p_563) :rule th_resolution :premises (t401 t403)) +(step t405 (cl @p_444) :rule resolution :premises (t271 t400 t278)) +(step t406 (cl @p_464) :rule resolution :premises (t281 t400)) +(step t407 (cl @p_469) :rule resolution :premises (t285 t405)) +(step t408 (cl @p_427) :rule resolution :premises (t241 t406 t405)) +(step t409 (cl @p_470) :rule resolution :premises (t287 t406)) +(step t410 (cl @p_466) :rule resolution :premises (t283 t409 t407)) +(step t411 (cl @p_475) :rule resolution :premises (t301 t410 t400 t406 t408)) +(step t412 (cl @p_486) :rule resolution :premises (t304 t410 t400 t405 t408)) +(step t413 (cl @p_472) :rule resolution :premises (t289 t412 t411)) +(step t414 (cl (not (! (= @p_2 0.0) :named @p_568)) (! (= @p_36 @p_322) :named @p_573)) :rule eq_congruent) +(step t415 (cl (! (not @p_521) :named @p_571) @p_567 @p_485 (! (not @p_472) :named @p_572) @p_483 @p_568) :rule eq_transitive) +(step t416 (cl @p_483 @p_484 @p_569 @p_557 @p_554 (! (not @p_566) :named @p_570) @p_482) :rule eq_transitive) +(step t417 (cl @p_483 @p_484 @p_557 @p_554 @p_570 @p_482 @p_558) :rule th_resolution :premises (t416 t389)) +(step t418 (cl @p_483 @p_484 @p_557 @p_554 @p_482 @p_558 @p_562 @p_563) :rule th_resolution :premises (t417 t404)) +(step t419 (cl @p_480 @p_483 @p_484 @p_557 @p_554 @p_558 @p_562 @p_563) :rule th_resolution :premises (t293 t418)) +(step t420 (cl @p_571 @p_485 @p_572 @p_483 @p_568 @p_483 @p_484 @p_557 @p_554 @p_558 @p_562 @p_563) :rule th_resolution :premises (t415 t419)) +(step t421 (cl @p_571 @p_485 @p_572 @p_483 @p_568 @p_484 @p_557 @p_554 @p_558 @p_562 @p_563) :rule contraction :premises (t420)) +(step t422 (cl @p_573 @p_571 @p_485 @p_572 @p_483 @p_484 @p_557 @p_554 @p_558 @p_562 @p_563) :rule th_resolution :premises (t414 t421)) +(step t423 (cl @p_571 @p_567 @p_485 @p_572 @p_484 (not @p_573) @p_574 @p_415 @p_413) :rule eq_transitive) +(step t424 (cl @p_571 @p_567 @p_485 @p_572 @p_484 @p_574 @p_415 @p_413 @p_571 @p_485 @p_572 @p_483 @p_484 @p_557 @p_554 @p_558 @p_562 @p_563) :rule th_resolution :premises (t423 t422)) +(step t425 (cl @p_571 @p_567 @p_485 @p_572 @p_484 @p_574 @p_415 @p_413 @p_483 @p_557 @p_554 @p_558 @p_562 @p_563) :rule contraction :premises (t424)) +(step t426 (cl @p_571 @p_485 @p_572 @p_484 @p_574 @p_415 @p_413 @p_483 @p_557 @p_554 @p_558 @p_562 @p_563 @p_483 @p_484 @p_557 @p_554 @p_558 @p_562 @p_563) :rule th_resolution :premises (t425 t419)) +(step t427 (cl @p_571 @p_485 @p_572 @p_484 @p_574 @p_415 @p_413 @p_483 @p_557 @p_554 @p_558 @p_562 @p_563) :rule contraction :premises (t426)) +(step t428 (cl) :rule resolution :premises (t427 t172 t176 t232 t226 t204 t408 t410 t413 t385 t387 t386 t384 t383)) +76e04d808f9049c53498efd03433d66318b777cb 26 0 +unsat +(assume a0 (! (forall ((?v0 A$) (?v1 B$) (?v2 C$)) (! (p$ ?v0 ?v1) :named @p_2)) :named @p_1)) +(assume a1 (! (=> (! (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 (a0 t3 t4)) +(anchor :step t6 :args ((:= ?v0 veriT_vr0) (:= ?v1 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 veriT_vr2) (:= veriT_vr1 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 (a1 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)) +1fd041856d406153ff429b6c1d40eb368a3738f2 23 0 +unsat +(assume a0 (! (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 (a0 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)) +e6caa829a170f55dd7b52a2f7910337007b8ab18 484 0 +unsat +(assume a0 (! (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 a3 (! (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 a4 (! (< 0.0 d$) :named @p_279)) +(assume a5 (! (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 a7 (! (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 a8 (! (forall ((?v0 Real) (?v1 Real)) (! (= (! (divide$ @p_4 ?v1) :named @p_142) (! (- @p_5) :named @p_147)) :named @p_149)) :named @p_140)) +(assume a9 (! (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 a10 (! (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 a11 (! (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 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 (a0 t10 t11)) +(anchor :step t13 :args ((:= veriT_vr0 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 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 (a3 t16 t17)) +(anchor :step t19 :args ((:= veriT_vr6 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 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 (a5 t22 t23)) +(anchor :step t25 :args ((:= veriT_vr8 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 veriT_vr10) (:= ?v1 veriT_vr11) (:= ?v2 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 (a7 t28 t29)) +(anchor :step t31 :args (veriT_vr10 veriT_vr11 veriT_vr12)) +(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 veriT_vr13) (:= veriT_vr11 veriT_vr14) (:= veriT_vr12 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 veriT_vr16) (:= ?v1 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 (a8 t37 t38)) +(anchor :step t40 :args ((:= veriT_vr16 veriT_vr18) (:= veriT_vr17 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 veriT_vr20) (:= ?v1 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 (a9 t43 t44)) +(anchor :step t46 :args ((:= veriT_vr20 veriT_vr22) (:= veriT_vr21 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 veriT_vr24) (:= ?v1 veriT_vr25) (:= ?v2 veriT_vr26) (:= ?v3 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 (a10 t49 t50)) +(anchor :step t52 :args ((:= veriT_vr24 veriT_vr28) (:= veriT_vr25 veriT_vr29) (:= veriT_vr26 veriT_vr30) (:= veriT_vr27 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 (a11 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_291)) :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_369))) :rule qnt_cnf) +(step t71 (cl (not (! (not (! (not @p_161) :named @p_245)) :named @p_254)) @p_161) :rule not_not) +(step t72 (cl (not (! (not (! (not @p_139) :named @p_256)) :named @p_276)) @p_139) :rule not_not) +(step t73 (cl (not (! (not @p_233) :named @p_374)) @p_218) :rule not_not) +(step t74 (cl (not (! (not (! (not @p_184) :named @p_234)) :named @p_243)) @p_184) :rule not_not) +(step t75 (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 t76) +(assume t76.h1 @p_235) +(step t76.t2 (cl (! (= 0.0 @p_238) :named @p_249)) :rule minus_simplify) +(step t76.t3 (cl (= @p_239 @p_236)) :rule cong :premises (t76.t2)) +(step t76.t4 (cl (! (= @p_235 (! (= @p_236 @p_237) :named @p_240)) :named @p_241)) :rule cong :premises (t76.t3)) +(step t76.t5 (cl (not @p_241) (! (not @p_235) :named @p_242) @p_240) :rule equiv_pos2) +(step t76.t6 (cl @p_240) :rule th_resolution :premises (t76.h1 t76.t4 t76.t5)) +(step t76 (cl @p_242 @p_240) :rule subproof :discharge (h1)) +(step t77 (cl @p_234 @p_235) :rule or :premises (t75)) +(step t78 (cl (! (or @p_234 @p_240) :named @p_244) @p_243) :rule or_neg) +(step t79 (cl @p_244 @p_184) :rule th_resolution :premises (t74 t78)) +(step t80 (cl @p_244 (! (not @p_240) :named @p_367)) :rule or_neg) +(step t81 (cl @p_244) :rule th_resolution :premises (t77 t76 t79 t80)) +(step t82 (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 t83) +(assume t83.h1 @p_246) +(step t83.t2 (cl @p_249) :rule minus_simplify) +(step t83.t3 (cl (= @p_250 @p_247)) :rule cong :premises (t83.t2)) +(step t83.t4 (cl (! (= @p_246 (! (= @p_247 @p_248) :named @p_251)) :named @p_252)) :rule cong :premises (t83.t3)) +(step t83.t5 (cl (not @p_252) (! (not @p_246) :named @p_253) @p_251) :rule equiv_pos2) +(step t83.t6 (cl @p_251) :rule th_resolution :premises (t83.h1 t83.t4 t83.t5)) +(step t83 (cl @p_253 @p_251) :rule subproof :discharge (h1)) +(step t84 (cl @p_245 @p_246) :rule or :premises (t82)) +(step t85 (cl (! (or @p_245 @p_251) :named @p_255) @p_254) :rule or_neg) +(step t86 (cl @p_255 @p_161) :rule th_resolution :premises (t71 t85)) +(step t87 (cl @p_255 (! (not @p_251) :named @p_350)) :rule or_neg) +(step t88 (cl @p_255) :rule th_resolution :premises (t84 t83 t86 t87)) +(step t89 (cl (or @p_256 (! (= (! (< (! (divide$ @p_36 @p_36) :named @p_257) @p_257) :named @p_263) (! (and (! (=> (! (< 0.0 @p_36) :named @p_260) (! (< @p_36 @p_36) :named @p_258)) :named @p_264) (! (=> (! (< @p_36 0.0) :named @p_261) @p_258) :named @p_267) (! (not (! (= @p_36 0.0) :named @p_342)) :named @p_262)) :named @p_270)) :named @p_259))) :rule forall_inst :args ((:= veriT_vr13 @p_36) (:= veriT_vr14 @p_36) (:= veriT_vr15 @p_36))) +(anchor :step t90) +(assume t90.h1 @p_259) +(step t90.t2 (cl (= @p_263 false)) :rule comp_simplify) +(step t90.t3 (cl (= @p_258 false)) :rule comp_simplify) +(step t90.t4 (cl (= @p_264 (! (=> @p_260 false) :named @p_265))) :rule cong :premises (t90.t3)) +(step t90.t5 (cl (= @p_265 (! (not @p_260) :named @p_266))) :rule implies_simplify) +(step t90.t6 (cl (= @p_264 @p_266)) :rule trans :premises (t90.t4 t90.t5)) +(step t90.t7 (cl (= @p_267 (! (=> @p_261 false) :named @p_268))) :rule cong :premises (t90.t3)) +(step t90.t8 (cl (= @p_268 (! (not @p_261) :named @p_269))) :rule implies_simplify) +(step t90.t9 (cl (= @p_267 @p_269)) :rule trans :premises (t90.t7 t90.t8)) +(step t90.t10 (cl (= @p_270 (! (and @p_266 @p_269 @p_262) :named @p_271))) :rule cong :premises (t90.t6 t90.t9)) +(step t90.t11 (cl (= @p_259 (! (= false @p_271) :named @p_272))) :rule cong :premises (t90.t2 t90.t10)) +(step t90.t12 (cl (= @p_272 (! (not @p_271) :named @p_273))) :rule equiv_simplify) +(step t90.t13 (cl (! (= @p_259 @p_273) :named @p_274)) :rule trans :premises (t90.t11 t90.t12)) +(step t90.t14 (cl (not @p_274) (! (not @p_259) :named @p_275) @p_273) :rule equiv_pos2) +(step t90.t15 (cl @p_273) :rule th_resolution :premises (t90.h1 t90.t13 t90.t14)) +(step t90 (cl @p_275 @p_273) :rule subproof :discharge (h1)) +(step t91 (cl @p_256 @p_259) :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 (t72 t92)) +(step t94 (cl @p_277 (! (not @p_273) :named @p_278)) :rule or_neg) +(step t95 (cl (not @p_278) @p_271) :rule not_not) +(step t96 (cl @p_277 @p_271) :rule th_resolution :premises (t95 t94)) +(step t97 (cl @p_277) :rule th_resolution :premises (t91 t90 t93 t96)) +(step t98 (cl (or @p_256 (! (= (! (< @p_247 @p_3) :named @p_281) (! (and (! (=> (! (< 0.0 2.0) :named @p_282) @p_279) :named @p_283) (! (=> (! (< 2.0 0.0) :named @p_285) (! (< d$ 0.0) :named @p_287)) :named @p_286) (! (not (! (= 2.0 0.0) :named @p_289)) :named @p_290)) :named @p_292)) :named @p_280))) :rule forall_inst :args ((:= veriT_vr13 0.0) (:= veriT_vr14 2.0) (:= veriT_vr15 d$))) +(anchor :step t99) +(assume t99.h1 @p_280) +(step t99.t2 (cl (= @p_282 true)) :rule comp_simplify) +(step t99.t3 (cl (= @p_283 (! (=> true @p_279) :named @p_284))) :rule cong :premises (t99.t2)) +(step t99.t4 (cl (= @p_284 @p_279)) :rule implies_simplify) +(step t99.t5 (cl (= @p_283 @p_279)) :rule trans :premises (t99.t3 t99.t4)) +(step t99.t6 (cl (= @p_285 false)) :rule comp_simplify) +(step t99.t7 (cl (= @p_286 (! (=> false @p_287) :named @p_288))) :rule cong :premises (t99.t6)) +(step t99.t8 (cl (= @p_288 true)) :rule implies_simplify) +(step t99.t9 (cl (= @p_286 true)) :rule trans :premises (t99.t7 t99.t8)) +(step t99.t10 (cl (= @p_289 false)) :rule eq_simplify) +(step t99.t11 (cl (= @p_290 @p_228)) :rule cong :premises (t99.t10)) +(step t99.t12 (cl @p_291) :rule not_simplify) +(step t99.t13 (cl (= @p_290 true)) :rule trans :premises (t99.t11 t99.t12)) +(step t99.t14 (cl (= @p_292 (! (and @p_279 true true) :named @p_293))) :rule cong :premises (t99.t5 t99.t9 t99.t13)) +(step t99.t15 (cl (= @p_293 (! (and @p_279) :named @p_294))) :rule and_simplify) +(step t99.t16 (cl (= @p_294 @p_279)) :rule and_simplify) +(step t99.t17 (cl (= @p_292 @p_279)) :rule trans :premises (t99.t14 t99.t15 t99.t16)) +(step t99.t18 (cl (! (= @p_280 (! (= @p_281 @p_279) :named @p_295)) :named @p_296)) :rule cong :premises (t99.t17)) +(step t99.t19 (cl (not @p_296) (! (not @p_280) :named @p_297) @p_295) :rule equiv_pos2) +(step t99.t20 (cl @p_295) :rule th_resolution :premises (t99.h1 t99.t18 t99.t19)) +(step t99 (cl @p_297 @p_295) :rule subproof :discharge (h1)) +(step t100 (cl @p_256 @p_280) :rule or :premises (t98)) +(step t101 (cl (! (or @p_256 @p_295) :named @p_298) @p_276) :rule or_neg) +(step t102 (cl @p_298 @p_139) :rule th_resolution :premises (t72 t101)) +(step t103 (cl @p_298 (! (not @p_295) :named @p_343)) :rule or_neg) +(step t104 (cl @p_298) :rule th_resolution :premises (t100 t99 t102 t103)) +(step t105 (cl (not (! (not (! (not @p_84) :named @p_299)) :named @p_312)) @p_84) :rule not_not) +(step t106 (cl (or @p_299 (! (= @p_36 (! (- @p_3 (! (ite (! (< @p_2 0.0) :named @p_301) (! (- @p_2) :named @p_303) @p_2) :named @p_304)) :named @p_305)) :named @p_300))) :rule forall_inst :args ((:= veriT_vr9 @p_2))) +(anchor :step t107) +(assume t107.h1 @p_300) +(step t107.t2 (cl (= @p_3 @p_303)) :rule minus_simplify) +(step t107.t3 (cl (= @p_304 (! (ite @p_301 @p_3 @p_2) :named @p_302))) :rule cong :premises (t107.t2)) +(step t107.t4 (cl (= @p_305 (! (- @p_3 @p_302) :named @p_306))) :rule cong :premises (t107.t3)) +(step t107.t5 (cl (! (= @p_300 (! (= @p_36 @p_306) :named @p_309)) :named @p_307)) :rule cong :premises (t107.t4)) +(step t107.t6 (cl (not @p_307) (! (not @p_300) :named @p_308) @p_309) :rule equiv_pos2) +(step t107.t7 (cl @p_309) :rule th_resolution :premises (t107.h1 t107.t5 t107.t6)) +(step t107.t8 (cl (! (= @p_309 (! (and (! (= @p_36 (- @p_3 @p_302)) :named @p_345) (! (ite @p_301 (! (= @p_3 @p_302) :named @p_347) (= @p_2 @p_302)) :named @p_346)) :named @p_310)) :named @p_311)) :rule ite_intro) +(step t107.t9 (cl (not @p_311) (not @p_309) @p_310) :rule equiv_pos2) +(step t107.t10 (cl @p_310) :rule th_resolution :premises (t107.t7 t107.t8 t107.t9)) +(step t107 (cl @p_308 @p_310) :rule subproof :discharge (h1)) +(step t108 (cl @p_299 @p_300) :rule or :premises (t106)) +(step t109 (cl (! (or @p_299 @p_310) :named @p_313) @p_312) :rule or_neg) +(step t110 (cl @p_313 @p_84) :rule th_resolution :premises (t105 t109)) +(step t111 (cl @p_313 (! (not @p_310) :named @p_344)) :rule or_neg) +(step t112 (cl @p_313) :rule th_resolution :premises (t108 t107 t110 t111)) +(step t113 (cl (or @p_299 (! (= (! (diamond_y$ 0.0) :named @p_315) (! (- @p_3 (! (ite (! (< 0.0 0.0) :named @p_316) @p_238 0.0) :named @p_317)) :named @p_319)) :named @p_314))) :rule forall_inst :args ((:= veriT_vr9 0.0))) +(anchor :step t114) +(assume t114.h1 @p_314) +(step t114.t2 (cl (= @p_316 false)) :rule comp_simplify) +(step t114.t3 (cl @p_249) :rule minus_simplify) +(step t114.t4 (cl (= @p_317 (! (ite false 0.0 0.0) :named @p_318))) :rule cong :premises (t114.t2 t114.t3)) +(step t114.t5 (cl (= 0.0 @p_318)) :rule ite_simplify) +(step t114.t6 (cl (= 0.0 @p_317)) :rule trans :premises (t114.t4 t114.t5)) +(step t114.t7 (cl (= @p_319 (! (- @p_3 0.0) :named @p_320))) :rule cong :premises (t114.t6)) +(step t114.t8 (cl (= @p_3 @p_320)) :rule minus_simplify) +(step t114.t9 (cl (= @p_3 @p_319)) :rule trans :premises (t114.t7 t114.t8)) +(step t114.t10 (cl (! (= @p_314 (! (= @p_3 @p_315) :named @p_321)) :named @p_322)) :rule cong :premises (t114.t9)) +(step t114.t11 (cl (not @p_322) (! (not @p_314) :named @p_323) @p_321) :rule equiv_pos2) +(step t114.t12 (cl @p_321) :rule th_resolution :premises (t114.h1 t114.t10 t114.t11)) +(step t114 (cl @p_323 @p_321) :rule subproof :discharge (h1)) +(step t115 (cl @p_299 @p_314) :rule or :premises (t113)) +(step t116 (cl (! (or @p_299 @p_321) :named @p_324) @p_312) :rule or_neg) +(step t117 (cl @p_324 @p_84) :rule th_resolution :premises (t105 t116)) +(step t118 (cl @p_324 (! (not @p_321) :named @p_359)) :rule or_neg) +(step t119 (cl @p_324) :rule th_resolution :premises (t115 t114 t117 t118)) +(step t120 (cl (or (! (not @p_60) :named @p_348) (! (= (! (fun_app$ uub$ @p_7) :named @p_379) (! (pair$ @p_2 (! (times$ (- (* 2.0 @p_7) 1.0) @p_36) :named @p_371)) :named @p_370)) :named @p_349))) :rule forall_inst :args ((:= veriT_vr7 @p_7))) +(step t121 (cl (or (! (not @p_34) :named @p_336) (! (= (! (fun_app$ uuc$ @p_7) :named @p_327) (! (pair$ (! (times$ (! (- @p_7 @p_7) :named @p_328) d$) :named @p_325) (! (diamond_y$ @p_325) :named @p_329)) :named @p_331)) :named @p_326))) :rule forall_inst :args ((:= veriT_vr1 @p_7))) +(anchor :step t122) +(assume t122.h1 @p_326) +(step t122.t2 (cl (= 0.0 @p_328)) :rule minus_simplify) +(step t122.t3 (cl (= @p_236 @p_325)) :rule cong :premises (t122.t2)) +(step t122.t4 (cl (= @p_329 (! (diamond_y$ @p_236) :named @p_330))) :rule cong :premises (t122.t3)) +(step t122.t5 (cl (= @p_331 (! (pair$ @p_236 @p_330) :named @p_332))) :rule cong :premises (t122.t3 t122.t4)) +(step t122.t6 (cl (! (= @p_326 (! (= @p_327 @p_332) :named @p_333)) :named @p_334)) :rule cong :premises (t122.t5)) +(step t122.t7 (cl (not @p_334) (! (not @p_326) :named @p_335) @p_333) :rule equiv_pos2) +(step t122.t8 (cl @p_333) :rule th_resolution :premises (t122.h1 t122.t6 t122.t7)) +(step t122 (cl @p_335 @p_333) :rule subproof :discharge (h1)) +(step t123 (cl @p_336 @p_326) :rule or :premises (t121)) +(step t124 (cl (! (or @p_336 @p_333) :named @p_338) (! (not @p_336) :named @p_337)) :rule or_neg) +(step t125 (cl (not @p_337) @p_34) :rule not_not) +(step t126 (cl @p_338 @p_34) :rule th_resolution :premises (t125 t124)) +(step t127 (cl @p_338 (! (not @p_333) :named @p_380)) :rule or_neg) +(step t128 (cl @p_338) :rule th_resolution :premises (t123 t122 t126 t127)) +(step t129 (cl @p_234 @p_240) :rule or :premises (t81)) +(step t130 (cl @p_240) :rule resolution :premises (t129 t48)) +(step t131 (cl @p_245 @p_251) :rule or :premises (t88)) +(step t132 (cl @p_251) :rule resolution :premises (t131 t42)) +(step t133 (cl @p_271 (! (not @p_266) :named @p_339) (! (not @p_269) :named @p_340) (! (not @p_262) :named @p_341)) :rule and_neg) +(step t134 (cl (not @p_339) @p_260) :rule not_not) +(step t135 (cl @p_271 @p_260 @p_340 @p_341) :rule th_resolution :premises (t134 t133)) +(step t136 (cl (not @p_340) @p_261) :rule not_not) +(step t137 (cl @p_271 @p_260 @p_261 @p_341) :rule th_resolution :premises (t136 t135)) +(step t138 (cl (not @p_341) @p_342) :rule not_not) +(step t139 (cl @p_271 @p_260 @p_261 @p_342) :rule th_resolution :premises (t138 t137)) +(step t140 (cl @p_256 @p_273) :rule or :premises (t97)) +(step t141 (cl @p_273) :rule resolution :premises (t140 t36)) +(step t142 (cl @p_343 @p_281 (not @p_279)) :rule equiv_pos1) +(step t143 (cl @p_256 @p_295) :rule or :premises (t104)) +(step t144 (cl @p_343 @p_281) :rule resolution :premises (t142 a4)) +(step t145 (cl @p_295) :rule resolution :premises (t143 t36)) +(step t146 (cl @p_281) :rule resolution :premises (t144 t145)) +(step t147 (cl @p_344 @p_345) :rule and_pos) +(step t148 (cl (not @p_346) (not @p_301) @p_347) :rule ite_pos2) +(step t149 (cl @p_344 @p_346) :rule and_pos) +(step t150 (cl @p_299 @p_310) :rule or :premises (t112)) +(step t151 (cl @p_310) :rule resolution :premises (t150 t27)) +(step t152 (cl @p_345) :rule resolution :premises (t147 t151)) +(step t153 (cl @p_346) :rule resolution :premises (t149 t151)) +(step t154 (cl @p_299 @p_321) :rule or :premises (t119)) +(step t155 (cl @p_321) :rule resolution :premises (t154 t27)) +(step t156 (cl @p_348 @p_349) :rule or :premises (t120)) +(step t157 (cl @p_349) :rule resolution :premises (t156 t21)) +(step t158 (cl @p_336 @p_333) :rule or :premises (t128)) +(step t159 (cl @p_333) :rule resolution :premises (t158 t15)) +(step t160 (cl (! (not @p_281) :named @p_352) (! (<= @p_2 @p_3) :named @p_351) @p_350) :rule la_generic :args (2.0 1.0 1)) +(step t161 (cl @p_351) :rule resolution :premises (t160 t132 t146)) +(step t162 (cl @p_352 (! (not (! (<= @p_3 @p_2) :named @p_355)) :named @p_353) @p_350) :rule la_generic :args (2.0 1.0 1)) +(step t163 (cl @p_353) :rule resolution :premises (t162 t132 t146)) +(step t164 (cl (! (not (! (= @p_3 @p_2) :named @p_360)) :named @p_354) @p_354 (! (not @p_351) :named @p_356) @p_355) :rule eq_congruent_pred) +(step t165 (cl @p_354 @p_356 @p_355) :rule contraction :premises (t164)) +(step t166 (cl @p_354) :rule resolution :premises (t165 t163 t161)) +(step t167 (cl @p_352 @p_301 @p_350) :rule la_generic :args (2.0 2.0 1)) +(step t168 (cl @p_301) :rule resolution :premises (t167 t132 t146)) +(step t169 (cl @p_347) :rule resolution :premises (t148 t168 t153)) +(step t170 (cl (! (= @p_2 @p_2) :named @p_386)) :rule eq_reflexive) +(step t171 (cl @p_266 (! (not @p_345) :named @p_357) (! (not @p_347) :named @p_358)) :rule la_generic :args (1.0 (- 1) (- 1))) +(step t172 (cl @p_266) :rule resolution :premises (t171 t152 t169)) +(step t173 (cl @p_269 @p_357 @p_358) :rule la_generic :args (1.0 1 1)) +(step t174 (cl @p_269) :rule resolution :premises (t173 t152 t169)) +(step t175 (cl @p_342) :rule resolution :premises (t139 t174 t141 t172)) +(step t176 (cl (! (= @p_7 @p_7) :named @p_382)) :rule eq_reflexive) +(step t177 (cl @p_359 (not (! (= @p_36 @p_315) :named @p_362)) @p_262 (! (not (! (= @p_2 0.0) :named @p_388)) :named @p_361) @p_360) :rule eq_transitive) +(step t178 (cl @p_361 @p_362) :rule eq_congruent) +(step t179 (cl @p_359 @p_262 @p_361 @p_360 @p_361) :rule th_resolution :premises (t177 t178)) +(step t180 (cl @p_359 @p_262 @p_361 @p_360) :rule contraction :premises (t179)) +(step t181 (cl @p_361) :rule resolution :premises (t180 t175 t155 t166)) +(step t182 (cl (or (! (= @p_36 @p_236) :named @p_363) (! (not (! (<= @p_36 @p_236) :named @p_366)) :named @p_364) (! (not (! (<= @p_236 @p_36) :named @p_368)) :named @p_365))) :rule la_disequality) +(step t183 (cl @p_363 @p_364 @p_365) :rule or :premises (t182)) +(step t184 (cl @p_366 @p_367 @p_357 @p_358) :rule la_generic :args (2.0 1 (- 2) (- 2))) +(step t185 (cl @p_366) :rule resolution :premises (t184 t130 t152 t169)) +(step t186 (cl @p_368 @p_367 @p_357 @p_358) :rule la_generic :args (2.0 (- 1) 2 2)) +(step t187 (cl @p_368) :rule resolution :premises (t186 t130 t152 t169)) +(step t188 (cl @p_363) :rule resolution :premises (t183 t187 t185)) +(step t189 (cl (! (not @p_363) :named @p_391) @p_262 (! (= 0.0 @p_236) :named @p_387)) :rule eq_transitive) +(step t190 (cl @p_233 @p_369) :rule or :premises (t70)) +(step t191 (cl (or (! (not @p_369) :named @p_372) (! (or (! (not (! (= @p_370 @p_332) :named @p_381)) :named @p_377) (! (= @p_2 @p_236) :named @p_378)) :named @p_373))) :rule forall_inst :args ((:= veriT_vr28 @p_236) (:= veriT_vr29 @p_330) (:= veriT_vr30 @p_2) (:= veriT_vr31 @p_371))) +(step t192 (cl @p_372 @p_373) :rule or :premises (t191)) +(step t193 (cl (! (or @p_233 @p_373) :named @p_375) @p_374) :rule or_neg) +(step t194 (cl @p_375 @p_218) :rule th_resolution :premises (t73 t193)) +(step t195 (cl @p_375 (! (not @p_373) :named @p_376)) :rule or_neg) +(step t196 (cl @p_375) :rule th_resolution :premises (t190 t192 t194 t195)) +(step t197 (cl @p_376 @p_377 @p_378) :rule or_pos) +(step t198 (cl @p_233 @p_373) :rule or :premises (t196)) +(step t199 (cl @p_373) :rule resolution :premises (t198 t54)) +(step t200 (cl (! (not @p_349) :named @p_385) (not (! (= @p_379 @p_327) :named @p_383)) @p_380 @p_381) :rule eq_transitive) +(step t201 (cl (! (not @p_227) :named @p_384) (not @p_382) @p_383) :rule eq_congruent) +(step t202 (cl @p_384 @p_383) :rule th_resolution :premises (t201 t176)) +(step t203 (cl @p_385 @p_380 @p_381 @p_384) :rule th_resolution :premises (t200 t202)) +(step t204 (cl @p_381) :rule resolution :premises (t203 t69 t157 t159)) +(step t205 (cl (not @p_386) (! (not @p_378) :named @p_389) (! (not @p_387) :named @p_390) @p_388) :rule eq_transitive) +(step t206 (cl @p_389 @p_390 @p_388) :rule th_resolution :premises (t205 t170)) +(step t207 (cl @p_389 @p_388 @p_391 @p_262) :rule th_resolution :premises (t206 t189)) +(step t208 (cl @p_389) :rule resolution :premises (t207 t175 t181 t188)) +(step t209 (cl) :rule resolution :premises (t197 t204 t199 t208)) +a2ae6f83b2ea3d3da88443f025d092ee4567250b 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 a29 (! (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 a31 (! (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 a33 (! (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 a36 (! (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 a40 (! (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 a44 (! (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 a48 (! (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 a50 (! (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 a51 (! (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 a23 (! (inj_on$ f$ @p_6) :named @p_353)) +(assume a24 (! (finite$a b$) :named @p_332)) +(assume a25 (not (! (= bot$a b$) :named @p_331))) +(assume a26 (! (member$a (! (arg_min_on$a (! (comp$ f$ g$) :named @p_4) b$) :named @p_5) b$) :named @p_423)) +(assume a27 (! (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 a52 (not (! (= @p_336 (! (fun_app$b g$ @p_5) :named @p_333)) :named @p_355))) +(anchor :step t16 :args ((:= ?v0 veriT_vr0) (:= ?v1 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 (a29 t16 t17)) +(anchor :step t19 :args ((:= veriT_vr0 veriT_vr2) (:= veriT_vr1 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 veriT_vr8) (:= ?v1 veriT_vr9) (:= ?v2 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 (a31 t22 t23)) +(anchor :step t25 :args (veriT_vr8 veriT_vr9 veriT_vr10)) +(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 veriT_vr11) (:= veriT_vr9 veriT_vr12) (:= veriT_vr10 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 veriT_vr20) (:= ?v1 veriT_vr21) (:= ?v2 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 (a33 t31 t32)) +(anchor :step t34 :args ((:= veriT_vr20 veriT_vr23) (:= veriT_vr21 veriT_vr24) (:= veriT_vr22 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 veriT_vr34) (:= ?v1 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 (a36 t37 t38)) +(anchor :step t40 :args ((:= veriT_vr34 veriT_vr36) (:= veriT_vr35 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 veriT_vr58) (:= ?v1 veriT_vr59) (:= ?v2 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 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 (a40 t43 t44)) +(anchor :step t46 :args (veriT_vr58 veriT_vr59 veriT_vr60)) +(anchor :step t46.t1 :args (veriT_vr61)) +(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 veriT_vr62) (:= veriT_vr59 veriT_vr63) (:= veriT_vr60 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 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 veriT_vr90) (:= ?v1 veriT_vr91) (:= ?v2 veriT_vr92) (:= ?v3 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 (a44 t52 t53)) +(anchor :step t55 :args ((:= veriT_vr90 veriT_vr94) (:= veriT_vr91 veriT_vr95) (:= veriT_vr92 veriT_vr96) (:= veriT_vr93 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 veriT_vr114) (:= ?v1 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 (a48 t58 t59)) +(anchor :step t61 :args ((:= veriT_vr114 veriT_vr116) (:= veriT_vr115 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 veriT_vr122) (:= ?v1 veriT_vr123) (:= ?v2 veriT_vr124) (:= ?v3 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 (a50 t64 t65)) +(anchor :step t67 :args (veriT_vr122 veriT_vr123 veriT_vr124 veriT_vr125)) +(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 veriT_vr126) (:= veriT_vr123 veriT_vr127) (:= veriT_vr124 veriT_vr128) (:= veriT_vr125 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 veriT_vr130) (:= ?v1 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 (a51 t73 t74)) +(anchor :step t76 :args ((:= veriT_vr130 veriT_vr132) (:= veriT_vr131 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 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 (a27 t80 t83)) +(anchor :step t85 :args ((:= veriT_vr134 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 a25)) +(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 a24)) +(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 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 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 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 a23 t125)) +(step t144 (cl @p_413 @p_414) :rule resolution :premises (t141 a52)) +(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 a26)) +(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)) +3e665c1d341eae599d476b91a009646c76f6e530 323 0 +unsat +(assume a1 (! (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 (:= ?v1 @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)) +(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 ?v3 ?v4)) +(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 (a1 t5 t8)) +(anchor :step t10 :args ((:= ?v0 veriT_vr2) (:= ?v1 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 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 veriT_vr4) (:= ?v3 veriT_vr5) (:= ?v4 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 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 veriT_vr4) (:= ?v3 veriT_vr5) (:= ?v4 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 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 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 veriT_vr4) (:= ?v3 veriT_vr5) (:= ?v4 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 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 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 veriT_vr4) (:= ?v2 veriT_vr5) (:= ?v3 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 veriT_vr3)) +(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)) +(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 veriT_vr3)) +(step t32.t1 (cl (= @p_77 (! (not @p_75) :named @p_209))) :rule implies_simplify) +(anchor :step t32.t2 :args (veriT_vr4)) +(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 veriT_vr5 veriT_vr6)) +(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)) +(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 veriT_vr5 veriT_vr6)) +(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)) +(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)) +(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 veriT_vr5 veriT_vr6)) +(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)) +dcd554acdbd4961913628ae6c525018e8ccd853b 36 0 +unsat +(assume a4 (! (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 a5 (not (! (pred$e 1) :named @p_26))) +(anchor :step t3 :args ((:= ?v0 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 (a4 t3 t4)) +(anchor :step t6 :args (veriT_vr8)) +(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 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 a5)) +cec5afc0b0ddc7f0bb72ad78466b1f47a6cdaa4d 158 0 +unsat +(assume a0 (! (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 a1 (forall ((?v0 Bool)) (= (g$b (some$a ?v0)) (g$c (cons$a ?v0 nil$a))))) +(assume a6 (! (forall ((?v0 Bool_list$)) (! (= (! (g$c ?v0) :named @p_27) (! (size$ ?v0) :named @p_30)) :named @p_32)) :named @p_26)) +(assume a7 (! (forall ((?v0 Int_list$)) (! (= (! (g$a ?v0) :named @p_43) (! (size$a ?v0) :named @p_46)) :named @p_48)) :named @p_42)) +(assume a12 (! (= zero$ (! (size$ nil$a) :named @p_118)) :named @p_143)) +(assume a13 (! (= zero$ (! (size$a nil$) :named @p_133)) :named @p_144)) +(assume a14 (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 a15 (! (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 a16 (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 (a1)) +(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 (a14)) +(anchor :step t12 :args ((:= ?v0 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 (a0 t12 t13)) +(anchor :step t15 :args ((:= veriT_vr0 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 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 (a6 t18 t19)) +(anchor :step t21 :args ((:= veriT_vr2 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 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 (a7 t24 t25)) +(anchor :step t27 :args ((:= veriT_vr4 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 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 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 veriT_vr22) (:= ?v1 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 (a15 t36 t37)) +(anchor :step t39 :args ((:= veriT_vr22 veriT_vr24) (:= veriT_vr23 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 a12 a13 a16 t49 t51 t53 t57 t59)) +339ffce021197fcaff7dca24e8371a4573b5ea1c 873 0 +unsat +(define-fun veriT_sk0 () V$ (! (choice ((veriT_vr65 V$)) (not (! (not (! (= x2$ (! (rraise$ veriT_vr65) :named @p_369)) :named @p_370)) :named @p_368))) :named @p_382)) +(define-fun veriT_sk1 () Abort$ (! (choice ((veriT_vr66 Abort$)) (not (! (not (! (= x2$ (! (rabort$ veriT_vr66) :named @p_372)) :named @p_373)) :named @p_371))) :named @p_386)) +(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_1) (! (pair$ (! (fst$ @p_1) :named @p_346) veriT_vr73) :named @p_429)) :named @p_428)) :named @p_433)) +(define-fun veriT_sk11 () V_list_v_result$ (! (choice ((veriT_vr108 V_list_v_result$)) (! (= (! (fix_clock$ st$a @p_1) :named @p_438) (! (pair$ st$ veriT_vr108) :named @p_471)) :named @p_470)) :named @p_483)) +(define-fun veriT_sk615 () V_list_v_result$ (! (choice ((veriT_vr3122 V_list_v_result$)) (! (= (! (case_error_result$ uua$ uub$ x2$) :named @p_547) (! (pair$ (! (update_clock$ (uu$ st$a @p_346) @p_346) :named @p_504) veriT_vr3122) :named @p_617)) :named @p_616)) :named @p_629)) +(assume a0 (! (forall ((?v0 V$)) (! (= (! (fun_app$ uua$ ?v0) :named @p_6) (! (fun_app$ (! (fun_evaluate_match$ st$ env$ ?v0 pes$) :named @p_9) ?v0) :named @p_11)) :named @p_13)) :named @p_5)) +(assume a1 (! (forall ((?v0 Abort$)) (! (= (! (fun_app$a uub$ ?v0) :named @p_25) (! (pair$ st$ (! (rerr$ (! (rabort$ ?v0) :named @p_28)) :named @p_30)) :named @p_32)) :named @p_34)) :named @p_24)) +(assume a3 (! (= @p_438 (! (pair$ st$ r$) :named @p_545)) :named @p_576)) +(assume a4 (! (less_eq$ (! (clock$ @p_346) :named @p_339) (! (clock$ st$a) :named @p_4)) :named @p_337)) +(assume a5 (! (forall ((?v0 Nat$) (?v1 Nat$) (?v2 Nat$)) (! (=> (! (and (! (less_eq$ ?v0 ?v1) :named @p_47) (! (less_eq$ ?v2 ?v0) :named @p_50)) :named @p_52) (! (less_eq$ ?v2 ?v1) :named @p_56)) :named @p_58)) :named @p_46)) +(assume a6 (! (forall ((?v0 Astate$) (?v1 Astate_v_list_v_result_prod$)) (! (= (! (= ?v0 (! (fst$ ?v1) :named @p_73)) :named @p_75) (! (exists ((?v2 V_list_v_result$)) (! (= ?v1 (! (pair$ ?v0 ?v2) :named @p_81)) :named @p_83)) :named @p_77)) :named @p_85)) :named @p_72)) +(assume a7 (! (forall ((?v0 V_error_result$)) (! (=> (! (and (! (forall ((?v1 V$)) (! (=> (! (= ?v0 (! (rraise$ ?v1) :named @p_140)) :named @p_3) false) :named @p_143)) :named @p_138) (! (forall ((?v1 Abort$)) (! (=> (! (= ?v0 (! (rabort$ ?v1) :named @p_148)) :named @p_150) false) :named @p_152)) :named @p_145)) :named @p_154) false) :named @p_156)) :named @p_137)) +(assume a8 (! (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_183)) :named @p_185) (! (fun_app$ ?v0 ?v2) :named @p_189)) :named @p_191)) :named @p_182)) +(assume a9 (! (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_204)) :named @p_206) (! (fun_app$a ?v1 ?v2) :named @p_210)) :named @p_212)) :named @p_203)) +(assume a10 (! (forall ((?v0 Astate$) (?v1 Astate$) (?v2 V_list_v_result$) (?v3 Astate$)) (! (=> (! (= (! (fix_clock$ ?v0 (! (pair$ ?v1 ?v2) :named @p_225)) :named @p_2) (! (pair$ ?v3 ?v2) :named @p_229)) :named @p_231) (! (less_eq$ (! (clock$ ?v3) :named @p_234) (! (clock$ ?v1) :named @p_237)) :named @p_239)) :named @p_241)) :named @p_224)) +(assume a11 (! (forall ((?v0 Astate$) (?v1 Astate$) (?v2 V_list_v_result$)) (! (= @p_2 (! (pair$ (! (update_clock$ (! (uu$ ?v0 ?v1) :named @p_263) ?v1) :named @p_265) ?v2) :named @p_268)) :named @p_270)) :named @p_258)) +(assume a12 (! (forall ((?v0 V_error_result$) (?v1 V$)) (! (=> (! (and (! (= r$ (! (rerr$ ?v0) :named @p_287)) :named @p_289) @p_3) :named @p_294) (! (less_eq$ (! (clock$ (! (fst$ (! (fun_app$ (! (fun_evaluate_match$ st$ env$ ?v1 pes$) :named @p_297) ?v1) :named @p_299)) :named @p_301)) :named @p_303) (! (clock$ st$) :named @p_286)) :named @p_305)) :named @p_307)) :named @p_285)) +(assume a13 (! (not (! (=> (! (= r$ (! (rerr$ x2$) :named @p_562)) :named @p_327) (! (less_eq$ (! (clock$ (! (fst$ @p_547) :named @p_565)) :named @p_338) @p_4) :named @p_328)) :named @p_332)) :named @p_326)) +(anchor :step t14 :args ((:= ?v0 veriT_vr0))) +(step t14.t1 (cl (! (= ?v0 veriT_vr0) :named @p_8)) :rule refl) +(step t14.t2 (cl (= @p_6 (! (fun_app$ uua$ veriT_vr0) :named @p_7))) :rule cong :premises (t14.t1)) +(step t14.t3 (cl @p_8) :rule refl) +(step t14.t4 (cl (= @p_9 (! (fun_evaluate_match$ st$ env$ veriT_vr0 pes$) :named @p_10))) :rule cong :premises (t14.t3)) +(step t14.t5 (cl @p_8) :rule refl) +(step t14.t6 (cl (= @p_11 (! (fun_app$ @p_10 veriT_vr0) :named @p_12))) :rule cong :premises (t14.t4 t14.t5)) +(step t14.t7 (cl (= @p_13 (! (= @p_7 @p_12) :named @p_14))) :rule cong :premises (t14.t2 t14.t6)) +(step t14 (cl (! (= @p_5 (! (forall ((veriT_vr0 V$)) @p_14) :named @p_16)) :named @p_15)) :rule bind) +(step t15 (cl (not @p_15) (not @p_5) @p_16) :rule equiv_pos2) +(step t16 (cl @p_16) :rule th_resolution :premises (a0 t14 t15)) +(anchor :step t17 :args ((:= veriT_vr0 veriT_vr1))) +(step t17.t1 (cl (! (= veriT_vr0 veriT_vr1) :named @p_18)) :rule refl) +(step t17.t2 (cl (= @p_7 (! (fun_app$ uua$ veriT_vr1) :named @p_17))) :rule cong :premises (t17.t1)) +(step t17.t3 (cl @p_18) :rule refl) +(step t17.t4 (cl (= @p_10 (! (fun_evaluate_match$ st$ env$ veriT_vr1 pes$) :named @p_19))) :rule cong :premises (t17.t3)) +(step t17.t5 (cl @p_18) :rule refl) +(step t17.t6 (cl (= @p_12 (! (fun_app$ @p_19 veriT_vr1) :named @p_20))) :rule cong :premises (t17.t4 t17.t5)) +(step t17.t7 (cl (= @p_14 (! (= @p_17 @p_20) :named @p_21))) :rule cong :premises (t17.t2 t17.t6)) +(step t17 (cl (! (= @p_16 (! (forall ((veriT_vr1 V$)) @p_21) :named @p_23)) :named @p_22)) :rule bind) +(step t18 (cl (not @p_22) (not @p_16) @p_23) :rule equiv_pos2) +(step t19 (cl @p_23) :rule th_resolution :premises (t16 t17 t18)) +(anchor :step t20 :args ((:= ?v0 veriT_vr2))) +(step t20.t1 (cl (! (= ?v0 veriT_vr2) :named @p_27)) :rule refl) +(step t20.t2 (cl (= @p_25 (! (fun_app$a uub$ veriT_vr2) :named @p_26))) :rule cong :premises (t20.t1)) +(step t20.t3 (cl @p_27) :rule refl) +(step t20.t4 (cl (= @p_28 (! (rabort$ veriT_vr2) :named @p_29))) :rule cong :premises (t20.t3)) +(step t20.t5 (cl (= @p_30 (! (rerr$ @p_29) :named @p_31))) :rule cong :premises (t20.t4)) +(step t20.t6 (cl (= @p_32 (! (pair$ st$ @p_31) :named @p_33))) :rule cong :premises (t20.t5)) +(step t20.t7 (cl (= @p_34 (! (= @p_26 @p_33) :named @p_35))) :rule cong :premises (t20.t2 t20.t6)) +(step t20 (cl (! (= @p_24 (! (forall ((veriT_vr2 Abort$)) @p_35) :named @p_37)) :named @p_36)) :rule bind) +(step t21 (cl (not @p_36) (not @p_24) @p_37) :rule equiv_pos2) +(step t22 (cl @p_37) :rule th_resolution :premises (a1 t20 t21)) +(anchor :step t23 :args ((:= veriT_vr2 veriT_vr3))) +(step t23.t1 (cl (! (= veriT_vr2 veriT_vr3) :named @p_39)) :rule refl) +(step t23.t2 (cl (= @p_26 (! (fun_app$a uub$ veriT_vr3) :named @p_38))) :rule cong :premises (t23.t1)) +(step t23.t3 (cl @p_39) :rule refl) +(step t23.t4 (cl (= @p_29 (! (rabort$ veriT_vr3) :named @p_40))) :rule cong :premises (t23.t3)) +(step t23.t5 (cl (= @p_31 (! (rerr$ @p_40) :named @p_41))) :rule cong :premises (t23.t4)) +(step t23.t6 (cl (= @p_33 (! (pair$ st$ @p_41) :named @p_42))) :rule cong :premises (t23.t5)) +(step t23.t7 (cl (= @p_35 (! (= @p_38 @p_42) :named @p_43))) :rule cong :premises (t23.t2 t23.t6)) +(step t23 (cl (! (= @p_37 (! (forall ((veriT_vr3 Abort$)) @p_43) :named @p_45)) :named @p_44)) :rule bind) +(step t24 (cl (not @p_44) (not @p_37) @p_45) :rule equiv_pos2) +(step t25 (cl @p_45) :rule th_resolution :premises (t22 t23 t24)) +(anchor :step t26 :args ((:= ?v0 veriT_vr10) (:= ?v1 veriT_vr11) (:= ?v2 veriT_vr12))) +(step t26.t1 (cl (! (= ?v0 veriT_vr10) :named @p_49)) :rule refl) +(step t26.t2 (cl (! (= ?v1 veriT_vr11) :named @p_55)) :rule refl) +(step t26.t3 (cl (= @p_47 (! (less_eq$ veriT_vr10 veriT_vr11) :named @p_48))) :rule cong :premises (t26.t1 t26.t2)) +(step t26.t4 (cl (! (= ?v2 veriT_vr12) :named @p_54)) :rule refl) +(step t26.t5 (cl @p_49) :rule refl) +(step t26.t6 (cl (= @p_50 (! (less_eq$ veriT_vr12 veriT_vr10) :named @p_51))) :rule cong :premises (t26.t4 t26.t5)) +(step t26.t7 (cl (= @p_52 (! (and @p_48 @p_51) :named @p_53))) :rule cong :premises (t26.t3 t26.t6)) +(step t26.t8 (cl @p_54) :rule refl) +(step t26.t9 (cl @p_55) :rule refl) +(step t26.t10 (cl (= @p_56 (! (less_eq$ veriT_vr12 veriT_vr11) :named @p_57))) :rule cong :premises (t26.t8 t26.t9)) +(step t26.t11 (cl (= @p_58 (! (=> @p_53 @p_57) :named @p_59))) :rule cong :premises (t26.t7 t26.t10)) +(step t26 (cl (! (= @p_46 (! (forall ((veriT_vr10 Nat$) (veriT_vr11 Nat$) (veriT_vr12 Nat$)) @p_59) :named @p_61)) :named @p_60)) :rule bind) +(step t27 (cl (not @p_60) (not @p_46) @p_61) :rule equiv_pos2) +(step t28 (cl @p_61) :rule th_resolution :premises (a5 t26 t27)) +(anchor :step t29 :args ((:= veriT_vr10 veriT_vr13) (:= veriT_vr11 veriT_vr14) (:= veriT_vr12 veriT_vr15))) +(step t29.t1 (cl (! (= veriT_vr10 veriT_vr13) :named @p_63)) :rule refl) +(step t29.t2 (cl (! (= veriT_vr11 veriT_vr14) :named @p_67)) :rule refl) +(step t29.t3 (cl (= @p_48 (! (less_eq$ veriT_vr13 veriT_vr14) :named @p_62))) :rule cong :premises (t29.t1 t29.t2)) +(step t29.t4 (cl (! (= veriT_vr12 veriT_vr15) :named @p_66)) :rule refl) +(step t29.t5 (cl @p_63) :rule refl) +(step t29.t6 (cl (= @p_51 (! (less_eq$ veriT_vr15 veriT_vr13) :named @p_64))) :rule cong :premises (t29.t4 t29.t5)) +(step t29.t7 (cl (= @p_53 (! (and @p_62 @p_64) :named @p_65))) :rule cong :premises (t29.t3 t29.t6)) +(step t29.t8 (cl @p_66) :rule refl) +(step t29.t9 (cl @p_67) :rule refl) +(step t29.t10 (cl (= @p_57 (! (less_eq$ veriT_vr15 veriT_vr14) :named @p_68))) :rule cong :premises (t29.t8 t29.t9)) +(step t29.t11 (cl (= @p_59 (! (=> @p_65 @p_68) :named @p_69))) :rule cong :premises (t29.t7 t29.t10)) +(step t29 (cl (! (= @p_61 (! (forall ((veriT_vr13 Nat$) (veriT_vr14 Nat$) (veriT_vr15 Nat$)) @p_69) :named @p_71)) :named @p_70)) :rule bind) +(step t30 (cl (not @p_70) (not @p_61) @p_71) :rule equiv_pos2) +(step t31 (cl @p_71) :rule th_resolution :premises (t28 t29 t30)) +(anchor :step t32 :args ((:= ?v0 veriT_vr16) (:= ?v1 veriT_vr17))) +(step t32.t1 (cl (! (= ?v0 veriT_vr16) :named @p_80)) :rule refl) +(step t32.t2 (cl (! (= ?v1 veriT_vr17) :named @p_79)) :rule refl) +(step t32.t3 (cl (= @p_73 (! (fst$ veriT_vr17) :named @p_74))) :rule cong :premises (t32.t2)) +(step t32.t4 (cl (= @p_75 (! (= veriT_vr16 @p_74) :named @p_76))) :rule cong :premises (t32.t1 t32.t3)) +(anchor :step t32.t5 :args ((:= ?v2 veriT_vr18))) +(step t32.t5.t1 (cl @p_79) :rule refl) +(step t32.t5.t2 (cl @p_80) :rule refl) +(step t32.t5.t3 (cl (= ?v2 veriT_vr18)) :rule refl) +(step t32.t5.t4 (cl (= @p_81 (! (pair$ veriT_vr16 veriT_vr18) :named @p_82))) :rule cong :premises (t32.t5.t2 t32.t5.t3)) +(step t32.t5.t5 (cl (= @p_83 (! (= veriT_vr17 @p_82) :named @p_84))) :rule cong :premises (t32.t5.t1 t32.t5.t4)) +(step t32.t5 (cl (= @p_77 (! (exists ((veriT_vr18 V_list_v_result$)) @p_84) :named @p_78))) :rule bind) +(step t32.t6 (cl (= @p_85 (! (= @p_76 @p_78) :named @p_86))) :rule cong :premises (t32.t4 t32.t5)) +(step t32 (cl (! (= @p_72 (! (forall ((veriT_vr16 Astate$) (veriT_vr17 Astate_v_list_v_result_prod$)) @p_86) :named @p_88)) :named @p_87)) :rule bind) +(step t33 (cl (not @p_87) (not @p_72) @p_88) :rule equiv_pos2) +(step t34 (cl @p_88) :rule th_resolution :premises (a6 t32 t33)) +(anchor :step t35 :args (veriT_vr16 veriT_vr17)) +(step t35.t1 (cl (= @p_86 (! (and (! (=> @p_76 @p_78) :named @p_99) (! (=> @p_78 @p_76) :named @p_106)) :named @p_89))) :rule connective_def) +(step t35 (cl (! (= @p_88 (! (forall ((veriT_vr16 Astate$) (veriT_vr17 Astate_v_list_v_result_prod$)) @p_89) :named @p_91)) :named @p_90)) :rule bind) +(step t36 (cl (not @p_90) (not @p_88) @p_91) :rule equiv_pos2) +(step t37 (cl @p_91) :rule th_resolution :premises (t34 t35 t36)) +(anchor :step t38 :args ((:= veriT_vr16 veriT_vr19) (:= veriT_vr17 veriT_vr20))) +(step t38.t1 (cl (! (= veriT_vr16 veriT_vr19) :named @p_96)) :rule refl) +(step t38.t2 (cl (! (= veriT_vr17 veriT_vr20) :named @p_95)) :rule refl) +(step t38.t3 (cl (! (= @p_74 (! (fst$ veriT_vr20) :named @p_93)) :named @p_104)) :rule cong :premises (t38.t2)) +(step t38.t4 (cl (! (= @p_76 (! (= veriT_vr19 @p_93) :named @p_92)) :named @p_105)) :rule cong :premises (t38.t1 t38.t3)) +(anchor :step t38.t5 :args ((:= veriT_vr18 veriT_vr21))) +(step t38.t5.t1 (cl @p_95) :rule refl) +(step t38.t5.t2 (cl @p_96) :rule refl) +(step t38.t5.t3 (cl (= veriT_vr18 veriT_vr21)) :rule refl) +(step t38.t5.t4 (cl (= @p_82 (! (pair$ veriT_vr19 veriT_vr21) :named @p_97))) :rule cong :premises (t38.t5.t2 t38.t5.t3)) +(step t38.t5.t5 (cl (= @p_84 (! (= veriT_vr20 @p_97) :named @p_98))) :rule cong :premises (t38.t5.t1 t38.t5.t4)) +(step t38.t5 (cl (= @p_78 (! (exists ((veriT_vr21 V_list_v_result$)) @p_98) :named @p_94))) :rule bind) +(step t38.t6 (cl (= @p_99 (! (=> @p_92 @p_94) :named @p_100))) :rule cong :premises (t38.t4 t38.t5)) +(anchor :step t38.t7 :args ((:= veriT_vr18 veriT_vr22))) +(step t38.t7.t1 (cl @p_95) :rule refl) +(step t38.t7.t2 (cl @p_96) :rule refl) +(step t38.t7.t3 (cl (= veriT_vr18 veriT_vr22)) :rule refl) +(step t38.t7.t4 (cl (= @p_82 (! (pair$ veriT_vr19 veriT_vr22) :named @p_102))) :rule cong :premises (t38.t7.t2 t38.t7.t3)) +(step t38.t7.t5 (cl (= @p_84 (! (= veriT_vr20 @p_102) :named @p_103))) :rule cong :premises (t38.t7.t1 t38.t7.t4)) +(step t38.t7 (cl (= @p_78 (! (exists ((veriT_vr22 V_list_v_result$)) @p_103) :named @p_101))) :rule bind) +(step t38.t8 (cl @p_96) :rule refl) +(step t38.t9 (cl @p_95) :rule refl) +(step t38.t10 (cl @p_104) :rule cong :premises (t38.t9)) +(step t38.t11 (cl @p_105) :rule cong :premises (t38.t8 t38.t10)) +(step t38.t12 (cl (= @p_106 (! (=> @p_101 @p_92) :named @p_107))) :rule cong :premises (t38.t7 t38.t11)) +(step t38.t13 (cl (= @p_89 (! (and @p_100 @p_107) :named @p_108))) :rule cong :premises (t38.t6 t38.t12)) +(step t38 (cl (! (= @p_91 (! (forall ((veriT_vr19 Astate$) (veriT_vr20 Astate_v_list_v_result_prod$)) @p_108) :named @p_110)) :named @p_109)) :rule bind) +(step t39 (cl (not @p_109) (not @p_91) @p_110) :rule equiv_pos2) +(step t40 (cl @p_110) :rule th_resolution :premises (t37 t38 t39)) +(anchor :step t41 :args ((:= veriT_vr19 veriT_vr23) (:= veriT_vr20 veriT_vr24))) +(step t41.t1 (cl (! (= veriT_vr19 veriT_vr23) :named @p_115)) :rule refl) +(step t41.t2 (cl (! (= veriT_vr20 veriT_vr24) :named @p_114)) :rule refl) +(step t41.t3 (cl (! (= @p_93 (! (fst$ veriT_vr24) :named @p_113)) :named @p_119)) :rule cong :premises (t41.t2)) +(step t41.t4 (cl (! (= @p_92 (! (= veriT_vr23 @p_113) :named @p_112)) :named @p_120)) :rule cong :premises (t41.t1 t41.t3)) +(anchor :step t41.t5 :args ((:= veriT_vr21 veriT_vr25))) +(step t41.t5.t1 (cl @p_114) :rule refl) +(step t41.t5.t2 (cl @p_115) :rule refl) +(step t41.t5.t3 (cl (= veriT_vr21 veriT_vr25)) :rule refl) +(step t41.t5.t4 (cl (= @p_97 (! (pair$ veriT_vr23 veriT_vr25) :named @p_116))) :rule cong :premises (t41.t5.t2 t41.t5.t3)) +(step t41.t5.t5 (cl (= @p_98 (! (= veriT_vr24 @p_116) :named @p_117))) :rule cong :premises (t41.t5.t1 t41.t5.t4)) +(step t41.t5 (cl (= @p_94 (! (exists ((veriT_vr25 V_list_v_result$)) @p_117) :named @p_111))) :rule bind) +(step t41.t6 (cl (= @p_100 (! (=> @p_112 @p_111) :named @p_118))) :rule cong :premises (t41.t4 t41.t5)) +(anchor :step t41.t7 :args ((:= veriT_vr22 veriT_vr25))) +(step t41.t7.t1 (cl @p_114) :rule refl) +(step t41.t7.t2 (cl @p_115) :rule refl) +(step t41.t7.t3 (cl (= veriT_vr22 veriT_vr25)) :rule refl) +(step t41.t7.t4 (cl (= @p_102 @p_116)) :rule cong :premises (t41.t7.t2 t41.t7.t3)) +(step t41.t7.t5 (cl (= @p_103 @p_117)) :rule cong :premises (t41.t7.t1 t41.t7.t4)) +(step t41.t7 (cl (= @p_101 @p_111)) :rule bind) +(step t41.t8 (cl @p_115) :rule refl) +(step t41.t9 (cl @p_114) :rule refl) +(step t41.t10 (cl @p_119) :rule cong :premises (t41.t9)) +(step t41.t11 (cl @p_120) :rule cong :premises (t41.t8 t41.t10)) +(step t41.t12 (cl (= @p_107 (! (=> @p_111 @p_112) :named @p_121))) :rule cong :premises (t41.t7 t41.t11)) +(step t41.t13 (cl (= @p_108 (! (and @p_118 @p_121) :named @p_122))) :rule cong :premises (t41.t6 t41.t12)) +(step t41 (cl (! (= @p_110 (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$)) @p_122) :named @p_124)) :named @p_123)) :rule bind) +(step t42 (cl (not @p_123) (not @p_110) @p_124) :rule equiv_pos2) +(step t43 (cl @p_124) :rule th_resolution :premises (t40 t41 t42)) +(anchor :step t44 :args ((:= veriT_vr23 veriT_vr23) (:= veriT_vr24 veriT_vr24))) +(anchor :step t44.t1 :args ((:= veriT_vr25 veriT_vr26))) +(step t44.t1.t1 (cl (= veriT_vr25 veriT_vr26)) :rule refl) +(step t44.t1.t2 (cl (= @p_116 (! (pair$ veriT_vr23 veriT_vr26) :named @p_126))) :rule cong :premises (t44.t1.t1)) +(step t44.t1.t3 (cl (= @p_117 (! (= veriT_vr24 @p_126) :named @p_127))) :rule cong :premises (t44.t1.t2)) +(step t44.t1 (cl (= @p_111 (! (exists ((veriT_vr26 V_list_v_result$)) @p_127) :named @p_125))) :rule bind) +(step t44.t2 (cl (= @p_121 (! (=> @p_125 @p_112) :named @p_128))) :rule cong :premises (t44.t1)) +(step t44.t3 (cl (= @p_122 (! (and @p_118 @p_128) :named @p_129))) :rule cong :premises (t44.t2)) +(step t44 (cl (! (= @p_124 (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$)) @p_129) :named @p_131)) :named @p_130)) :rule bind) +(step t45 (cl (not @p_130) (not @p_124) @p_131) :rule equiv_pos2) +(step t46 (cl @p_131) :rule th_resolution :premises (t43 t44 t45)) +(anchor :step t47 :args (veriT_vr23 veriT_vr24)) +(step t47.t1 (cl (= @p_125 (! (not (forall ((veriT_vr26 V_list_v_result$)) (! (not @p_127) :named @p_335))) :named @p_132))) :rule connective_def) +(step t47.t2 (cl (= @p_128 (! (=> @p_132 @p_112) :named @p_133))) :rule cong :premises (t47.t1)) +(step t47.t3 (cl (= @p_129 (! (and @p_118 @p_133) :named @p_134))) :rule cong :premises (t47.t2)) +(step t47 (cl (! (= @p_131 (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$)) @p_134) :named @p_136)) :named @p_135)) :rule bind) +(step t48 (cl (not @p_135) (not @p_131) @p_136) :rule equiv_pos2) +(step t49 (cl @p_136) :rule th_resolution :premises (t46 t47 t48)) +(anchor :step t50 :args ((:= ?v0 veriT_vr27))) +(anchor :step t50.t1 :args ((:= ?v1 veriT_vr28))) +(step t50.t1.t1 (cl (! (= ?v0 veriT_vr27) :named @p_147)) :rule refl) +(step t50.t1.t2 (cl (= ?v1 veriT_vr28)) :rule refl) +(step t50.t1.t3 (cl (= @p_140 (! (rraise$ veriT_vr28) :named @p_141))) :rule cong :premises (t50.t1.t2)) +(step t50.t1.t4 (cl (= @p_3 (! (= veriT_vr27 @p_141) :named @p_142))) :rule cong :premises (t50.t1.t1 t50.t1.t3)) +(step t50.t1.t5 (cl (= @p_143 (! (=> @p_142 false) :named @p_144))) :rule cong :premises (t50.t1.t4)) +(step t50.t1 (cl (= @p_138 (! (forall ((veriT_vr28 V$)) @p_144) :named @p_139))) :rule bind) +(anchor :step t50.t2 :args ((:= ?v1 veriT_vr29))) +(step t50.t2.t1 (cl @p_147) :rule refl) +(step t50.t2.t2 (cl (= ?v1 veriT_vr29)) :rule refl) +(step t50.t2.t3 (cl (= @p_148 (! (rabort$ veriT_vr29) :named @p_149))) :rule cong :premises (t50.t2.t2)) +(step t50.t2.t4 (cl (= @p_150 (! (= veriT_vr27 @p_149) :named @p_151))) :rule cong :premises (t50.t2.t1 t50.t2.t3)) +(step t50.t2.t5 (cl (= @p_152 (! (=> @p_151 false) :named @p_153))) :rule cong :premises (t50.t2.t4)) +(step t50.t2 (cl (= @p_145 (! (forall ((veriT_vr29 Abort$)) @p_153) :named @p_146))) :rule bind) +(step t50.t3 (cl (= @p_154 (! (and @p_139 @p_146) :named @p_155))) :rule cong :premises (t50.t1 t50.t2)) +(step t50.t4 (cl (= @p_156 (! (=> @p_155 false) :named @p_157))) :rule cong :premises (t50.t3)) +(step t50 (cl (! (= @p_137 (! (forall ((veriT_vr27 V_error_result$)) @p_157) :named @p_159)) :named @p_158)) :rule bind) +(step t51 (cl (not @p_158) (not @p_137) @p_159) :rule equiv_pos2) +(step t52 (cl @p_159) :rule th_resolution :premises (a7 t50 t51)) +(anchor :step t53 :args (veriT_vr27)) +(anchor :step t53.t1 :args (veriT_vr28)) +(step t53.t1.t1 (cl (= @p_144 (! (not @p_142) :named @p_161))) :rule implies_simplify) +(step t53.t1 (cl (= @p_139 (! (forall ((veriT_vr28 V$)) @p_161) :named @p_160))) :rule bind) +(anchor :step t53.t2 :args (veriT_vr29)) +(step t53.t2.t1 (cl (= @p_153 (! (not @p_151) :named @p_163))) :rule implies_simplify) +(step t53.t2 (cl (= @p_146 (! (forall ((veriT_vr29 Abort$)) @p_163) :named @p_162))) :rule bind) +(step t53.t3 (cl (= @p_155 (! (and @p_160 @p_162) :named @p_164))) :rule cong :premises (t53.t1 t53.t2)) +(step t53.t4 (cl (= @p_157 (! (=> @p_164 false) :named @p_165))) :rule cong :premises (t53.t3)) +(step t53.t5 (cl (= @p_165 (! (not @p_164) :named @p_166))) :rule implies_simplify) +(step t53.t6 (cl (= @p_157 @p_166)) :rule trans :premises (t53.t4 t53.t5)) +(step t53 (cl (! (= @p_159 (! (forall ((veriT_vr27 V_error_result$)) @p_166) :named @p_168)) :named @p_167)) :rule bind) +(step t54 (cl (not @p_167) (not @p_159) @p_168) :rule equiv_pos2) +(step t55 (cl @p_168) :rule th_resolution :premises (t52 t53 t54)) +(anchor :step t56 :args ((:= veriT_vr27 veriT_vr30))) +(anchor :step t56.t1 :args ((:= veriT_vr28 veriT_vr31))) +(step t56.t1.t1 (cl (! (= veriT_vr27 veriT_vr30) :named @p_174)) :rule refl) +(step t56.t1.t2 (cl (= veriT_vr28 veriT_vr31)) :rule refl) +(step t56.t1.t3 (cl (= @p_141 (! (rraise$ veriT_vr31) :named @p_170))) :rule cong :premises (t56.t1.t2)) +(step t56.t1.t4 (cl (= @p_142 (! (= veriT_vr30 @p_170) :named @p_171))) :rule cong :premises (t56.t1.t1 t56.t1.t3)) +(step t56.t1.t5 (cl (= @p_161 (! (not @p_171) :named @p_172))) :rule cong :premises (t56.t1.t4)) +(step t56.t1 (cl (= @p_160 (! (forall ((veriT_vr31 V$)) @p_172) :named @p_169))) :rule bind) +(anchor :step t56.t2 :args ((:= veriT_vr29 veriT_vr32))) +(step t56.t2.t1 (cl @p_174) :rule refl) +(step t56.t2.t2 (cl (= veriT_vr29 veriT_vr32)) :rule refl) +(step t56.t2.t3 (cl (= @p_149 (! (rabort$ veriT_vr32) :named @p_175))) :rule cong :premises (t56.t2.t2)) +(step t56.t2.t4 (cl (= @p_151 (! (= veriT_vr30 @p_175) :named @p_176))) :rule cong :premises (t56.t2.t1 t56.t2.t3)) +(step t56.t2.t5 (cl (= @p_163 (! (not @p_176) :named @p_177))) :rule cong :premises (t56.t2.t4)) +(step t56.t2 (cl (= @p_162 (! (forall ((veriT_vr32 Abort$)) @p_177) :named @p_173))) :rule bind) +(step t56.t3 (cl (= @p_164 (! (and @p_169 @p_173) :named @p_178))) :rule cong :premises (t56.t1 t56.t2)) +(step t56.t4 (cl (= @p_166 (! (not @p_178) :named @p_179))) :rule cong :premises (t56.t3)) +(step t56 (cl (! (= @p_168 (! (forall ((veriT_vr30 V_error_result$)) @p_179) :named @p_181)) :named @p_180)) :rule bind) +(step t57 (cl (not @p_180) (not @p_168) @p_181) :rule equiv_pos2) +(step t58 (cl @p_181) :rule th_resolution :premises (t55 t56 t57)) +(anchor :step t59 :args ((:= ?v0 veriT_vr33) (:= ?v1 veriT_vr34) (:= ?v2 veriT_vr35))) +(step t59.t1 (cl (! (= ?v0 veriT_vr33) :named @p_187)) :rule refl) +(step t59.t2 (cl (= ?v1 veriT_vr34)) :rule refl) +(step t59.t3 (cl (! (= ?v2 veriT_vr35) :named @p_188)) :rule refl) +(step t59.t4 (cl (= @p_183 (! (rraise$ veriT_vr35) :named @p_184))) :rule cong :premises (t59.t3)) +(step t59.t5 (cl (= @p_185 (! (case_error_result$ veriT_vr33 veriT_vr34 @p_184) :named @p_186))) :rule cong :premises (t59.t1 t59.t2 t59.t4)) +(step t59.t6 (cl @p_187) :rule refl) +(step t59.t7 (cl @p_188) :rule refl) +(step t59.t8 (cl (= @p_189 (! (fun_app$ veriT_vr33 veriT_vr35) :named @p_190))) :rule cong :premises (t59.t6 t59.t7)) +(step t59.t9 (cl (= @p_191 (! (= @p_186 @p_190) :named @p_192))) :rule cong :premises (t59.t5 t59.t8)) +(step t59 (cl (! (= @p_182 (! (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_192) :named @p_194)) :named @p_193)) :rule bind) +(step t60 (cl (not @p_193) (not @p_182) @p_194) :rule equiv_pos2) +(step t61 (cl @p_194) :rule th_resolution :premises (a8 t59 t60)) +(anchor :step t62 :args ((:= veriT_vr33 veriT_vr36) (:= veriT_vr34 veriT_vr37) (:= veriT_vr35 veriT_vr38))) +(step t62.t1 (cl (! (= veriT_vr33 veriT_vr36) :named @p_197)) :rule refl) +(step t62.t2 (cl (= veriT_vr34 veriT_vr37)) :rule refl) +(step t62.t3 (cl (! (= veriT_vr35 veriT_vr38) :named @p_198)) :rule refl) +(step t62.t4 (cl (= @p_184 (! (rraise$ veriT_vr38) :named @p_195))) :rule cong :premises (t62.t3)) +(step t62.t5 (cl (= @p_186 (! (case_error_result$ veriT_vr36 veriT_vr37 @p_195) :named @p_196))) :rule cong :premises (t62.t1 t62.t2 t62.t4)) +(step t62.t6 (cl @p_197) :rule refl) +(step t62.t7 (cl @p_198) :rule refl) +(step t62.t8 (cl (= @p_190 (! (fun_app$ veriT_vr36 veriT_vr38) :named @p_199))) :rule cong :premises (t62.t6 t62.t7)) +(step t62.t9 (cl (= @p_192 (! (= @p_196 @p_199) :named @p_200))) :rule cong :premises (t62.t5 t62.t8)) +(step t62 (cl (! (= @p_194 (! (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_200) :named @p_202)) :named @p_201)) :rule bind) +(step t63 (cl (not @p_201) (not @p_194) @p_202) :rule equiv_pos2) +(step t64 (cl @p_202) :rule th_resolution :premises (t61 t62 t63)) +(anchor :step t65 :args ((:= ?v0 veriT_vr39) (:= ?v1 veriT_vr40) (:= ?v2 veriT_vr41))) +(step t65.t1 (cl (= ?v0 veriT_vr39)) :rule refl) +(step t65.t2 (cl (! (= ?v1 veriT_vr40) :named @p_208)) :rule refl) +(step t65.t3 (cl (! (= ?v2 veriT_vr41) :named @p_209)) :rule refl) +(step t65.t4 (cl (= @p_204 (! (rabort$ veriT_vr41) :named @p_205))) :rule cong :premises (t65.t3)) +(step t65.t5 (cl (= @p_206 (! (case_error_result$ veriT_vr39 veriT_vr40 @p_205) :named @p_207))) :rule cong :premises (t65.t1 t65.t2 t65.t4)) +(step t65.t6 (cl @p_208) :rule refl) +(step t65.t7 (cl @p_209) :rule refl) +(step t65.t8 (cl (= @p_210 (! (fun_app$a veriT_vr40 veriT_vr41) :named @p_211))) :rule cong :premises (t65.t6 t65.t7)) +(step t65.t9 (cl (= @p_212 (! (= @p_207 @p_211) :named @p_213))) :rule cong :premises (t65.t5 t65.t8)) +(step t65 (cl (! (= @p_203 (! (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_213) :named @p_215)) :named @p_214)) :rule bind) +(step t66 (cl (not @p_214) (not @p_203) @p_215) :rule equiv_pos2) +(step t67 (cl @p_215) :rule th_resolution :premises (a9 t65 t66)) +(anchor :step t68 :args ((:= veriT_vr39 veriT_vr42) (:= veriT_vr40 veriT_vr43) (:= veriT_vr41 veriT_vr44))) +(step t68.t1 (cl (= veriT_vr39 veriT_vr42)) :rule refl) +(step t68.t2 (cl (! (= veriT_vr40 veriT_vr43) :named @p_218)) :rule refl) +(step t68.t3 (cl (! (= veriT_vr41 veriT_vr44) :named @p_219)) :rule refl) +(step t68.t4 (cl (= @p_205 (! (rabort$ veriT_vr44) :named @p_216))) :rule cong :premises (t68.t3)) +(step t68.t5 (cl (= @p_207 (! (case_error_result$ veriT_vr42 veriT_vr43 @p_216) :named @p_217))) :rule cong :premises (t68.t1 t68.t2 t68.t4)) +(step t68.t6 (cl @p_218) :rule refl) +(step t68.t7 (cl @p_219) :rule refl) +(step t68.t8 (cl (= @p_211 (! (fun_app$a veriT_vr43 veriT_vr44) :named @p_220))) :rule cong :premises (t68.t6 t68.t7)) +(step t68.t9 (cl (= @p_213 (! (= @p_217 @p_220) :named @p_221))) :rule cong :premises (t68.t5 t68.t8)) +(step t68 (cl (! (= @p_215 (! (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_221) :named @p_223)) :named @p_222)) :rule bind) +(step t69 (cl (not @p_222) (not @p_215) @p_223) :rule equiv_pos2) +(step t70 (cl @p_223) :rule th_resolution :premises (t67 t68 t69)) +(anchor :step t71 :args ((:= ?v0 veriT_vr45) (:= ?v1 veriT_vr46) (:= ?v2 veriT_vr47) (:= ?v3 veriT_vr48))) +(step t71.t1 (cl (= ?v0 veriT_vr45)) :rule refl) +(step t71.t2 (cl (! (= ?v1 veriT_vr46) :named @p_236)) :rule refl) +(step t71.t3 (cl (! (= ?v2 veriT_vr47) :named @p_228)) :rule refl) +(step t71.t4 (cl (= @p_225 (! (pair$ veriT_vr46 veriT_vr47) :named @p_226))) :rule cong :premises (t71.t2 t71.t3)) +(step t71.t5 (cl (= @p_2 (! (fix_clock$ veriT_vr45 @p_226) :named @p_227))) :rule cong :premises (t71.t1 t71.t4)) +(step t71.t6 (cl (! (= ?v3 veriT_vr48) :named @p_233)) :rule refl) +(step t71.t7 (cl @p_228) :rule refl) +(step t71.t8 (cl (= @p_229 (! (pair$ veriT_vr48 veriT_vr47) :named @p_230))) :rule cong :premises (t71.t6 t71.t7)) +(step t71.t9 (cl (= @p_231 (! (= @p_227 @p_230) :named @p_232))) :rule cong :premises (t71.t5 t71.t8)) +(step t71.t10 (cl @p_233) :rule refl) +(step t71.t11 (cl (= @p_234 (! (clock$ veriT_vr48) :named @p_235))) :rule cong :premises (t71.t10)) +(step t71.t12 (cl @p_236) :rule refl) +(step t71.t13 (cl (= @p_237 (! (clock$ veriT_vr46) :named @p_238))) :rule cong :premises (t71.t12)) +(step t71.t14 (cl (= @p_239 (! (less_eq$ @p_235 @p_238) :named @p_240))) :rule cong :premises (t71.t11 t71.t13)) +(step t71.t15 (cl (= @p_241 (! (=> @p_232 @p_240) :named @p_242))) :rule cong :premises (t71.t9 t71.t14)) +(step t71 (cl (! (= @p_224 (! (forall ((veriT_vr45 Astate$) (veriT_vr46 Astate$) (veriT_vr47 V_list_v_result$) (veriT_vr48 Astate$)) @p_242) :named @p_244)) :named @p_243)) :rule bind) +(step t72 (cl (not @p_243) (not @p_224) @p_244) :rule equiv_pos2) +(step t73 (cl @p_244) :rule th_resolution :premises (a10 t71 t72)) +(anchor :step t74 :args ((:= veriT_vr45 veriT_vr49) (:= veriT_vr46 veriT_vr50) (:= veriT_vr47 veriT_vr51) (:= veriT_vr48 veriT_vr52))) +(step t74.t1 (cl (= veriT_vr45 veriT_vr49)) :rule refl) +(step t74.t2 (cl (! (= veriT_vr46 veriT_vr50) :named @p_252)) :rule refl) +(step t74.t3 (cl (! (= veriT_vr47 veriT_vr51) :named @p_247)) :rule refl) +(step t74.t4 (cl (= @p_226 (! (pair$ veriT_vr50 veriT_vr51) :named @p_245))) :rule cong :premises (t74.t2 t74.t3)) +(step t74.t5 (cl (= @p_227 (! (fix_clock$ veriT_vr49 @p_245) :named @p_246))) :rule cong :premises (t74.t1 t74.t4)) +(step t74.t6 (cl (! (= veriT_vr48 veriT_vr52) :named @p_250)) :rule refl) +(step t74.t7 (cl @p_247) :rule refl) +(step t74.t8 (cl (= @p_230 (! (pair$ veriT_vr52 veriT_vr51) :named @p_248))) :rule cong :premises (t74.t6 t74.t7)) +(step t74.t9 (cl (= @p_232 (! (= @p_246 @p_248) :named @p_249))) :rule cong :premises (t74.t5 t74.t8)) +(step t74.t10 (cl @p_250) :rule refl) +(step t74.t11 (cl (= @p_235 (! (clock$ veriT_vr52) :named @p_251))) :rule cong :premises (t74.t10)) +(step t74.t12 (cl @p_252) :rule refl) +(step t74.t13 (cl (= @p_238 (! (clock$ veriT_vr50) :named @p_253))) :rule cong :premises (t74.t12)) +(step t74.t14 (cl (= @p_240 (! (less_eq$ @p_251 @p_253) :named @p_254))) :rule cong :premises (t74.t11 t74.t13)) +(step t74.t15 (cl (= @p_242 (! (=> @p_249 @p_254) :named @p_255))) :rule cong :premises (t74.t9 t74.t14)) +(step t74 (cl (! (= @p_244 (! (forall ((veriT_vr49 Astate$) (veriT_vr50 Astate$) (veriT_vr51 V_list_v_result$) (veriT_vr52 Astate$)) @p_255) :named @p_257)) :named @p_256)) :rule bind) +(step t75 (cl (not @p_256) (not @p_244) @p_257) :rule equiv_pos2) +(step t76 (cl @p_257) :rule th_resolution :premises (t73 t74 t75)) +(anchor :step t77 :args ((:= ?v0 veriT_vr53) (:= ?v1 veriT_vr54) (:= ?v2 veriT_vr55))) +(step t77.t1 (cl (! (= ?v0 veriT_vr53) :named @p_261)) :rule refl) +(step t77.t2 (cl (! (= ?v1 veriT_vr54) :named @p_262)) :rule refl) +(step t77.t3 (cl (! (= ?v2 veriT_vr55) :named @p_267)) :rule refl) +(step t77.t4 (cl (= @p_225 (! (pair$ veriT_vr54 veriT_vr55) :named @p_259))) :rule cong :premises (t77.t2 t77.t3)) +(step t77.t5 (cl (= @p_2 (! (fix_clock$ veriT_vr53 @p_259) :named @p_260))) :rule cong :premises (t77.t1 t77.t4)) +(step t77.t6 (cl @p_261) :rule refl) +(step t77.t7 (cl @p_262) :rule refl) +(step t77.t8 (cl (= @p_263 (! (uu$ veriT_vr53 veriT_vr54) :named @p_264))) :rule cong :premises (t77.t6 t77.t7)) +(step t77.t9 (cl @p_262) :rule refl) +(step t77.t10 (cl (= @p_265 (! (update_clock$ @p_264 veriT_vr54) :named @p_266))) :rule cong :premises (t77.t8 t77.t9)) +(step t77.t11 (cl @p_267) :rule refl) +(step t77.t12 (cl (= @p_268 (! (pair$ @p_266 veriT_vr55) :named @p_269))) :rule cong :premises (t77.t10 t77.t11)) +(step t77.t13 (cl (= @p_270 (! (= @p_260 @p_269) :named @p_271))) :rule cong :premises (t77.t5 t77.t12)) +(step t77 (cl (! (= @p_258 (! (forall ((veriT_vr53 Astate$) (veriT_vr54 Astate$) (veriT_vr55 V_list_v_result$)) @p_271) :named @p_273)) :named @p_272)) :rule bind) +(step t78 (cl (not @p_272) (not @p_258) @p_273) :rule equiv_pos2) +(step t79 (cl @p_273) :rule th_resolution :premises (a11 t77 t78)) +(anchor :step t80 :args ((:= veriT_vr53 veriT_vr56) (:= veriT_vr54 veriT_vr57) (:= veriT_vr55 veriT_vr58))) +(step t80.t1 (cl (! (= veriT_vr53 veriT_vr56) :named @p_276)) :rule refl) +(step t80.t2 (cl (! (= veriT_vr54 veriT_vr57) :named @p_277)) :rule refl) +(step t80.t3 (cl (! (= veriT_vr55 veriT_vr58) :named @p_280)) :rule refl) +(step t80.t4 (cl (= @p_259 (! (pair$ veriT_vr57 veriT_vr58) :named @p_274))) :rule cong :premises (t80.t2 t80.t3)) +(step t80.t5 (cl (= @p_260 (! (fix_clock$ veriT_vr56 @p_274) :named @p_275))) :rule cong :premises (t80.t1 t80.t4)) +(step t80.t6 (cl @p_276) :rule refl) +(step t80.t7 (cl @p_277) :rule refl) +(step t80.t8 (cl (= @p_264 (! (uu$ veriT_vr56 veriT_vr57) :named @p_278))) :rule cong :premises (t80.t6 t80.t7)) +(step t80.t9 (cl @p_277) :rule refl) +(step t80.t10 (cl (= @p_266 (! (update_clock$ @p_278 veriT_vr57) :named @p_279))) :rule cong :premises (t80.t8 t80.t9)) +(step t80.t11 (cl @p_280) :rule refl) +(step t80.t12 (cl (= @p_269 (! (pair$ @p_279 veriT_vr58) :named @p_281))) :rule cong :premises (t80.t10 t80.t11)) +(step t80.t13 (cl (= @p_271 (! (= @p_275 @p_281) :named @p_282))) :rule cong :premises (t80.t5 t80.t12)) +(step t80 (cl (! (= @p_273 (! (forall ((veriT_vr56 Astate$) (veriT_vr57 Astate$) (veriT_vr58 V_list_v_result$)) @p_282) :named @p_284)) :named @p_283)) :rule bind) +(step t81 (cl (not @p_283) (not @p_273) @p_284) :rule equiv_pos2) +(step t82 (cl @p_284) :rule th_resolution :premises (t79 t80 t81)) +(anchor :step t83 :args ((:= ?v0 veriT_vr59) (:= ?v1 veriT_vr60))) +(step t83.t1 (cl (! (= ?v0 veriT_vr59) :named @p_291)) :rule refl) +(step t83.t2 (cl (= @p_287 (! (rerr$ veriT_vr59) :named @p_288))) :rule cong :premises (t83.t1)) +(step t83.t3 (cl (= @p_289 (! (= r$ @p_288) :named @p_290))) :rule cong :premises (t83.t2)) +(step t83.t4 (cl @p_291) :rule refl) +(step t83.t5 (cl (! (= ?v1 veriT_vr60) :named @p_296)) :rule refl) +(step t83.t6 (cl (= @p_140 (! (rraise$ veriT_vr60) :named @p_292))) :rule cong :premises (t83.t5)) +(step t83.t7 (cl (= @p_3 (! (= veriT_vr59 @p_292) :named @p_293))) :rule cong :premises (t83.t4 t83.t6)) +(step t83.t8 (cl (= @p_294 (! (and @p_290 @p_293) :named @p_295))) :rule cong :premises (t83.t3 t83.t7)) +(step t83.t9 (cl @p_296) :rule refl) +(step t83.t10 (cl (= @p_297 (! (fun_evaluate_match$ st$ env$ veriT_vr60 pes$) :named @p_298))) :rule cong :premises (t83.t9)) +(step t83.t11 (cl @p_296) :rule refl) +(step t83.t12 (cl (= @p_299 (! (fun_app$ @p_298 veriT_vr60) :named @p_300))) :rule cong :premises (t83.t10 t83.t11)) +(step t83.t13 (cl (= @p_301 (! (fst$ @p_300) :named @p_302))) :rule cong :premises (t83.t12)) +(step t83.t14 (cl (= @p_303 (! (clock$ @p_302) :named @p_304))) :rule cong :premises (t83.t13)) +(step t83.t15 (cl (= @p_305 (! (less_eq$ @p_304 @p_286) :named @p_306))) :rule cong :premises (t83.t14)) +(step t83.t16 (cl (= @p_307 (! (=> @p_295 @p_306) :named @p_308))) :rule cong :premises (t83.t8 t83.t15)) +(step t83 (cl (! (= @p_285 (! (forall ((veriT_vr59 V_error_result$) (veriT_vr60 V$)) @p_308) :named @p_310)) :named @p_309)) :rule bind) +(step t84 (cl (not @p_309) (not @p_285) @p_310) :rule equiv_pos2) +(step t85 (cl @p_310) :rule th_resolution :premises (a12 t83 t84)) +(anchor :step t86 :args ((:= veriT_vr59 veriT_vr61) (:= veriT_vr60 veriT_vr62))) +(step t86.t1 (cl (! (= veriT_vr59 veriT_vr61) :named @p_313)) :rule refl) +(step t86.t2 (cl (= @p_288 (! (rerr$ veriT_vr61) :named @p_311))) :rule cong :premises (t86.t1)) +(step t86.t3 (cl (= @p_290 (! (= r$ @p_311) :named @p_312))) :rule cong :premises (t86.t2)) +(step t86.t4 (cl @p_313) :rule refl) +(step t86.t5 (cl (! (= veriT_vr60 veriT_vr62) :named @p_317)) :rule refl) +(step t86.t6 (cl (= @p_292 (! (rraise$ veriT_vr62) :named @p_314))) :rule cong :premises (t86.t5)) +(step t86.t7 (cl (= @p_293 (! (= veriT_vr61 @p_314) :named @p_315))) :rule cong :premises (t86.t4 t86.t6)) +(step t86.t8 (cl (= @p_295 (! (and @p_312 @p_315) :named @p_316))) :rule cong :premises (t86.t3 t86.t7)) +(step t86.t9 (cl @p_317) :rule refl) +(step t86.t10 (cl (= @p_298 (! (fun_evaluate_match$ st$ env$ veriT_vr62 pes$) :named @p_318))) :rule cong :premises (t86.t9)) +(step t86.t11 (cl @p_317) :rule refl) +(step t86.t12 (cl (= @p_300 (! (fun_app$ @p_318 veriT_vr62) :named @p_319))) :rule cong :premises (t86.t10 t86.t11)) +(step t86.t13 (cl (= @p_302 (! (fst$ @p_319) :named @p_320))) :rule cong :premises (t86.t12)) +(step t86.t14 (cl (= @p_304 (! (clock$ @p_320) :named @p_321))) :rule cong :premises (t86.t13)) +(step t86.t15 (cl (= @p_306 (! (less_eq$ @p_321 @p_286) :named @p_322))) :rule cong :premises (t86.t14)) +(step t86.t16 (cl (= @p_308 (! (=> @p_316 @p_322) :named @p_323))) :rule cong :premises (t86.t8 t86.t15)) +(step t86 (cl (! (= @p_310 (! (forall ((veriT_vr61 V_error_result$) (veriT_vr62 V$)) @p_323) :named @p_325)) :named @p_324)) :rule bind) +(step t87 (cl (not @p_324) (not @p_310) @p_325) :rule equiv_pos2) +(step t88 (cl @p_325) :rule th_resolution :premises (t85 t86 t87)) +(step t89 (cl (! (= @p_326 (! (and @p_327 (! (not @p_328) :named @p_334)) :named @p_330)) :named @p_329)) :rule bool_simplify) +(step t90 (cl (! (not @p_329) :named @p_333) (! (not @p_326) :named @p_331) @p_330) :rule equiv_pos2) +(step t91 (cl (not @p_331) @p_332) :rule not_not) +(step t92 (cl @p_333 @p_332 @p_330) :rule th_resolution :premises (t91 t90)) +(step t93 (cl @p_330) :rule th_resolution :premises (a13 t89 t92)) +(step t94 (cl @p_327) :rule and :premises (t93)) +(step t95 (cl @p_334) :rule and :premises (t93)) +(step t96 (cl (or (! (not @p_71) :named @p_336) (! (forall ((veriT_vr13 Nat$) (veriT_vr14 Nat$) (veriT_vr15 Nat$)) (or (not @p_62) (not @p_64) @p_68)) :named @p_516))) :rule qnt_cnf) +(step t97 (cl (or (! (not @p_136) :named @p_399) (! (forall ((veriT_vr23 Astate$) (veriT_vr24 Astate_v_list_v_result_prod$) (veriT_vr26 V_list_v_result$)) (or @p_335 @p_112)) :named @p_577))) :rule qnt_cnf) +(step t98 (cl (or @p_336 (! (=> (! (and @p_337 (! (less_eq$ @p_338 @p_339) :named @p_341)) :named @p_340) @p_328) :named @p_342))) :rule forall_inst :args ((:= veriT_vr13 @p_339) (:= veriT_vr14 @p_4) (:= veriT_vr15 @p_338))) +(step t99 (cl @p_340 (! (not @p_337) :named @p_517) (! (not @p_341) :named @p_343)) :rule and_neg) +(step t100 (cl (! (not @p_342) :named @p_344) (! (not @p_340) :named @p_345) @p_328) :rule implies_pos) +(step t101 (cl @p_336 @p_342) :rule or :premises (t98)) +(step t102 (cl @p_340 @p_343) :rule resolution :premises (t99 a4)) +(step t103 (cl @p_344 @p_345) :rule resolution :premises (t100 t95)) +(step t104 (cl @p_342) :rule resolution :premises (t101 t31)) +(step t105 (cl @p_345) :rule resolution :premises (t103 t104)) +(step t106 (cl @p_343) :rule resolution :premises (t102 t105)) +(step t107 (cl (not (! (not @p_336) :named @p_521)) @p_71) :rule not_not) +(step t108 (cl (or (! (not @p_284) :named @p_497) (! (= (fix_clock$ st$a (pair$ @p_346 r$)) (pair$ @p_504 r$)) :named @p_498))) :rule forall_inst :args ((:= veriT_vr56 st$a) (:= veriT_vr57 @p_346) (:= veriT_vr58 r$))) +(step t109 (cl (or (! (not @p_181) :named @p_395) (! (not (! (and (! (forall ((veriT_vr31 V$)) (! (not (! (= x2$ @p_170) :named @p_350)) :named @p_352)) :named @p_348) (! (forall ((veriT_vr32 Abort$)) (! (not (! (= x2$ @p_175) :named @p_356)) :named @p_358)) :named @p_354)) :named @p_360)) :named @p_347))) :rule forall_inst :args ((:= veriT_vr30 x2$))) +(anchor :step t110) +(assume t110.h1 @p_347) +(anchor :step t110.t2 :args ((:= veriT_vr31 veriT_vr63))) +(step t110.t2.t1 (cl (= veriT_vr31 veriT_vr63)) :rule refl) +(step t110.t2.t2 (cl (= @p_170 (! (rraise$ veriT_vr63) :named @p_349))) :rule cong :premises (t110.t2.t1)) +(step t110.t2.t3 (cl (= @p_350 (! (= x2$ @p_349) :named @p_351))) :rule cong :premises (t110.t2.t2)) +(step t110.t2.t4 (cl (= @p_352 (! (not @p_351) :named @p_353))) :rule cong :premises (t110.t2.t3)) +(step t110.t2 (cl (= @p_348 (! (forall ((veriT_vr63 V$)) @p_353) :named @p_361))) :rule bind) +(anchor :step t110.t3 :args ((:= veriT_vr32 veriT_vr64))) +(step t110.t3.t1 (cl (= veriT_vr32 veriT_vr64)) :rule refl) +(step t110.t3.t2 (cl (= @p_175 (! (rabort$ veriT_vr64) :named @p_355))) :rule cong :premises (t110.t3.t1)) +(step t110.t3.t3 (cl (= @p_356 (! (= x2$ @p_355) :named @p_357))) :rule cong :premises (t110.t3.t2)) +(step t110.t3.t4 (cl (= @p_358 (! (not @p_357) :named @p_359))) :rule cong :premises (t110.t3.t3)) +(step t110.t3 (cl (= @p_354 (! (forall ((veriT_vr64 Abort$)) @p_359) :named @p_362))) :rule bind) +(step t110.t4 (cl (= @p_360 (! (and @p_361 @p_362) :named @p_363))) :rule cong :premises (t110.t2 t110.t3)) +(step t110.t5 (cl (! (= @p_347 (! (not @p_363) :named @p_366)) :named @p_364)) :rule cong :premises (t110.t4)) +(step t110.t6 (cl (! (not @p_364) :named @p_367) (! (not @p_347) :named @p_365) @p_366) :rule equiv_pos2) +(step t110.t7 (cl (! (not @p_365) :named @p_394) @p_360) :rule not_not) +(step t110.t8 (cl @p_367 @p_360 @p_366) :rule th_resolution :premises (t110.t7 t110.t6)) +(step t110.t9 (cl @p_366) :rule th_resolution :premises (t110.h1 t110.t5 t110.t8)) +(anchor :step t110.t10 :args ((:= veriT_vr63 veriT_vr65))) +(step t110.t10.t1 (cl (= veriT_vr63 veriT_vr65)) :rule refl) +(step t110.t10.t2 (cl (= @p_349 @p_369)) :rule cong :premises (t110.t10.t1)) +(step t110.t10.t3 (cl (= @p_351 @p_370)) :rule cong :premises (t110.t10.t2)) +(step t110.t10.t4 (cl (= @p_353 @p_368)) :rule cong :premises (t110.t10.t3)) +(step t110.t10 (cl (= @p_361 (! (forall ((veriT_vr65 V$)) @p_368) :named @p_374))) :rule bind) +(anchor :step t110.t11 :args ((:= veriT_vr64 veriT_vr66))) +(step t110.t11.t1 (cl (= veriT_vr64 veriT_vr66)) :rule refl) +(step t110.t11.t2 (cl (= @p_355 @p_372)) :rule cong :premises (t110.t11.t1)) +(step t110.t11.t3 (cl (= @p_357 @p_373)) :rule cong :premises (t110.t11.t2)) +(step t110.t11.t4 (cl (= @p_359 @p_371)) :rule cong :premises (t110.t11.t3)) +(step t110.t11 (cl (= @p_362 (! (forall ((veriT_vr66 Abort$)) @p_371) :named @p_375))) :rule bind) +(step t110.t12 (cl (= @p_363 (! (and @p_374 @p_375) :named @p_376))) :rule cong :premises (t110.t10 t110.t11)) +(step t110.t13 (cl (! (= @p_366 (! (not @p_376) :named @p_378)) :named @p_377)) :rule cong :premises (t110.t12)) +(step t110.t14 (cl (! (not @p_377) :named @p_380) (! (not @p_366) :named @p_379) @p_378) :rule equiv_pos2) +(step t110.t15 (cl (not @p_379) @p_363) :rule not_not) +(step t110.t16 (cl @p_380 @p_363 @p_378) :rule th_resolution :premises (t110.t15 t110.t14)) +(step t110.t17 (cl @p_378) :rule th_resolution :premises (t110.t9 t110.t13 t110.t16)) +(anchor :step t110.t18 :args ((:= veriT_vr65 veriT_sk0))) +(step t110.t18.t1 (cl (= veriT_vr65 veriT_sk0)) :rule refl) +(step t110.t18.t2 (cl (= @p_369 (! (rraise$ veriT_sk0) :named @p_383))) :rule cong :premises (t110.t18.t1)) +(step t110.t18.t3 (cl (= @p_370 (! (= x2$ @p_383) :named @p_384))) :rule cong :premises (t110.t18.t2)) +(step t110.t18.t4 (cl (= @p_368 (! (not @p_384) :named @p_381))) :rule cong :premises (t110.t18.t3)) +(step t110.t18 (cl (= @p_374 @p_381)) :rule sko_forall) +(anchor :step t110.t19 :args ((:= veriT_vr66 veriT_sk1))) +(step t110.t19.t1 (cl (= veriT_vr66 veriT_sk1)) :rule refl) +(step t110.t19.t2 (cl (= @p_372 (! (rabort$ veriT_sk1) :named @p_387))) :rule cong :premises (t110.t19.t1)) +(step t110.t19.t3 (cl (= @p_373 (! (= x2$ @p_387) :named @p_388))) :rule cong :premises (t110.t19.t2)) +(step t110.t19.t4 (cl (= @p_371 (! (not @p_388) :named @p_385))) :rule cong :premises (t110.t19.t3)) +(step t110.t19 (cl (= @p_375 @p_385)) :rule sko_forall) +(step t110.t20 (cl (= @p_376 (! (and @p_381 @p_385) :named @p_389))) :rule cong :premises (t110.t18 t110.t19)) +(step t110.t21 (cl (! (= @p_378 (! (not @p_389) :named @p_390)) :named @p_391)) :rule cong :premises (t110.t20)) +(step t110.t22 (cl (! (not @p_391) :named @p_393) (! (not @p_378) :named @p_392) @p_390) :rule equiv_pos2) +(step t110.t23 (cl (not @p_392) @p_376) :rule not_not) +(step t110.t24 (cl @p_393 @p_376 @p_390) :rule th_resolution :premises (t110.t23 t110.t22)) +(step t110.t25 (cl @p_390) :rule th_resolution :premises (t110.t17 t110.t21 t110.t24)) +(step t110 (cl @p_365 @p_390) :rule subproof :discharge (h1)) +(step t111 (cl @p_394 @p_360) :rule not_not) +(step t112 (cl @p_360 @p_390) :rule th_resolution :premises (t111 t110)) +(step t113 (cl @p_395 @p_347) :rule or :premises (t109)) +(step t114 (cl (! (or @p_395 @p_390) :named @p_397) (! (not @p_395) :named @p_396)) :rule or_neg) +(step t115 (cl (not @p_396) @p_181) :rule not_not) +(step t116 (cl @p_397 @p_181) :rule th_resolution :premises (t115 t114)) +(step t117 (cl @p_397 (! (not @p_390) :named @p_398)) :rule or_neg) +(step t118 (cl (not @p_398) @p_389) :rule not_not) +(step t119 (cl @p_397 @p_389) :rule th_resolution :premises (t118 t117)) +(step t120 (cl @p_397) :rule th_resolution :premises (t113 t112 t116 t119)) +(step t121 (cl (not (! (not @p_399) :named @p_436)) @p_136) :rule not_not) +(step t122 (cl (or @p_399 (! (and (! (=> (! (= @p_346 @p_346) :named @p_400) (! (exists ((veriT_vr25 V_list_v_result$)) (! (= @p_1 (! (pair$ @p_346 veriT_vr25) :named @p_403)) :named @p_405)) :named @p_402)) :named @p_407) (! (=> (! (not (! (forall ((veriT_vr26 V_list_v_result$)) (! (not (! (= @p_1 (! (pair$ @p_346 veriT_vr26) :named @p_410)) :named @p_411)) :named @p_412)) :named @p_409)) :named @p_414) @p_400) :named @p_416)) :named @p_401))) :rule forall_inst :args ((:= veriT_vr23 @p_346) (:= veriT_vr24 @p_1))) +(anchor :step t123) +(assume t123.h1 @p_401) +(anchor :step t123.t2 :args ((:= veriT_vr25 veriT_vr72))) +(step t123.t2.t1 (cl (= veriT_vr25 veriT_vr72)) :rule refl) +(step t123.t2.t2 (cl (= @p_403 (! (pair$ @p_346 veriT_vr72) :named @p_404))) :rule cong :premises (t123.t2.t1)) +(step t123.t2.t3 (cl (= @p_405 (! (= @p_1 @p_404) :named @p_406))) :rule cong :premises (t123.t2.t2)) +(step t123.t2 (cl (= @p_402 (! (exists ((veriT_vr72 V_list_v_result$)) @p_406) :named @p_408))) :rule bind) +(step t123.t3 (cl (= @p_407 (! (=> @p_400 @p_408) :named @p_418))) :rule cong :premises (t123.t2)) +(anchor :step t123.t4 :args ((:= veriT_vr26 veriT_vr72))) +(step t123.t4.t1 (cl (= veriT_vr26 veriT_vr72)) :rule refl) +(step t123.t4.t2 (cl (= @p_410 @p_404)) :rule cong :premises (t123.t4.t1)) +(step t123.t4.t3 (cl (= @p_411 @p_406)) :rule cong :premises (t123.t4.t2)) +(step t123.t4.t4 (cl (= @p_412 (! (not @p_406) :named @p_413))) :rule cong :premises (t123.t4.t3)) +(step t123.t4 (cl (= @p_409 (! (forall ((veriT_vr72 V_list_v_result$)) @p_413) :named @p_415))) :rule bind) +(step t123.t5 (cl (= @p_414 (! (not @p_415) :named @p_417))) :rule cong :premises (t123.t4)) +(step t123.t6 (cl (= @p_416 (! (=> @p_417 @p_400) :named @p_419))) :rule cong :premises (t123.t5)) +(step t123.t7 (cl (! (= @p_401 (! (and @p_418 @p_419) :named @p_422)) :named @p_420)) :rule cong :premises (t123.t3 t123.t6)) +(step t123.t8 (cl (not @p_420) (! (not @p_401) :named @p_421) @p_422) :rule equiv_pos2) +(step t123.t9 (cl @p_422) :rule th_resolution :premises (t123.h1 t123.t7 t123.t8)) +(step t123.t10 (cl (= @p_400 true)) :rule eq_simplify) +(step t123.t11 (cl (= @p_418 (! (=> true @p_408) :named @p_423))) :rule cong :premises (t123.t10)) +(step t123.t12 (cl (= @p_423 @p_408)) :rule implies_simplify) +(step t123.t13 (cl (= @p_418 @p_408)) :rule trans :premises (t123.t11 t123.t12)) +(step t123.t14 (cl (= @p_419 (! (=> @p_417 true) :named @p_424))) :rule cong :premises (t123.t10)) +(step t123.t15 (cl (= @p_424 true)) :rule implies_simplify) +(step t123.t16 (cl (= @p_419 true)) :rule trans :premises (t123.t14 t123.t15)) +(step t123.t17 (cl (= @p_422 (! (and @p_408 true) :named @p_425))) :rule cong :premises (t123.t13 t123.t16)) +(step t123.t18 (cl (= @p_425 (! (and @p_408) :named @p_426))) :rule and_simplify) +(step t123.t19 (cl (= @p_426 @p_408)) :rule and_simplify) +(step t123.t20 (cl (! (= @p_422 @p_408) :named @p_427)) :rule trans :premises (t123.t17 t123.t18 t123.t19)) +(step t123.t21 (cl (not @p_427) (not @p_422) @p_408) :rule equiv_pos2) +(step t123.t22 (cl @p_408) :rule th_resolution :premises (t123.t9 t123.t20 t123.t21)) +(anchor :step t123.t23 :args ((:= veriT_vr72 veriT_vr73))) +(step t123.t23.t1 (cl (= veriT_vr72 veriT_vr73)) :rule refl) +(step t123.t23.t2 (cl (= @p_404 @p_429)) :rule cong :premises (t123.t23.t1)) +(step t123.t23.t3 (cl (= @p_406 @p_428)) :rule cong :premises (t123.t23.t2)) +(step t123.t23 (cl (! (= @p_408 (! (exists ((veriT_vr73 V_list_v_result$)) @p_428) :named @p_431)) :named @p_430)) :rule bind) +(step t123.t24 (cl (not @p_430) (not @p_408) @p_431) :rule equiv_pos2) +(step t123.t25 (cl @p_431) :rule th_resolution :premises (t123.t22 t123.t23 t123.t24)) +(anchor :step t123.t26 :args ((:= veriT_vr73 veriT_sk3))) +(step t123.t26.t1 (cl (= veriT_vr73 veriT_sk3)) :rule refl) +(step t123.t26.t2 (cl (= @p_429 (! (pair$ @p_346 veriT_sk3) :named @p_434))) :rule cong :premises (t123.t26.t1)) +(step t123.t26.t3 (cl (= @p_428 (! (= @p_1 @p_434) :named @p_432))) :rule cong :premises (t123.t26.t2)) +(step t123.t26 (cl (! (= @p_431 @p_432) :named @p_435)) :rule sko_ex) +(step t123.t27 (cl (not @p_435) (not @p_431) @p_432) :rule equiv_pos2) +(step t123.t28 (cl @p_432) :rule th_resolution :premises (t123.t25 t123.t26 t123.t27)) +(step t123 (cl @p_421 @p_432) :rule subproof :discharge (h1)) +(step t124 (cl @p_399 @p_401) :rule or :premises (t122)) +(step t125 (cl (! (or @p_399 @p_432) :named @p_437) @p_436) :rule or_neg) +(step t126 (cl @p_437 @p_136) :rule th_resolution :premises (t121 t125)) +(step t127 (cl @p_437 (! (not @p_432) :named @p_538)) :rule or_neg) +(step t128 (cl @p_437) :rule th_resolution :premises (t124 t123 t126 t127)) +(step t129 (cl (or @p_399 (! (and (! (=> (! (= st$ (! (fst$ @p_438) :named @p_674)) :named @p_439) (! (exists ((veriT_vr25 V_list_v_result$)) (! (= @p_438 (! (pair$ st$ veriT_vr25) :named @p_442)) :named @p_444)) :named @p_441)) :named @p_446) (! (=> (! (not (! (forall ((veriT_vr26 V_list_v_result$)) (! (not (! (= @p_438 (! (pair$ st$ veriT_vr26) :named @p_449)) :named @p_450)) :named @p_451)) :named @p_448)) :named @p_453) @p_439) :named @p_455)) :named @p_440))) :rule forall_inst :args ((:= veriT_vr23 st$) (:= veriT_vr24 @p_438))) +(anchor :step t130) +(assume t130.h1 @p_440) +(anchor :step t130.t2 :args ((:= veriT_vr25 veriT_vr106))) +(step t130.t2.t1 (cl (= veriT_vr25 veriT_vr106)) :rule refl) +(step t130.t2.t2 (cl (= @p_442 (! (pair$ st$ veriT_vr106) :named @p_443))) :rule cong :premises (t130.t2.t1)) +(step t130.t2.t3 (cl (= @p_444 (! (= @p_438 @p_443) :named @p_445))) :rule cong :premises (t130.t2.t2)) +(step t130.t2 (cl (= @p_441 (! (exists ((veriT_vr106 V_list_v_result$)) @p_445) :named @p_447))) :rule bind) +(step t130.t3 (cl (= @p_446 (! (=> @p_439 @p_447) :named @p_457))) :rule cong :premises (t130.t2)) +(anchor :step t130.t4 :args ((:= veriT_vr26 veriT_vr106))) +(step t130.t4.t1 (cl (= veriT_vr26 veriT_vr106)) :rule refl) +(step t130.t4.t2 (cl (= @p_449 @p_443)) :rule cong :premises (t130.t4.t1)) +(step t130.t4.t3 (cl (= @p_450 @p_445)) :rule cong :premises (t130.t4.t2)) +(step t130.t4.t4 (cl (= @p_451 (! (not @p_445) :named @p_452))) :rule cong :premises (t130.t4.t3)) +(step t130.t4 (cl (= @p_448 (! (forall ((veriT_vr106 V_list_v_result$)) @p_452) :named @p_454))) :rule bind) +(step t130.t5 (cl (= @p_453 (! (not @p_454) :named @p_456))) :rule cong :premises (t130.t4)) +(step t130.t6 (cl (= @p_455 (! (=> @p_456 @p_439) :named @p_458))) :rule cong :premises (t130.t5)) +(step t130.t7 (cl (! (= @p_440 (! (and @p_457 @p_458) :named @p_461)) :named @p_459)) :rule cong :premises (t130.t3 t130.t6)) +(step t130.t8 (cl (not @p_459) (! (not @p_440) :named @p_460) @p_461) :rule equiv_pos2) +(step t130.t9 (cl @p_461) :rule th_resolution :premises (t130.h1 t130.t7 t130.t8)) +(anchor :step t130.t10 :args ((:= veriT_vr106 veriT_vr107))) +(step t130.t10.t1 (cl (= veriT_vr106 veriT_vr107)) :rule refl) +(step t130.t10.t2 (cl (= @p_443 (! (pair$ st$ veriT_vr107) :named @p_462))) :rule cong :premises (t130.t10.t1)) +(step t130.t10.t3 (cl (= @p_445 (! (= @p_438 @p_462) :named @p_463))) :rule cong :premises (t130.t10.t2)) +(step t130.t10.t4 (cl (= @p_452 (! (not @p_463) :named @p_464))) :rule cong :premises (t130.t10.t3)) +(step t130.t10 (cl (= @p_454 (! (forall ((veriT_vr107 V_list_v_result$)) @p_464) :named @p_465))) :rule bind) +(step t130.t11 (cl (= @p_456 (! (not @p_465) :named @p_466))) :rule cong :premises (t130.t10)) +(step t130.t12 (cl (= @p_458 (! (=> @p_466 @p_439) :named @p_467))) :rule cong :premises (t130.t11)) +(step t130.t13 (cl (! (= @p_461 (! (and @p_457 @p_467) :named @p_469)) :named @p_468)) :rule cong :premises (t130.t12)) +(step t130.t14 (cl (not @p_468) (not @p_461) @p_469) :rule equiv_pos2) +(step t130.t15 (cl @p_469) :rule th_resolution :premises (t130.t9 t130.t13 t130.t14)) +(anchor :step t130.t16 :args ((:= veriT_vr106 veriT_vr108))) +(step t130.t16.t1 (cl (= veriT_vr106 veriT_vr108)) :rule refl) +(step t130.t16.t2 (cl (= @p_443 @p_471)) :rule cong :premises (t130.t16.t1)) +(step t130.t16.t3 (cl (= @p_445 @p_470)) :rule cong :premises (t130.t16.t2)) +(step t130.t16 (cl (= @p_447 (! (exists ((veriT_vr108 V_list_v_result$)) @p_470) :named @p_472))) :rule bind) +(step t130.t17 (cl (= @p_457 (! (=> @p_439 @p_472) :named @p_478))) :rule cong :premises (t130.t16)) +(anchor :step t130.t18 :args ((:= veriT_vr107 veriT_vr109))) +(step t130.t18.t1 (cl (= veriT_vr107 veriT_vr109)) :rule refl) +(step t130.t18.t2 (cl (= @p_462 (! (pair$ st$ veriT_vr109) :named @p_473))) :rule cong :premises (t130.t18.t1)) +(step t130.t18.t3 (cl (= @p_463 (! (= @p_438 @p_473) :named @p_474))) :rule cong :premises (t130.t18.t2)) +(step t130.t18.t4 (cl (= @p_464 (! (not @p_474) :named @p_475))) :rule cong :premises (t130.t18.t3)) +(step t130.t18 (cl (= @p_465 (! (forall ((veriT_vr109 V_list_v_result$)) @p_475) :named @p_476))) :rule bind) +(step t130.t19 (cl (= @p_466 (! (not @p_476) :named @p_477))) :rule cong :premises (t130.t18)) +(step t130.t20 (cl (= @p_467 (! (=> @p_477 @p_439) :named @p_479))) :rule cong :premises (t130.t19)) +(step t130.t21 (cl (! (= @p_469 (! (and @p_478 @p_479) :named @p_481)) :named @p_480)) :rule cong :premises (t130.t17 t130.t20)) +(step t130.t22 (cl (not @p_480) (not @p_469) @p_481) :rule equiv_pos2) +(step t130.t23 (cl @p_481) :rule th_resolution :premises (t130.t15 t130.t21 t130.t22)) +(anchor :step t130.t24 :args ((:= veriT_vr108 veriT_sk11))) +(step t130.t24.t1 (cl (= veriT_vr108 veriT_sk11)) :rule refl) +(step t130.t24.t2 (cl (= @p_471 (! (pair$ st$ veriT_sk11) :named @p_484))) :rule cong :premises (t130.t24.t1)) +(step t130.t24.t3 (cl (= @p_470 (! (= @p_438 @p_484) :named @p_482))) :rule cong :premises (t130.t24.t2)) +(step t130.t24 (cl (= @p_472 @p_482)) :rule sko_ex) +(step t130.t25 (cl (= @p_478 (! (=> @p_439 @p_482) :named @p_485))) :rule cong :premises (t130.t24)) +(step t130.t26 (cl (! (= @p_481 (! (and @p_485 @p_479) :named @p_487)) :named @p_486)) :rule cong :premises (t130.t25)) +(step t130.t27 (cl (not @p_486) (not @p_481) @p_487) :rule equiv_pos2) +(step t130.t28 (cl @p_487) :rule th_resolution :premises (t130.t23 t130.t26 t130.t27)) +(anchor :step t130.t29 :args ((:= veriT_vr109 veriT_vr110))) +(step t130.t29.t1 (cl (= veriT_vr109 veriT_vr110)) :rule refl) +(step t130.t29.t2 (cl (= @p_473 (! (pair$ st$ veriT_vr110) :named @p_489))) :rule cong :premises (t130.t29.t1)) +(step t130.t29.t3 (cl (= @p_474 (! (= @p_438 @p_489) :named @p_490))) :rule cong :premises (t130.t29.t2)) +(step t130.t29.t4 (cl (= @p_475 (! (not @p_490) :named @p_491))) :rule cong :premises (t130.t29.t3)) +(step t130.t29 (cl (= @p_476 (! (forall ((veriT_vr110 V_list_v_result$)) @p_491) :named @p_488))) :rule bind) +(step t130.t30 (cl (= @p_477 (! (not @p_488) :named @p_492))) :rule cong :premises (t130.t29)) +(step t130.t31 (cl (= @p_479 (! (=> @p_492 @p_439) :named @p_493))) :rule cong :premises (t130.t30)) +(step t130.t32 (cl (! (= @p_487 (! (and @p_485 @p_493) :named @p_494)) :named @p_495)) :rule cong :premises (t130.t31)) +(step t130.t33 (cl (not @p_495) (not @p_487) @p_494) :rule equiv_pos2) +(step t130.t34 (cl @p_494) :rule th_resolution :premises (t130.t28 t130.t32 t130.t33)) +(step t130 (cl @p_460 @p_494) :rule subproof :discharge (h1)) +(step t131 (cl @p_399 @p_440) :rule or :premises (t129)) +(step t132 (cl (! (or @p_399 @p_494) :named @p_496) @p_436) :rule or_neg) +(step t133 (cl @p_496 @p_136) :rule th_resolution :premises (t121 t132)) +(step t134 (cl @p_496 (! (not @p_494) :named @p_503)) :rule or_neg) +(step t135 (cl @p_496) :rule th_resolution :premises (t131 t130 t133 t134)) +(step t136 (cl @p_497 @p_498) :rule or :premises (t108)) +(step t137 (cl @p_498) :rule resolution :premises (t136 t82)) +(step t138 (cl @p_389 (! (not @p_381) :named @p_499) (! (not @p_385) :named @p_500)) :rule and_neg) +(step t139 (cl (not @p_499) @p_384) :rule not_not) +(step t140 (cl @p_389 @p_384 @p_500) :rule th_resolution :premises (t139 t138)) +(step t141 (cl (not @p_500) @p_388) :rule not_not) +(step t142 (cl @p_389 @p_384 @p_388) :rule th_resolution :premises (t141 t140)) +(step t143 (cl @p_395 @p_390) :rule or :premises (t120)) +(step t144 (cl @p_390) :rule resolution :premises (t143 t58)) +(step t145 (cl @p_399 @p_432) :rule or :premises (t128)) +(step t146 (cl @p_432) :rule resolution :premises (t145 t49)) +(step t147 (cl (! (not @p_493) :named @p_502) (! (not @p_492) :named @p_501) @p_439) :rule implies_pos) +(step t148 (cl (not @p_501) @p_488) :rule not_not) +(step t149 (cl @p_502 @p_488 @p_439) :rule th_resolution :premises (t148 t147)) +(step t150 (cl @p_503 @p_493) :rule and_pos) +(step t151 (cl @p_399 @p_494) :rule or :premises (t135)) +(step t152 (cl @p_494) :rule resolution :premises (t151 t49)) +(step t153 (cl @p_493) :rule resolution :premises (t150 t152)) +(step t154 (cl (or (! (not @p_325) :named @p_508) (! (=> (! (and @p_327 @p_384) :named @p_505) (! (less_eq$ (! (clock$ (! (fst$ (! (fun_app$ (fun_evaluate_match$ st$ env$ veriT_sk0 pes$) veriT_sk0) :named @p_526)) :named @p_566)) :named @p_567) @p_286) :named @p_507)) :named @p_506))) :rule forall_inst :args ((:= veriT_vr61 x2$) (:= veriT_vr62 veriT_sk0))) +(step t155 (cl (or @p_497 (! (= (! (fix_clock$ st$a @p_434) :named @p_539) (! (pair$ @p_504 veriT_sk3) :named @p_652)) :named @p_509))) :rule forall_inst :args ((:= veriT_vr56 st$a) (:= veriT_vr57 @p_346) (:= veriT_vr58 veriT_sk3))) +(step t156 (cl (or (! (not @p_257) :named @p_512) (! (=> @p_498 (! (less_eq$ (! (clock$ @p_504) :named @p_518) @p_339) :named @p_511)) :named @p_510))) :rule forall_inst :args ((:= veriT_vr49 st$a) (:= veriT_vr50 @p_346) (:= veriT_vr51 r$) (:= veriT_vr52 @p_504))) +(step t157 (cl (or (! (not @p_202) :named @p_514) (! (= (! (case_error_result$ uua$ uub$ @p_383) :named @p_569) (! (fun_app$ uua$ veriT_sk0) :named @p_527)) :named @p_515))) :rule forall_inst :args ((:= veriT_vr36 uua$) (:= veriT_vr37 uub$) (:= veriT_vr38 veriT_sk0))) +(step t158 (cl @p_505 (! (not @p_327) :named @p_561) @p_381) :rule and_neg) +(step t159 (cl (not @p_506) (not @p_505) @p_507) :rule implies_pos) +(step t160 (cl @p_508 @p_506) :rule or :premises (t154)) +(step t161 (cl @p_505 @p_381) :rule resolution :premises (t158 t94)) +(step t162 (cl @p_506) :rule resolution :premises (t160 t88)) +(step t163 (cl @p_497 @p_509) :rule or :premises (t155)) +(step t164 (cl @p_509) :rule resolution :premises (t163 t82)) +(step t165 (cl (! (not @p_510) :named @p_513) (not @p_498) @p_511) :rule implies_pos) +(step t166 (cl @p_512 @p_510) :rule or :premises (t156)) +(step t167 (cl @p_513 @p_511) :rule resolution :premises (t165 t137)) +(step t168 (cl @p_510) :rule resolution :premises (t166 t76)) +(step t169 (cl @p_511) :rule resolution :premises (t167 t168)) +(step t170 (cl @p_514 @p_515) :rule or :premises (t157)) +(step t171 (cl @p_515) :rule resolution :premises (t170 t64)) +(step t172 (cl @p_336 @p_516) :rule or :premises (t96)) +(step t173 (cl (or (! (not @p_516) :named @p_519) (! (or @p_517 (! (not @p_511) :named @p_524) (! (less_eq$ @p_518 @p_4) :named @p_525)) :named @p_520))) :rule forall_inst :args ((:= veriT_vr13 @p_339) (:= veriT_vr14 @p_4) (:= veriT_vr15 @p_518))) +(step t174 (cl @p_519 @p_520) :rule or :premises (t173)) +(step t175 (cl (! (or @p_336 @p_520) :named @p_522) @p_521) :rule or_neg) +(step t176 (cl @p_522 @p_71) :rule th_resolution :premises (t107 t175)) +(step t177 (cl @p_522 (! (not @p_520) :named @p_523)) :rule or_neg) +(step t178 (cl @p_522) :rule th_resolution :premises (t172 t174 t176 t177)) +(step t179 (cl @p_523 @p_517 @p_524 @p_525) :rule or_pos) +(step t180 (cl @p_336 @p_520) :rule or :premises (t178)) +(step t181 (cl @p_523 @p_525) :rule resolution :premises (t179 a4 t169)) +(step t182 (cl @p_520) :rule resolution :premises (t180 t31)) +(step t183 (cl @p_525) :rule resolution :premises (t181 t182)) +(step t184 (cl (or (! (not @p_223) :named @p_528) (! (= (! (case_error_result$ uua$ uub$ @p_387) :named @p_548) (! (fun_app$a uub$ veriT_sk1) :named @p_542)) :named @p_529))) :rule forall_inst :args ((:= veriT_vr42 uua$) (:= veriT_vr43 uub$) (:= veriT_vr44 veriT_sk1))) +(step t185 (cl (or @p_336 (! (=> (! (and @p_525 (! (less_eq$ @p_338 @p_518) :named @p_531)) :named @p_530) @p_328) :named @p_532))) :rule forall_inst :args ((:= veriT_vr13 @p_518) (:= veriT_vr14 @p_4) (:= veriT_vr15 @p_338))) +(step t186 (cl (or (! (not @p_23) :named @p_536) (! (= @p_526 @p_527) :named @p_537))) :rule forall_inst :args ((:= veriT_vr1 veriT_sk0))) +(step t187 (cl @p_528 @p_529) :rule or :premises (t184)) +(step t188 (cl @p_529) :rule resolution :premises (t187 t70)) +(step t189 (cl @p_530 (not @p_525) (! (not @p_531) :named @p_533)) :rule and_neg) +(step t190 (cl (! (not @p_532) :named @p_534) (! (not @p_530) :named @p_535) @p_328) :rule implies_pos) +(step t191 (cl @p_336 @p_532) :rule or :premises (t185)) +(step t192 (cl @p_530 @p_533) :rule resolution :premises (t189 t183)) +(step t193 (cl @p_534 @p_535) :rule resolution :premises (t190 t95)) +(step t194 (cl @p_532) :rule resolution :premises (t191 t31)) +(step t195 (cl @p_535) :rule resolution :premises (t193 t194)) +(step t196 (cl @p_533) :rule resolution :premises (t192 t195)) +(step t197 (cl @p_536 @p_537) :rule or :premises (t186)) +(step t198 (cl @p_537) :rule resolution :premises (t197 t19)) +(step t199 (cl (not (! (= st$a st$a) :named @p_540)) @p_538 (! (= @p_438 @p_539) :named @p_541)) :rule eq_congruent) +(step t200 (cl @p_540) :rule eq_reflexive) +(step t201 (cl @p_538 @p_541) :rule th_resolution :premises (t199 t200)) +(step t202 (cl (or (! (not @p_45) :named @p_543) (! (= @p_542 (! (pair$ st$ (! (rerr$ @p_387) :named @p_557)) :named @p_546)) :named @p_544))) :rule forall_inst :args ((:= veriT_vr3 veriT_sk1))) +(step t203 (cl @p_543 @p_544) :rule or :premises (t202)) +(step t204 (cl @p_544) :rule resolution :premises (t203 t25)) +(step t205 (cl (! (not (! (= @p_545 @p_546) :named @p_558)) :named @p_553) (! (not @p_544) :named @p_554) (! (not @p_529) :named @p_555) (not (! (= @p_547 @p_548) :named @p_549)) (! (= @p_545 @p_547) :named @p_556)) :rule eq_transitive) +(step t206 (cl (! (not (! (= uua$ uua$) :named @p_550)) :named @p_571) (! (not (! (= uub$ uub$) :named @p_552)) :named @p_551) @p_385 @p_549) :rule eq_congruent) +(step t207 (cl @p_550) :rule eq_reflexive) +(step t208 (cl @p_551 @p_385 @p_549) :rule th_resolution :premises (t206 t207)) +(step t209 (cl @p_552) :rule eq_reflexive) +(step t210 (cl @p_385 @p_549) :rule th_resolution :premises (t208 t209)) +(step t211 (cl @p_553 @p_554 @p_555 @p_556 @p_385) :rule th_resolution :premises (t205 t210)) +(step t212 (cl (not (! (= st$ st$) :named @p_559)) (! (not (! (= r$ @p_557) :named @p_563)) :named @p_560) @p_558) :rule eq_congruent) +(step t213 (cl @p_559) :rule eq_reflexive) +(step t214 (cl @p_560 @p_558) :rule th_resolution :premises (t212 t213)) +(step t215 (cl @p_561 (not (! (= @p_562 @p_557) :named @p_564)) @p_563) :rule eq_transitive) +(step t216 (cl @p_385 @p_564) :rule eq_congruent) +(step t217 (cl @p_561 @p_563 @p_385) :rule th_resolution :premises (t215 t216)) +(step t218 (cl @p_558 @p_561 @p_385) :rule th_resolution :premises (t214 t217)) +(step t219 (cl @p_554 @p_555 @p_556 @p_385 @p_561 @p_385) :rule th_resolution :premises (t211 t218)) +(step t220 (cl @p_554 @p_555 @p_556 @p_385 @p_561) :rule contraction :premises (t219)) +(step t221 (cl @p_556 @p_385) :rule resolution :premises (t220 t94 t188 t204)) +(step t222 (cl (not (! (= @p_565 @p_566) :named @p_568)) (! (= @p_338 @p_567) :named @p_575)) :rule eq_congruent) +(step t223 (cl (not (! (= @p_547 @p_526) :named @p_570)) @p_568) :rule eq_congruent) +(step t224 (cl (not (! (= @p_547 @p_569) :named @p_572)) (! (not @p_515) :named @p_573) (! (not @p_537) :named @p_574) @p_570) :rule eq_transitive) +(step t225 (cl @p_571 @p_551 @p_381 @p_572) :rule eq_congruent) +(step t226 (cl @p_551 @p_381 @p_572) :rule th_resolution :premises (t225 t207)) +(step t227 (cl @p_381 @p_572) :rule th_resolution :premises (t226 t209)) +(step t228 (cl @p_573 @p_574 @p_570 @p_381) :rule th_resolution :premises (t224 t227)) +(step t229 (cl @p_568 @p_573 @p_574 @p_381) :rule th_resolution :premises (t223 t228)) +(step t230 (cl @p_575 @p_573 @p_574 @p_381) :rule th_resolution :premises (t222 t229)) +(step t231 (cl (or @p_492 (! (not @p_576) :named @p_578))) :rule forall_inst :args ((:= veriT_vr110 r$))) +(step t232 (cl @p_399 @p_577) :rule or :premises (t97)) +(step t233 (cl @p_492 @p_578) :rule or :premises (t231)) +(step t234 (cl @p_492) :rule resolution :premises (t233 a3)) +(step t235 (cl @p_439) :rule resolution :premises (t149 t234 t153)) +(step t236 (cl (or @p_497 (! (= (! (pair$ @p_504 @p_557) :named @p_647) (! (fix_clock$ st$a (pair$ @p_346 @p_557)) :named @p_580)) :named @p_579))) :rule forall_inst :args ((:= veriT_vr56 st$a) (:= veriT_vr57 @p_346) (:= veriT_vr58 @p_557))) +(step t237 (cl @p_497 @p_579) :rule or :premises (t236)) +(step t238 (cl @p_579) :rule resolution :premises (t237 t82)) +(step t239 (cl (or @p_512 (! (=> (! (= (! (pair$ @p_565 @p_557) :named @p_646) @p_580) :named @p_582) @p_341) :named @p_581))) :rule forall_inst :args ((:= veriT_vr49 st$a) (:= veriT_vr50 @p_346) (:= veriT_vr51 @p_557) (:= veriT_vr52 @p_565))) +(step t240 (cl (! (not @p_581) :named @p_583) (! (not @p_582) :named @p_584) @p_341) :rule implies_pos) +(step t241 (cl @p_512 @p_581) :rule or :premises (t239)) +(step t242 (cl @p_583 @p_584) :rule resolution :premises (t240 t106)) +(step t243 (cl @p_581) :rule resolution :premises (t241 t76)) +(step t244 (cl @p_584) :rule resolution :premises (t242 t243)) +(step t245 (cl (or @p_399 (! (and (! (=> (! (= @p_565 @p_504) :named @p_585) (! (exists ((veriT_vr25 V_list_v_result$)) (! (= @p_547 (! (pair$ @p_504 veriT_vr25) :named @p_588)) :named @p_590)) :named @p_587)) :named @p_592) (! (=> (! (not (! (forall ((veriT_vr26 V_list_v_result$)) (! (not (! (= @p_547 (! (pair$ @p_504 veriT_vr26) :named @p_595)) :named @p_596)) :named @p_597)) :named @p_594)) :named @p_599) @p_585) :named @p_601)) :named @p_586))) :rule forall_inst :args ((:= veriT_vr23 @p_504) (:= veriT_vr24 @p_547))) +(anchor :step t246) +(assume t246.h1 @p_586) +(anchor :step t246.t2 :args ((:= veriT_vr25 veriT_vr3120))) +(step t246.t2.t1 (cl (= veriT_vr25 veriT_vr3120)) :rule refl) +(step t246.t2.t2 (cl (= @p_588 (! (pair$ @p_504 veriT_vr3120) :named @p_589))) :rule cong :premises (t246.t2.t1)) +(step t246.t2.t3 (cl (= @p_590 (! (= @p_547 @p_589) :named @p_591))) :rule cong :premises (t246.t2.t2)) +(step t246.t2 (cl (= @p_587 (! (exists ((veriT_vr3120 V_list_v_result$)) @p_591) :named @p_593))) :rule bind) +(step t246.t3 (cl (= @p_592 (! (=> @p_585 @p_593) :named @p_603))) :rule cong :premises (t246.t2)) +(anchor :step t246.t4 :args ((:= veriT_vr26 veriT_vr3120))) +(step t246.t4.t1 (cl (= veriT_vr26 veriT_vr3120)) :rule refl) +(step t246.t4.t2 (cl (= @p_595 @p_589)) :rule cong :premises (t246.t4.t1)) +(step t246.t4.t3 (cl (= @p_596 @p_591)) :rule cong :premises (t246.t4.t2)) +(step t246.t4.t4 (cl (= @p_597 (! (not @p_591) :named @p_598))) :rule cong :premises (t246.t4.t3)) +(step t246.t4 (cl (= @p_594 (! (forall ((veriT_vr3120 V_list_v_result$)) @p_598) :named @p_600))) :rule bind) +(step t246.t5 (cl (= @p_599 (! (not @p_600) :named @p_602))) :rule cong :premises (t246.t4)) +(step t246.t6 (cl (= @p_601 (! (=> @p_602 @p_585) :named @p_604))) :rule cong :premises (t246.t5)) +(step t246.t7 (cl (! (= @p_586 (! (and @p_603 @p_604) :named @p_607)) :named @p_605)) :rule cong :premises (t246.t3 t246.t6)) +(step t246.t8 (cl (not @p_605) (! (not @p_586) :named @p_606) @p_607) :rule equiv_pos2) +(step t246.t9 (cl @p_607) :rule th_resolution :premises (t246.h1 t246.t7 t246.t8)) +(anchor :step t246.t10 :args ((:= veriT_vr3120 veriT_vr3121))) +(step t246.t10.t1 (cl (= veriT_vr3120 veriT_vr3121)) :rule refl) +(step t246.t10.t2 (cl (= @p_589 (! (pair$ @p_504 veriT_vr3121) :named @p_608))) :rule cong :premises (t246.t10.t1)) +(step t246.t10.t3 (cl (= @p_591 (! (= @p_547 @p_608) :named @p_609))) :rule cong :premises (t246.t10.t2)) +(step t246.t10.t4 (cl (= @p_598 (! (not @p_609) :named @p_610))) :rule cong :premises (t246.t10.t3)) +(step t246.t10 (cl (= @p_600 (! (forall ((veriT_vr3121 V_list_v_result$)) @p_610) :named @p_611))) :rule bind) +(step t246.t11 (cl (= @p_602 (! (not @p_611) :named @p_612))) :rule cong :premises (t246.t10)) +(step t246.t12 (cl (= @p_604 (! (=> @p_612 @p_585) :named @p_613))) :rule cong :premises (t246.t11)) +(step t246.t13 (cl (! (= @p_607 (! (and @p_603 @p_613) :named @p_615)) :named @p_614)) :rule cong :premises (t246.t12)) +(step t246.t14 (cl (not @p_614) (not @p_607) @p_615) :rule equiv_pos2) +(step t246.t15 (cl @p_615) :rule th_resolution :premises (t246.t9 t246.t13 t246.t14)) +(anchor :step t246.t16 :args ((:= veriT_vr3120 veriT_vr3122))) +(step t246.t16.t1 (cl (= veriT_vr3120 veriT_vr3122)) :rule refl) +(step t246.t16.t2 (cl (= @p_589 @p_617)) :rule cong :premises (t246.t16.t1)) +(step t246.t16.t3 (cl (= @p_591 @p_616)) :rule cong :premises (t246.t16.t2)) +(step t246.t16 (cl (= @p_593 (! (exists ((veriT_vr3122 V_list_v_result$)) @p_616) :named @p_618))) :rule bind) +(step t246.t17 (cl (= @p_603 (! (=> @p_585 @p_618) :named @p_624))) :rule cong :premises (t246.t16)) +(anchor :step t246.t18 :args ((:= veriT_vr3121 veriT_vr3123))) +(step t246.t18.t1 (cl (= veriT_vr3121 veriT_vr3123)) :rule refl) +(step t246.t18.t2 (cl (= @p_608 (! (pair$ @p_504 veriT_vr3123) :named @p_619))) :rule cong :premises (t246.t18.t1)) +(step t246.t18.t3 (cl (= @p_609 (! (= @p_547 @p_619) :named @p_620))) :rule cong :premises (t246.t18.t2)) +(step t246.t18.t4 (cl (= @p_610 (! (not @p_620) :named @p_621))) :rule cong :premises (t246.t18.t3)) +(step t246.t18 (cl (= @p_611 (! (forall ((veriT_vr3123 V_list_v_result$)) @p_621) :named @p_622))) :rule bind) +(step t246.t19 (cl (= @p_612 (! (not @p_622) :named @p_623))) :rule cong :premises (t246.t18)) +(step t246.t20 (cl (= @p_613 (! (=> @p_623 @p_585) :named @p_625))) :rule cong :premises (t246.t19)) +(step t246.t21 (cl (! (= @p_615 (! (and @p_624 @p_625) :named @p_627)) :named @p_626)) :rule cong :premises (t246.t17 t246.t20)) +(step t246.t22 (cl (not @p_626) (not @p_615) @p_627) :rule equiv_pos2) +(step t246.t23 (cl @p_627) :rule th_resolution :premises (t246.t15 t246.t21 t246.t22)) +(anchor :step t246.t24 :args ((:= veriT_vr3122 veriT_sk615))) +(step t246.t24.t1 (cl (= veriT_vr3122 veriT_sk615)) :rule refl) +(step t246.t24.t2 (cl (= @p_617 (! (pair$ @p_504 veriT_sk615) :named @p_630))) :rule cong :premises (t246.t24.t1)) +(step t246.t24.t3 (cl (= @p_616 (! (= @p_547 @p_630) :named @p_628))) :rule cong :premises (t246.t24.t2)) +(step t246.t24 (cl (= @p_618 @p_628)) :rule sko_ex) +(step t246.t25 (cl (= @p_624 (! (=> @p_585 @p_628) :named @p_631))) :rule cong :premises (t246.t24)) +(step t246.t26 (cl (! (= @p_627 (! (and @p_631 @p_625) :named @p_633)) :named @p_632)) :rule cong :premises (t246.t25)) +(step t246.t27 (cl (not @p_632) (not @p_627) @p_633) :rule equiv_pos2) +(step t246.t28 (cl @p_633) :rule th_resolution :premises (t246.t23 t246.t26 t246.t27)) +(anchor :step t246.t29 :args ((:= veriT_vr3123 veriT_vr3124))) +(step t246.t29.t1 (cl (= veriT_vr3123 veriT_vr3124)) :rule refl) +(step t246.t29.t2 (cl (= @p_619 (! (pair$ @p_504 veriT_vr3124) :named @p_635))) :rule cong :premises (t246.t29.t1)) +(step t246.t29.t3 (cl (= @p_620 (! (= @p_547 @p_635) :named @p_636))) :rule cong :premises (t246.t29.t2)) +(step t246.t29.t4 (cl (= @p_621 (! (not @p_636) :named @p_637))) :rule cong :premises (t246.t29.t3)) +(step t246.t29 (cl (= @p_622 (! (forall ((veriT_vr3124 V_list_v_result$)) @p_637) :named @p_634))) :rule bind) +(step t246.t30 (cl (= @p_623 (! (not @p_634) :named @p_638))) :rule cong :premises (t246.t29)) +(step t246.t31 (cl (= @p_625 (! (=> @p_638 @p_585) :named @p_639))) :rule cong :premises (t246.t30)) +(step t246.t32 (cl (! (= @p_633 (! (and @p_631 @p_639) :named @p_640)) :named @p_641)) :rule cong :premises (t246.t31)) +(step t246.t33 (cl (not @p_641) (not @p_633) @p_640) :rule equiv_pos2) +(step t246.t34 (cl @p_640) :rule th_resolution :premises (t246.t28 t246.t32 t246.t33)) +(step t246 (cl @p_606 @p_640) :rule subproof :discharge (h1)) +(step t247 (cl @p_399 @p_586) :rule or :premises (t245)) +(step t248 (cl (! (or @p_399 @p_640) :named @p_642) @p_436) :rule or_neg) +(step t249 (cl @p_642 @p_136) :rule th_resolution :premises (t121 t248)) +(step t250 (cl @p_642 (! (not @p_640) :named @p_645)) :rule or_neg) +(step t251 (cl @p_642) :rule th_resolution :premises (t247 t246 t249 t250)) +(step t252 (cl (! (not @p_639) :named @p_644) (! (not @p_638) :named @p_643) @p_585) :rule implies_pos) +(step t253 (cl (not @p_643) @p_634) :rule not_not) +(step t254 (cl @p_644 @p_634 @p_585) :rule th_resolution :premises (t253 t252)) +(step t255 (cl @p_645 @p_639) :rule and_pos) +(step t256 (cl @p_399 @p_640) :rule or :premises (t251)) +(step t257 (cl @p_640) :rule resolution :premises (t256 t49)) +(step t258 (cl @p_639) :rule resolution :premises (t255 t257)) +(step t259 (cl (! (= @p_557 @p_557) :named @p_648)) :rule eq_reflexive) +(step t260 (cl (not (! (= @p_646 @p_647) :named @p_649)) (! (not @p_579) :named @p_651) @p_582) :rule eq_transitive) +(step t261 (cl (! (not @p_585) :named @p_650) (not @p_648) @p_649) :rule eq_congruent) +(step t262 (cl @p_650 @p_649) :rule th_resolution :premises (t261 t259)) +(step t263 (cl @p_651 @p_582 @p_650) :rule th_resolution :premises (t260 t262)) +(step t264 (cl @p_650) :rule resolution :premises (t263 t238 t244)) +(step t265 (cl @p_634) :rule resolution :premises (t254 t264 t258)) +(step t266 (cl (or @p_638 (! (not (! (= @p_547 @p_652) :named @p_671)) :named @p_665))) :rule forall_inst :args ((:= veriT_vr3124 veriT_sk3))) +(step t267 (cl (or (! (not @p_577) :named @p_653) (! (or (! (not (! (= @p_652 @p_652) :named @p_658)) :named @p_659) (! (= @p_504 (! (fst$ @p_652) :named @p_675)) :named @p_657)) :named @p_654))) :rule forall_inst :args ((:= veriT_vr23 @p_504) (:= veriT_vr24 @p_652) (:= veriT_vr26 veriT_sk3))) +(step t268 (cl @p_653 @p_654) :rule or :premises (t267)) +(step t269 (cl (! (or @p_399 @p_654) :named @p_655) @p_436) :rule or_neg) +(step t270 (cl @p_655 @p_136) :rule th_resolution :premises (t121 t269)) +(step t271 (cl @p_655 (! (not @p_654) :named @p_656)) :rule or_neg) +(step t272 (cl @p_655) :rule th_resolution :premises (t232 t268 t270 t271)) +(anchor :step t273) +(assume t273.h1 @p_654) +(step t273.t2 (cl (= @p_658 true)) :rule eq_simplify) +(step t273.t3 (cl (= @p_659 (! (not true) :named @p_660))) :rule cong :premises (t273.t2)) +(step t273.t4 (cl (= @p_660 false)) :rule not_simplify) +(step t273.t5 (cl (= @p_659 false)) :rule trans :premises (t273.t3 t273.t4)) +(step t273.t6 (cl (= @p_654 (! (or false @p_657) :named @p_661))) :rule cong :premises (t273.t5)) +(step t273.t7 (cl (= @p_661 (! (or @p_657) :named @p_662))) :rule or_simplify) +(step t273.t8 (cl (= @p_662 @p_657)) :rule or_simplify) +(step t273.t9 (cl (! (= @p_654 @p_657) :named @p_663)) :rule trans :premises (t273.t6 t273.t7 t273.t8)) +(step t273.t10 (cl (not @p_663) @p_656 @p_657) :rule equiv_pos2) +(step t273.t11 (cl @p_657) :rule th_resolution :premises (t273.h1 t273.t9 t273.t10)) +(step t273 (cl @p_656 @p_657) :rule subproof :discharge (h1)) +(step t274 (cl @p_399 @p_654) :rule or :premises (t272)) +(step t275 (cl (! (or @p_399 @p_657) :named @p_664) @p_436) :rule or_neg) +(step t276 (cl @p_664 @p_136) :rule th_resolution :premises (t121 t275)) +(step t277 (cl @p_664 (! (not @p_657) :named @p_677)) :rule or_neg) +(step t278 (cl @p_664) :rule th_resolution :premises (t274 t273 t276 t277)) +(step t279 (cl @p_638 @p_665) :rule or :premises (t266)) +(step t280 (cl @p_665) :rule resolution :premises (t279 t265)) +(step t281 (cl @p_399 @p_657) :rule or :premises (t278)) +(step t282 (cl @p_657) :rule resolution :premises (t281 t49)) +(step t283 (cl (! (not @p_541) :named @p_668) (! (not @p_509) :named @p_666) (! (= @p_438 @p_652) :named @p_667)) :rule eq_transitive) +(step t284 (cl @p_666 @p_667 @p_538) :rule th_resolution :premises (t283 t201)) +(step t285 (cl @p_578 @p_668 @p_666 (! (= @p_545 @p_652) :named @p_669)) :rule eq_transitive) +(step t286 (cl @p_578 @p_666 @p_669 @p_538) :rule th_resolution :premises (t285 t201)) +(step t287 (cl (! (= @p_547 @p_547) :named @p_670)) :rule eq_reflexive) +(step t288 (cl (not @p_669) (! (not @p_556) :named @p_672) (! (not @p_670) :named @p_673) @p_671) :rule eq_transitive) +(step t289 (cl @p_672 @p_673 @p_671 @p_578 @p_666 @p_538) :rule th_resolution :premises (t288 t286)) +(step t290 (cl @p_672 @p_671 @p_578 @p_666 @p_538) :rule th_resolution :premises (t289 t287)) +(step t291 (cl @p_672) :rule resolution :premises (t290 a3 t146 t164 t280)) +(step t292 (cl (not @p_667) (! (= @p_674 @p_675) :named @p_676)) :rule eq_congruent) +(step t293 (cl @p_676 @p_666 @p_538) :rule th_resolution :premises (t292 t284)) +(step t294 (cl (! (not @p_439) :named @p_678) (not @p_676) @p_677 (! (= st$ @p_504) :named @p_679)) :rule eq_transitive) +(step t295 (cl @p_678 @p_677 @p_679 @p_666 @p_538) :rule th_resolution :premises (t294 t293)) +(step t296 (cl (not @p_679) (! (= @p_286 @p_518) :named @p_680)) :rule eq_congruent) +(step t297 (cl @p_680 @p_678 @p_677 @p_666 @p_538) :rule th_resolution :premises (t296 t295)) +(step t298 (cl @p_385) :rule resolution :premises (t221 t291)) +(step t299 (cl @p_384) :rule resolution :premises (t142 t298 t144)) +(step t300 (cl @p_505) :rule resolution :premises (t161 t299)) +(step t301 (cl @p_507) :rule resolution :premises (t159 t300 t162)) +(step t302 (cl (not @p_575) (! (not @p_680) :named @p_681) (! (not @p_507) :named @p_682) @p_531) :rule eq_congruent_pred) +(step t303 (cl @p_681 @p_682 @p_531 @p_573 @p_574 @p_381) :rule th_resolution :premises (t302 t230)) +(step t304 (cl @p_682 @p_531 @p_573 @p_574 @p_381 @p_678 @p_677 @p_666 @p_538) :rule th_resolution :premises (t303 t297)) +(step t305 (cl) :rule resolution :premises (t304 t299 t146 t235 t301 t164 t171 t196 t198 t282)) diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/SMT_Examples/SMT_Examples_Verit.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT_Examples/SMT_Examples_Verit.thy Mon Oct 12 18:59:44 2020 +0200 @@ -0,0 +1,741 @@ +(* 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]] + + +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)" + supply [[smt_timeout=360]] + 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) + +end \ No newline at end of file diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/SMT_Examples/SMT_Tests_Verit.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT_Examples/SMT_Tests_Verit.thy Mon Oct 12 18:59:44 2020 +0200 @@ -0,0 +1,717 @@ +(* Title: HOL/SMT_Examples/SMT_Tests.thy + Author: Sascha Boehme, TU Muenchen + Author: Mathias Fleury, MPII, JKU +*) + +section \Tests for the SMT binding\ + +theory SMT_Tests_Verit +imports Complex_Main +begin + +declare [[smt_solver=verit]] +smt_status + +text \Most examples are taken from the equivalent Z3 theory called \<^file>\SMT_Tests.thy\, +and have been taken from various Isabelle and HOL4 developments.\ + + +section \Propositional logic\ + +lemma + "True" + "\ False" + "\ \ True" + "True \ True" + "True \ False" + "False \ True" + "\ (False \ True)" + by smt+ + +lemma + "P \ \ P" + "\ (P \ \ P)" + "(True \ P) \ \ P \ (False \ P) \ P" + "P \ P" + "P \ \ P \ False" + "P \ Q \ Q \ P" + "P \ Q \ Q \ P" + "P \ Q \ P \ Q" + "\ (P \ Q) \ \ P" + "\ (P \ Q) \ \ Q" + "\ P \ \ (P \ Q)" + "\ Q \ \ (P \ Q)" + "(P \ Q) \ (\ (\ P \ \ Q))" + "(P \ Q) \ R \ P \ (Q \ R)" + "(P \ Q) \ R \ P \ (Q \ R)" + "(P \ Q) \ R \ (P \ R) \ (Q \ R)" + "(P \ R) \ (Q \ R) \ (P \ Q) \ R" + "(P \ Q) \ R \ (P \ R) \ (Q \ R)" + "(P \ R) \ (Q \ R) \ (P \ Q) \ R" + "((P \ Q) \ P) \ P" + "(P \ R) \ (Q \ R) \ (P \ Q \ R)" + "(P \ Q \ R) \ (P \ (Q \ R))" + "((P \ R) \ R) \ ((Q \ R) \ R) \ (P \ Q \ R) \ R" + "\ (P \ R) \ \ (Q \ R) \ \ (P \ Q \ R)" + "(P \ Q \ R) \ (P \ Q) \ (P \ R)" + "P \ (Q \ P)" + "(P \ Q \ R) \ (P \ Q)\ (P \ R)" + "(P \ Q) \ (P \ R) \ (P \ Q \ R)" + "((((P \ Q) \ P) \ P) \ Q) \ Q" + "(P \ Q) \ (\ Q \ \ P)" + "(P \ Q \ R) \ (P \ Q) \ (P \ R)" + "(P \ Q) \ (Q \ P) \ (P \ Q)" + "(P \ Q) \ (Q \ P)" + "\ (P \ \ P)" + "(P \ Q) \ (\ Q \ \ P)" + "P \ P \ P \ P \ P \ P \ P \ P \ P \ P" + by smt+ + +lemma + "(if P then Q1 else Q2) \ ((P \ Q1) \ (\ P \ Q2))" + "if P then (Q \ P) else (P \ Q)" + "(if P1 \ P2 then Q1 else Q2) \ (if P1 then Q1 else if P2 then Q1 else Q2)" + "(if P1 \ P2 then Q1 else Q2) \ (if P1 then if P2 then Q1 else Q2 else Q2)" + "(P1 \ (if P2 then Q1 else Q2)) \ + (if P1 \ P2 then P1 \ Q1 else P1 \ Q2)" + by smt+ + +lemma + "case P of True \ P | False \ \ P" + "case P of False \ \ P | True \ P" + "case \ P of True \ \ P | False \ P" + "case P of True \ (Q \ P) | False \ (P \ Q)" + by smt+ + + +section \First-order logic with equality\ + +lemma + "x = x" + "x = y \ y = x" + "x = y \ y = z \ x = z" + "x = y \ f x = f y" + "x = y \ g x y = g y x" + "f (f x) = x \ f (f (f (f (f x)))) = x \ f x = x" + "((if a then b else c) = d) = ((a \ (b = d)) \ (\ a \ (c = d)))" + by smt+ + +lemma + "\x. x = x" + "(\x. P x) \ (\y. P y)" + "\x. P x \ (\y. P x \ P y)" + "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" + "(\x. P x) \ R \ (\x. P x \ R)" + "(\x y z. S x z) \ (\x z. S x z)" + "(\x y. S x y \ S y x) \ (\x. S x y) \ S y x" + "(\x. P x \ P (f x)) \ P d \ P (f(f(f(d))))" + "(\x y. s x y = s y x) \ a = a \ s a b = s b a" + "(\s. q s \ r s) \ \ r s \ (\s. \ r s \ \ q s \ p t \ q t) \ p t \ r t" + by smt+ + +lemma + "(\x. P x) \ R \ (\x. P x \ R)" + by smt + +lemma + "\x. x = x" + "(\x. P x) \ (\y. P y)" + "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" + "(\x. P x) \ R \ (\x. P x \ R)" + "(\x y z. S x z) \ (\x z. S x z)" + "\ ((\x. \ P x) \ ((\x. P x) \ (\x. P x \ Q x)) \ \ (\x. P x))" + by smt+ + +lemma + "\x y. x = y" + "(\x. P x) \ R \ (\x. P x \ R)" + "\x. P x \ P a \ P b" + "(\x. Q \ P x) \ (Q \ (\x. P x))" + supply[[smt_trace]] + by smt+ + +lemma + "(\ (\x. P x)) \ (\x. \ P x)" + "(\x. P x \ Q) \ (\x. P x) \ Q" + "(\x y. R x y = x) \ (\y. R x y) = R x c" + "(if P x then \ (\y. P y) else (\y. \ P y)) \ P x \ P y" + "(\x y. R x y = x) \ (\x. \y. R x y) = (\x. R x c) \ (\y. R x y) = R x c" + by smt+ + +lemma + "\x. \y. f x y = f x (g x)" + "(\ \ (\x. P x)) \ (\ (\x. \ P x))" + "\u. \v. \w. \x. f u v w x = f u (g u) w (h u w)" + "\x. if x = y then (\y. y = x \ y \ x) else (\y. y = (x, x) \ y \ (x, x))" + "\x. if x = y then (\y. y = x \ y \ x) else (\y. y = (x, x) \ y \ (x, x))" + "(\x. \y. P x \ P y) \ ((\x. P x) \ (\y. P y))" + "(\y. \x. R x y) \ (\x. \y. R x y)" + by smt+ + +lemma + "(\!x. P x) \ (\x. P x)" + "(\!x. P x) \ (\x. P x \ (\y. y \ x \ \ P y))" + "P a \ (\x. P x \ x = a) \ (\!x. P x)" + "(\x. P x) \ (\x y. P x \ P y \ x = y) \ (\!x. P x)" + "(\!x. P x) \ (\x. P x \ (\y. P y \ y = x) \ R) \ R" + by smt+ + +lemma + "(\x\M. P x) \ c \ M \ P c" + "(\x\M. P x) \ \ (P c \ c \ M)" + by smt+ + +lemma + "let P = True in P" + "let P = P1 \ P2 in P \ \ P" + "let P1 = True; P2 = False in P1 \ P2 \ P2 \ P1" + "(let x = y in x) = y" + "(let x = y in Q x) \ (let z = y in Q z)" + "(let x = y1; z = y2 in R x z) \ (let z = y2; x = y1 in R x z)" + "(let x = y1; z = y2 in R x z) \ (let z = y1; x = y2 in R z x)" + "let P = (\x. Q x) in if P then P else \ P" + by smt+ + +lemma + "a \ b \ a \ c \ b \ c \ (\x y. f x = f y \ y = x) \ f a \ f b" + by smt + +lemma + "(\x y z. f x y = f x z \ y = z) \ b \ c \ f a b \ f a c" + "(\x y z. f x y = f z y \ x = z) \ a \ d \ f a b \ f d b" + by smt+ + + +section \Guidance for quantifier heuristics: patterns\ + +lemma + assumes "\x. + SMT.trigger (SMT.Symb_Cons (SMT.Symb_Cons (SMT.pat (f x)) SMT.Symb_Nil) SMT.Symb_Nil) + (f x = x)" + shows "f 1 = 1" + using assms by smt + +lemma + assumes "\x y. + SMT.trigger (SMT.Symb_Cons (SMT.Symb_Cons (SMT.pat (f x)) + (SMT.Symb_Cons (SMT.pat (g y)) SMT.Symb_Nil)) SMT.Symb_Nil) (f x = g y)" + shows "f a = g b" + using assms by smt + + +section \Meta-logical connectives\ + +lemma + "True \ True" + "False \ True" + "False \ False" + "P' x \ P' x" + "P \ P \ Q" + "Q \ P \ Q" + "\ P \ P \ Q" + "Q \ P \ Q" + "\P; \ Q\ \ \ (P \ Q)" + "P' x \ P' x" + "P' x \ Q' x \ P' x = Q' x" + "P' x = Q' x \ P' x \ Q' x" + "x \ y \ y \ z \ x \ (z::'a::type)" + "x \ y \ (f x :: 'b::type) \ f y" + "(\x. g x) \ g a \ a" + "(\x y. h x y \ h y x) \ \x. h x x" + "(p \ q) \ \ p \ q" + "(a \ b) \ (c \ d) \ (a \ b) \ (c \ d)" + by smt+ + + +section \Natural numbers\ + +declare [[smt_nat_as_int]] + +lemma + "(0::nat) = 0" + "(1::nat) = 1" + "(0::nat) < 1" + "(0::nat) \ 1" + "(123456789::nat) < 2345678901" + by smt+ + +lemma + "Suc 0 = 1" + "Suc x = x + 1" + "x < Suc x" + "(Suc x = Suc y) = (x = y)" + "Suc (x + y) < Suc x + Suc y" + by smt+ + +lemma + "(x::nat) + 0 = x" + "0 + x = x" + "x + y = y + x" + "x + (y + z) = (x + y) + z" + "(x + y = 0) = (x = 0 \ y = 0)" + by smt+ + +lemma + "(x::nat) - 0 = x" + "x < y \ x - y = 0" + "x - y = 0 \ y - x = 0" + "(x - y) + y = (if x < y then y else x)" + "x - y - z = x - (y + z)" + by smt+ + +lemma + "(x::nat) * 0 = 0" + "0 * x = 0" + "x * 1 = x" + "1 * x = x" + "3 * x = x * 3" + by smt+ + +lemma + "min (x::nat) y \ x" + "min x y \ y" + "min x y \ x + y" + "z < x \ z < y \ z < min x y" + "min x y = min y x" + "min x 0 = 0" + by smt+ + +lemma + "max (x::nat) y \ x" + "max x y \ y" + "max x y \ (x - y) + (y - x)" + "z > x \ z > y \ z > max x y" + "max x y = max y x" + "max x 0 = x" + by smt+ + +lemma + "0 \ (x::nat)" + "0 < x \ x \ 1 \ x = 1" + "x \ x" + "x \ y \ 3 * x \ 3 * y" + "x < y \ 3 * x < 3 * y" + "x < y \ x \ y" + "(x < y) = (x + 1 \ y)" + "\ (x < x)" + "x \ y \ y \ z \ x \ z" + "x < y \ y \ z \ x \ z" + "x \ y \ y < z \ x \ z" + "x < y \ y < z \ x < z" + "x < y \ y < z \ \ (z < x)" + by smt+ + +declare [[smt_nat_as_int = false]] + + +section \Integers\ + +lemma + "(0::int) = 0" + "(0::int) = (- 0)" + "(1::int) = 1" + "\ (-1 = (1::int))" + "(0::int) < 1" + "(0::int) \ 1" + "-123 + 345 < (567::int)" + "(123456789::int) < 2345678901" + "(-123456789::int) < 2345678901" + by smt+ + +lemma + "(x::int) + 0 = x" + "0 + x = x" + "x + y = y + x" + "x + (y + z) = (x + y) + z" + "(x + y = 0) = (x = -y)" + by smt+ + +lemma + "(-1::int) = - 1" + "(-3::int) = - 3" + "-(x::int) < 0 \ x > 0" + "x > 0 \ -x < 0" + "x < 0 \ -x > 0" + by smt+ + +lemma + "(x::int) - 0 = x" + "0 - x = -x" + "x < y \ x - y < 0" + "x - y = -(y - x)" + "x - y = -y + x" + "x - y - z = x - (y + z)" + by smt+ + +lemma + "(x::int) * 0 = 0" + "0 * x = 0" + "x * 1 = x" + "1 * x = x" + "x * -1 = -x" + "-1 * x = -x" + "3 * x = x * 3" + by smt+ + +lemma + "\x::int\ \ 0" + "(\x\ = 0) = (x = 0)" + "(x \ 0) = (\x\ = x)" + "(x \ 0) = (\x\ = -x)" + "\\x\\ = \x\" + by smt+ + +lemma + "min (x::int) y \ x" + "min x y \ y" + "z < x \ z < y \ z < min x y" + "min x y = min y x" + "x \ 0 \ min x 0 = 0" + "min x y \ \x + y\" + by smt+ + +lemma + "max (x::int) y \ x" + "max x y \ y" + "z > x \ z > y \ z > max x y" + "max x y = max y x" + "x \ 0 \ max x 0 = x" + "max x y \ - \x\ - \y\" + by smt+ + +lemma + "0 < (x::int) \ x \ 1 \ x = 1" + "x \ x" + "x \ y \ 3 * x \ 3 * y" + "x < y \ 3 * x < 3 * y" + "x < y \ x \ y" + "(x < y) = (x + 1 \ y)" + "\ (x < x)" + "x \ y \ y \ z \ x \ z" + "x < y \ y \ z \ x \ z" + "x \ y \ y < z \ x \ z" + "x < y \ y < z \ x < z" + "x < y \ y < z \ \ (z < x)" + by smt+ + + +section \Reals\ + +lemma + "(0::real) = 0" + "(0::real) = -0" + "(0::real) = (- 0)" + "(1::real) = 1" + "\ (-1 = (1::real))" + "(0::real) < 1" + "(0::real) \ 1" + "-123 + 345 < (567::real)" + "(123456789::real) < 2345678901" + "(-123456789::real) < 2345678901" + by smt+ + +lemma + "(x::real) + 0 = x" + "0 + x = x" + "x + y = y + x" + "x + (y + z) = (x + y) + z" + "(x + y = 0) = (x = -y)" + by smt+ + +lemma + "(-1::real) = - 1" + "(-3::real) = - 3" + "-(x::real) < 0 \ x > 0" + "x > 0 \ -x < 0" + "x < 0 \ -x > 0" + by smt+ + +lemma + "(x::real) - 0 = x" + "0 - x = -x" + "x < y \ x - y < 0" + "x - y = -(y - x)" + "x - y = -y + x" + "x - y - z = x - (y + z)" + by smt+ + +lemma + "(x::real) * 0 = 0" + "0 * x = 0" + "x * 1 = x" + "1 * x = x" + "x * -1 = -x" + "-1 * x = -x" + "3 * x = x * 3" + by smt+ + +lemma + "\x::real\ \ 0" + "(\x\ = 0) = (x = 0)" + "(x \ 0) = (\x\ = x)" + "(x \ 0) = (\x\ = -x)" + "\\x\\ = \x\" + by smt+ + +lemma + "min (x::real) y \ x" + "min x y \ y" + "z < x \ z < y \ z < min x y" + "min x y = min y x" + "x \ 0 \ min x 0 = 0" + "min x y \ \x + y\" + by smt+ + +lemma + "max (x::real) y \ x" + "max x y \ y" + "z > x \ z > y \ z > max x y" + "max x y = max y x" + "x \ 0 \ max x 0 = x" + "max x y \ - \x\ - \y\" + by smt+ + +lemma + "x \ (x::real)" + "x \ y \ 3 * x \ 3 * y" + "x < y \ 3 * x < 3 * y" + "x < y \ x \ y" + "\ (x < x)" + "x \ y \ y \ z \ x \ z" + "x < y \ y \ z \ x \ z" + "x \ y \ y < z \ x \ z" + "x < y \ y < z \ x < z" + "x < y \ y < z \ \ (z < x)" + by smt+ + + +section \Datatypes, records, and typedefs\ + +subsection \Without support by the SMT solver\ + +subsubsection \Algebraic datatypes\ + +lemma + "x = fst (x, y)" + "y = snd (x, y)" + "((x, y) = (y, x)) = (x = y)" + "((x, y) = (u, v)) = (x = u \ y = v)" + "(fst (x, y, z) = fst (u, v, w)) = (x = u)" + "(snd (x, y, z) = snd (u, v, w)) = (y = v \ z = w)" + "(fst (snd (x, y, z)) = fst (snd (u, v, w))) = (y = v)" + "(snd (snd (x, y, z)) = snd (snd (u, v, w))) = (z = w)" + "(fst (x, y) = snd (x, y)) = (x = y)" + "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" + "(fst (x, y) = snd (x, y)) = (x = y)" + "(fst p = snd p) = (p = (snd p, fst p))" + using fst_conv snd_conv prod.collapse + by smt+ + +lemma + "[x] \ Nil" + "[x, y] \ Nil" + "x \ y \ [x] \ [y]" + "hd (x # xs) = x" + "tl (x # xs) = xs" + "hd [x, y, z] = x" + "tl [x, y, z] = [y, z]" + "hd (tl [x, y, z]) = y" + "tl (tl [x, y, z]) = [z]" + using list.sel(1,3) list.simps + by smt+ + +lemma + "fst (hd [(a, b)]) = a" + "snd (hd [(a, b)]) = b" + using fst_conv snd_conv prod.collapse list.sel(1,3) list.simps + by smt+ + + +subsubsection \Records\ + +record point = + cx :: int + cy :: int + +record bw_point = point + + black :: bool + +lemma + "\cx = x, cy = y\ = \cx = x', cy = y'\ \ x = x' \ y = y'" + using point.simps + by smt + +lemma + "cx \ cx = 3, cy = 4 \ = 3" + "cy \ cx = 3, cy = 4 \ = 4" + "cx \ cx = 3, cy = 4 \ \ cy \ cx = 3, cy = 4 \" + "\ cx = 3, cy = 4 \ \ cx := 5 \ = \ cx = 5, cy = 4 \" + "\ cx = 3, cy = 4 \ \ cy := 6 \ = \ cx = 3, cy = 6 \" + "p = \ cx = 3, cy = 4 \ \ p \ cx := 3 \ \ cy := 4 \ = p" + "p = \ cx = 3, cy = 4 \ \ p \ cy := 4 \ \ cx := 3 \ = p" + using point.simps + by smt+ + +lemma + "\cx = x, cy = y, black = b\ = \cx = x', cy = y', black = b'\ \ x = x' \ y = y' \ b = b'" + using point.simps bw_point.simps + by smt + +lemma + "cx \ cx = 3, cy = 4, black = b \ = 3" + "cy \ cx = 3, cy = 4, black = b \ = 4" + "black \ cx = 3, cy = 4, black = b \ = b" + "cx \ cx = 3, cy = 4, black = b \ \ cy \ cx = 3, cy = 4, black = b \" + "\ cx = 3, cy = 4, black = b \ \ cx := 5 \ = \ cx = 5, cy = 4, black = b \" + "\ cx = 3, cy = 4, black = b \ \ cy := 6 \ = \ cx = 3, cy = 6, black = b \" + "p = \ cx = 3, cy = 4, black = True \ \ + p \ cx := 3 \ \ cy := 4 \ \ black := True \ = p" + "p = \ cx = 3, cy = 4, black = True \ \ + p \ cy := 4 \ \ black := True \ \ cx := 3 \ = p" + "p = \ cx = 3, cy = 4, black = True \ \ + p \ black := True \ \ cx := 3 \ \ cy := 4 \ = p" + using point.simps bw_point.simps + by smt+ + +lemma + "\ cx = 3, cy = 4, black = b \ \ black := w \ = \ cx = 3, cy = 4, black = w \" + "\ cx = 3, cy = 4, black = True \ \ black := False \ = + \ cx = 3, cy = 4, black = False \" + "p \ cx := 3 \ \ cy := 4 \ \ black := True \ = + p \ black := True \ \ cy := 4 \ \ cx := 3 \" + apply (smt add_One add_inc bw_point.update_convs(1) default_unit_def inc.simps(2) one_plus_BitM + semiring_norm(6,26)) + apply (smt bw_point.update_convs(1)) + apply (smt bw_point.cases_scheme bw_point.update_convs(1) point.update_convs(1,2)) + done + + +subsubsection \Type definitions\ + +typedef int' = "UNIV::int set" by (rule UNIV_witness) + +definition n0 where "n0 = Abs_int' 0" +definition n1 where "n1 = Abs_int' 1" +definition n2 where "n2 = Abs_int' 2" +definition plus' where "plus' n m = Abs_int' (Rep_int' n + Rep_int' m)" + +lemma + "n0 \ n1" + "plus' n1 n1 = n2" + "plus' n0 n2 = n2" + by (smt n0_def n1_def n2_def plus'_def Abs_int'_inverse Rep_int'_inverse UNIV_I)+ + + +subsection \With support by the SMT solver (but without proofs)\ + +subsubsection \Algebraic datatypes\ + +lemma + "x = fst (x, y)" + "y = snd (x, y)" + "((x, y) = (y, x)) = (x = y)" + "((x, y) = (u, v)) = (x = u \ y = v)" + "(fst (x, y, z) = fst (u, v, w)) = (x = u)" + "(snd (x, y, z) = snd (u, v, w)) = (y = v \ z = w)" + "(fst (snd (x, y, z)) = fst (snd (u, v, w))) = (y = v)" + "(snd (snd (x, y, z)) = snd (snd (u, v, w))) = (z = w)" + "(fst (x, y) = snd (x, y)) = (x = y)" + "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" + "(fst (x, y) = snd (x, y)) = (x = y)" + "(fst p = snd p) = (p = (snd p, fst p))" + using fst_conv snd_conv prod.collapse + by smt+ + +lemma + "x \ y \ [x] \ [y]" + "hd (x # xs) = x" + "tl (x # xs) = xs" + "hd [x, y, z] = x" + "tl [x, y, z] = [y, z]" + "hd (tl [x, y, z]) = y" + "tl (tl [x, y, z]) = [z]" + using list.sel(1,3) + by smt+ + +lemma + "fst (hd [(a, b)]) = a" + "snd (hd [(a, b)]) = b" + using fst_conv snd_conv prod.collapse list.sel(1,3) + by smt+ + + +subsubsection \Records\ +text \The equivalent theory for Z3 contains more example, but unlike Z3, we are able +to reconstruct the proofs.\ + +lemma + "cx \ cx = 3, cy = 4 \ = 3" + "cy \ cx = 3, cy = 4 \ = 4" + "cx \ cx = 3, cy = 4 \ \ cy \ cx = 3, cy = 4 \" + "\ cx = 3, cy = 4 \ \ cx := 5 \ = \ cx = 5, cy = 4 \" + "\ cx = 3, cy = 4 \ \ cy := 6 \ = \ cx = 3, cy = 6 \" + "p = \ cx = 3, cy = 4 \ \ p \ cx := 3 \ \ cy := 4 \ = p" + "p = \ cx = 3, cy = 4 \ \ p \ cy := 4 \ \ cx := 3 \ = p" + using point.simps + by smt+ + + +lemma + "cx \ cx = 3, cy = 4, black = b \ = 3" + "cy \ cx = 3, cy = 4, black = b \ = 4" + "black \ cx = 3, cy = 4, black = b \ = b" + "cx \ cx = 3, cy = 4, black = b \ \ cy \ cx = 3, cy = 4, black = b \" + "\ cx = 3, cy = 4, black = b \ \ cx := 5 \ = \ cx = 5, cy = 4, black = b \" + "\ cx = 3, cy = 4, black = b \ \ cy := 6 \ = \ cx = 3, cy = 6, black = b \" + "p = \ cx = 3, cy = 4, black = True \ \ + p \ cx := 3 \ \ cy := 4 \ \ black := True \ = p" + "p = \ cx = 3, cy = 4, black = True \ \ + p \ cy := 4 \ \ black := True \ \ cx := 3 \ = p" + "p = \ cx = 3, cy = 4, black = True \ \ + p \ black := True \ \ cx := 3 \ \ cy := 4 \ = p" + using point.simps bw_point.simps + by smt+ + + +section \Functions\ + +lemma "\f. map_option f (Some x) = Some (y + x)" + by (smt option.map(2)) + +lemma + "(f (i := v)) i = v" + "i1 \ i2 \ (f (i1 := v)) i2 = f i2" + "i1 \ i2 \ (f (i1 := v1, i2 := v2)) i1 = v1" + "i1 \ i2 \ (f (i1 := v1, i2 := v2)) i2 = v2" + "i1 = i2 \ (f (i1 := v1, i2 := v2)) i1 = v2" + "i1 = i2 \ (f (i1 := v1, i2 := v2)) i1 = v2" + "i1 \ i2 \i1 \ i3 \ i2 \ i3 \ (f (i1 := v1, i2 := v2)) i3 = f i3" + using fun_upd_same fun_upd_apply + by smt+ + + +section \Sets\ + +lemma Empty: "x \ {}" by simp + +lemmas smt_sets = Empty UNIV_I Un_iff Int_iff + +lemma + "x \ {}" + "x \ UNIV" + "x \ A \ B \ x \ A \ x \ B" + "x \ P \ {} \ x \ P" + "x \ P \ UNIV" + "x \ P \ Q \ x \ Q \ P" + "x \ P \ P \ x \ P" + "x \ P \ (Q \ R) \ x \ (P \ Q) \ R" + "x \ A \ B \ x \ A \ x \ B" + "x \ P \ {}" + "x \ P \ UNIV \ x \ P" + "x \ P \ Q \ x \ Q \ P" + "x \ P \ P \ x \ P" + "x \ P \ (Q \ R) \ x \ (P \ Q) \ R" + "{x. x \ P} = {y. y \ P}" + by (smt smt_sets)+ + +end diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smt_config.ML --- a/src/HOL/Tools/SMT/smt_config.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/smt_config.ML Mon Oct 12 18:59:44 2020 +0200 @@ -38,6 +38,7 @@ val infer_triggers: bool Config.T val debug_files: string Config.T val sat_solver: string Config.T + val compress_verit_proofs: Proof.context -> bool (*tools*) val with_time_limit: Proof.context -> real Config.T -> ('a -> 'b) -> 'a -> 'b @@ -47,7 +48,10 @@ val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit val verbose_msg: Proof.context -> ('a -> string) -> 'a -> unit val statistics_msg: Proof.context -> ('a -> string) -> 'a -> unit - val veriT_msg: Proof.context -> (unit -> 'a) -> unit + val verit_msg: Proof.context -> (unit -> 'a) -> unit + val verit_arith_msg: Proof.context -> (unit -> 'a) -> unit + val spy_verit: Proof.context -> bool + val spy_Z3: Proof.context -> bool (*certificates*) val select_certificates: string -> Context.generic -> Context.generic @@ -180,7 +184,11 @@ val read_only_certificates = Attrib.setup_config_bool \<^binding>\smt_read_only_certificates\ (K false) val verbose = Attrib.setup_config_bool \<^binding>\smt_verbose\ (K true) val trace = Attrib.setup_config_bool \<^binding>\smt_trace\ (K false) -val trace_veriT = Attrib.setup_config_bool \<^binding>\smt_debug_verit\ (K false) +val trace_verit = Attrib.setup_config_bool \<^binding>\smt_debug_verit\ (K false) +val spy_verit_attr = Attrib.setup_config_bool \<^binding>\smt_spy_verit\ (K false) +val spy_z3 = Attrib.setup_config_bool \<^binding>\smt_spy_z3\ (K false) +val trace_arith_verit = Attrib.setup_config_bool \<^binding>\smt_debug_arith_verit\ (K false) +val trace_verit_compression = Attrib.setup_config_bool \<^binding>\verit_compress_proofs\ (K false) val statistics = Attrib.setup_config_bool \<^binding>\smt_statistics\ (K false) val monomorph_limit = Attrib.setup_config_int \<^binding>\smt_monomorph_limit\ (K 10) val monomorph_instances = Attrib.setup_config_int \<^binding>\smt_monomorph_instances\ (K 500) @@ -200,7 +208,12 @@ fun trace_msg ctxt = cond_trace (Config.get ctxt trace) fun statistics_msg ctxt = cond_trace (Config.get ctxt statistics) -fun veriT_msg ctxt (x : unit -> 'a) = if (Config.get ctxt trace_veriT) then ignore(x ()) else () +fun verit_msg ctxt (x : unit -> 'a) = if (Config.get ctxt trace_verit) then ignore(x ()) else () +fun verit_arith_msg ctxt (x : unit -> 'a) = if (Config.get ctxt trace_arith_verit) then ignore(x ()) else () + +fun spy_verit ctxt = Config.get ctxt spy_verit_attr +fun spy_Z3 ctxt = Config.get ctxt spy_z3 +fun compress_verit_proofs ctxt = Config.get ctxt trace_verit_compression (* tools *) diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smt_replay.ML --- a/src/HOL/Tools/SMT/smt_replay.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/smt_replay.ML Mon Oct 12 18:59:44 2020 +0200 @@ -39,6 +39,10 @@ (*theorem transformation*) val varify: Proof.context -> thm -> thm val params_of: term -> (string * typ) list + + (*spy*) + val spying: bool -> Proof.context -> (unit -> string) -> string -> unit + val print_stats: (string * int list) list -> string end; structure SMT_Replay : SMT_REPLAY = @@ -249,13 +253,14 @@ fun pretty_statistics solver total stats = let + val stats = Symtab.map (K (map (fn i => curry Int.div i 1000000))) stats fun mean_of is = let val len = length is val mid = len div 2 in if len mod 2 = 0 then (nth is (mid - 1) + nth is mid) div 2 else nth is mid end fun pretty_item name p = Pretty.item (Pretty.separate ":" [Pretty.str name, p]) - fun pretty (name, milliseconds) = pretty_item name (Pretty.block (Pretty.separate "," [ + fun pretty (name, milliseconds) = (Pretty.block (Pretty.str (name ^": ") :: Pretty.separate "," [ Pretty.str (string_of_int (length milliseconds) ^ " occurrences") , Pretty.str (string_of_int (mean_of milliseconds) ^ " ms mean time"), Pretty.str (string_of_int (fold Integer.max milliseconds 0) ^ " ms maximum time"), @@ -266,4 +271,26 @@ map pretty (Symtab.dest stats)) end +fun timestamp_format time = + Date.fmt "%Y-%m-%d %H:%M:%S." (Date.fromTimeLocal time) ^ + (StringCvt.padLeft #"0" 3 (string_of_int (Time.toMilliseconds time - 1000 * Time.toSeconds time))) + +fun print_stats stats = + let + fun print_list xs = fold (fn x => fn msg => msg ^ string_of_int x ^ ",") xs "" + in + fold (fn (x,y) => fn msg => msg ^ x ^ ": " ^ print_list y ^ "\n") stats "" + end + +fun spying false _ _ _ = () + | spying true ctxt f filename = + let + val message = f () + val thy = Context.theory_long_name ((Context.theory_of o Context.Proof) ctxt) + val spying_version = "1" + in + File.append (Path.explode ("$ISABELLE_HOME_USER/" ^ filename)) + (spying_version ^ "; " ^ thy ^ "; " ^ (timestamp_format (Time.now ())) ^ ";\n" ^ message ^ "\n") + end + end; diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smt_replay_arith.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/SMT/smt_replay_arith.ML Mon Oct 12 18:59:44 2020 +0200 @@ -0,0 +1,117 @@ +(* Title: HOL/Tools/SMT/smt_replay_arith.ML + Author: Mathias Fleury, MPII, JKU + +Proof methods for replaying arithmetic steps with Farkas Lemma. +*) + +signature SMT_REPLAY_ARITH = +sig + val la_farkas_tac: Subgoal.focus -> term list -> thm -> thm Seq.seq; + val la_farkas: term list -> Proof.context -> int -> tactic; +end; + +structure SMT_Replay_Arith: SMT_REPLAY_ARITH = +struct + +fun extract_number (Const ("SMT.z3div", _) $ (Const ("Groups.uminus_class.uminus", _) $ n1) $ n2) = + (n1, n2, false) + | extract_number (Const ("SMT.z3div", _) $ n1 $ n2) = (n1, n2, true) + | extract_number (Const ("Groups.uminus_class.uminus", _) $ n) = + (n, HOLogic.numeral_const (fastype_of n) $ HOLogic.one_const, false) + | extract_number n = (n, HOLogic.numeral_const (fastype_of n) $ HOLogic.one_const, true) + +fun try_OF _ thm1 thm2 = + (thm1 OF [thm2]) + |> single + handle THM _ => [] + +fun try_OF_inst ctxt thm1 ((r, q, pos), thm2) = + map (fn thm => + (Drule.infer_instantiate' ctxt (map SOME [Thm.cterm_of ctxt q, Thm.cterm_of ctxt r]) thm, + pos)) + (try_OF ctxt thm1 thm2) + handle THM _ => [] + +fun try_OF_insts ctxt thms thm = map (fn thm1 => try_OF_inst ctxt thm1 thm) thms |> flat + +fun try_OFs ctxt thms thm = map (fn thm1 => try_OF ctxt thm1 thm) thms |> flat + +fun la_farkas_tac ({context = ctxt, prems, concl, ...}: Subgoal.focus) args thm = + let + fun trace v = (SMT_Config.verit_arith_msg ctxt v; v) + val _ = trace (fn () => @{print} (prems, concl)) + val arith_thms = Named_Theorems.get ctxt @{named_theorems smt_arith_simplify} + val verit_arith = Named_Theorems.get ctxt @{named_theorems smt_arith_combine} + val verit_arith_multiplication = Named_Theorems.get ctxt @{named_theorems smt_arith_multiplication} + val _ = trace (fn () => @{print} "******************************************************") + val _ = trace (fn () => @{print} " Multiply by constants ") + val _ = trace (fn () => @{print} "******************************************************") + val coeff_prems = + map extract_number args + |> (fn xs => ListPair.zip (xs, prems)) + val _ = trace (fn () => @{print} coeff_prems) + fun negate_if_needed (thm, pos) = + if pos then thm + else map (fn thm0 => try_OF ctxt thm0 thm) @{thms verit_negate_coefficient} + |> flat |> the_single + val normalized = + map (try_OF_insts ctxt verit_arith_multiplication) coeff_prems + |> flat + |> map negate_if_needed + fun import_assumption thm (Trues, ctxt) = + let val (assms, ctxt) = Assumption.add_assumes ((map (Thm.cterm_of ctxt) o Thm.prems_of) thm) ctxt in + (thm OF assms, (Trues + length assms, ctxt)) end + val (n :: normalized, (Trues, ctxt_with_assms)) = fold_map import_assumption normalized (0, ctxt) + + val _ = trace (fn () => @{print} (n :: normalized)) + val _ = trace (fn () => @{print} "*****************************************************") + val _ = trace (fn () => @{print} " Combine equalities ") + val _ = trace (fn () => @{print} "******************************************************") + + val combined = + List.foldl + (fn (thm1, thm2) => + try_OFs ctxt verit_arith thm1 + |> (fn thms => try_OFs ctxt thms thm2) + |> the_single) + n + normalized + |> singleton (Proof_Context.export ctxt_with_assms ctxt) + val _ = trace (fn () => @{print} combined) + fun arith_full_simps ctxt thms = + ctxt + |> empty_simpset + |> put_simpset HOL_basic_ss + |> (fn ctxt => ctxt addsimps thms + addsimps arith_thms + addsimprocs [@{simproc int_div_cancel_numeral_factors}, @{simproc int_combine_numerals}, + @{simproc divide_cancel_numeral_factor}, @{simproc intle_cancel_numerals}, + @{simproc field_combine_numerals}, @{simproc intless_cancel_numerals}]) + |> asm_full_simplify + val final_False = combined + |> arith_full_simps ctxt (Named_Theorems.get ctxt @{named_theorems ac_simps}) + val _ = trace (fn () => @{print} final_False) + val final_theorem = try_OF ctxt thm final_False + in + (case final_theorem of + [thm] => Seq.single (thm OF replicate Trues @{thm TrueI}) + | _ => Seq.empty) + end + +fun TRY' tac = fn i => TRY (tac i) +fun REPEAT' tac = fn i => REPEAT (tac i) + +fun rewrite_only_thms ctxt thms = + ctxt + |> empty_simpset + |> put_simpset HOL_basic_ss + |> (fn ctxt => ctxt addsimps thms) + |> Simplifier.full_simp_tac + +fun la_farkas args ctxt = + REPEAT' (resolve_tac ctxt @{thms verit_and_pos(2,3)}) + THEN' TRY' (resolve_tac ctxt @{thms ccontr}) + THEN' TRY' (rewrite_only_thms ctxt @{thms linorder_class.not_less linorder_class.not_le not_not}) + THEN' (Subgoal.FOCUS (fn focus => la_farkas_tac focus args) ctxt) + +end \ No newline at end of file diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smt_replay_methods.ML --- a/src/HOL/Tools/SMT/smt_replay_methods.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/smt_replay_methods.ML Mon Oct 12 18:59:44 2020 +0200 @@ -8,12 +8,14 @@ signature SMT_REPLAY_METHODS = sig + (*Printing*) val pretty_goal: Proof.context -> string -> string -> thm list -> term -> Pretty.T val trace_goal: Proof.context -> string -> thm list -> term -> unit val trace: Proof.context -> (unit -> string) -> unit + (*Replaying*) val replay_error: Proof.context -> string -> string -> thm list -> term -> 'a - val replay_rule_error: Proof.context -> string -> thm list -> term -> 'a + val replay_rule_error: string -> Proof.context -> string -> thm list -> term -> 'a (*theory lemma methods*) type th_lemma_method = Proof.context -> thm list -> term -> thm @@ -24,6 +26,10 @@ val match_instantiate: Proof.context -> term -> thm -> thm val prove: Proof.context -> term -> (Proof.context -> int -> tactic) -> thm + val prove_arith_rewrite: ((term -> int * term Termtab.table -> + term * (int * term Termtab.table)) -> term -> int * term Termtab.table -> + term * (int * term Termtab.table)) -> Proof.context -> term -> thm + (*abstraction*) type abs_context = int * term Termtab.table type 'a abstracter = term -> abs_context -> 'a * abs_context @@ -36,24 +42,44 @@ val abstract_not: (term -> abs_context -> term * abs_context) -> term -> abs_context -> term * abs_context val abstract_unit: term -> abs_context -> term * abs_context + val abstract_bool: term -> abs_context -> term * abs_context + (* also abstracts over equivalences and conjunction and implication*) + val abstract_bool_shallow: term -> abs_context -> term * abs_context + (* abstracts down to equivalence symbols *) + val abstract_bool_shallow_equivalence: term -> abs_context -> term * abs_context val abstract_prop: term -> abs_context -> term * abs_context val abstract_term: term -> abs_context -> term * abs_context + val abstract_eq: (term -> int * term Termtab.table -> term * (int * term Termtab.table)) -> + term -> int * term Termtab.table -> term * (int * term Termtab.table) + val abstract_neq: (term -> int * term Termtab.table -> term * (int * term Termtab.table)) -> + term -> int * term Termtab.table -> term * (int * term Termtab.table) val abstract_arith: Proof.context -> term -> abs_context -> term * abs_context + (* also abstracts over if-then-else *) + val abstract_arith_shallow: Proof.context -> term -> abs_context -> term * abs_context val prove_abstract: Proof.context -> thm list -> term -> (Proof.context -> thm list -> int -> tactic) -> (abs_context -> (term list * term) * abs_context) -> thm val prove_abstract': Proof.context -> term -> (Proof.context -> thm list -> int -> tactic) -> (abs_context -> term * abs_context) -> thm - val try_provers: Proof.context -> string -> (string * (term -> 'a)) list -> thm list -> term -> - 'a + val try_provers: string -> Proof.context -> string -> (string * (term -> 'a)) list -> thm list -> + term -> 'a (*shared tactics*) + val cong_unfolding_trivial: Proof.context -> thm list -> term -> thm val cong_basic: Proof.context -> thm list -> term -> thm val cong_full: Proof.context -> thm list -> term -> thm val cong_unfolding_first: Proof.context -> thm list -> term -> thm + val arith_th_lemma: Proof.context -> thm list -> term -> thm + val arith_th_lemma_wo: Proof.context -> thm list -> term -> thm + val arith_th_lemma_wo_shallow: Proof.context -> thm list -> term -> thm + val arith_th_lemma_tac_with_wo: Proof.context -> thm list -> int -> tactic + + val dest_thm: thm -> term + val prop_tac: Proof.context -> thm list -> int -> tactic val certify_prop: Proof.context -> term -> cterm + val dest_prop: term -> term end; @@ -77,7 +103,7 @@ fun replay_error ctxt msg rule thms t = error (Pretty.string_of (pretty_goal ctxt msg rule thms t)) -fun replay_rule_error ctxt = replay_error ctxt "Failed to replay Z3 proof step" +fun replay_rule_error name ctxt = replay_error ctxt ("Failed to replay" ^ name ^ " proof step") fun trace_goal ctxt rule thms t = trace ctxt (fn () => Pretty.string_of (pretty_goal ctxt "Goal" rule thms t)) @@ -221,7 +247,7 @@ fun abstract_arith ctxt u = let - fun abs (t as (c as Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ Abs (s, T, t'))) = + fun abs (t as (Const (\<^const_name>\HOL.The\, _) $ Abs (_, _, _))) = abstract_sub t (abstract_term t) | abs (t as (c as Const _) $ Abs (s, T, t')) = abstract_sub t (abs t' #>> (fn u' => c $ Abs (s, T, u'))) @@ -267,7 +293,7 @@ abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>> HOLogic.mk_eq) else abstract_lit t - | abstract_unit (t as (\<^const>\HOL.Not\ $ Const(\<^const_name>\HOL.eq\, _) $ t1 $ t2)) = + | abstract_unit (t as (\<^const>\HOL.Not\ $ (Const(\<^const_name>\HOL.eq\, _) $ t1 $ t2))) = if fastype_of t1 = \<^typ>\bool\ then abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>> HOLogic.mk_eq #>> HOLogic.mk_not) @@ -276,14 +302,123 @@ abstract_sub t (abstract_unit t1 #>> HOLogic.mk_not) | abstract_unit t = abstract_lit t +fun abstract_bool (t as (@{const HOL.disj} $ t1 $ t2)) = + abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> + HOLogic.mk_disj) + | abstract_bool (t as (@{const HOL.conj} $ t1 $ t2)) = + abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> + HOLogic.mk_conj) + | abstract_bool (t as (Const(@{const_name HOL.eq}, _) $ t1 $ t2)) = + if fastype_of t1 = @{typ bool} then + abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> + HOLogic.mk_eq) + else abstract_lit t + | abstract_bool (t as (@{const HOL.Not} $ t1)) = + abstract_sub t (abstract_bool t1 #>> HOLogic.mk_not) + | abstract_bool (t as (@{const HOL.implies} $ t1 $ t2)) = + abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> + HOLogic.mk_imp) + | abstract_bool t = abstract_lit t + +fun abstract_bool_shallow (t as (@{const HOL.disj} $ t1 $ t2)) = + abstract_sub t (abstract_bool_shallow t1 ##>> abstract_bool_shallow t2 #>> + HOLogic.mk_disj) + | abstract_bool_shallow (t as (@{const HOL.Not} $ t1)) = + abstract_sub t (abstract_bool_shallow t1 #>> HOLogic.mk_not) + | abstract_bool_shallow t = abstract_term t + +fun abstract_bool_shallow_equivalence (t as (@{const HOL.disj} $ t1 $ t2)) = + abstract_sub t (abstract_bool_shallow_equivalence t1 ##>> abstract_bool_shallow_equivalence t2 #>> + HOLogic.mk_disj) + | abstract_bool_shallow_equivalence (t as (Const(@{const_name HOL.eq}, _) $ t1 $ t2)) = + if fastype_of t1 = @{typ bool} then + abstract_sub t (abstract_lit t1 ##>> abstract_lit t2 #>> + HOLogic.mk_eq) + else abstract_lit t + | abstract_bool_shallow_equivalence (t as (@{const HOL.Not} $ t1)) = + abstract_sub t (abstract_bool_shallow_equivalence t1 #>> HOLogic.mk_not) + | abstract_bool_shallow_equivalence t = abstract_lit t + +fun abstract_arith_shallow ctxt u = + let + fun abs (t as (Const (\<^const_name>\HOL.The\, _) $ Abs (_, _, _))) = + abstract_sub t (abstract_term t) + | abs (t as (c as Const _) $ Abs (s, T, t')) = + abstract_sub t (abs t' #>> (fn u' => c $ Abs (s, T, u'))) + | abs (t as (Const (\<^const_name>\If\, _)) $ _ $ _ $ _) = + abstract_sub t (abstract_term t) + | abs (t as \<^const>\HOL.Not\ $ t1) = abstract_sub t (abs t1 #>> HOLogic.mk_not) + | abs (t as \<^const>\HOL.disj\ $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> HOLogic.mk_disj) + | abs (t as (c as Const (\<^const_name>\uminus_class.uminus\, _)) $ t1) = + abstract_sub t (abs t1 #>> (fn u => c $ u)) + | abs (t as (c as Const (\<^const_name>\plus_class.plus\, _)) $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) + | abs (t as (c as Const (\<^const_name>\minus_class.minus\, _)) $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) + | abs (t as (c as Const (\<^const_name>\times_class.times\, _)) $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) + | abs (t as (Const (\<^const_name>\z3div\, _)) $ _ $ _) = + abstract_sub t (abstract_term t) + | abs (t as (Const (\<^const_name>\z3mod\, _)) $ _ $ _) = + abstract_sub t (abstract_term t) + | abs (t as (c as Const (\<^const_name>\HOL.eq\, _)) $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) + | abs (t as (c as Const (\<^const_name>\ord_class.less\, _)) $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) + | abs (t as (c as Const (\<^const_name>\ord_class.less_eq\, _)) $ t1 $ t2) = + abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) + | abs t = abstract_sub t (fn cx => + if can HOLogic.dest_number t then (t, cx) + else + (case apply_abstracters abs (get_arith_abstracters ctxt) t cx of + (SOME u, cx') => (u, cx') + | (NONE, _) => abstract_term t cx)) + in abs u end (* theory lemmas *) -fun try_provers ctxt rule [] thms t = replay_rule_error ctxt rule thms t - | try_provers ctxt rule ((name, prover) :: named_provers) thms t = +fun try_provers prover_name ctxt rule [] thms t = replay_rule_error prover_name ctxt rule thms t + | try_provers prover_name ctxt rule ((name, prover) :: named_provers) thms t = (case (trace ctxt (K ("Trying prover " ^ quote name)); try prover t) of SOME thm => thm - | NONE => try_provers ctxt rule named_provers thms t) + | NONE => try_provers prover_name ctxt rule named_provers thms t) + +(*theory-lemma*) + +fun arith_th_lemma_tac ctxt prems = + Method.insert_tac ctxt prems + THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms z3div_def z3mod_def}) + THEN' Arith_Data.arith_tac ctxt + +fun arith_th_lemma ctxt thms t = + prove_abstract ctxt thms t arith_th_lemma_tac ( + fold_map (abstract_arith ctxt o dest_thm) thms ##>> + abstract_arith ctxt (dest_prop t)) + +fun arith_th_lemma_tac_with_wo ctxt prems = + Method.insert_tac ctxt prems + THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms z3div_def z3mod_def int_distrib}) + THEN' Simplifier.asm_full_simp_tac + (empty_simpset ctxt addsimprocs [(*@{simproc linordered_ring_strict_eq_cancel_factor_poly}, + @{simproc linordered_ring_strict_le_cancel_factor_poly}, + @{simproc linordered_ring_strict_less_cancel_factor_poly}, + @{simproc nat_eq_cancel_factor_poly}, + @{simproc nat_le_cancel_factor_poly}, + @{simproc nat_less_cancel_factor_poly}*)]) + THEN' (fn i => TRY (Arith_Data.arith_tac ctxt i)) + +fun arith_th_lemma_wo ctxt thms t = + prove_abstract ctxt thms t arith_th_lemma_tac_with_wo ( + fold_map (abstract_arith ctxt o dest_thm) thms ##>> + abstract_arith ctxt (dest_prop t)) + +fun arith_th_lemma_wo_shallow ctxt thms t = + prove_abstract ctxt thms t arith_th_lemma_tac_with_wo ( + fold_map (abstract_arith_shallow ctxt o dest_thm) thms ##>> + abstract_arith_shallow ctxt (dest_prop t)) + +val _ = Theory.setup (Context.theory_map (add_th_lemma_method ("arith", arith_th_lemma))) (* congruence *) @@ -318,20 +453,66 @@ ORELSE' dresolve_tac ctxt cong_dest_rules THEN' cong_full_core_tac ctxt')) +local + val reorder_for_simp = try (fn thm => + let val t = Thm.prop_of (@{thm eq_reflection} OF [thm]) + val thm = + (case Logic.dest_equals t of + (t1, t2) => + if t1 aconv t2 then raise TERM("identical terms", [t1, t2]) + else if Term.size_of_term t1 > Term.size_of_term t2 then @{thm eq_reflection} OF [thm] + else @{thm eq_reflection} OF [@{thm sym} OF [thm]]) + handle TERM("dest_equals", _) => @{thm eq_reflection} OF [thm] + in thm end) +in +fun cong_unfolding_trivial ctxt thms t = + prove ctxt t (fn _ => + EVERY' (map (fn thm => K (unfold_tac ctxt [thm])) ((map_filter reorder_for_simp thms))) + THEN' (fn i => TRY (resolve_tac ctxt @{thms refl} i))) + fun cong_unfolding_first ctxt thms t = - let val reorder_for_simp = try (fn thm => - let val t = Thm.prop_of ( @{thm eq_reflection} OF [thm]) - val thm = (case Logic.dest_equals t of - (t1, t2) => if Term.size_of_term t1 > Term.size_of_term t2 then @{thm eq_reflection} OF [thm] - else @{thm eq_reflection} OF [thm OF @{thms sym}]) - handle TERM("dest_equals", _) => @{thm eq_reflection} OF [thm] - in thm end) + let val ctxt = + ctxt + |> empty_simpset + |> put_simpset HOL_basic_ss + |> (fn ctxt => ctxt addsimps @{thms not_not eq_commute}) in - prove ctxt t (fn ctxt => - Raw_Simplifier.rewrite_goal_tac ctxt - (map_filter reorder_for_simp thms) + prove ctxt t (fn _ => + EVERY' (map (fn thm => K (unfold_tac ctxt [thm])) ((map_filter reorder_for_simp thms))) THEN' Method.insert_tac ctxt thms - THEN' K (Clasimp.auto_tac ctxt)) + THEN' (full_simp_tac ctxt) + THEN' K (ALLGOALS (K (Clasimp.auto_tac ctxt)))) end +end + + +fun arith_rewrite_tac ctxt _ = + let val backup_tac = Arith_Data.arith_tac ctxt ORELSE' Clasimp.force_tac ctxt in + (TRY o Simplifier.simp_tac ctxt) THEN_ALL_NEW backup_tac + ORELSE' backup_tac + end + +fun abstract_eq f (Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = + f t1 ##>> f t2 #>> HOLogic.mk_eq + | abstract_eq _ t = abstract_term t + +fun abstract_neq f (Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = + f t1 ##>> f t2 #>> HOLogic.mk_eq + | abstract_neq f (Const (\<^const_name>\HOL.Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2)) = + f t1 ##>> f t2 #>> HOLogic.mk_eq #>> curry (op $) HOLogic.Not + | abstract_neq f (Const (\<^const_name>\HOL.disj\, _) $ t1 $ t2) = + f t1 ##>> f t2 #>> HOLogic.mk_disj + | abstract_neq _ t = abstract_term t + +fun prove_arith_rewrite abstracter ctxt t = + prove_abstract' ctxt t arith_rewrite_tac ( + abstracter (abstract_arith ctxt) (dest_prop t)) + +fun prop_tac ctxt prems = + Method.insert_tac ctxt prems + THEN' SUBGOAL (fn (prop, i) => + if Term.size_of_term prop > 100 then SAT.satx_tac ctxt i + else (Classical.fast_tac ctxt ORELSE' Clasimp.force_tac ctxt) i) + end; diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smt_solver.ML --- a/src/HOL/Tools/SMT/smt_solver.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/smt_solver.ML Mon Oct 12 18:59:44 2020 +0200 @@ -292,9 +292,10 @@ SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) => (SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE) | SMT_Failure.SMT (fail as SMT_Failure.Time_Out) => - error ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^ + (SMT_Config.verbose_msg ctxt (K ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^ SMT_Failure.string_of_failure fail ^ " (setting the " ^ - "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)") + "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")) (); + NONE) | SMT_Failure.SMT fail => error (str_of ctxt fail) fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1 diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smt_systems.ML --- a/src/HOL/Tools/SMT/smt_systems.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/smt_systems.ML Mon Oct 12 18:59:44 2020 +0200 @@ -120,12 +120,14 @@ avail = make_avail "VERIT", command = make_command "VERIT", options = (fn ctxt => [ - "--proof-version=2", + "--proof-with-sharing", + "--proof-define-skolems", "--proof-prune", "--proof-merge", "--disable-print-success", "--disable-banner", - "--max-time=" ^ string_of_int (Real.ceil (Config.get ctxt SMT_Config.timeout))]), + "--max-time=" ^ string_of_int (1000 * Real.ceil (Config.get ctxt SMT_Config.timeout))] @ + Verit_Proof.veriT_current_strategy (Context.Proof ctxt)), smt_options = [(":produce-proofs", "true")], default_max_relevant = 200 (* FUDGE *), outcome = on_first_non_unsupported_line (outcome_of "unsat" "sat" "unknown" "timeout"), @@ -164,6 +166,35 @@ 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) + |> 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 *) diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/smtlib_proof.ML --- a/src/HOL/Tools/SMT/smtlib_proof.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/smtlib_proof.ML Mon Oct 12 18:59:44 2020 +0200 @@ -37,6 +37,16 @@ val type_of: ('a, 'b) context -> SMTLIB.tree -> typ val term_of: SMTLIB.tree -> ((string * (string * typ)) list, 'a) context -> term * ((string * (string * typ)) list, 'a) context + + (*name bindings*) + type name_bindings + val empty_name_binding: name_bindings + val update_name_binding: Symtab.key * SMTLIB.tree -> name_bindings -> name_bindings + val extract_name_bindings: SMTLIB.tree -> name_bindings -> name_bindings * string list + val expand_name_bindings: SMTLIB.tree -> name_bindings -> SMTLIB.tree * name_bindings + val extract_and_update_name_bindings: SMTLIB.tree -> name_bindings -> + (SMTLIB.tree * name_bindings) * string list + val remove_name_bindings: SMTLIB.tree -> SMTLIB.tree end; structure SMTLIB_Proof: SMTLIB_PROOF = @@ -265,7 +275,7 @@ fun dest_binding (SMTLIB.S [SMTLIB.Sym name, t]) = (name, Tree t) | dest_binding b = raise SMTLIB_PARSE ("bad SMT let binding format", b) -fun mk_choice (x, T, P) = HOLogic.choice_const T $ absfree (x, T) P +fun mk_choice (x, T, P) = HOLogic.choice_const T $ absfree (x, T) P fun term_of t cx = (case t of @@ -298,4 +308,68 @@ |>> fold_rev (fn (_, (n, T)) => fn t => q (n, T, t)) vs end + +(* Sharing of terms via "! _ :named name"*) + +type name_bindings = SMTLIB.tree Symtab.table * SMTLIB.tree Symtab.table + +val empty_name_binding = (Symtab.empty, Symtab.empty) + +fun update_name_binding (name, t) (tab, exp_tab) = + (tab, Symtab.update (name, t) exp_tab) + +fun remove_name_bindings (SMTLIB.S (SMTLIB.Sym "!" :: body :: _)) = + remove_name_bindings body + | remove_name_bindings (SMTLIB.S body) = + SMTLIB.S (map remove_name_bindings body) + | remove_name_bindings t = t + +fun extract_name_bindings t (tab, exp_tab) = + let + fun extract t (tab, new) = + (case t of + SMTLIB.S [SMTLIB.Sym "forall", SMTLIB.S _, body] => extract body (tab, new) + | SMTLIB.S [SMTLIB.Sym "exists", SMTLIB.S _, body] => extract body (tab, new) + | SMTLIB.S [SMTLIB.Sym "choice", SMTLIB.S _, body] => extract body (tab, new) + | SMTLIB.S [SMTLIB.Sym "let", SMTLIB.S bindings, body] => + raise (Fail "unsupported let construction in name extraction") + | SMTLIB.S (SMTLIB.Sym "!" :: t :: [SMTLIB.Key _, SMTLIB.Sym name]) => + extract t (tab, new) + |>> Symtab.update (name, remove_name_bindings t) + ||> curry (op :: ) name + | SMTLIB.S args => + fold extract args (tab, new) + | _ => ((tab, new))) + val (tab, new) = extract t (tab, []) + in ((tab, exp_tab), new) end + +fun expand_name_bindings t (tab, exp_tab) = + let + fun expand_update (SMTLIB.Sym sym) exp_tab = + (case Symtab.lookup exp_tab sym of + SOME t => (t, exp_tab) + | NONE => + (case (Symtab.lookup tab sym) of + NONE => (SMTLIB.Sym sym, exp_tab) + | SOME u => + let val (u2, exp_tab) = expand_update u exp_tab + in (u2, Symtab.update (sym, u2) exp_tab) end + ) + ) + | expand_update (t as SMTLIB.Key _) exp_tab = (t, exp_tab) + | expand_update (t as SMTLIB.Num _) exp_tab = (t, exp_tab) + | expand_update (t as SMTLIB.Dec _) exp_tab = (t, exp_tab) + | expand_update (t as SMTLIB.Str _) exp_tab = (t, exp_tab) + | expand_update (SMTLIB.S s) (exp_tab : SMTLIB.tree Symtab.table) = + fold_map expand_update s exp_tab + |>> (fn s => SMTLIB.S s) + in + expand_update t exp_tab + |> (fn (t, exp_tab) => (remove_name_bindings t, (tab, exp_tab))) + end + +fun extract_and_update_name_bindings t tab = + extract_name_bindings t tab + |>> expand_name_bindings (remove_name_bindings t) + end; diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/verit_isar.ML --- a/src/HOL/Tools/SMT/verit_isar.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/verit_isar.ML Mon Oct 12 18:59:44 2020 +0200 @@ -1,4 +1,4 @@ -(* Title: HOL/Tools/SMT/verit_isar.ML +(* Title: HOL/Tools/SMT/lethe_isar.ML Author: Mathias Fleury, TU Muenchen Author: Jasmin Blanchette, TU Muenchen @@ -9,7 +9,7 @@ sig type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step val atp_proof_of_veriT_proof: Proof.context -> term list -> thm list -> term list -> term -> - (string * term) list -> int list -> int -> (int * string) list -> VeriT_Proof.veriT_step list -> + (string * term) list -> int list -> int -> (int * string) list -> Verit_Proof.veriT_step list -> (term, string) ATP_Proof.atp_step list end; @@ -22,17 +22,17 @@ open ATP_Proof_Reconstruct open SMTLIB_Interface open SMTLIB_Isar -open VeriT_Proof +open Verit_Proof fun atp_proof_of_veriT_proof ctxt ll_defs rewrite_rules hyp_ts concl_t fact_helper_ts prem_ids conjecture_id fact_helper_ids = let - fun steps_of (VeriT_Proof.VeriT_Step {id, rule, prems, concl, ...}) = + fun steps_of (Verit_Proof.VeriT_Step {id, rule, prems, concl, ...}) = let val concl' = postprocess_step_conclusion ctxt rewrite_rules ll_defs concl fun standard_step role = ((id, []), role, concl', rule, map (rpair []) prems) in - if rule = veriT_input_rule then + if rule = input_rule then let val id_num = the (Int.fromString (unprefix assert_prefix id)) val ss = the_list (AList.lookup (op =) fact_helper_ids id_num) @@ -46,7 +46,7 @@ val concl0 = unskolemize_names ctxt concl00 in [(name0, role0, concl0, rule, []), - ((id, []), Plain, concl', veriT_rewrite_rule, + ((id, []), Plain, concl', rewrite_rule, name0 :: normalizing_prems ctxt concl0)] end) end diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/verit_proof.ML --- a/src/HOL/Tools/SMT/verit_proof.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/verit_proof.ML Mon Oct 12 18:59:44 2020 +0200 @@ -1,4 +1,4 @@ -(* Title: HOL/Tools/SMT/verit_proof.ML +(* Title: HOL/Tools/SMT/Verit_Proof.ML Author: Mathias Fleury, ENS Rennes Author: Sascha Boehme, TU Muenchen @@ -24,7 +24,9 @@ proof_ctxt: term list, concl: term, bounds: (string * typ) list, - subproof: (string * typ) list * term list * veriT_replay_node 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 -> @@ -32,34 +34,153 @@ val parse_replay: typ Symtab.table -> term Symtab.table -> string list -> Proof.context -> veriT_replay_node list * Proof.context - val map_replay_prems: (string list -> string list) -> veriT_replay_node -> veriT_replay_node - val veriT_step_prefix : string - val veriT_input_rule: string - val veriT_normalized_input_rule: string - val veriT_la_generic_rule : string - val veriT_rewrite_rule : string - val veriT_simp_arith_rule : string - val veriT_tmp_skolemize_rule : string - val veriT_subproof_rule : string - val veriT_local_input_rule : string + 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 is_skolemisation: string -> bool + val is_skolemisation_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_Proof: VERIT_PROOF = 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 extend = I + 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 + +fun verit_tac ctxt = SMT_Solver.smt_tac (Context.proof_map (SMT_Config.select_solver "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 concl subproof = - Raw_VeriT_Node {id = id, rule = rule, args = args, prems = prems, concl = concl, - subproof = subproof} +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, @@ -81,11 +202,14 @@ proof_ctxt: term list, concl: term, bounds: (string * typ) list, - subproof: (string * typ) list * term list * veriT_replay_node 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 subproof = +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, subproof = subproof} + concl = concl, bounds = bounds, insts = insts, declarations = declarations, + subproof = subproof} datatype veriT_step = VeriT_Step of { id: string, @@ -99,17 +223,28 @@ VeriT_Step {id = id, rule = rule, prems = prems, proof_ctxt = proof_ctxt, concl = concl, fixes = fixes} -val veriT_step_prefix = ".c" -val veriT_input_rule = "input" -val veriT_la_generic_rule = "la_generic" -val veriT_normalized_input_rule = "__normalized_input" (* arbitrary *) -val veriT_rewrite_rule = "__rewrite" (* arbitrary *) -val veriT_subproof_rule = "subproof" -val veriT_local_input_rule = "__local_input" (* arbitrary *) -val veriT_simp_arith_rule = "simp_arith" +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*) -(* Even the veriT developer do not know if the following rule can still appear in proofs: *) -val veriT_tmp_skolemize_rule = "tmp_skolemize" +val is_skolemisation = member (op =) ["sko_forall", "sko_ex"] +val keep_app_symbols = member (op =) ["eq_congruent_pred", "eq_congruent", "ite_intro"] +val keep_raw_lifting = member (op =) ["eq_congruent_pred", "eq_congruent", "ite_intro"] + +fun is_skolemisation_step (VeriT_Replay_Node {id, ...}) = is_skolemisation 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 *) @@ -128,47 +263,79 @@ if String.isPrefix var_name v then SOME T else NONE | find_type_in_formula _ _ = NONE -fun find_type_of_free_in_formula (Free (v, T) $ u) var_name = - if String.isPrefix var_name v then SOME T else find_type_in_formula u var_name - | find_type_of_free_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_of_free_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_of_free_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 concl = fold (update_binding o (fn s => (s, Term (Free (s, the_default dummyT (find_type_in_formula concl s)))))) - local fun remove_Sym (SMTLIB.Sym y) = y + | remove_Sym y = (@{print} y; raise (Fail "failed to match")) fun extract_symbols bds = bds - |> map (fn SMTLIB.S [SMTLIB.Key _, SMTLIB.Sym x, SMTLIB.Sym y] => [x, y] - | SMTLIB.S syms => map remove_Sym syms) + |> map (fn SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, SMTLIB.Sym y] => [[x, y]] + | SMTLIB.S [SMTLIB.Key "=", SMTLIB.Sym x, SMTLIB.Sym y] => [[x, y]] + | SMTLIB.S syms => map (single o remove_Sym) syms + | SMTLIB.Sym x => [[x]] + | 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] => + (case node_of (SMTLIB.Sym y) cx of + ((_, []), _) => [[x]] + | _ => [[x, y]]) + | SMTLIB.S [SMTLIB.Key "=", SMTLIB.Sym x, SMTLIB.Sym y] => + (case node_of (SMTLIB.Sym y) cx of + ((_, []), _) => [[x]] + | _ => [[x, y]]) + | SMTLIB.S (SMTLIB.Sym "=" :: SMTLIB.Sym x :: _) => [[x]] + | SMTLIB.S (SMTLIB.Key "=" :: SMTLIB.Sym x :: _) => [[x]] + | SMTLIB.S syms => map (single o remove_Sym) syms + | SMTLIB.Sym x => [[x]] + | t => raise (Fail ("match error " ^ @{make_string} t))) |> flat fun extract_symbols_map bds = bds - |> map (fn SMTLIB.S [SMTLIB.Key _, SMTLIB.Sym x, _] => [x] - | SMTLIB.S syms => map remove_Sym syms) + |> map (fn SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym x, _] => [[x]] + | SMTLIB.S syms => map (single o remove_Sym) syms) |> flat in -fun bound_vars_by_rule "bind" (SMTLIB.S bds) = extract_symbols bds - | bound_vars_by_rule "qnt_simplify" (SMTLIB.S bds) = extract_symbols_map bds - | bound_vars_by_rule "sko_forall" (SMTLIB.S bds) = extract_symbols_map bds - | bound_vars_by_rule "sko_ex" (SMTLIB.S bds) = extract_symbols_map bds - | bound_vars_by_rule _ _ = [] +fun declared_csts _ "__skolem_definition" (SMTLIB.S [SMTLIB.Sym x, typ, _]) = [(x, typ)] + | declared_csts _ _ _ = [] + +fun skolems_introduced_by_rule (SMTLIB.S bds) = + fold (fn (SMTLIB.S [SMTLIB.Sym "=", SMTLIB.Sym _, SMTLIB.Sym y]) => curry (op ::) y) bds [] -fun global_bound_vars_by_rule _ _ = [] +(*FIXME there is probably a way to use the information given by onepoint*) +fun bound_vars_by_rule _ "bind" (SMTLIB.S bds) = extract_symbols bds + | bound_vars_by_rule cx "onepoint" (SMTLIB.S bds) = extract_qnt_symbols cx bds + | bound_vars_by_rule _ "sko_forall" (SMTLIB.S bds) = extract_symbols_map bds + | bound_vars_by_rule _ "sko_ex" (SMTLIB.S bds) = extract_symbols_map bds + | bound_vars_by_rule _ "__skolem_definition" (SMTLIB.S [SMTLIB.Sym x, _, _]) = [[x]] + | bound_vars_by_rule _ _ _ = [] -(* VeriT adds "?" before some variable. *) +(* 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' @@ -181,200 +348,336 @@ | remove_all_qm2 (SMTLIB.Key v) = SMTLIB.Key v | remove_all_qm2 v = v -val parse_rule_and_args = - let - fun parse_rule_name (SMTLIB.Sym rule :: l) = (rule, l) - | parse_rule_name l = (veriT_subproof_rule, l) - fun parse_args (SMTLIB.Key "args" :: args :: l) = (remove_all_qm2 args, l) - | parse_args l = (SMTLIB.S [], l) - in - parse_rule_name - ##> parse_args - end - end -fun parse_raw_proof_step (p : SMTLIB.tree) : raw_veriT_node = +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 get_id (SMTLIB.S [SMTLIB.Sym "set", SMTLIB.Sym id, SMTLIB.S l]) = (id, l) + 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 "clauses" :: SMTLIB.S source ::l) = - (SOME (map (fn (SMTLIB.Sym id) => id) source), l) - | parse_source l = (NONE, l) - fun parse_subproof rule args id_of_father_step ((subproof_step as SMTLIB.S (SMTLIB.Sym "set" :: _)) :: l) = - let - val subproof_steps = parse_raw_proof_step subproof_step - in - apfst (curry (op ::) subproof_steps) (parse_subproof rule args id_of_father_step l) - end - | parse_subproof _ _ _ l = ([], l) + 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 parse_and_clausify_conclusion (SMTLIB.Key "conclusion" :: SMTLIB.S [] :: []) = - SMTLIB.Sym "false" - | parse_and_clausify_conclusion (SMTLIB.Key "conclusion" :: SMTLIB.S concl :: []) = - (SMTLIB.S (remove_all_qm (SMTLIB.Sym "or" :: concl))) - - fun to_raw_node ((((((id, rule), args), prems), subproof), concl)) = - (mk_raw_node id rule args (the_default [] prems) concl subproof) + 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 - (get_id - ##> parse_rule_and_args - #> rotate_pair - #> rotate_pair - ##> parse_source - #> rotate_pair - #> (fn ((((id, rule), args), prems), sub) => - ((((id, rule), args), prems), parse_subproof rule args id sub)) - #> rotate_pair - ##> parse_and_clausify_conclusion - #> to_raw_node) - p + 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 "qnt_simplify" t = t + | proof_ctxt_of_rule "onepoint" t = t | proof_ctxt_of_rule _ _ = [] -fun args_of_rule "forall_inst" t = t +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 map_replay_prems f (VeriT_Replay_Node {id, rule, args, prems, proof_ctxt, concl, bounds, - subproof = (bound, assms, subproof)}) = - (VeriT_Replay_Node {id = id, rule = rule, args = args, prems = f prems, proof_ctxt = proof_ctxt, - concl = concl, bounds = bounds, subproof = (bound, assms, map (map_replay_prems f) subproof)}) - -fun map_replay_id f (VeriT_Replay_Node {id, rule, args, prems, proof_ctxt, concl, bounds, - subproof = (bound, assms, subproof)}) = - (VeriT_Replay_Node {id = f id, rule = rule, args = args, prems = prems, proof_ctxt = proof_ctxt, - concl = concl, bounds = bounds, subproof = (bound, assms, map (map_replay_id f) subproof)}) +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 -val extract_assumptions_from_subproof = - let fun extract_assumptions_from_subproof (VeriT_Replay_Node {rule, concl, ...}) = - if rule = veriT_local_input_rule then [concl] else [] +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 - map extract_assumptions_from_subproof - #> flat + fold extract_assumptions_from_subproof subproof [] end fun normalized_rule_name id rule = - (case (rule = veriT_input_rule, can (unprefix SMTLIB_Interface.assert_prefix) id) of - (true, true) => veriT_normalized_input_rule - | (true, _) => veriT_local_input_rule + (case (rule = input_rule, can (unprefix SMTLIB_Interface.assert_prefix) id) of + (true, true) => normalized_input_rule + | (true, _) => local_input_rule | _ => rule) fun is_assm_repetition id rule = - rule = veriT_input_rule andalso can (unprefix SMTLIB_Interface.assert_prefix) id + rule = input_rule andalso can (unprefix SMTLIB_Interface.assert_prefix) 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 (SMTLIB.S S) = map (fn x => SMTLIB.S [SMTLIB.Sym "=", x, x]) S + | expand_lonely_arguments (x as SMTLIB.Sym _) = [SMTLIB.S [SMTLIB.Sym "=", x, x]] + | expand_lonely_arguments t = [t] + + fun preprocess (Raw_VeriT_Node {id, rule, args, prems, concl, subproof, ...}) (cx, remap_assms) = + let + val 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 -fun postprocess_proof ctxt step = - let fun postprocess (Raw_VeriT_Node {id = id, rule = rule, args = args, - prems = prems, concl = concl, subproof = subproof}) cx = + val (subproof, (cx, _)) = fold_map preprocess subproof (cx, remap_assms) |> apfst flat + val (skolem_names, stripped_args) = stripped_args + 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_skolemisation rule then map (fn id => id ^ veriT_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_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 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 rhs_vars = fold (fn [t', t] => t <> t' ? (curry (op ::) t) | _ => fn x => x) bound_vars [] + 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, _] => not_already_bound cx t | _ => true) + |>> map (single o hd) + |>> (fn vars => vars @ map (fn [_, t] => [t] | _ => []) bound_vars) + |>> flat + val subproof_rew = fold (fn [t, t'] => curry (op ::) (t, t ^ t')) + 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 bound_vars = bound_vars_by_rule rule args + val extra_lhs_vars = map (fn [a,b] => (a, a^b)) rebound_lhs_vars (* postprocess conclusion *) - val new_global_bounds = global_bound_vars_by_rule rule args - val concl = SMTLIB_Isar.unskolemize_names ctxt concl + 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', + 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_vars = filter_out (member ((op =)) new_global_bounds) bound_vars + val bound_tvars = - map (fn s => (s, the (find_type_in_formula concl s))) bound_vars - val subproof_cx = add_bound_variables_to_ctxt concl bound_vars cx - val (p : veriT_replay_node list list, _) = - fold_map postprocess subproof subproof_cx + map (fn s => (s, the (find_type_in_formula concl s))) + (shadowing_vars @ map snd extra_lhs_vars) + val subproof_cx = + add_bound_variables_to_ctxt concl (shadowing_vars @ map snd extra_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_skolemisation 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 SMTLIB.S stripped_args = args - val sanitized_args = - proof_ctxt_of_rule rule stripped_args - |> map - (fn SMTLIB.S [SMTLIB.Key "=", x, y] => SMTLIB.S [SMTLIB.Sym "=", x, y] - | SMTLIB.S syms => - SMTLIB.S (SMTLIB.Sym "and" :: map (fn x => SMTLIB.S [SMTLIB.Sym "=", x, x]) syms) - | x => x) - val (termified_args, _) = fold_map node_of sanitized_args subproof_cx |> apfst (map fst) - val normalized_args = map (SMTLIB_Isar.unskolemize_names ctxt) termified_args + val stripped_args = args |> (fn SMTLIB.S S => S) + val sanitized_args = proof_ctxt_of_rule rule stripped_args + + val arg_cx = + subproof_cx + |> add_bound_variables_to_ctxt concl (shadowing_vars @ map fst extra_lhs_vars) + 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 (SMTLIB_Isar.unskolemize_names ctxt) termified_args - val rule_args = normalized_args + val normalized_args = map unsk_and_rewrite termified_args + val rule_args = map subproof_rewriter normalized_args - (* fix subproof *) - val p = flat p - val p = map (map_replay_prems (map (curry (op ^) id))) p - val p = map (map_replay_id (curry (op ^) id)) p + 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) - val extra_assms2 = - (if rule = veriT_subproof_rule then extract_assumptions_from_subproof p else []) + (* declarations *) + val (declarations, _) = fold_map termify_term declarations cx + |> apfst (map (apsnd unsk_and_rewrite)) (* fix step *) - val bound_t = - bounds - |> map (fn s => (s, the_default dummyT (find_type_of_free_in_formula concl s))) + val bound_t = bounds + |> map (fn s => (s, the (find_type_in_formula concl s))) + + val skolem_defs = (if is_skolemisation rule + then map (fn id => id ^ veriT_def) (skolems_introduced_by_rule args) else []) + val skolems_of_subproof = (if is_skolemisation rule + then flat (map collect_skolem_defs subproof) else []) val fixed_prems = - (if null subproof then prems else map (curry (op ^) id) prems) @ - (if is_assm_repetition id rule then [id] else []) @ - id_of_last_step p + 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 - val step = mk_replay_node id normalized_rule rule_args fixed_prems subproof_assms concl - bound_t (bound_tvars, subproof_assms @ extra_assms2, p) - in - ([step], cx') - end - in postprocess step end - + |> (if compress andalso alpha_conversion then K "refl" else I) -(*subproofs are written on multiple lines: SMTLIB can not parse then, because parentheses are -unbalanced on each line*) -fun seperate_into_steps lines = - let - fun count ("(" :: l) n = count l (n + 1) - | count (")" :: l) n = count l (n - 1) - | count (_ :: l) n = count l n - | count [] n = n - fun seperate (line :: l) actual_lines m = - let val n = count (raw_explode line) 0 in - if m + n = 0 then - [actual_lines ^ line] :: seperate l "" 0 - else - seperate l (actual_lines ^ line) (m + n) - end - | seperate [] _ 0 = [] + 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 + bound_t insts declarations (bound_tvars, subproof_assms, extra_assms2, fixed_subproof) + + in + (step, (cx', rew)) + end in - seperate lines "" 0 - end + postprocess step (cx, []) + |> (fn (step, (cx, _)) => (step, cx)) + end -fun unprefix_all_syms c (SMTLIB.Sym v :: l) = - SMTLIB.Sym (perhaps (try (unprefix c)) v) :: unprefix_all_syms c l - | unprefix_all_syms c (SMTLIB.S l :: l') = SMTLIB.S (unprefix_all_syms c l) :: unprefix_all_syms c l' - | unprefix_all_syms c (SMTLIB.Key v :: l) = SMTLIB.Key v :: unprefix_all_syms c l - | unprefix_all_syms c (v :: l) = v :: unprefix_all_syms c l - | unprefix_all_syms _ [] = [] - -(* VeriT adds "@" before every variable. *) -val remove_all_ats = unprefix_all_syms "@" val linearize_proof = let fun linearize (VeriT_Replay_Node {id = id, rule = rule, args = _, prems = prems, - proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, subproof = (_, _, subproof)}) = + proof_ctxt = proof_ctxt, concl = concl, bounds = bounds, insts = _, declarations = _, + subproof = (_, _, _, subproof)}) = let fun mk_prop_of_term concl = concl |> fastype_of concl = \<^typ>\bool\ ? curry (op $) \<^term>\Trueprop\ @@ -387,7 +690,7 @@ fun find_input_steps_and_inline [] = [] | find_input_steps_and_inline (VeriT_Node {id = id', rule, prems, concl, bounds, ...} :: steps) = - if rule = veriT_input_rule then + if rule = input_rule then find_input_steps_and_inline (map (inline_assumption concl id') steps) else mk_node (id') rule prems [] concl bounds :: find_input_steps_and_inline steps @@ -402,12 +705,25 @@ local fun import_proof_and_post_process typs funs lines ctxt = let - val smtlib_lines_without_at = - seperate_into_steps lines - |> map SMTLIB.parse - |> remove_all_ats - in apfst flat (fold_map (fn l => postprocess_proof ctxt (parse_raw_proof_step l)) - smtlib_lines_without_at (empty_context ctxt typs funs)) end + 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)) + in step end in fun parse typs funs lines ctxt = @@ -423,7 +739,7 @@ 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) + val _ = (SMT_Config.verit_msg ctxt) (fn () => \<^print> u) in (u, ctxt_of env) end diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/verit_proof_parse.ML --- a/src/HOL/Tools/SMT/verit_proof_parse.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/verit_proof_parse.ML Mon Oct 12 18:59:44 2020 +0200 @@ -21,9 +21,9 @@ open ATP_Proof open ATP_Proof_Reconstruct open VeriT_Isar -open VeriT_Proof +open Verit_Proof -fun add_used_asserts_in_step (VeriT_Proof.VeriT_Step {prems, ...}) = +fun add_used_asserts_in_step (Verit_Proof.VeriT_Step {prems, ...}) = union (op =) (map_filter (try SMTLIB_Interface.assert_index_of_name) prems) fun parse_proof @@ -36,10 +36,10 @@ val index_of_id = Integer.add (~ num_ll_defs) fun step_of_assume j (_, th) = - VeriT_Proof.VeriT_Step {id = SMTLIB_Interface.assert_name_of_index (id_of_index j), - rule = veriT_input_rule, prems = [], proof_ctxt = [], concl = Thm.prop_of th, fixes = []} + Verit_Proof.VeriT_Step {id = SMTLIB_Interface.assert_name_of_index (id_of_index j), + rule = input_rule, prems = [], proof_ctxt = [], concl = Thm.prop_of th, fixes = []} - val (actual_steps, _) = VeriT_Proof.parse typs terms output ctxt + val (actual_steps, _) = Verit_Proof.parse typs terms output ctxt val used_assert_ids = fold add_used_asserts_in_step actual_steps [] val used_assm_js = map_filter (fn id => let val i = index_of_id id in if i >= 0 then SOME i else NONE end) diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/verit_replay.ML --- a/src/HOL/Tools/SMT/verit_replay.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/verit_replay.ML Mon Oct 12 18:59:44 2020 +0200 @@ -12,64 +12,86 @@ structure Verit_Replay: VERIT_REPLAY = struct -fun under_fixes f unchanged_prems (prems, nthms) names args (concl, ctxt) = +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 _ = 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 concl) end + 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 - (VeriT_Proof.VeriT_Replay_Node {id, rule, concl, bounds, ...}) = + concl_transformation global_transformation args insts + (Verit_Proof.VeriT_Replay_Node {id, rule, concl, bounds, declarations = decls, ...}) = let - val _ = SMT_Config.veriT_msg ctxt (fn () => \<^print> id) + 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) ? SMTLIB_Isar.unlift_term ll_defs + #> not (null ll_defs andalso Verit_Proof.keep_raw_lifting rule) ? SMTLIB_Isar.unlift_term ll_defs end + val rewrite_concl = if Verit_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_rules [] + 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 - |> concl_transformation - |> global_transformation + |> Term.subst_free concl_transformation + |> subst_only_free global_transformation |> post -in - if rule = VeriT_Proof.veriT_input_rule then - (case Symtab.lookup assumed id of - SOME (_, thm) => thm) - else - under_fixes (method_for rule) unchanged_prems - (prems, nthms) (map fst bounds) - (map rewrite args) (concl, ctxt) -end + in + if rule = Verit_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, - subproof = (_, _, subproof), ...}) = +fun add_used_asserts_in_step (Verit_Proof.VeriT_Replay_Node {prems, + subproof = (_, _, _, subproof), ...}) = union (op =) (map_filter (try SMTLIB_Interface.assert_index_of_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 Verit_Proof.VeriT_Replay_Node {id, ...}) => (case try SMTLIB_Interface.assert_index_of_name id of NONE => SOME step | SOME a => if a < n then NONE else SOME step)) -fun replay_step rewrite_rules ll_defs assumed proof_prems - (step as VeriT_Proof.VeriT_Replay_Node {id, rule, prems, bounds, args, - subproof = (fixes, assms, subproof), concl, ...}) state = + +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, + 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 @@ -78,55 +100,132 @@ val (names, sub_ctxt) = Variable.variant_fixes (map fst fixes) ctxt ||> fold Variable.declare_term (map Free fixes) - val export_vars = - Term.subst_free (ListPair.zip (map Free fixes, map Free (ListPair.zip (names, map snd fixes)))) - o concl_tranformation + 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 rewrite_rules [] + Raw_Simplifier.rewrite_term thy ((if Verit_Proof.keep_raw_lifting rule then tl rewrite_rules else rewrite_rules)) [] #> Object_Logic.atomize_term ctxt - #> not (null ll_defs) ? SMTLIB_Isar.unlift_term ll_defs + #> not (null ll_defs andalso Verit_Proof.keep_raw_lifting rule) ? SMTLIB_Isar.unlift_term ll_defs #> SMTLIB_Isar.unskolemize_names ctxt #> HOLogic.mk_Trueprop end - val assms = map (export_vars o global_transformation o post) assms - val (proof_prems', sub_ctxt2) = Assumption.add_assumes (map (Thm.cterm_of sub_ctxt) assms) + 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_step rewrite_rules ll_defs assumed all_proof_prems) subproof - (assumed, stats, sub_ctxt2, export_vars, global_transformation) + 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 - |> map (apsnd export_thm o the o (Symtab.lookup (if null subproof then proofs else proofs'))) + |> map (apsnd export_thm) o map_filter (Symtab.lookup (if (null subproof) then proofs else proofs')) + val nthms' = (if Verit_Proof.is_skolemisation rule + then prems else []) + |> 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 Verit_Replay_Methods.veriT_step_requires_subproof_assms rule then proof_prems else [] - val replay = Timing.timing (replay_thm Verit_Replay_Methods.method_for rewrite_rules ll_defs - ctxt assumed [] (proof_prems) nthms concl_tranformation global_transformation args) + 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 Lethe_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.toMilliseconds elapsed) stats - in (Symtab.update (id, (map fst bounds, thm)) proofs, stats', ctxt, + 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_ll_def assms ll_defs rewrite_rules stats ctxt term = +fun replay_definition_step rewrite_rules ll_defs _ _ _ + (Verit_Proof.VeriT_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 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 ("ll_defs", Time.toMilliseconds elapsed) stats + val stats' = Symtab.cons_list (Verit_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 + 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 = @@ -139,21 +238,25 @@ 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 + Verit_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 { + Verit_Proof.VeriT_Replay_Node { id = SMTLIB_Interface.assert_name_of_index (id_of_index j), - rule = VeriT_Proof.veriT_input_rule, + rule = Verit_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)) [] [], + (empty_simpset ctxt addsimps rewrite_rules)) rewrite_rules [], bounds = [], - subproof = ([], [], [])} + 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 @@ -163,45 +266,54 @@ used_assert_ids val assm_steps = map step_of_assume used_assm_js - val steps = assm_steps @ actual_steps - fun extract (VeriT_Proof.VeriT_Replay_Node {id, rule, concl, bounds, ...}) = + fun extract (Verit_Proof.VeriT_Replay_Node {id, rule, concl, bounds, ...}) = (id, rule, concl, map fst bounds) - fun cond rule = rule = VeriT_Proof.veriT_input_rule + fun cond rule = rule = Verit_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 assms - (map_filter (remove_rewrite_rules_from_rules num_ll_defs) steps) ctxt2 + (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_ll_def assms ll_defs rewrite_rules stats ctxt thm + let val (thm, stats) = replay_assumed assms ll_defs rewrite_rules stats ctxt thm in (Symtab.update (SMTLIB_Interface.assert_name_of_index id, ([], thm)) assumed, stats) end) - used_rew_js (assumed, Symtab.empty) + 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 = length steps + val len = Verit_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 + 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) y = - (if i > 0 andalso i mod 100 = 0 then print_runtime_statistics i else (); f x y) - val (proofs, stats, ctxt5, _, _) = - fold_index (blockwise (replay_step rewrite_rules ll_defs assumed [])) steps - (assumed, stats, ctxt4, fn x => x, fn x => x) + 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 _ = print_runtime_statistics len - val total = Time.toMilliseconds (#elapsed (Timing.result start)) - val (_, VeriT_Proof.VeriT_Replay_Node {id, ...}) = split_last steps + 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 - Symtab.lookup proofs id |> the |> snd |> singleton (Proof_Context.export ctxt5 outer_ctxt) + Lethe_Replay_Methods.discharge ctxt [thm_with_defs] @{term False} end end diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/verit_replay_methods.ML --- a/src/HOL/Tools/SMT/verit_replay_methods.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/verit_replay_methods.ML Mon Oct 12 18:59:44 2020 +0200 @@ -1,25 +1,23 @@ (* Title: HOL/Tools/SMT/verit_replay_methods.ML - Author: Mathias Fleury, MPII + Author: Mathias Fleury, MPII, JKU -Proof methods for replaying veriT proofs. +Proof method for replaying veriT proofs. *) -signature VERIT_REPLAY_METHODS = +signature LETHE_REPLAY_METHODS = sig - - val is_skolemisation: string -> bool - val is_skolemisation_step: VeriT_Proof.veriT_replay_node -> bool + (*methods for verit proof rules*) + val method_for: string -> Proof.context -> thm list -> term list -> term Symtab.table -> + (string * term) list -> term -> thm - (* methods for veriT proof rules *) - val method_for: string -> Proof.context -> thm list -> term list -> term -> - thm - - val veriT_step_requires_subproof_assms : string -> bool + val requires_subproof_assms : string list -> string -> bool + val requires_local_input: string list -> string -> bool val eq_congruent_pred: Proof.context -> 'a -> term -> thm + val discharge: Proof.context -> thm list -> term -> thm end; -structure Verit_Replay_Methods: VERIT_REPLAY_METHODS = +structure Lethe_Replay_Methods: LETHE_REPLAY_METHODS = struct (*Some general comments on the proof format: @@ -30,51 +28,53 @@ 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. We currently do not care - about it, since in all cases we have met, a rule like tmp_AC_simp is called to do the - simplification. + 3. Duplicate literals are sometimes removed, mostly by the SAT solver. Rules unsupported on purpose: * Distinct_Elim, XOR, let (we don't need them). - * tmp_skolemize (because it is not clear if veriT still generates using it). + * 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 *) + (*input: a repeated (normalized) assumption of assumption of in a subproof*) Normalized_Input | Local_Input | - (* Subproof: *) + (*Subproof:*) Subproof | - (* Conjunction: *) + (*Conjunction:*) And | Not_And | And_Pos | And_Neg | - (* Disjunction"" *) + (*Disjunction""*) Or | Or_Pos | Not_Or | Or_Neg | - (* Disjunction: *) + (*Disjunction:*) Implies | Implies_Neg1 | Implies_Neg2 | Implies_Pos | Not_Implies1 | Not_Implies2 | - (* Equivalence: *) + (*Equivalence:*) Equiv_neg1 | Equiv_pos1 | Equiv_pos2 | Equiv_neg2 | Not_Equiv1 | Not_Equiv2 | Equiv1 | Equiv2 | - (* If-then-else: *) + (*If-then-else:*) ITE_Pos1 | ITE_Pos2 | ITE_Neg1 | ITE_Neg2 | Not_ITE1 | Not_ITE2 | ITE_Intro | ITE1 | ITE2 | - (* Equality: *) + (*Equality:*) Eq_Congruent | Eq_Reflexive | Eq_Transitive | Eq_Congruent_Pred | Trans | Refl | Cong | - (* Arithmetics: *) + (*Arithmetics:*) LA_Disequality | LA_Generic | LA_Tautology | LIA_Generic | LA_Totality | LA_RW_Eq | NLA_Generic | - (* Quantifiers: *) - Forall_Inst | Qnt_Rm_Unused | Qnt_Join | Qnt_Simplify | Bind | Skolem_Forall | Skolem_Ex | - (* Resolution: *) + (*Quantifiers:*) + Forall_Inst | Qnt_Rm_Unused | Qnt_Join | Onepoint | Bind | Skolem_Forall | Skolem_Ex | + (*Resolution:*) Theory_Resolution | Resolution | - (* Various transformation: *) - Connective_Equiv | - (* Temporary rules, that the veriT developpers want to remove: *) - Tmp_AC_Simp | - Tmp_Bfun_Elim | - (* Unsupported rule *) + (*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 | + Qnt_Simplify | Not_Not | Tautological_Clause | Duplicate_Literals | + (*Unsupported rule*) Unsupported -val is_skolemisation = member (op =) ["sko_forall", "sko_ex"] -fun is_skolemisation_step (VeriT_Proof.VeriT_Replay_Node {id, ...}) = is_skolemisation id fun verit_rule_of "bind" = Bind | verit_rule_of "cong" = Cong @@ -95,10 +95,9 @@ | verit_rule_of "not_or" = Not_Or | verit_rule_of "resolution" = Resolution | verit_rule_of "eq_congruent" = Eq_Congruent - | verit_rule_of "connective_equiv" = Connective_Equiv | verit_rule_of "trans" = Trans | verit_rule_of "false" = False - | verit_rule_of "tmp_AC_simp" = Tmp_AC_Simp + | verit_rule_of "ac_simp" = AC_Simp | verit_rule_of "and" = And | verit_rule_of "not_and" = Not_And | verit_rule_of "and_pos" = And_Pos @@ -112,7 +111,7 @@ | verit_rule_of "implies_neg1" = Implies_Neg1 | verit_rule_of "implies_neg2" = Implies_Neg2 | verit_rule_of "implies" = Implies - | verit_rule_of "tmp_bfun_elim" = Tmp_Bfun_Elim + | verit_rule_of "bfun_elim" = Bfun_Elim | verit_rule_of "ite1" = ITE1 | verit_rule_of "ite2" = ITE2 | verit_rule_of "not_ite1" = Not_ITE1 @@ -131,14 +130,33 @@ | verit_rule_of "nla_generic"= NLA_Generic | verit_rule_of "eq_transitive" = Eq_Transitive | verit_rule_of "qnt_rm_unused" = Qnt_Rm_Unused - | verit_rule_of "qnt_simplify" = Qnt_Simplify + | verit_rule_of "onepoint" = Onepoint | verit_rule_of "qnt_join" = Qnt_Join | verit_rule_of "eq_congruent_pred" = Eq_Congruent_Pred | verit_rule_of "subproof" = Subproof + | verit_rule_of "bool_simplify" = Bool_Simplify + | verit_rule_of "or_simplify" = Or_Simplify + | verit_rule_of "ite_simplify" = ITE_Simplify + | verit_rule_of "and_simplify" = And_Simplify + | verit_rule_of "not_simplify" = Not_Simplify + | verit_rule_of "equiv_simplify" = Equiv_Simplify + | verit_rule_of "eq_simplify" = Eq_Simplify + | verit_rule_of "implies_simplify" = Implies_Simplify + | verit_rule_of "connective_def" = Connective_Def + | verit_rule_of "minus_simplify" = Minus_Simplify + | verit_rule_of "sum_simplify" = Sum_Simplify + | verit_rule_of "prod_simplify" = Prod_Simplify + | verit_rule_of "comp_simplify" = Comp_Simplify + | verit_rule_of "qnt_simplify" = Qnt_Simplify + | verit_rule_of "not_not" = Not_Not + | verit_rule_of "tautology" = Tautological_Clause + | verit_rule_of "contraction" = Duplicate_Literals + | verit_rule_of "qnt_cnf" = Qnt_CNF | verit_rule_of r = - if r = VeriT_Proof.veriT_normalized_input_rule then Normalized_Input - else if r = VeriT_Proof.veriT_local_input_rule then Local_Input - else Unsupported + if r = Verit_Proof.normalized_input_rule then Normalized_Input + else if r = Verit_Proof.local_input_rule then Local_Input + else if r = Verit_Proof.veriT_def then Definition + else (@{print} ("Unsupport rule", r); Unsupported) fun string_of_verit_rule Bind = "Bind" | string_of_verit_rule Cong = "Cong" @@ -158,7 +176,6 @@ | 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 Connective_Equiv = "connective_equiv" | string_of_verit_rule Trans = "trans" | string_of_verit_rule False = "false" | string_of_verit_rule And = "and" @@ -167,7 +184,7 @@ | 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 Tmp_AC_Simp = "tmp_AC_simp" + | 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" @@ -175,7 +192,7 @@ | 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 Tmp_Bfun_Elim = "tmp_bfun_elim" + | 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" @@ -194,480 +211,658 @@ | 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 Qnt_Simplify = "qnt_simplify" + | 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 = VeriT_Proof.veriT_normalized_input_rule - | string_of_verit_rule Local_Input = VeriT_Proof.veriT_normalized_input_rule + | string_of_verit_rule Normalized_Input = Verit_Proof.normalized_input_rule + | string_of_verit_rule Local_Input = Verit_Proof.local_input_rule | string_of_verit_rule Subproof = "subproof" - | string_of_verit_rule r = "Unsupported rule: " ^ \<^make_string> r + | 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 Comp_Simplify = "comp_simplify" + | string_of_verit_rule Qnt_Simplify = "qnt_simplify" + | string_of_verit_rule Not_Not = "not_not" + | string_of_verit_rule Tautological_Clause = "tautology" + | string_of_verit_rule Duplicate_Literals = "contraction" + | string_of_verit_rule Qnt_CNF = "qnt_cnf" + | string_of_verit_rule r = "Unknown rule: " ^ \<^make_string> r -(*** Methods to Replay Normal steps ***) -(* sko_forall requires the assumptions to be able to SMT_Replay_Methods.prove the equivalence in case of double -skolemization. See comment below. *) -fun veriT_step_requires_subproof_assms t = - member (op =) ["refl", "cong", VeriT_Proof.veriT_local_input_rule, "sko_forall", - "sko_ex"] t +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 not_not eq_commute} addsimps thms) + |> (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 =) [Verit_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 bind_thms = - [@{lemma "(\x x'. P x = Q x) \ (\x. P x) = (\y. Q y)" - by blast}, - @{lemma "(\x x'. P x = Q x) \ (\x. P x) = (\y. Q y)" - by blast}, - @{lemma "(\x x'. P x = Q x) \ (\x. P x = Q x)" - by blast}, - @{lemma "(\x x'. P x = Q x) \ (\x. P x = Q x)" - by blast}] +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)) -fun bind ctxt [prems] t = SMT_Replay_Methods.prove ctxt t (fn _ => - REPEAT' (resolve_tac ctxt bind_thms) - THEN' match_tac ctxt [prems] - THEN' simplify_tac ctxt [] - THEN' REPEAT' (match_tac ctxt [@{thm refl}])) + +(* 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 + + +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 (Z3_Replay_Methods.refl ctxt thm) t of - NONE => - ( Z3_Replay_Methods.cong ctxt thm t) + (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 equiv_pos_neg_term ctxt thm (\<^term>\Trueprop\ $ - (\<^term>\HOL.disj\ $ (_) $ - ((\<^term>\HOL.disj\ $ a $ b)))) = - Drule.infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) [a, b]) thm - - fun prove_equiv_pos_neg thm ctxt _ t = - let val thm = equiv_pos_neg_term ctxt thm t - in - SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [thm] - THEN' simplify_tac ctxt []) - end +fun dropWhile _ [] = [] + | dropWhile f (x :: xs) = if f x then dropWhile f xs else x :: xs in -val equiv_pos1_thm = - @{lemma "\(a \ ~b) \ a \ b" - by blast+} - -val equiv_pos1 = prove_equiv_pos_neg equiv_pos1_thm - -val equiv_pos2_thm = - @{lemma "\a b. ((\a) \ b) \ a \ b" - by blast+} - -val equiv_pos2 = prove_equiv_pos_neg equiv_pos2_thm - -val equiv_neg1_thm = - @{lemma "(~a \ ~b) \ a \ b" - by blast} - -val equiv_neg1 = prove_equiv_pos_neg equiv_neg1_thm - -val equiv_neg2_thm = - @{lemma "(a \ b) \ a \ b" - by blast} - -val equiv_neg2 = prove_equiv_pos_neg equiv_neg2_thm - +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 + 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 -(* Most of the code below is due to the proof output of veriT: The description of the rule is wrong -(and according to Pascal Fontaine, it is a bug). Anyway, currently, forall_inst does: - 1. swapping out the forall quantifiers - 2. instantiation - 3. boolean. + +(* Or *) -However, types can mess-up things: - lemma \(0 < degree a) = (0 \ degree a) \ 0 = degree a \ 0 < degree a\ - by fast -works unlike - lemma \((0::nat) < degree a) = (0 \ degree a) \ 0 = degree a \ 0 < degree a\ - by fast. -Therefore, we use fast and auto as fall-back. -*) -fun forall_inst ctxt _ args t = - let - val instantiate = - fold (fn inst => fn tac => - let val thm = Drule.infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt inst)] @{thm spec} - in tac THEN' dmatch_tac ctxt [thm]end) - args - (K all_tac) - in - SMT_Replay_Methods.prove ctxt t (fn _ => - resolve_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF [@{thm impI}]] - THEN' TRY' (Raw_Simplifier.rewrite_goal_tac ctxt @{thms all_simps[symmetric] not_all}) - THEN' TRY' instantiate - THEN' TRY' (simplify_tac ctxt []) - THEN' TRY' (SOLVED' (fn _ => HEADGOAL ( (assume_tac ctxt) - ORELSE' - TRY' (dresolve_tac ctxt @{thms conjE} - THEN_ALL_NEW assume_tac ctxt) - ORELSE' - TRY' (dresolve_tac ctxt @{thms verit_forall_inst} - THEN_ALL_NEW assume_tac ctxt)))) - THEN' (TRY' (Classical.fast_tac ctxt)) - THEN' (TRY' (K (Clasimp.auto_tac ctxt)))) - end +fun or _ (thm :: _) _ = thm + | or ctxt thms t = replay_error ctxt "invalid bind application" Or thms t -fun or _ [thm] _ = thm + +(* Implication *) val implies_pos_thm = - [@{lemma "\(A \ B) \ \A \ B" - by blast}, - @{lemma "\(\A \ B) \ A \ B" - by blast}] + [@{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) -fun extract_rewrite_rule_assumption thms = +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 + \<^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 = + thms + |> filter is_context_rule + val rew = + thms + |> filter_out is_context_rule + |> filter is_rewrite_rule in - thms - |> filter is_rewrite_rule - |> map (fn thm => thm COMP @{thm eq_reflection}) + (ctxt_eq, rew) end -(* We need to unfold the assumptions if we are in a subproof: For multiple skolemization, the context -contains a mapping "verit_vrX \ Eps f". The variable "verit_vrX" must be unfolded to "Eps f". -Otherwise, the proof cannot be done. *) -fun skolem_forall ctxt (thms) t = +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 + | 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 ts = extract_rewrite_rule_assumption thms + val (ctxt_eq, ts) = 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 SMT_Replay_Methods.prove ctxt t (fn _ => - REPEAT_CHANGED (resolve_tac ctxt @{thms verit_sko_forall'}) - THEN' TRY' (simplify_tac ctxt ts) - THEN' TRY'(resolve_tac ctxt thms THEN_ALL_NEW resolve_tac ctxt @{thms refl}) - THEN' TRY' (resolve_tac ctxt @{thms refl})) + 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}))) + end +in + +val skolem_forall = + skolem_tac extract_all_forall_quantified_names @{thm verit_sko_forall_indirect} + @{thm verit_sko_forall_indirect2} + +val skolem_ex = + skolem_tac extract_all_exists_quantified_names @{thm verit_sko_ex_indirect} + @{thm verit_sko_ex_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) +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 -fun skolem_ex ctxt (thms) t = - let - val ts = extract_rewrite_rule_assumption thms - in - SMT_Replay_Methods.prove ctxt t (fn _ => - Raw_Simplifier.rewrite_goal_tac ctxt ts - THEN' REPEAT_CHANGED (resolve_tac ctxt @{thms verit_sko_ex'}) - THEN' REPEAT_CHANGED (resolve_tac ctxt thms THEN_ALL_NEW resolve_tac ctxt @{thms refl}) - THEN' TRY' (resolve_tac ctxt @{thms refl})) - 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 -fun eq_reflexive ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - resolve_tac ctxt [@{thm refl}]) +val duplicate_literals = + prove_abstract SMT_Replay_Methods.abstract_bool_shallow duplicate_literals_prove -fun connective_equiv ctxt thms t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt thms - THEN' K (Clasimp.auto_tac ctxt)) +val unit_res = prove_abstract SMT_Replay_Methods.abstract_bool_shallow SMT_Replay_Methods.prop_tac + +end fun normalized_input ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt prems - THEN' TRY' (simplify_tac ctxt []) - THEN' TRY' (K (Clasimp.auto_tac ctxt))) + 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} +val false_rule_thm = @{lemma \\False\ by blast} -fun false_rule ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - resolve_tac ctxt [false_rule_thm]) +fun false_rule ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t false_rule_thm -(* transitivity *) +(* Transitivity *) val trans_bool_thm = - @{lemma "P = Q \ Q \ P" by blast} -fun trans _ [thm1, thm2] _ = - (case (Thm.full_prop_of thm1, Thm.full_prop_of thm2) of - (\<^term>\Trueprop\ $ (Const(\<^const_name>\HOL.eq\, _) $ _ $ t2), - \<^term>\Trueprop\ $ (Const(\<^const_name>\HOL.eq\, _) $ t3 $ _)) => - if t2 = t3 then thm1 RSN (1, thm2 RSN (2, @{thm trans})) - else thm1 RSN (1, (thm2 RS sym) RSN (2, @{thm trans})) - | _ => trans_bool_thm OF [thm1, thm2]) - | trans ctxt (thm1 :: thm2 :: thms) t = - trans ctxt (trans ctxt [thm1, thm2] t :: thms) t + @{lemma \P = Q \ Q \ P\ by blast} -fun tmp_AC_rule ctxt _ t = - let - val simplify = - ctxt - |> empty_simpset - |> put_simpset HOL_basic_ss - |> (fn ctxt => ctxt addsimps @{thms ac_simps conj_ac}) - |> Simplifier.full_simp_tac - in SMT_Replay_Methods.prove ctxt t (fn _ => - REPEAT_ALL_NEW (simplify_tac ctxt [] - THEN' TRY' simplify - THEN' TRY' (Classical.fast_tac ctxt))) end +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 = t3 then thm1 RSN (1, thm2 RSN (2, @{thm trans})) + else if t2 = t4 then thm1 RSN ((1, ((thm2 RS sym)) RSN (2, @{thm trans}))) + else if t1 = 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 and_rule ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => - 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) - THEN' TRY' (simplify_tac ctxt [])) -fun not_and_rule ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt prems THEN' - Classical.fast_tac ctxt ORELSE' Clasimp.force_tac ctxt) +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))) -fun not_or_rule ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt prems THEN' - Classical.fast_tac ctxt ORELSE' Clasimp.force_tac ctxt) + +(* And/Or *) local - fun simplify_and_pos ctxt = - ctxt - |> empty_simpset - |> put_simpset HOL_basic_ss - |> (fn ctxt => ctxt addsimps @{thms eq_commute verit_ite_intro_simp if_cancel} - addsimps @{thms simp_thms de_Morgan_conj}) - |> Simplifier.full_simp_tac + 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 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_pos} + 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' (simplify_and_pos ctxt) - THEN' TRY' (assume_tac ctxt) - THEN' TRY' (Classical.fast_tac ctxt)) + 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 -fun and_neg_rule ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - REPEAT_CHANGED (resolve_tac ctxt @{thms verit_and_neg}) - THEN' simplify_tac ctxt @{thms de_Morgan_conj[symmetric] excluded_middle - excluded_middle[of \\_\, unfolded not_not]}) -fun or_pos_rule ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - simplify_tac ctxt @{thms simp_thms}) - -fun or_neg_rule ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - resolve_tac ctxt @{thms verit_or_neg} - THEN' (fn i => dresolve_tac ctxt @{thms verit_subst_bool} i - THEN assume_tac ctxt (i+1)) - THEN' simplify_tac ctxt @{thms simp_thms}) - -val not_equiv1_thm = - @{lemma "\(A \ B) \ A \ B" - by blast} - -fun not_equiv1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [not_equiv1_thm OF [thm]] - THEN' simplify_tac ctxt []) - -val not_equiv2_thm = - @{lemma "\(A \ B) \ \A \ \B" - by blast} - -fun not_equiv2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [not_equiv2_thm OF [thm]] - THEN' simplify_tac ctxt []) +(* 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 equiv1_thm = - @{lemma "(A \ B) \ \A \ B" - by blast} - -fun equiv1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [equiv1_thm OF [thm]] - THEN' simplify_tac ctxt []) +val equiv_pos1_thm = @{lemma \\(a \ b) \ a \ \b\ by blast+} +val equiv_pos1 = prove_equiv_pos_neg2 equiv_pos1_thm -val equiv2_thm = - @{lemma "(A \ B) \ A \ \B" - by blast} - -fun equiv2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [equiv2_thm OF [thm]] - THEN' simplify_tac ctxt []) +val equiv_pos2_thm = @{lemma \\a b. (a \ b) \ \a \ b\ by blast+} +val equiv_pos2 = prove_equiv_pos_neg2 equiv_pos2_thm - -val not_implies1_thm = - @{lemma "\(A \ B) \ A" - by blast} +val equiv_neg1_thm = @{lemma \(a \ b) \ \a \ \b\ by blast+} +val equiv_neg1 = prove_equiv_pos_neg2 equiv_neg1_thm -fun not_implies1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [not_implies1_thm OF [thm]] - THEN' simplify_tac ctxt []) +val equiv_neg2_thm = @{lemma \(a \ b) \ a \ b\ by blast} +val equiv_neg2 = prove_equiv_pos_neg2 equiv_neg2_thm -val not_implies2_thm = - @{lemma "\(A \B) \ \B" - by blast} - -fun not_implies2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt [not_implies2_thm OF [thm]] - THEN' simplify_tac ctxt []) +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 + 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' simplify_tac ctxt []) + THEN' TRY' (assume_tac ctxt)) end in -val implies_neg1_thm = - @{lemma "(a \ b) \ a" - by blast} - +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_thm = @{lemma \(a \ b) \ \b\ by blast} +val implies_neg2 = prove_implies_pos_neg implies_neg2_thm -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 -val implies_thm = - @{lemma "(~a \ b) \ a \ b" - "(a \ b) \ \a \ b" - by blast+} + +(* BFun *) + +local + val bfun_theorems = @{thms verit_bfun_elim} +in -fun implies_rules ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => +fun bfun_elim ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => Method.insert_tac ctxt prems - THEN' resolve_tac ctxt implies_thm - THEN' assume_tac ctxt) + THEN' TRY' (eqsubst_all ctxt bfun_theorems) + THEN' TRY' (simplify_tac ctxt @{thms eq_commute all_conj_distrib ex_disj_distrib})) + +end -(* -Here is a case where force_tac fails, but auto_tac succeeds: - Ex (P x) \ P x c \ - (\v0. if x then P True v0 else P False v0) \ (if x then P True c else P False c) +(* If-Then-Else *) -(this was before we added the eqsubst_tac). Therefore, to be safe, we add the fast, auto, and force. -*) -fun tmp_bfun_elim ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => - Method.insert_tac ctxt prems - THEN' REPEAT_CHANGED (EqSubst.eqsubst_tac ctxt [0] @{thms verit_tmp_bfun_elim}) - THEN' TRY' (simplify_tac ctxt []) - THEN' (Classical.fast_tac ctxt - ORELSE' K (Clasimp.auto_tac ctxt) - ORELSE' Clasimp.force_tac ctxt)) +local + fun prove_ite ite_thm thm ctxt = + resolve_tac ctxt [ite_thm OF [thm]] + THEN' K (unfold_tac ctxt @{thms not_not}) + THEN' TRY' (assume_tac ctxt) +in val ite_pos1_thm = - @{lemma "\(if x then P else Q) \ x \ Q" - by auto} + @{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} + @{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) +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} + @{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) +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" + @{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) +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) } + @{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 _ => - resolve_tac ctxt [ite1_thm OF [thm]]) +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) } + @{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 _ => - resolve_tac ctxt [ite2_thm OF [thm]]) - +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) } + @{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 _ => - resolve_tac ctxt [not_ite1_thm OF [thm]]) +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 _ => - resolve_tac ctxt [not_ite2_thm OF [thm]]) - + @{lemma \\(if x then P else Q) \ \x \ \P\ by (auto split: if_splits) } -fun unit_res 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 = Z3_Replay_Methods.unit_res ctxt thms t2 - in - @{thm verit_Pure_trans} OF [t', thm] - end +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 = + fun simplify_ite ctxt thms = ctxt |> empty_simpset |> put_simpset HOL_basic_ss - |> (fn ctxt => ctxt addsimps @{thms eq_commute verit_ite_intro_simp if_cancel} - addsimps @{thms simp_thms}) + |> (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' (Blast.blast_tac ctxt - ORELSE' K (Clasimp.auto_tac ctxt) - ORELSE' Clasimp.force_tac ctxt))) + 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 _ => - Classical.fast_tac ctxt) + 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 _ => - Classical.fast_tac ctxt) + partial_simplify_tac ctxt []) -fun qnt_join ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - Classical.fast_tac ctxt) - +end (* Equality *) @@ -679,165 +874,306 @@ 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)) - (* Rewrite might apply below choice. As we do not want to change them (it can break other - rewriting steps), we cannot use Term.lambda *) - fun abstract_over_no_choice (v, body) = - let - fun abs lev tm = - if v aconv tm then Bound lev - else - (case tm of - Abs (a, T, t) => Abs (a, T, abs (lev + 1) t) - | t as (Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ _) => t - | t $ u => - (abs lev t $ (abs lev u handle Same.SAME => u) - handle Same.SAME => t $ abs lev u) - | _ => raise Same.SAME); - in abs 0 body handle Same.SAME => body end; - - fun lambda_name (x, v) t = - Abs (if x = "" then Term.term_name v else x, fastype_of v, abstract_over_no_choice (v, t)); + 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 + | _ => Conv.all_conv ctrm)) - fun lambda v t = lambda_name ("", v) t; - - fun extract_equal_terms (Const(\<^const_name>\Trueprop\, _) $ t) = - let fun ext (Const(\<^const_name>\HOL.disj\, _) $ (Const(\<^const_name>\HOL.Not\, _) $ - (Const(\<^const_name>\HOL.eq\, _) $ t1 $ t2)) $ t) = - apfst (curry (op ::) (t1, t2)) (ext t) - | ext t = ([], t) - in ext t end - fun eq_congruent_tac ctxt t = + fun eq_pred_rewrite_tac ({context = ctxt, prems, concl, ...}: Subgoal.focus) = let - val (eqs, g) = extract_equal_terms t - fun replace1 (t1, t2) (g, tac) = - let - val abs_t1 = lambda t2 g - val subst = Drule.infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) [t1, t2, abs_t1]) - @{thm subst} - in (Term.betapply (abs_t1, t1), - tac THEN' resolve_tac ctxt [subst] - THEN' TRY' (assume_tac ctxt)) end - val (_, tac) = fold replace1 eqs (g, K all_tac) + 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 - tac + 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' REPEAT' (eresolve_tac ctxt @{thms conjI}) - THEN' eq_congruent_tac ctxt t - THEN' resolve_tac ctxt @{thms refl excluded_middle - excluded_middle[of \\_\, unfolded not_not]}) + 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 *) +(* Subproof *) fun subproof ctxt [prem] t = - SMT_Replay_Methods.prove ctxt t (fn _ => - (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)) + SMT_Replay_Methods.prove ctxt t (fn _ => + (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 -(* la_rw_eq *) +(* 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_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 -val la_rw_eq_thm = @{lemma \(a :: nat) = b \ (a \ b) \ (a \ b)\ - by auto} + fun chain_rewrite_thms ctxt thms = + TRY' (rewrite_only_thms ctxt thms) + THEN' TRY' (rewrite_thms ctxt thms) + + fun chain_rewrite_two_step_with_ac_simps ctxt thms1 thms2 = + TRY' (rewrite_only_thms ctxt thms1) + THEN' TRY' (rewrite_thms ctxt thms2) + + 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 la_rw_eq ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => - resolve_tac ctxt [la_rw_eq_thm]) +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})) + +fun rewrite_not_simplify ctxt _ t = + SMT_Replay_Methods.prove ctxt t (fn _ => + (rewrite_only_thms ctxt @{thms verit_not_simplify})) + +fun rewrite_equiv_simplify ctxt _ t = + SMT_Replay_Methods.prove ctxt t (fn _ => + (rewrite_only_thms 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 ctxt @{thms verit_implies_simplify})) -(* congruence *) -fun cong ctxt thms = SMT_Replay_Methods.try_provers ctxt - (string_of_verit_rule Cong) [ - ("basic", SMT_Replay_Methods.cong_basic ctxt thms), - ("full", SMT_Replay_Methods.cong_full ctxt thms), - ("unfolding then auto", SMT_Replay_Methods.cong_unfolding_first ctxt thms)] thms +(* 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 ctxt @{thms verit_ite_simplify})) +end -fun unsupported rule ctxt thms _ t = SMT_Replay_Methods.replay_error ctxt "Unsupported verit rule" - rule thms t +(* 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}) + +fun prod_simplify ctxt _ t = + SMT_Replay_Methods.prove ctxt t (fn _ => + simplify_arith ctxt @{thms verit_prod_simplify}) -fun ignore_args f ctxt thm _ t = f ctxt thm t +end + + +(* 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 choose Bind = ignore_args bind - | choose Refl = ignore_args refl +fun la_disequality ctxt _ t = + SMT_Replay_Methods.match_instantiate ctxt t @{thm verit_la_disequality} + +end + +(* Combining all together *) + +fun unsupported rule ctxt thms _ _ _ = SMT_Replay_Methods.replay_error ctxt "Unsupported verit rule" + rule thms + +fun ignore_args f ctxt thm _ _ _ t = f ctxt thm t +fun ignore_decls f ctxt thm args insts _ t = f ctxt thm args insts t +fun ignore_insts f ctxt thm args _ _ t = f ctxt thm args t + +fun choose And = ignore_args and_rule + | choose And_Neg = ignore_args and_neg_rule | choose And_Pos = ignore_args and_pos - | choose And_Neg = ignore_args and_neg_rule + | choose And_Simplify = ignore_args rewrite_and_simplify + | choose Bind = ignore_insts bind + | choose Bool_Simplify = ignore_args rewrite_bool_simplify + | choose Comp_Simplify = ignore_args rewrite_comp_simplify | choose Cong = ignore_args cong + | choose Connective_Def = ignore_args rewrite_connective_def + | choose Duplicate_Literals = ignore_args duplicate_literals + | choose Eq_Congruent = ignore_args eq_congruent + | choose Eq_Congruent_Pred = ignore_args eq_congruent_pred + | choose Eq_Reflexive = ignore_args eq_reflexive + | choose Eq_Simplify = ignore_args rewrite_eq_simplify + | choose Eq_Transitive = ignore_args eq_transitive + | choose Equiv1 = ignore_args equiv1 + | choose Equiv2 = ignore_args equiv2 + | choose Equiv_neg1 = ignore_args equiv_neg1 + | choose Equiv_neg2 = ignore_args equiv_neg2 | choose Equiv_pos1 = ignore_args equiv_pos1 | choose Equiv_pos2 = ignore_args equiv_pos2 - | choose Equiv_neg1 = ignore_args equiv_neg1 - | choose Equiv_neg2 = ignore_args equiv_neg2 - | choose Equiv1 = ignore_args equiv1 - | choose Equiv2 = ignore_args equiv2 + | choose Equiv_Simplify = ignore_args rewrite_equiv_simplify + | choose False = ignore_args false_rule + | choose Forall_Inst = ignore_decls forall_inst + | choose Implies = ignore_args implies_rules + | choose Implies_Neg1 = ignore_args implies_neg1 + | choose Implies_Neg2 = ignore_args implies_neg2 + | choose Implies_Pos = ignore_args implies_pos + | choose Implies_Simplify = ignore_args rewrite_implies_simplify + | choose ITE1 = ignore_args ite1 + | choose ITE2 = ignore_args ite2 + | choose ITE_Intro = ignore_args ite_intro + | choose ITE_Neg1 = ignore_args ite_neg1 + | choose ITE_Neg2 = ignore_args ite_neg2 + | choose ITE_Pos1 = ignore_args ite_pos1 + | choose ITE_Pos2 = ignore_args ite_pos2 + | choose ITE_Simplify = ignore_args rewrite_ite_simplify + | choose LA_Disequality = ignore_args la_disequality + | choose LA_Generic = ignore_insts la_generic + | choose LA_RW_Eq = ignore_args la_rw_eq + | choose LA_Tautology = ignore_args SMT_Replay_Methods.arith_th_lemma_wo_shallow + | choose LA_Totality = ignore_args SMT_Replay_Methods.arith_th_lemma_wo_shallow + | choose LIA_Generic = ignore_args lia_generic + | choose Local_Input = ignore_args refl + | choose Minus_Simplify = ignore_args rewrite_minus_simplify + | choose NLA_Generic = ignore_args SMT_Replay_Methods.arith_th_lemma_wo_shallow + | choose Normalized_Input = ignore_args normalized_input + | choose Not_And = ignore_args not_and_rule | choose Not_Equiv1 = ignore_args not_equiv1 | choose Not_Equiv2 = ignore_args not_equiv2 | choose Not_Implies1 = ignore_args not_implies1 | choose Not_Implies2 = ignore_args not_implies2 - | choose Implies_Neg1 = ignore_args implies_neg1 - | choose Implies_Neg2 = ignore_args implies_neg2 - | choose Implies_Pos = ignore_args implies_pos - | choose Implies = ignore_args implies_rules - | choose Forall_Inst = forall_inst - | choose Skolem_Forall = ignore_args skolem_forall - | choose Skolem_Ex = ignore_args skolem_ex - | choose Or = ignore_args or - | choose Theory_Resolution = ignore_args unit_res - | choose Resolution = ignore_args unit_res - | choose Eq_Reflexive = ignore_args eq_reflexive - | choose Connective_Equiv = ignore_args connective_equiv - | choose Trans = ignore_args trans - | choose False = ignore_args false_rule - | choose Tmp_AC_Simp = ignore_args tmp_AC_rule - | choose And = ignore_args and_rule - | choose Not_And = ignore_args not_and_rule - | choose Not_Or = ignore_args not_or_rule - | choose Or_Pos = ignore_args or_pos_rule - | choose Or_Neg = ignore_args or_neg_rule - | choose Tmp_Bfun_Elim = ignore_args tmp_bfun_elim - | choose ITE1 = ignore_args ite1 - | choose ITE2 = ignore_args ite2 | choose Not_ITE1 = ignore_args not_ite1 | choose Not_ITE2 = ignore_args not_ite2 - | choose ITE_Pos1 = ignore_args ite_pos1 - | choose ITE_Pos2 = ignore_args ite_pos2 - | choose ITE_Neg1 = ignore_args ite_neg1 - | choose ITE_Neg2 = ignore_args ite_neg2 - | choose ITE_Intro = ignore_args ite_intro - | choose LA_Disequality = ignore_args Z3_Replay_Methods.arith_th_lemma - | choose LIA_Generic = ignore_args Z3_Replay_Methods.arith_th_lemma - | choose LA_Generic = ignore_args Z3_Replay_Methods.arith_th_lemma - | choose LA_Totality = ignore_args Z3_Replay_Methods.arith_th_lemma - | choose LA_Tautology = ignore_args Z3_Replay_Methods.arith_th_lemma - | choose LA_RW_Eq = ignore_args la_rw_eq - | choose NLA_Generic = ignore_args Z3_Replay_Methods.arith_th_lemma - | choose Normalized_Input = ignore_args normalized_input + | choose Not_Not = ignore_args not_not + | choose Not_Or = ignore_args not_or_rule + | choose Not_Simplify = ignore_args rewrite_not_simplify + | choose Or = ignore_args or + | choose Or_Neg = ignore_args or_neg_rule + | choose Or_Pos = ignore_args or_pos_rule + | choose Or_Simplify = ignore_args rewrite_or_simplify + | choose Prod_Simplify = ignore_args prod_simplify + | choose Qnt_Join = ignore_args qnt_join | choose Qnt_Rm_Unused = ignore_args qnt_rm_unused + | choose Onepoint = ignore_args onepoint | choose Qnt_Simplify = ignore_args qnt_simplify - | choose Qnt_Join = ignore_args qnt_join - | choose Eq_Congruent_Pred = ignore_args eq_congruent_pred - | choose Eq_Congruent = ignore_args eq_congruent_pred - | choose Eq_Transitive = ignore_args eq_transitive - | choose Local_Input = ignore_args refl + | choose Refl = ignore_args refl + | choose Resolution = ignore_args unit_res + | choose Skolem_Ex = ignore_args skolem_ex + | choose Skolem_Forall = ignore_args skolem_forall | choose Subproof = ignore_args subproof + | choose Sum_Simplify = ignore_args sum_simplify + | choose Tautological_Clause = ignore_args tautological_clause + | choose Theory_Resolution = ignore_args unit_res + | choose AC_Simp = ignore_args tmp_AC_rule + | choose Bfun_Elim = ignore_args bfun_elim + | choose Qnt_CNF = ignore_args qnt_cnf + | choose Trans = ignore_args trans | choose r = unsupported (string_of_verit_rule r) -type Verit_method = Proof.context -> thm list -> term -> thm +type Lethe_method = Proof.context -> thm list -> term -> thm type abs_context = int * term Termtab.table -fun with_tracing rule method ctxt thms args t = +fun with_tracing rule method ctxt thms args insts decls t = let val _ = SMT_Replay_Methods.trace_goal ctxt rule thms t - in method ctxt thms args t end + in method ctxt thms args insts decls t end fun method_for rule = with_tracing rule (choose (verit_rule_of rule)) +fun discharge ctxt [thm] t = + SMT_Replay_Methods.prove ctxt t (fn _ => + resolve_tac ctxt [thm] THEN_ALL_NEW (resolve_tac ctxt @{thms refl})) + end; diff -r 2c7f0ef8323a -r b44e894796d5 src/HOL/Tools/SMT/z3_replay_methods.ML --- a/src/HOL/Tools/SMT/z3_replay_methods.ML Mon Oct 12 17:42:15 2020 +0200 +++ b/src/HOL/Tools/SMT/z3_replay_methods.ML Mon Oct 12 18:59:44 2020 +0200 @@ -55,19 +55,18 @@ SMT_Replay_Methods.replay_error ctxt msg (Z3_Proof.string_of_rule rule) thms t fun replay_rule_error ctxt rule thms t = - SMT_Replay_Methods.replay_rule_error ctxt (Z3_Proof.string_of_rule rule) thms t + SMT_Replay_Methods.replay_rule_error "Z3" ctxt (Z3_Proof.string_of_rule rule) thms t fun trace_goal ctxt rule thms t = SMT_Replay_Methods.trace_goal ctxt (Z3_Proof.string_of_rule rule) thms t -fun as_prop (t as Const (\<^const_name>\Trueprop\, _) $ _) = t - | as_prop t = HOLogic.mk_Trueprop t - fun dest_prop (Const (\<^const_name>\Trueprop\, _) $ t) = t | dest_prop t = t fun dest_thm thm = dest_prop (Thm.concl_of thm) +val try_provers = SMT_Replay_Methods.try_provers "Z3" + fun prop_tac ctxt prems = Method.insert_tac ctxt prems THEN' SUBGOAL (fn (prop, i) => @@ -95,9 +94,6 @@ fold_map (SMT_Replay_Methods.abstract_arith ctxt o dest_thm) thms ##>> SMT_Replay_Methods.abstract_arith ctxt (dest_prop t)) -val _ = Theory.setup (Context.theory_map ( - SMT_Replay_Methods.add_th_lemma_method ("arith", arith_th_lemma))) - fun th_lemma name ctxt thms = (case Symtab.lookup (SMT_Replay_Methods.get_th_lemma_method ctxt) name of SOME method => method ctxt thms @@ -136,7 +132,7 @@ (* congruence *) -fun cong ctxt thms = SMT_Replay_Methods.try_provers ctxt +fun cong ctxt thms = try_provers ctxt (Z3_Proof.string_of_rule Z3_Proof.Monotonicity) [ ("basic", SMT_Replay_Methods.cong_basic ctxt thms), ("full", SMT_Replay_Methods.cong_full ctxt thms)] thms @@ -252,13 +248,13 @@ | _ => Conv.sub_conv (Conv.top_sweep_conv if_context_conv) ctxt) ct fun lift_ite_rewrite ctxt t = - SMT_Replay_Methods.prove ctxt t (fn ctxt => + SMT_Replay_Methods.prove ctxt t (fn ctxt => CONVERSION (HOLogic.Trueprop_conv (Conv.binop_conv (if_context_conv ctxt))) THEN' resolve_tac ctxt @{thms refl}) fun prove_conj_disj_perm ctxt t = SMT_Replay_Methods.prove ctxt t Conj_Disj_Perm.conj_disj_perm_tac -fun rewrite ctxt _ = SMT_Replay_Methods.try_provers ctxt +fun rewrite ctxt _ = try_provers ctxt (Z3_Proof.string_of_rule Z3_Proof.Rewrite) [ ("rules", apply_rule ctxt), ("conj_disj_perm", prove_conj_disj_perm ctxt), @@ -386,7 +382,7 @@ HOLogic.mk_disj o swap) | u => SMT_Replay_Methods.prove_abstract' ctxt t prop_tac (SMT_Replay_Methods.abstract_prop u)) -fun def_axiom ctxt _ = SMT_Replay_Methods.try_provers ctxt +fun def_axiom ctxt _ = try_provers ctxt (Z3_Proof.string_of_rule Z3_Proof.Def_Axiom) [ ("rules", apply_rule ctxt), ("disj", def_axiom_disj ctxt)] [] @@ -410,7 +406,7 @@ fold_map (SMT_Replay_Methods.abstract_prop o dest_thm) thms ##>> SMT_Replay_Methods.abstract_prop (dest_prop t)) -fun nnf ctxt rule thms = SMT_Replay_Methods.try_provers ctxt (Z3_Proof.string_of_rule rule) [ +fun nnf ctxt rule thms = try_provers ctxt (Z3_Proof.string_of_rule rule) [ ("prop", nnf_prop ctxt thms), ("quant", quant_intro ctxt [hd thms])] thms